-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathLas_funciones_de_extraccion_no_estan_acotadas.thy
101 lines (91 loc) · 2.98 KB
/
Las_funciones_de_extraccion_no_estan_acotadas.thy
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
(* Las_funciones_de_extraccion_no_estan_acotadas.thy
-- Las funciones de extracción no están acotadas.
-- José A. Alonso Jiménez <https://jaalonso.github.io>
-- Sevilla, 11-julio-2024
-- ------------------------------------------------------------------ *)
(* ---------------------------------------------------------------------
-- Para extraer una subsucesión se aplica una función de extracción que
-- conserva el orden; por ejemplo, la subsucesión
-- uₒ, u₂, u₄, u₆, ...
-- se ha obtenido con la función de extracción \<phi> tal que \<phi>(n) = 2*n.
--
-- En Isabelle/HOL, se puede definir que \<phi> es una función de
-- extracción por
-- definition extraccion :: "(nat \<Rightarrow> nat) \<Rightarrow> bool" where
-- "extraccion \<phi> \<longleftrightarrow> (\<forall> n m. n < m \<longrightarrow> \<phi> n < \<phi> m)"
--
-- Demostrar que las funciones de extracción no está acotadas; es decir,
-- que si \<phi> es una función de extracción, entonces
-- \<forall> N N', \<exists> k \<ge> N', \<phi> k \<ge> N
-- ------------------------------------------------------------------ *)
theory Las_funciones_de_extraccion_no_estan_acotadas
imports Main
begin
definition extraccion :: "(nat \<Rightarrow> nat) \<Rightarrow> bool" where
"extraccion \<phi> \<longleftrightarrow> (\<forall> n m. n < m \<longrightarrow> \<phi> n < \<phi> m)"
(* En la demostración se usará el siguiente lema *)
lemma aux :
assumes "extraccion \<phi>"
shows "n \<le> \<phi> n"
proof (induct n)
show "0 \<le> \<phi> 0"
by simp
next
fix n
assume HI : "n \<le> \<phi> n"
also have "\<phi> n < \<phi> (Suc n)"
using assms extraccion_def by blast
finally show "Suc n \<le> \<phi> (Suc n)"
by simp
qed
(* 1\<ordfeminine> demostración *)
lemma
assumes "extraccion \<phi>"
shows "\<forall> N N'. \<exists> k \<ge> N'. \<phi> k \<ge> N"
proof (intro allI)
fix N N' :: nat
let ?k = "max N N'"
have "max N N' \<le> ?k"
by (rule le_refl)
then have hk : "N \<le> ?k \<and> N' \<le> ?k"
by (simp only: max.bounded_iff)
then have "?k \<ge> N'"
by (rule conjunct2)
moreover
have "N \<le> \<phi> ?k"
proof -
have "N \<le> ?k"
using hk by (rule conjunct1)
also have "\<dots> \<le> \<phi> ?k"
using assms by (rule aux)
finally show "N \<le> \<phi> ?k"
by this
qed
ultimately have "?k \<ge> N' \<and> \<phi> ?k \<ge> N"
by (rule conjI)
then show "\<exists>k \<ge> N'. \<phi> k \<ge> N"
by (rule exI)
qed
(* 2\<ordfeminine> demostración *)
lemma
assumes "extraccion \<phi>"
shows "\<forall> N N'. \<exists> k \<ge> N'. \<phi> k \<ge> N"
proof (intro allI)
fix N N' :: nat
let ?k = "max N N'"
have "?k \<ge> N'"
by simp
moreover
have "N \<le> \<phi> ?k"
proof -
have "N \<le> ?k"
by simp
also have "\<dots> \<le> \<phi> ?k"
using assms by (rule aux)
finally show "N \<le> \<phi> ?k"
by this
qed
ultimately show "\<exists>k \<ge> N'. \<phi> k \<ge> N"
by blast
qed
end