author | haftmann |
Sun, 27 Jul 2025 17:52:06 +0200 | |
changeset 82909 | e4fae2227594 |
parent 76217 | 8655344f1cf6 |
permissions | -rw-r--r-- |
13505 | 1 |
(* Title: ZF/Constructible/L_axioms.thy |
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
3 |
*) |
|
13429 | 4 |
|
60770 | 5 |
section \<open>The ZF Axioms (Except Separation) in L\<close> |
13223 | 6 |
|
16417 | 7 |
theory L_axioms imports Formula Relative Reflection MetaExists begin |
13223 | 8 |
|
61798 | 9 |
text \<open>The class L satisfies the premises of locale \<open>M_trivial\<close>\<close> |
13223 | 10 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
11 |
lemma transL: "\<lbrakk>y\<in>x; L(x)\<rbrakk> \<Longrightarrow> L(y)" |
13429 | 12 |
apply (insert Transset_Lset) |
13 |
apply (simp add: Transset_def L_def, blast) |
|
13223 | 14 |
done |
15 |
||
16 |
lemma nonempty: "L(0)" |
|
13429 | 17 |
apply (simp add: L_def) |
18 |
apply (blast intro: zero_in_Lset) |
|
13223 | 19 |
done |
20 |
||
13563 | 21 |
theorem upair_ax: "upair_ax(L)" |
13223 | 22 |
apply (simp add: upair_ax_def upair_def, clarify) |
13429 | 23 |
apply (rule_tac x="{x,y}" in rexI) |
24 |
apply (simp_all add: doubleton_in_L) |
|
13223 | 25 |
done |
26 |
||
13563 | 27 |
theorem Union_ax: "Union_ax(L)" |
13223 | 28 |
apply (simp add: Union_ax_def big_union_def, clarify) |
46823 | 29 |
apply (rule_tac x="\<Union>(x)" in rexI) |
13429 | 30 |
apply (simp_all add: Union_in_L, auto) |
31 |
apply (blast intro: transL) |
|
13223 | 32 |
done |
33 |
||
13563 | 34 |
theorem power_ax: "power_ax(L)" |
13223 | 35 |
apply (simp add: power_ax_def powerset_def Relative.subset_def, clarify) |
13429 | 36 |
apply (rule_tac x="{y \<in> Pow(x). L(y)}" in rexI) |
13299 | 37 |
apply (simp_all add: LPow_in_L, auto) |
13429 | 38 |
apply (blast intro: transL) |
13223 | 39 |
done |
40 |
||
69593 | 41 |
text\<open>We don't actually need \<^term>\<open>L\<close> to satisfy the foundation axiom.\<close> |
13563 | 42 |
theorem foundation_ax: "foundation_ax(L)" |
43 |
apply (simp add: foundation_ax_def) |
|
44 |
apply (rule rallI) |
|
45 |
apply (cut_tac A=x in foundation) |
|
46 |
apply (blast intro: transL) |
|
47 |
done |
|
48 |
||
60770 | 49 |
subsection\<open>For L to satisfy Replacement\<close> |
13223 | 50 |
|
51 |
(*Can't move these to Formula unless the definition of univalent is moved |
|
52 |
there too!*) |
|
53 |
||
54 |
lemma LReplace_in_Lset: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
55 |
"\<lbrakk>X \<in> Lset(i); univalent(L,X,Q); Ord(i)\<rbrakk> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
56 |
\<Longrightarrow> \<exists>j. Ord(j) \<and> Replace(X, \<lambda>x y. Q(x,y) \<and> L(y)) \<subseteq> Lset(j)" |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
57 |
apply (rule_tac x="\<Union>y \<in> Replace(X, \<lambda>x y. Q(x,y) \<and> L(y)). succ(lrank(y))" |
13223 | 58 |
in exI) |
59 |
apply simp |
|
13429 | 60 |
apply clarify |
61 |
apply (rule_tac a=x in UN_I) |
|
62 |
apply (simp_all add: Replace_iff univalent_def) |
|
63 |
apply (blast dest: transL L_I) |
|
13223 | 64 |
done |
65 |
||
13429 | 66 |
lemma LReplace_in_L: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
67 |
"\<lbrakk>L(X); univalent(L,X,Q)\<rbrakk> |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
68 |
\<Longrightarrow> \<exists>Y. L(Y) \<and> Replace(X, \<lambda>x y. Q(x,y) \<and> L(y)) \<subseteq> Y" |
13429 | 69 |
apply (drule L_D, clarify) |
13223 | 70 |
apply (drule LReplace_in_Lset, assumption+) |
71 |
apply (blast intro: L_I Lset_in_Lset_succ) |
|
72 |
done |
|
73 |
||
13563 | 74 |
theorem replacement: "replacement(L,P)" |
13223 | 75 |
apply (simp add: replacement_def, clarify) |
13429 | 76 |
apply (frule LReplace_in_L, assumption+, clarify) |
77 |
apply (rule_tac x=Y in rexI) |
|
78 |
apply (simp_all add: Replace_iff univalent_def, blast) |
|
13223 | 79 |
done |
80 |
||
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
81 |
lemma strong_replacementI [rule_format]: |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
82 |
"\<lbrakk>\<forall>B[L]. separation(L, \<lambda>u. \<exists>x[L]. x\<in>B \<and> P(x,u))\<rbrakk> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
83 |
\<Longrightarrow> strong_replacement(L,P)" |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
84 |
apply (simp add: strong_replacement_def, clarify) |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
85 |
apply (frule replacementD [OF replacement], assumption, clarify) |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
86 |
apply (drule_tac x=A in rspec, clarify) |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
87 |
apply (drule_tac z=Y in separationD, assumption, clarify) |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
88 |
apply (rule_tac x=y in rexI, force, assumption) |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
89 |
done |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
90 |
|
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
91 |
|
61798 | 92 |
subsection\<open>Instantiating the locale \<open>M_trivial\<close>\<close> |
60770 | 93 |
text\<open>No instances of Separation yet.\<close> |
13291 | 94 |
|
95 |
lemma Lset_mono_le: "mono_le_subset(Lset)" |
|
13429 | 96 |
by (simp add: mono_le_subset_def le_imp_subset Lset_mono) |
13291 | 97 |
|
98 |
lemma Lset_cont: "cont_Ord(Lset)" |
|
13429 | 99 |
by (simp add: cont_Ord_def Limit_Lset_eq OUnion_def Limit_is_Ord) |
13291 | 100 |
|
13428 | 101 |
lemmas L_nat = Ord_in_L [OF Ord_nat] |
13291 | 102 |
|
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
103 |
theorem M_trivial_L: "M_trivial(L)" |
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13563
diff
changeset
|
104 |
apply (rule M_trivial.intro) |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
105 |
apply (rule M_trans.intro) |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
106 |
apply (erule (1) transL) |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
107 |
apply(rule exI,rule nonempty) |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
108 |
apply (rule M_trivial_axioms.intro) |
13428 | 109 |
apply (rule upair_ax) |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
110 |
apply (rule Union_ax) |
13428 | 111 |
done |
13291 | 112 |
|
71568
1005c50b2750
prefer strict qualification (default for 'interpretation', see 461ee3e49ad3) as proposed by Pedro Sánchez Terraf;
wenzelm
parents:
71417
diff
changeset
|
113 |
interpretation L: M_trivial L by (rule M_trivial_L) |
15696 | 114 |
|
15764 | 115 |
(* Replaces the following declarations... |
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13563
diff
changeset
|
116 |
lemmas rall_abs = M_trivial.rall_abs [OF M_trivial_L] |
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13563
diff
changeset
|
117 |
and rex_abs = M_trivial.rex_abs [OF M_trivial_L] |
15764 | 118 |
... |
13429 | 119 |
declare rall_abs [simp] |
120 |
declare rex_abs [simp] |
|
15764 | 121 |
...and dozens of similar ones. |
15696 | 122 |
*) |
13291 | 123 |
|
61798 | 124 |
subsection\<open>Instantiation of the locale \<open>reflection\<close>\<close> |
13291 | 125 |
|
60770 | 126 |
text\<open>instances of locale constants\<close> |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
127 |
|
21233 | 128 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
129 |
L_F0 :: "[i\<Rightarrow>o,i] \<Rightarrow> i" where |
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
130 |
"L_F0(P,y) \<equiv> \<mu> b. (\<exists>z. L(z) \<and> P(\<langle>y,z\<rangle>)) \<longrightarrow> (\<exists>z\<in>Lset(b). P(\<langle>y,z\<rangle>))" |
13291 | 131 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
132 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
133 |
L_FF :: "[i\<Rightarrow>o,i] \<Rightarrow> i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
134 |
"L_FF(P) \<equiv> \<lambda>a. \<Union>y\<in>Lset(a). L_F0(P,y)" |
13291 | 135 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
136 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
137 |
L_ClEx :: "[i\<Rightarrow>o,i] \<Rightarrow> o" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
138 |
"L_ClEx(P) \<equiv> \<lambda>a. Limit(a) \<and> normalize(L_FF(P),a) = a" |
13291 | 139 |
|
140 |
||
60770 | 141 |
text\<open>We must use the meta-existential quantifier; otherwise the reflection |
142 |
terms become enormous!\<close> |
|
21233 | 143 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
144 |
L_Reflects :: "[i\<Rightarrow>o,[i,i]\<Rightarrow>o] \<Rightarrow> prop" (\<open>(3REFLECTS/ [_,/ _])\<close>) where |
76214 | 145 |
"REFLECTS[P,Q] \<equiv> (\<Or>Cl. Closed_Unbounded(Cl) \<and> |
46823 | 146 |
(\<forall>a. Cl(a) \<longrightarrow> (\<forall>x \<in> Lset(a). P(x) \<longleftrightarrow> Q(a,x))))" |
13291 | 147 |
|
148 |
||
13314 | 149 |
theorem Triv_reflection: |
150 |
"REFLECTS[P, \<lambda>a x. P(x)]" |
|
13429 | 151 |
apply (simp add: L_Reflects_def) |
152 |
apply (rule meta_exI) |
|
153 |
apply (rule Closed_Unbounded_Ord) |
|
13314 | 154 |
done |
155 |
||
156 |
theorem Not_reflection: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
157 |
"REFLECTS[P,Q] \<Longrightarrow> REFLECTS[\<lambda>x. \<not>P(x), \<lambda>a x. \<not>Q(a,x)]" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
158 |
unfolding L_Reflects_def |
13429 | 159 |
apply (erule meta_exE) |
160 |
apply (rule_tac x=Cl in meta_exI, simp) |
|
13314 | 161 |
done |
162 |
||
163 |
theorem And_reflection: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
164 |
"\<lbrakk>REFLECTS[P,Q]; REFLECTS[P',Q']\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
165 |
\<Longrightarrow> REFLECTS[\<lambda>x. P(x) \<and> P'(x), \<lambda>a x. Q(a,x) \<and> Q'(a,x)]" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
166 |
unfolding L_Reflects_def |
13429 | 167 |
apply (elim meta_exE) |
168 |
apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI) |
|
169 |
apply (simp add: Closed_Unbounded_Int, blast) |
|
13314 | 170 |
done |
171 |
||
172 |
theorem Or_reflection: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
173 |
"\<lbrakk>REFLECTS[P,Q]; REFLECTS[P',Q']\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
174 |
\<Longrightarrow> REFLECTS[\<lambda>x. P(x) \<or> P'(x), \<lambda>a x. Q(a,x) \<or> Q'(a,x)]" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
175 |
unfolding L_Reflects_def |
13429 | 176 |
apply (elim meta_exE) |
177 |
apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI) |
|
178 |
apply (simp add: Closed_Unbounded_Int, blast) |
|
13314 | 179 |
done |
180 |
||
181 |
theorem Imp_reflection: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
182 |
"\<lbrakk>REFLECTS[P,Q]; REFLECTS[P',Q']\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
183 |
\<Longrightarrow> REFLECTS[\<lambda>x. P(x) \<longrightarrow> P'(x), \<lambda>a x. Q(a,x) \<longrightarrow> Q'(a,x)]" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
184 |
unfolding L_Reflects_def |
13429 | 185 |
apply (elim meta_exE) |
186 |
apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI) |
|
187 |
apply (simp add: Closed_Unbounded_Int, blast) |
|
13314 | 188 |
done |
189 |
||
190 |
theorem Iff_reflection: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
191 |
"\<lbrakk>REFLECTS[P,Q]; REFLECTS[P',Q']\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
192 |
\<Longrightarrow> REFLECTS[\<lambda>x. P(x) \<longleftrightarrow> P'(x), \<lambda>a x. Q(a,x) \<longleftrightarrow> Q'(a,x)]" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
193 |
unfolding L_Reflects_def |
13429 | 194 |
apply (elim meta_exE) |
195 |
apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI) |
|
196 |
apply (simp add: Closed_Unbounded_Int, blast) |
|
13314 | 197 |
done |
198 |
||
199 |
||
13434 | 200 |
lemma reflection_Lset: "reflection(Lset)" |
13651
ac80e101306a
Cosmetic changes suggested by writing the paper. Deleted some
paulson
parents:
13634
diff
changeset
|
201 |
by (blast intro: reflection.intro Lset_mono_le Lset_cont |
ac80e101306a
Cosmetic changes suggested by writing the paper. Deleted some
paulson
parents:
13634
diff
changeset
|
202 |
Formula.Pair_in_LLimit)+ |
ac80e101306a
Cosmetic changes suggested by writing the paper. Deleted some
paulson
parents:
13634
diff
changeset
|
203 |
|
13434 | 204 |
|
13314 | 205 |
theorem Ex_reflection: |
206 |
"REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))] |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
207 |
\<Longrightarrow> REFLECTS[\<lambda>x. \<exists>z. L(z) \<and> P(x,z), \<lambda>a x. \<exists>z\<in>Lset(a). Q(a,x,z)]" |
76217 | 208 |
unfolding L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def |
13429 | 209 |
apply (elim meta_exE) |
13314 | 210 |
apply (rule meta_exI) |
13434 | 211 |
apply (erule reflection.Ex_reflection [OF reflection_Lset]) |
13291 | 212 |
done |
213 |
||
13314 | 214 |
theorem All_reflection: |
215 |
"REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))] |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
216 |
\<Longrightarrow> REFLECTS[\<lambda>x. \<forall>z. L(z) \<longrightarrow> P(x,z), \<lambda>a x. \<forall>z\<in>Lset(a). Q(a,x,z)]" |
76217 | 217 |
unfolding L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def |
13429 | 218 |
apply (elim meta_exE) |
13314 | 219 |
apply (rule meta_exI) |
13434 | 220 |
apply (erule reflection.All_reflection [OF reflection_Lset]) |
13291 | 221 |
done |
222 |
||
13314 | 223 |
theorem Rex_reflection: |
224 |
"REFLECTS[ \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))] |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
225 |
\<Longrightarrow> REFLECTS[\<lambda>x. \<exists>z[L]. P(x,z), \<lambda>a x. \<exists>z\<in>Lset(a). Q(a,x,z)]" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
226 |
unfolding rex_def |
13314 | 227 |
apply (intro And_reflection Ex_reflection, assumption) |
228 |
done |
|
13291 | 229 |
|
13314 | 230 |
theorem Rall_reflection: |
231 |
"REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))] |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
232 |
\<Longrightarrow> REFLECTS[\<lambda>x. \<forall>z[L]. P(x,z), \<lambda>a x. \<forall>z\<in>Lset(a). Q(a,x,z)]" |
76216
9fc34f76b4e8
getting rid of apply (unfold ...)
paulson <lp15@cam.ac.uk>
parents:
76215
diff
changeset
|
233 |
unfolding rall_def |
13314 | 234 |
apply (intro Imp_reflection All_reflection, assumption) |
235 |
done |
|
236 |
||
60770 | 237 |
text\<open>This version handles an alternative form of the bounded quantifier |
61798 | 238 |
in the second argument of \<open>REFLECTS\<close>.\<close> |
13440 | 239 |
theorem Rex_reflection': |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
240 |
"REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))] |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
241 |
\<Longrightarrow> REFLECTS[\<lambda>x. \<exists>z[L]. P(x,z), \<lambda>a x. \<exists>z[##Lset(a)]. Q(a,x,z)]" |
76217 | 242 |
unfolding setclass_def rex_def |
13440 | 243 |
apply (erule Rex_reflection [unfolded rex_def Bex_def]) |
244 |
done |
|
245 |
||
60770 | 246 |
text\<open>As above.\<close> |
13440 | 247 |
theorem Rall_reflection': |
248 |
"REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))] |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
249 |
\<Longrightarrow> REFLECTS[\<lambda>x. \<forall>z[L]. P(x,z), \<lambda>a x. \<forall>z[##Lset(a)]. Q(a,x,z)]" |
76217 | 250 |
unfolding setclass_def rall_def |
13440 | 251 |
apply (erule Rall_reflection [unfolded rall_def Ball_def]) |
252 |
done |
|
253 |
||
13429 | 254 |
lemmas FOL_reflections = |
13314 | 255 |
Triv_reflection Not_reflection And_reflection Or_reflection |
256 |
Imp_reflection Iff_reflection Ex_reflection All_reflection |
|
13440 | 257 |
Rex_reflection Rall_reflection Rex_reflection' Rall_reflection' |
13291 | 258 |
|
259 |
lemma ReflectsD: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
260 |
"\<lbrakk>REFLECTS[P,Q]; Ord(i)\<rbrakk> |
76214 | 261 |
\<Longrightarrow> \<exists>j. i<j \<and> (\<forall>x \<in> Lset(j). P(x) \<longleftrightarrow> Q(j,x))" |
76217 | 262 |
unfolding L_Reflects_def Closed_Unbounded_def |
13429 | 263 |
apply (elim meta_exE, clarify) |
264 |
apply (blast dest!: UnboundedD) |
|
13291 | 265 |
done |
266 |
||
267 |
lemma ReflectsE: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
268 |
"\<lbrakk>REFLECTS[P,Q]; Ord(i); |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
269 |
\<And>j. \<lbrakk>i<j; \<forall>x \<in> Lset(j). P(x) \<longleftrightarrow> Q(j,x)\<rbrakk> \<Longrightarrow> R\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
270 |
\<Longrightarrow> R" |
15764 | 271 |
by (drule ReflectsD, assumption, blast) |
13291 | 272 |
|
13428 | 273 |
lemma Collect_mem_eq: "{x\<in>A. x\<in>B} = A \<inter> B" |
13291 | 274 |
by blast |
275 |
||
276 |
||
60770 | 277 |
subsection\<open>Internalized Formulas for some Set-Theoretic Concepts\<close> |
13298 | 278 |
|
60770 | 279 |
subsubsection\<open>Some numbers to help write de Bruijn indices\<close> |
13306 | 280 |
|
21233 | 281 |
abbreviation |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
282 |
digit3 :: i (\<open>3\<close>) where "3 \<equiv> succ(2)" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
283 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
284 |
abbreviation |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
285 |
digit4 :: i (\<open>4\<close>) where "4 \<equiv> succ(3)" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
286 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
287 |
abbreviation |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
288 |
digit5 :: i (\<open>5\<close>) where "5 \<equiv> succ(4)" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
289 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
290 |
abbreviation |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
291 |
digit6 :: i (\<open>6\<close>) where "6 \<equiv> succ(5)" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
292 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
293 |
abbreviation |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
294 |
digit7 :: i (\<open>7\<close>) where "7 \<equiv> succ(6)" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
295 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
296 |
abbreviation |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
297 |
digit8 :: i (\<open>8\<close>) where "8 \<equiv> succ(7)" |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
298 |
|
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
299 |
abbreviation |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
300 |
digit9 :: i (\<open>9\<close>) where "9 \<equiv> succ(8)" |
13306 | 301 |
|
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
302 |
|
60770 | 303 |
subsubsection\<open>The Empty Set, Internalized\<close> |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
304 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
305 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
306 |
empty_fm :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
307 |
"empty_fm(x) \<equiv> Forall(Neg(Member(0,succ(x))))" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
308 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
309 |
lemma empty_type [TC]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
310 |
"x \<in> nat \<Longrightarrow> empty_fm(x) \<in> formula" |
13429 | 311 |
by (simp add: empty_fm_def) |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
312 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
313 |
lemma sats_empty_fm [simp]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
314 |
"\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
315 |
\<Longrightarrow> sats(A, empty_fm(x), env) \<longleftrightarrow> empty(##A, nth(x,env))" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
316 |
by (simp add: empty_fm_def empty_def) |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
317 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
318 |
lemma empty_iff_sats: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
319 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
320 |
i \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
321 |
\<Longrightarrow> empty(##A, x) \<longleftrightarrow> sats(A, empty_fm(i), env)" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
322 |
by simp |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
323 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
324 |
theorem empty_reflection: |
13429 | 325 |
"REFLECTS[\<lambda>x. empty(L,f(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
326 |
\<lambda>i x. empty(##Lset(i),f(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
327 |
apply (simp only: empty_def) |
13429 | 328 |
apply (intro FOL_reflections) |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
329 |
done |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
330 |
|
60770 | 331 |
text\<open>Not used. But maybe useful?\<close> |
13385
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
paulson
parents:
13363
diff
changeset
|
332 |
lemma Transset_sats_empty_fm_eq_0: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
333 |
"\<lbrakk>n \<in> nat; env \<in> list(A); Transset(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
334 |
\<Longrightarrow> sats(A, empty_fm(n), env) \<longleftrightarrow> nth(n,env) = 0" |
13385
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
paulson
parents:
13363
diff
changeset
|
335 |
apply (simp add: empty_fm_def empty_def Transset_def, auto) |
13429 | 336 |
apply (case_tac "n < length(env)") |
337 |
apply (frule nth_type, assumption+, blast) |
|
338 |
apply (simp_all add: not_lt_iff_le nth_eq_0) |
|
13385
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
paulson
parents:
13363
diff
changeset
|
339 |
done |
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
paulson
parents:
13363
diff
changeset
|
340 |
|
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
341 |
|
60770 | 342 |
subsubsection\<open>Unordered Pairs, Internalized\<close> |
13298 | 343 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
344 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
345 |
upair_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
346 |
"upair_fm(x,y,z) \<equiv> |
13429 | 347 |
And(Member(x,z), |
13298 | 348 |
And(Member(y,z), |
13429 | 349 |
Forall(Implies(Member(0,succ(z)), |
13298 | 350 |
Or(Equal(0,succ(x)), Equal(0,succ(y)))))))" |
351 |
||
352 |
lemma upair_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
353 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> upair_fm(x,y,z) \<in> formula" |
13429 | 354 |
by (simp add: upair_fm_def) |
13298 | 355 |
|
356 |
lemma sats_upair_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
357 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
358 |
\<Longrightarrow> sats(A, upair_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
359 |
upair(##A, nth(x,env), nth(y,env), nth(z,env))" |
13298 | 360 |
by (simp add: upair_fm_def upair_def) |
361 |
||
362 |
lemma upair_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
363 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
364 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
365 |
\<Longrightarrow> upair(##A, x, y, z) \<longleftrightarrow> sats(A, upair_fm(i,j,k), env)" |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
366 |
by (simp) |
13298 | 367 |
|
60770 | 368 |
text\<open>Useful? At least it refers to "real" unordered pairs\<close> |
13298 | 369 |
lemma sats_upair_fm2 [simp]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
370 |
"\<lbrakk>x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A); Transset(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
371 |
\<Longrightarrow> sats(A, upair_fm(x,y,z), env) \<longleftrightarrow> |
13298 | 372 |
nth(z,env) = {nth(x,env), nth(y,env)}" |
13429 | 373 |
apply (frule lt_length_in_nat, assumption) |
374 |
apply (simp add: upair_fm_def Transset_def, auto) |
|
375 |
apply (blast intro: nth_type) |
|
13298 | 376 |
done |
377 |
||
13314 | 378 |
theorem upair_reflection: |
13429 | 379 |
"REFLECTS[\<lambda>x. upair(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
380 |
\<lambda>i x. upair(##Lset(i),f(x),g(x),h(x))]" |
13314 | 381 |
apply (simp add: upair_def) |
13429 | 382 |
apply (intro FOL_reflections) |
13314 | 383 |
done |
13306 | 384 |
|
60770 | 385 |
subsubsection\<open>Ordered pairs, Internalized\<close> |
13298 | 386 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
387 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
388 |
pair_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
389 |
"pair_fm(x,y,z) \<equiv> |
13298 | 390 |
Exists(And(upair_fm(succ(x),succ(x),0), |
391 |
Exists(And(upair_fm(succ(succ(x)),succ(succ(y)),0), |
|
392 |
upair_fm(1,0,succ(succ(z)))))))" |
|
393 |
||
394 |
lemma pair_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
395 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> pair_fm(x,y,z) \<in> formula" |
13429 | 396 |
by (simp add: pair_fm_def) |
13298 | 397 |
|
398 |
lemma sats_pair_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
399 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
400 |
\<Longrightarrow> sats(A, pair_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
401 |
pair(##A, nth(x,env), nth(y,env), nth(z,env))" |
13298 | 402 |
by (simp add: pair_fm_def pair_def) |
403 |
||
404 |
lemma pair_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
405 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
406 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
407 |
\<Longrightarrow> pair(##A, x, y, z) \<longleftrightarrow> sats(A, pair_fm(i,j,k), env)" |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
408 |
by (simp) |
13298 | 409 |
|
13314 | 410 |
theorem pair_reflection: |
13429 | 411 |
"REFLECTS[\<lambda>x. pair(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
412 |
\<lambda>i x. pair(##Lset(i),f(x),g(x),h(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
413 |
apply (simp only: pair_def) |
13429 | 414 |
apply (intro FOL_reflections upair_reflection) |
13314 | 415 |
done |
13306 | 416 |
|
417 |
||
60770 | 418 |
subsubsection\<open>Binary Unions, Internalized\<close> |
13298 | 419 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
420 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
421 |
union_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
422 |
"union_fm(x,y,z) \<equiv> |
13306 | 423 |
Forall(Iff(Member(0,succ(z)), |
424 |
Or(Member(0,succ(x)),Member(0,succ(y)))))" |
|
425 |
||
426 |
lemma union_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
427 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> union_fm(x,y,z) \<in> formula" |
13429 | 428 |
by (simp add: union_fm_def) |
13306 | 429 |
|
430 |
lemma sats_union_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
431 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
432 |
\<Longrightarrow> sats(A, union_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
433 |
union(##A, nth(x,env), nth(y,env), nth(z,env))" |
13306 | 434 |
by (simp add: union_fm_def union_def) |
435 |
||
436 |
lemma union_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
437 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
438 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
439 |
\<Longrightarrow> union(##A, x, y, z) \<longleftrightarrow> sats(A, union_fm(i,j,k), env)" |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
440 |
by (simp) |
13298 | 441 |
|
13314 | 442 |
theorem union_reflection: |
13429 | 443 |
"REFLECTS[\<lambda>x. union(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
444 |
\<lambda>i x. union(##Lset(i),f(x),g(x),h(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
445 |
apply (simp only: union_def) |
13429 | 446 |
apply (intro FOL_reflections) |
13314 | 447 |
done |
13306 | 448 |
|
13298 | 449 |
|
60770 | 450 |
subsubsection\<open>Set ``Cons,'' Internalized\<close> |
13306 | 451 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
452 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
453 |
cons_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
454 |
"cons_fm(x,y,z) \<equiv> |
13306 | 455 |
Exists(And(upair_fm(succ(x),succ(x),0), |
456 |
union_fm(0,succ(y),succ(z))))" |
|
13298 | 457 |
|
458 |
||
13306 | 459 |
lemma cons_type [TC]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
460 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> cons_fm(x,y,z) \<in> formula" |
13429 | 461 |
by (simp add: cons_fm_def) |
13306 | 462 |
|
463 |
lemma sats_cons_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
464 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
465 |
\<Longrightarrow> sats(A, cons_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
466 |
is_cons(##A, nth(x,env), nth(y,env), nth(z,env))" |
13306 | 467 |
by (simp add: cons_fm_def is_cons_def) |
468 |
||
469 |
lemma cons_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
470 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
471 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
472 |
\<Longrightarrow> is_cons(##A, x, y, z) \<longleftrightarrow> sats(A, cons_fm(i,j,k), env)" |
13306 | 473 |
by simp |
474 |
||
13314 | 475 |
theorem cons_reflection: |
13429 | 476 |
"REFLECTS[\<lambda>x. is_cons(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
477 |
\<lambda>i x. is_cons(##Lset(i),f(x),g(x),h(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
478 |
apply (simp only: is_cons_def) |
13429 | 479 |
apply (intro FOL_reflections upair_reflection union_reflection) |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
480 |
done |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
481 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
482 |
|
60770 | 483 |
subsubsection\<open>Successor Function, Internalized\<close> |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
484 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
485 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
486 |
succ_fm :: "[i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
487 |
"succ_fm(x,y) \<equiv> cons_fm(x,x,y)" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
488 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
489 |
lemma succ_type [TC]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
490 |
"\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> succ_fm(x,y) \<in> formula" |
13429 | 491 |
by (simp add: succ_fm_def) |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
492 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
493 |
lemma sats_succ_fm [simp]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
494 |
"\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
495 |
\<Longrightarrow> sats(A, succ_fm(x,y), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
496 |
successor(##A, nth(x,env), nth(y,env))" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
497 |
by (simp add: succ_fm_def successor_def) |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
498 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
499 |
lemma successor_iff_sats: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
500 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
501 |
i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
502 |
\<Longrightarrow> successor(##A, x, y) \<longleftrightarrow> sats(A, succ_fm(i,j), env)" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
503 |
by simp |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
504 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
505 |
theorem successor_reflection: |
13429 | 506 |
"REFLECTS[\<lambda>x. successor(L,f(x),g(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
507 |
\<lambda>i x. successor(##Lset(i),f(x),g(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
508 |
apply (simp only: successor_def) |
13429 | 509 |
apply (intro cons_reflection) |
13314 | 510 |
done |
13298 | 511 |
|
512 |
||
60770 | 513 |
subsubsection\<open>The Number 1, Internalized\<close> |
13363 | 514 |
|
76214 | 515 |
(* "number1(M,a) \<equiv> (\<exists>x[M]. empty(M,x) \<and> successor(M,x,a))" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
516 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
517 |
number1_fm :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
518 |
"number1_fm(a) \<equiv> Exists(And(empty_fm(0), succ_fm(0,succ(a))))" |
13363 | 519 |
|
520 |
lemma number1_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
521 |
"x \<in> nat \<Longrightarrow> number1_fm(x) \<in> formula" |
13429 | 522 |
by (simp add: number1_fm_def) |
13363 | 523 |
|
524 |
lemma sats_number1_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
525 |
"\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
526 |
\<Longrightarrow> sats(A, number1_fm(x), env) \<longleftrightarrow> number1(##A, nth(x,env))" |
13363 | 527 |
by (simp add: number1_fm_def number1_def) |
528 |
||
529 |
lemma number1_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
530 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
531 |
i \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
532 |
\<Longrightarrow> number1(##A, x) \<longleftrightarrow> sats(A, number1_fm(i), env)" |
13363 | 533 |
by simp |
534 |
||
535 |
theorem number1_reflection: |
|
13429 | 536 |
"REFLECTS[\<lambda>x. number1(L,f(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
537 |
\<lambda>i x. number1(##Lset(i),f(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
538 |
apply (simp only: number1_def) |
13363 | 539 |
apply (intro FOL_reflections empty_reflection successor_reflection) |
540 |
done |
|
541 |
||
542 |
||
60770 | 543 |
subsubsection\<open>Big Union, Internalized\<close> |
13306 | 544 |
|
76214 | 545 |
(* "big_union(M,A,z) \<equiv> \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A \<and> x \<in> y)" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
546 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
547 |
big_union_fm :: "[i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
548 |
"big_union_fm(A,z) \<equiv> |
13352 | 549 |
Forall(Iff(Member(0,succ(z)), |
550 |
Exists(And(Member(0,succ(succ(A))), Member(1,0)))))" |
|
13298 | 551 |
|
13352 | 552 |
lemma big_union_type [TC]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
553 |
"\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> big_union_fm(x,y) \<in> formula" |
13429 | 554 |
by (simp add: big_union_fm_def) |
13306 | 555 |
|
13352 | 556 |
lemma sats_big_union_fm [simp]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
557 |
"\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
558 |
\<Longrightarrow> sats(A, big_union_fm(x,y), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
559 |
big_union(##A, nth(x,env), nth(y,env))" |
13352 | 560 |
by (simp add: big_union_fm_def big_union_def) |
13306 | 561 |
|
13352 | 562 |
lemma big_union_iff_sats: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
563 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
564 |
i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
565 |
\<Longrightarrow> big_union(##A, x, y) \<longleftrightarrow> sats(A, big_union_fm(i,j), env)" |
13306 | 566 |
by simp |
567 |
||
13352 | 568 |
theorem big_union_reflection: |
13429 | 569 |
"REFLECTS[\<lambda>x. big_union(L,f(x),g(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
570 |
\<lambda>i x. big_union(##Lset(i),f(x),g(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
571 |
apply (simp only: big_union_def) |
13429 | 572 |
apply (intro FOL_reflections) |
13314 | 573 |
done |
13298 | 574 |
|
575 |
||
60770 | 576 |
subsubsection\<open>Variants of Satisfaction Definitions for Ordinals, etc.\<close> |
13306 | 577 |
|
61798 | 578 |
text\<open>The \<open>sats\<close> theorems below are standard versions of the ones proved |
69593 | 579 |
in theory \<open>Formula\<close>. They relate elements of type \<^term>\<open>formula\<close> to |
580 |
relativized concepts such as \<^term>\<open>subset\<close> or \<^term>\<open>ordinal\<close> rather than to |
|
581 |
real concepts such as \<^term>\<open>Ord\<close>. Now that we have instantiated the locale |
|
61798 | 582 |
\<open>M_trivial\<close>, we no longer require the earlier versions.\<close> |
13306 | 583 |
|
584 |
lemma sats_subset_fm': |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
585 |
"\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
586 |
\<Longrightarrow> sats(A, subset_fm(x,y), env) \<longleftrightarrow> subset(##A, nth(x,env), nth(y,env))" |
13429 | 587 |
by (simp add: subset_fm_def Relative.subset_def) |
13298 | 588 |
|
13314 | 589 |
theorem subset_reflection: |
13429 | 590 |
"REFLECTS[\<lambda>x. subset(L,f(x),g(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
591 |
\<lambda>i x. subset(##Lset(i),f(x),g(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
592 |
apply (simp only: Relative.subset_def) |
13429 | 593 |
apply (intro FOL_reflections) |
13314 | 594 |
done |
13306 | 595 |
|
596 |
lemma sats_transset_fm': |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
597 |
"\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
598 |
\<Longrightarrow> sats(A, transset_fm(x), env) \<longleftrightarrow> transitive_set(##A, nth(x,env))" |
13429 | 599 |
by (simp add: sats_subset_fm' transset_fm_def transitive_set_def) |
13298 | 600 |
|
13314 | 601 |
theorem transitive_set_reflection: |
602 |
"REFLECTS[\<lambda>x. transitive_set(L,f(x)), |
|
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
603 |
\<lambda>i x. transitive_set(##Lset(i),f(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
604 |
apply (simp only: transitive_set_def) |
13429 | 605 |
apply (intro FOL_reflections subset_reflection) |
13314 | 606 |
done |
13306 | 607 |
|
608 |
lemma sats_ordinal_fm': |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
609 |
"\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
610 |
\<Longrightarrow> sats(A, ordinal_fm(x), env) \<longleftrightarrow> ordinal(##A,nth(x,env))" |
13306 | 611 |
by (simp add: sats_transset_fm' ordinal_fm_def ordinal_def) |
612 |
||
613 |
lemma ordinal_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
614 |
"\<lbrakk>nth(i,env) = x; i \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
615 |
\<Longrightarrow> ordinal(##A, x) \<longleftrightarrow> sats(A, ordinal_fm(i), env)" |
13306 | 616 |
by (simp add: sats_ordinal_fm') |
617 |
||
13314 | 618 |
theorem ordinal_reflection: |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
619 |
"REFLECTS[\<lambda>x. ordinal(L,f(x)), \<lambda>i x. ordinal(##Lset(i),f(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
620 |
apply (simp only: ordinal_def) |
13429 | 621 |
apply (intro FOL_reflections transitive_set_reflection) |
13314 | 622 |
done |
13298 | 623 |
|
624 |
||
60770 | 625 |
subsubsection\<open>Membership Relation, Internalized\<close> |
13298 | 626 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
627 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
628 |
Memrel_fm :: "[i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
629 |
"Memrel_fm(A,r) \<equiv> |
13306 | 630 |
Forall(Iff(Member(0,succ(r)), |
631 |
Exists(And(Member(0,succ(succ(A))), |
|
632 |
Exists(And(Member(0,succ(succ(succ(A)))), |
|
633 |
And(Member(1,0), |
|
634 |
pair_fm(1,0,2))))))))" |
|
635 |
||
636 |
lemma Memrel_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
637 |
"\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> Memrel_fm(x,y) \<in> formula" |
13429 | 638 |
by (simp add: Memrel_fm_def) |
13298 | 639 |
|
13306 | 640 |
lemma sats_Memrel_fm [simp]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
641 |
"\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
642 |
\<Longrightarrow> sats(A, Memrel_fm(x,y), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
643 |
membership(##A, nth(x,env), nth(y,env))" |
13306 | 644 |
by (simp add: Memrel_fm_def membership_def) |
13298 | 645 |
|
13306 | 646 |
lemma Memrel_iff_sats: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
647 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
648 |
i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
649 |
\<Longrightarrow> membership(##A, x, y) \<longleftrightarrow> sats(A, Memrel_fm(i,j), env)" |
13306 | 650 |
by simp |
13304 | 651 |
|
13314 | 652 |
theorem membership_reflection: |
13429 | 653 |
"REFLECTS[\<lambda>x. membership(L,f(x),g(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
654 |
\<lambda>i x. membership(##Lset(i),f(x),g(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
655 |
apply (simp only: membership_def) |
13429 | 656 |
apply (intro FOL_reflections pair_reflection) |
13314 | 657 |
done |
13304 | 658 |
|
60770 | 659 |
subsubsection\<open>Predecessor Set, Internalized\<close> |
13304 | 660 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
661 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
662 |
pred_set_fm :: "[i,i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
663 |
"pred_set_fm(A,x,r,B) \<equiv> |
13306 | 664 |
Forall(Iff(Member(0,succ(B)), |
665 |
Exists(And(Member(0,succ(succ(r))), |
|
666 |
And(Member(1,succ(succ(A))), |
|
667 |
pair_fm(1,succ(succ(x)),0))))))" |
|
668 |
||
669 |
||
670 |
lemma pred_set_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
671 |
"\<lbrakk>A \<in> nat; x \<in> nat; r \<in> nat; B \<in> nat\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
672 |
\<Longrightarrow> pred_set_fm(A,x,r,B) \<in> formula" |
13429 | 673 |
by (simp add: pred_set_fm_def) |
13304 | 674 |
|
13306 | 675 |
lemma sats_pred_set_fm [simp]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
676 |
"\<lbrakk>U \<in> nat; x \<in> nat; r \<in> nat; B \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
677 |
\<Longrightarrow> sats(A, pred_set_fm(U,x,r,B), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
678 |
pred_set(##A, nth(U,env), nth(x,env), nth(r,env), nth(B,env))" |
13306 | 679 |
by (simp add: pred_set_fm_def pred_set_def) |
680 |
||
681 |
lemma pred_set_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
682 |
"\<lbrakk>nth(i,env) = U; nth(j,env) = x; nth(k,env) = r; nth(l,env) = B; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
683 |
i \<in> nat; j \<in> nat; k \<in> nat; l \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
684 |
\<Longrightarrow> pred_set(##A,U,x,r,B) \<longleftrightarrow> sats(A, pred_set_fm(i,j,k,l), env)" |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
685 |
by (simp) |
13306 | 686 |
|
13314 | 687 |
theorem pred_set_reflection: |
13429 | 688 |
"REFLECTS[\<lambda>x. pred_set(L,f(x),g(x),h(x),b(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
689 |
\<lambda>i x. pred_set(##Lset(i),f(x),g(x),h(x),b(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
690 |
apply (simp only: pred_set_def) |
13429 | 691 |
apply (intro FOL_reflections pair_reflection) |
13314 | 692 |
done |
13304 | 693 |
|
694 |
||
13298 | 695 |
|
60770 | 696 |
subsubsection\<open>Domain of a Relation, Internalized\<close> |
13306 | 697 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
698 |
(* "is_domain(M,r,z) \<equiv> |
76214 | 699 |
\<forall>x[M]. (x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r \<and> (\<exists>y[M]. pair(M,x,y,w))))" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
700 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
701 |
domain_fm :: "[i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
702 |
"domain_fm(r,z) \<equiv> |
13306 | 703 |
Forall(Iff(Member(0,succ(z)), |
704 |
Exists(And(Member(0,succ(succ(r))), |
|
705 |
Exists(pair_fm(2,0,1))))))" |
|
706 |
||
707 |
lemma domain_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
708 |
"\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> domain_fm(x,y) \<in> formula" |
13429 | 709 |
by (simp add: domain_fm_def) |
13306 | 710 |
|
711 |
lemma sats_domain_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
712 |
"\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
713 |
\<Longrightarrow> sats(A, domain_fm(x,y), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
714 |
is_domain(##A, nth(x,env), nth(y,env))" |
13306 | 715 |
by (simp add: domain_fm_def is_domain_def) |
716 |
||
717 |
lemma domain_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
718 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
719 |
i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
720 |
\<Longrightarrow> is_domain(##A, x, y) \<longleftrightarrow> sats(A, domain_fm(i,j), env)" |
13306 | 721 |
by simp |
722 |
||
13314 | 723 |
theorem domain_reflection: |
13429 | 724 |
"REFLECTS[\<lambda>x. is_domain(L,f(x),g(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
725 |
\<lambda>i x. is_domain(##Lset(i),f(x),g(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
726 |
apply (simp only: is_domain_def) |
13429 | 727 |
apply (intro FOL_reflections pair_reflection) |
13314 | 728 |
done |
13306 | 729 |
|
730 |
||
60770 | 731 |
subsubsection\<open>Range of a Relation, Internalized\<close> |
13306 | 732 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
733 |
(* "is_range(M,r,z) \<equiv> |
76214 | 734 |
\<forall>y[M]. (y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r \<and> (\<exists>x[M]. pair(M,x,y,w))))" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
735 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
736 |
range_fm :: "[i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
737 |
"range_fm(r,z) \<equiv> |
13306 | 738 |
Forall(Iff(Member(0,succ(z)), |
739 |
Exists(And(Member(0,succ(succ(r))), |
|
740 |
Exists(pair_fm(0,2,1))))))" |
|
741 |
||
742 |
lemma range_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
743 |
"\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> range_fm(x,y) \<in> formula" |
13429 | 744 |
by (simp add: range_fm_def) |
13306 | 745 |
|
746 |
lemma sats_range_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
747 |
"\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
748 |
\<Longrightarrow> sats(A, range_fm(x,y), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
749 |
is_range(##A, nth(x,env), nth(y,env))" |
13306 | 750 |
by (simp add: range_fm_def is_range_def) |
751 |
||
752 |
lemma range_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
753 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
754 |
i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
755 |
\<Longrightarrow> is_range(##A, x, y) \<longleftrightarrow> sats(A, range_fm(i,j), env)" |
13306 | 756 |
by simp |
757 |
||
13314 | 758 |
theorem range_reflection: |
13429 | 759 |
"REFLECTS[\<lambda>x. is_range(L,f(x),g(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
760 |
\<lambda>i x. is_range(##Lset(i),f(x),g(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
761 |
apply (simp only: is_range_def) |
13429 | 762 |
apply (intro FOL_reflections pair_reflection) |
13314 | 763 |
done |
13306 | 764 |
|
13429 | 765 |
|
60770 | 766 |
subsubsection\<open>Field of a Relation, Internalized\<close> |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
767 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
768 |
(* "is_field(M,r,z) \<equiv> |
76214 | 769 |
\<exists>dr[M]. is_domain(M,r,dr) \<and> |
770 |
(\<exists>rr[M]. is_range(M,r,rr) \<and> union(M,dr,rr,z))" *) |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
771 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
772 |
field_fm :: "[i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
773 |
"field_fm(r,z) \<equiv> |
13429 | 774 |
Exists(And(domain_fm(succ(r),0), |
775 |
Exists(And(range_fm(succ(succ(r)),0), |
|
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
776 |
union_fm(1,0,succ(succ(z)))))))" |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
777 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
778 |
lemma field_type [TC]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
779 |
"\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> field_fm(x,y) \<in> formula" |
13429 | 780 |
by (simp add: field_fm_def) |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
781 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
782 |
lemma sats_field_fm [simp]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
783 |
"\<lbrakk>x \<in> nat; y \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
784 |
\<Longrightarrow> sats(A, field_fm(x,y), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
785 |
is_field(##A, nth(x,env), nth(y,env))" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
786 |
by (simp add: field_fm_def is_field_def) |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
787 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
788 |
lemma field_iff_sats: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
789 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
790 |
i \<in> nat; j \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
791 |
\<Longrightarrow> is_field(##A, x, y) \<longleftrightarrow> sats(A, field_fm(i,j), env)" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
792 |
by simp |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
793 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
794 |
theorem field_reflection: |
13429 | 795 |
"REFLECTS[\<lambda>x. is_field(L,f(x),g(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
796 |
\<lambda>i x. is_field(##Lset(i),f(x),g(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
797 |
apply (simp only: is_field_def) |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
798 |
apply (intro FOL_reflections domain_reflection range_reflection |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
799 |
union_reflection) |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
800 |
done |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
801 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
802 |
|
60770 | 803 |
subsubsection\<open>Image under a Relation, Internalized\<close> |
13306 | 804 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
805 |
(* "image(M,r,A,z) \<equiv> |
76214 | 806 |
\<forall>y[M]. (y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r \<and> (\<exists>x[M]. x\<in>A \<and> pair(M,x,y,w))))" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
807 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
808 |
image_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
809 |
"image_fm(r,A,z) \<equiv> |
13306 | 810 |
Forall(Iff(Member(0,succ(z)), |
811 |
Exists(And(Member(0,succ(succ(r))), |
|
812 |
Exists(And(Member(0,succ(succ(succ(A)))), |
|
13429 | 813 |
pair_fm(0,2,1)))))))" |
13306 | 814 |
|
815 |
lemma image_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
816 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> image_fm(x,y,z) \<in> formula" |
13429 | 817 |
by (simp add: image_fm_def) |
13306 | 818 |
|
819 |
lemma sats_image_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
820 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
821 |
\<Longrightarrow> sats(A, image_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
822 |
image(##A, nth(x,env), nth(y,env), nth(z,env))" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
823 |
by (simp add: image_fm_def Relative.image_def) |
13306 | 824 |
|
825 |
lemma image_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
826 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
827 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
828 |
\<Longrightarrow> image(##A, x, y, z) \<longleftrightarrow> sats(A, image_fm(i,j,k), env)" |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
829 |
by (simp) |
13306 | 830 |
|
13314 | 831 |
theorem image_reflection: |
13429 | 832 |
"REFLECTS[\<lambda>x. image(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
833 |
\<lambda>i x. image(##Lset(i),f(x),g(x),h(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
834 |
apply (simp only: Relative.image_def) |
13429 | 835 |
apply (intro FOL_reflections pair_reflection) |
13314 | 836 |
done |
13306 | 837 |
|
838 |
||
60770 | 839 |
subsubsection\<open>Pre-Image under a Relation, Internalized\<close> |
13348 | 840 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
841 |
(* "pre_image(M,r,A,z) \<equiv> |
76214 | 842 |
\<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r \<and> (\<exists>y[M]. y\<in>A \<and> pair(M,x,y,w)))" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
843 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
844 |
pre_image_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
845 |
"pre_image_fm(r,A,z) \<equiv> |
13348 | 846 |
Forall(Iff(Member(0,succ(z)), |
847 |
Exists(And(Member(0,succ(succ(r))), |
|
848 |
Exists(And(Member(0,succ(succ(succ(A)))), |
|
13429 | 849 |
pair_fm(2,0,1)))))))" |
13348 | 850 |
|
851 |
lemma pre_image_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
852 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> pre_image_fm(x,y,z) \<in> formula" |
13429 | 853 |
by (simp add: pre_image_fm_def) |
13348 | 854 |
|
855 |
lemma sats_pre_image_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
856 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
857 |
\<Longrightarrow> sats(A, pre_image_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
858 |
pre_image(##A, nth(x,env), nth(y,env), nth(z,env))" |
13348 | 859 |
by (simp add: pre_image_fm_def Relative.pre_image_def) |
860 |
||
861 |
lemma pre_image_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
862 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
863 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
864 |
\<Longrightarrow> pre_image(##A, x, y, z) \<longleftrightarrow> sats(A, pre_image_fm(i,j,k), env)" |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
865 |
by (simp) |
13348 | 866 |
|
867 |
theorem pre_image_reflection: |
|
13429 | 868 |
"REFLECTS[\<lambda>x. pre_image(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
869 |
\<lambda>i x. pre_image(##Lset(i),f(x),g(x),h(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
870 |
apply (simp only: Relative.pre_image_def) |
13429 | 871 |
apply (intro FOL_reflections pair_reflection) |
13348 | 872 |
done |
873 |
||
874 |
||
60770 | 875 |
subsubsection\<open>Function Application, Internalized\<close> |
13352 | 876 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
877 |
(* "fun_apply(M,f,x,y) \<equiv> |
13429 | 878 |
(\<exists>xs[M]. \<exists>fxs[M]. |
76214 | 879 |
upair(M,x,x,xs) \<and> image(M,f,xs,fxs) \<and> big_union(M,fxs,y))" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
880 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
881 |
fun_apply_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
882 |
"fun_apply_fm(f,x,y) \<equiv> |
13352 | 883 |
Exists(Exists(And(upair_fm(succ(succ(x)), succ(succ(x)), 1), |
13429 | 884 |
And(image_fm(succ(succ(f)), 1, 0), |
13352 | 885 |
big_union_fm(0,succ(succ(y)))))))" |
886 |
||
887 |
lemma fun_apply_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
888 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> fun_apply_fm(x,y,z) \<in> formula" |
13429 | 889 |
by (simp add: fun_apply_fm_def) |
13352 | 890 |
|
891 |
lemma sats_fun_apply_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
892 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
893 |
\<Longrightarrow> sats(A, fun_apply_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
894 |
fun_apply(##A, nth(x,env), nth(y,env), nth(z,env))" |
13352 | 895 |
by (simp add: fun_apply_fm_def fun_apply_def) |
896 |
||
897 |
lemma fun_apply_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
898 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
899 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
900 |
\<Longrightarrow> fun_apply(##A, x, y, z) \<longleftrightarrow> sats(A, fun_apply_fm(i,j,k), env)" |
13352 | 901 |
by simp |
902 |
||
903 |
theorem fun_apply_reflection: |
|
13429 | 904 |
"REFLECTS[\<lambda>x. fun_apply(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
905 |
\<lambda>i x. fun_apply(##Lset(i),f(x),g(x),h(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
906 |
apply (simp only: fun_apply_def) |
13352 | 907 |
apply (intro FOL_reflections upair_reflection image_reflection |
13429 | 908 |
big_union_reflection) |
13352 | 909 |
done |
910 |
||
911 |
||
60770 | 912 |
subsubsection\<open>The Concept of Relation, Internalized\<close> |
13306 | 913 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
914 |
(* "is_relation(M,r) \<equiv> |
46823 | 915 |
(\<forall>z[M]. z\<in>r \<longrightarrow> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
916 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
917 |
relation_fm :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
918 |
"relation_fm(r) \<equiv> |
13306 | 919 |
Forall(Implies(Member(0,succ(r)), Exists(Exists(pair_fm(1,0,2)))))" |
920 |
||
921 |
lemma relation_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
922 |
"\<lbrakk>x \<in> nat\<rbrakk> \<Longrightarrow> relation_fm(x) \<in> formula" |
13429 | 923 |
by (simp add: relation_fm_def) |
13306 | 924 |
|
925 |
lemma sats_relation_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
926 |
"\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
927 |
\<Longrightarrow> sats(A, relation_fm(x), env) \<longleftrightarrow> is_relation(##A, nth(x,env))" |
13306 | 928 |
by (simp add: relation_fm_def is_relation_def) |
929 |
||
930 |
lemma relation_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
931 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
932 |
i \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
933 |
\<Longrightarrow> is_relation(##A, x) \<longleftrightarrow> sats(A, relation_fm(i), env)" |
13306 | 934 |
by simp |
935 |
||
13314 | 936 |
theorem is_relation_reflection: |
13429 | 937 |
"REFLECTS[\<lambda>x. is_relation(L,f(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
938 |
\<lambda>i x. is_relation(##Lset(i),f(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
939 |
apply (simp only: is_relation_def) |
13429 | 940 |
apply (intro FOL_reflections pair_reflection) |
13314 | 941 |
done |
13306 | 942 |
|
943 |
||
60770 | 944 |
subsubsection\<open>The Concept of Function, Internalized\<close> |
13306 | 945 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
946 |
(* "is_function(M,r) \<equiv> |
13429 | 947 |
\<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M]. |
46823 | 948 |
pair(M,x,y,p) \<longrightarrow> pair(M,x,y',p') \<longrightarrow> p\<in>r \<longrightarrow> p'\<in>r \<longrightarrow> y=y'" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
949 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
950 |
function_fm :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
951 |
"function_fm(r) \<equiv> |
13306 | 952 |
Forall(Forall(Forall(Forall(Forall( |
953 |
Implies(pair_fm(4,3,1), |
|
954 |
Implies(pair_fm(4,2,0), |
|
955 |
Implies(Member(1,r#+5), |
|
956 |
Implies(Member(0,r#+5), Equal(3,2))))))))))" |
|
957 |
||
958 |
lemma function_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
959 |
"\<lbrakk>x \<in> nat\<rbrakk> \<Longrightarrow> function_fm(x) \<in> formula" |
13429 | 960 |
by (simp add: function_fm_def) |
13306 | 961 |
|
962 |
lemma sats_function_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
963 |
"\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
964 |
\<Longrightarrow> sats(A, function_fm(x), env) \<longleftrightarrow> is_function(##A, nth(x,env))" |
13306 | 965 |
by (simp add: function_fm_def is_function_def) |
966 |
||
13505 | 967 |
lemma is_function_iff_sats: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
968 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
969 |
i \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
970 |
\<Longrightarrow> is_function(##A, x) \<longleftrightarrow> sats(A, function_fm(i), env)" |
13306 | 971 |
by simp |
972 |
||
13314 | 973 |
theorem is_function_reflection: |
13429 | 974 |
"REFLECTS[\<lambda>x. is_function(L,f(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
975 |
\<lambda>i x. is_function(##Lset(i),f(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
976 |
apply (simp only: is_function_def) |
13429 | 977 |
apply (intro FOL_reflections pair_reflection) |
13314 | 978 |
done |
13298 | 979 |
|
980 |
||
60770 | 981 |
subsubsection\<open>Typed Functions, Internalized\<close> |
13309 | 982 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
983 |
(* "typed_function(M,A,B,r) \<equiv> |
76214 | 984 |
is_function(M,r) \<and> is_relation(M,r) \<and> is_domain(M,r,A) \<and> |
46823 | 985 |
(\<forall>u[M]. u\<in>r \<longrightarrow> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) \<longrightarrow> y\<in>B))" *) |
13309 | 986 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
987 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
988 |
typed_function_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
989 |
"typed_function_fm(A,B,r) \<equiv> |
13309 | 990 |
And(function_fm(r), |
991 |
And(relation_fm(r), |
|
992 |
And(domain_fm(r,A), |
|
993 |
Forall(Implies(Member(0,succ(r)), |
|
994 |
Forall(Forall(Implies(pair_fm(1,0,2),Member(0,B#+3)))))))))" |
|
995 |
||
996 |
lemma typed_function_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
997 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> typed_function_fm(x,y,z) \<in> formula" |
13429 | 998 |
by (simp add: typed_function_fm_def) |
13309 | 999 |
|
1000 |
lemma sats_typed_function_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1001 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1002 |
\<Longrightarrow> sats(A, typed_function_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1003 |
typed_function(##A, nth(x,env), nth(y,env), nth(z,env))" |
13309 | 1004 |
by (simp add: typed_function_fm_def typed_function_def) |
1005 |
||
1006 |
lemma typed_function_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1007 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1008 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1009 |
\<Longrightarrow> typed_function(##A, x, y, z) \<longleftrightarrow> sats(A, typed_function_fm(i,j,k), env)" |
13309 | 1010 |
by simp |
1011 |
||
13429 | 1012 |
lemmas function_reflections = |
13363 | 1013 |
empty_reflection number1_reflection |
13429 | 1014 |
upair_reflection pair_reflection union_reflection |
1015 |
big_union_reflection cons_reflection successor_reflection |
|
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1016 |
fun_apply_reflection subset_reflection |
13429 | 1017 |
transitive_set_reflection membership_reflection |
1018 |
pred_set_reflection domain_reflection range_reflection field_reflection |
|
13348 | 1019 |
image_reflection pre_image_reflection |
13429 | 1020 |
is_relation_reflection is_function_reflection |
13309 | 1021 |
|
13429 | 1022 |
lemmas function_iff_sats = |
1023 |
empty_iff_sats number1_iff_sats |
|
1024 |
upair_iff_sats pair_iff_sats union_iff_sats |
|
13505 | 1025 |
big_union_iff_sats cons_iff_sats successor_iff_sats |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1026 |
fun_apply_iff_sats Memrel_iff_sats |
13429 | 1027 |
pred_set_iff_sats domain_iff_sats range_iff_sats field_iff_sats |
1028 |
image_iff_sats pre_image_iff_sats |
|
13505 | 1029 |
relation_iff_sats is_function_iff_sats |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1030 |
|
13309 | 1031 |
|
13314 | 1032 |
theorem typed_function_reflection: |
13429 | 1033 |
"REFLECTS[\<lambda>x. typed_function(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1034 |
\<lambda>i x. typed_function(##Lset(i),f(x),g(x),h(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
1035 |
apply (simp only: typed_function_def) |
13429 | 1036 |
apply (intro FOL_reflections function_reflections) |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1037 |
done |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1038 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1039 |
|
60770 | 1040 |
subsubsection\<open>Composition of Relations, Internalized\<close> |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1041 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1042 |
(* "composition(M,r,s,t) \<equiv> |
46823 | 1043 |
\<forall>p[M]. p \<in> t \<longleftrightarrow> |
13429 | 1044 |
(\<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M]. |
76214 | 1045 |
pair(M,x,z,p) \<and> pair(M,x,y,xy) \<and> pair(M,y,z,yz) \<and> |
1046 |
xy \<in> s \<and> yz \<in> r)" *) |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
1047 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
1048 |
composition_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1049 |
"composition_fm(r,s,t) \<equiv> |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1050 |
Forall(Iff(Member(0,succ(t)), |
13429 | 1051 |
Exists(Exists(Exists(Exists(Exists( |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1052 |
And(pair_fm(4,2,5), |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1053 |
And(pair_fm(4,3,1), |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1054 |
And(pair_fm(3,2,0), |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1055 |
And(Member(1,s#+6), Member(0,r#+6))))))))))))" |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1056 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1057 |
lemma composition_type [TC]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1058 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> composition_fm(x,y,z) \<in> formula" |
13429 | 1059 |
by (simp add: composition_fm_def) |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1060 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1061 |
lemma sats_composition_fm [simp]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1062 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1063 |
\<Longrightarrow> sats(A, composition_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1064 |
composition(##A, nth(x,env), nth(y,env), nth(z,env))" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1065 |
by (simp add: composition_fm_def composition_def) |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1066 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1067 |
lemma composition_iff_sats: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1068 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1069 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1070 |
\<Longrightarrow> composition(##A, x, y, z) \<longleftrightarrow> sats(A, composition_fm(i,j,k), env)" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1071 |
by simp |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1072 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1073 |
theorem composition_reflection: |
13429 | 1074 |
"REFLECTS[\<lambda>x. composition(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1075 |
\<lambda>i x. composition(##Lset(i),f(x),g(x),h(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
1076 |
apply (simp only: composition_def) |
13429 | 1077 |
apply (intro FOL_reflections pair_reflection) |
13314 | 1078 |
done |
1079 |
||
13309 | 1080 |
|
60770 | 1081 |
subsubsection\<open>Injections, Internalized\<close> |
13309 | 1082 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1083 |
(* "injection(M,A,B,f) \<equiv> |
76214 | 1084 |
typed_function(M,A,B,f) \<and> |
13429 | 1085 |
(\<forall>x[M]. \<forall>x'[M]. \<forall>y[M]. \<forall>p[M]. \<forall>p'[M]. |
46823 | 1086 |
pair(M,x,y,p) \<longrightarrow> pair(M,x',y,p') \<longrightarrow> p\<in>f \<longrightarrow> p'\<in>f \<longrightarrow> x=x')" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
1087 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
1088 |
injection_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1089 |
"injection_fm(A,B,f) \<equiv> |
13309 | 1090 |
And(typed_function_fm(A,B,f), |
1091 |
Forall(Forall(Forall(Forall(Forall( |
|
1092 |
Implies(pair_fm(4,2,1), |
|
1093 |
Implies(pair_fm(3,2,0), |
|
1094 |
Implies(Member(1,f#+5), |
|
1095 |
Implies(Member(0,f#+5), Equal(4,3)))))))))))" |
|
1096 |
||
1097 |
||
1098 |
lemma injection_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1099 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> injection_fm(x,y,z) \<in> formula" |
13429 | 1100 |
by (simp add: injection_fm_def) |
13309 | 1101 |
|
1102 |
lemma sats_injection_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1103 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1104 |
\<Longrightarrow> sats(A, injection_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1105 |
injection(##A, nth(x,env), nth(y,env), nth(z,env))" |
13309 | 1106 |
by (simp add: injection_fm_def injection_def) |
1107 |
||
1108 |
lemma injection_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1109 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1110 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1111 |
\<Longrightarrow> injection(##A, x, y, z) \<longleftrightarrow> sats(A, injection_fm(i,j,k), env)" |
13309 | 1112 |
by simp |
1113 |
||
13314 | 1114 |
theorem injection_reflection: |
13429 | 1115 |
"REFLECTS[\<lambda>x. injection(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1116 |
\<lambda>i x. injection(##Lset(i),f(x),g(x),h(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
1117 |
apply (simp only: injection_def) |
13429 | 1118 |
apply (intro FOL_reflections function_reflections typed_function_reflection) |
13314 | 1119 |
done |
13309 | 1120 |
|
1121 |
||
60770 | 1122 |
subsubsection\<open>Surjections, Internalized\<close> |
13309 | 1123 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
1124 |
(* surjection :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1125 |
"surjection(M,A,B,f) \<equiv> |
76214 | 1126 |
typed_function(M,A,B,f) \<and> |
1127 |
(\<forall>y[M]. y\<in>B \<longrightarrow> (\<exists>x[M]. x\<in>A \<and> fun_apply(M,f,x,y)))" *) |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
1128 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
1129 |
surjection_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1130 |
"surjection_fm(A,B,f) \<equiv> |
13309 | 1131 |
And(typed_function_fm(A,B,f), |
1132 |
Forall(Implies(Member(0,succ(B)), |
|
1133 |
Exists(And(Member(0,succ(succ(A))), |
|
1134 |
fun_apply_fm(succ(succ(f)),0,1))))))" |
|
1135 |
||
1136 |
lemma surjection_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1137 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> surjection_fm(x,y,z) \<in> formula" |
13429 | 1138 |
by (simp add: surjection_fm_def) |
13309 | 1139 |
|
1140 |
lemma sats_surjection_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1141 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1142 |
\<Longrightarrow> sats(A, surjection_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1143 |
surjection(##A, nth(x,env), nth(y,env), nth(z,env))" |
13309 | 1144 |
by (simp add: surjection_fm_def surjection_def) |
1145 |
||
1146 |
lemma surjection_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1147 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1148 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1149 |
\<Longrightarrow> surjection(##A, x, y, z) \<longleftrightarrow> sats(A, surjection_fm(i,j,k), env)" |
13309 | 1150 |
by simp |
1151 |
||
13314 | 1152 |
theorem surjection_reflection: |
13429 | 1153 |
"REFLECTS[\<lambda>x. surjection(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1154 |
\<lambda>i x. surjection(##Lset(i),f(x),g(x),h(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
1155 |
apply (simp only: surjection_def) |
13429 | 1156 |
apply (intro FOL_reflections function_reflections typed_function_reflection) |
13314 | 1157 |
done |
13309 | 1158 |
|
1159 |
||
1160 |
||
60770 | 1161 |
subsubsection\<open>Bijections, Internalized\<close> |
13309 | 1162 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
1163 |
(* bijection :: "[i\<Rightarrow>o,i,i,i] \<Rightarrow> o" |
76214 | 1164 |
"bijection(M,A,B,f) \<equiv> injection(M,A,B,f) \<and> surjection(M,A,B,f)" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
1165 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
1166 |
bijection_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1167 |
"bijection_fm(A,B,f) \<equiv> And(injection_fm(A,B,f), surjection_fm(A,B,f))" |
13309 | 1168 |
|
1169 |
lemma bijection_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1170 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> bijection_fm(x,y,z) \<in> formula" |
13429 | 1171 |
by (simp add: bijection_fm_def) |
13309 | 1172 |
|
1173 |
lemma sats_bijection_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1174 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1175 |
\<Longrightarrow> sats(A, bijection_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1176 |
bijection(##A, nth(x,env), nth(y,env), nth(z,env))" |
13309 | 1177 |
by (simp add: bijection_fm_def bijection_def) |
1178 |
||
1179 |
lemma bijection_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1180 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1181 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1182 |
\<Longrightarrow> bijection(##A, x, y, z) \<longleftrightarrow> sats(A, bijection_fm(i,j,k), env)" |
13309 | 1183 |
by simp |
1184 |
||
13314 | 1185 |
theorem bijection_reflection: |
13429 | 1186 |
"REFLECTS[\<lambda>x. bijection(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1187 |
\<lambda>i x. bijection(##Lset(i),f(x),g(x),h(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
1188 |
apply (simp only: bijection_def) |
13429 | 1189 |
apply (intro And_reflection injection_reflection surjection_reflection) |
13314 | 1190 |
done |
13309 | 1191 |
|
1192 |
||
60770 | 1193 |
subsubsection\<open>Restriction of a Relation, Internalized\<close> |
13348 | 1194 |
|
1195 |
||
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1196 |
(* "restriction(M,r,A,z) \<equiv> |
76214 | 1197 |
\<forall>x[M]. x \<in> z \<longleftrightarrow> (x \<in> r \<and> (\<exists>u[M]. u\<in>A \<and> (\<exists>v[M]. pair(M,u,v,x))))" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
1198 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
1199 |
restriction_fm :: "[i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1200 |
"restriction_fm(r,A,z) \<equiv> |
13348 | 1201 |
Forall(Iff(Member(0,succ(z)), |
1202 |
And(Member(0,succ(r)), |
|
1203 |
Exists(And(Member(0,succ(succ(A))), |
|
1204 |
Exists(pair_fm(1,0,2)))))))" |
|
1205 |
||
1206 |
lemma restriction_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1207 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> restriction_fm(x,y,z) \<in> formula" |
13429 | 1208 |
by (simp add: restriction_fm_def) |
13348 | 1209 |
|
1210 |
lemma sats_restriction_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1211 |
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1212 |
\<Longrightarrow> sats(A, restriction_fm(x,y,z), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1213 |
restriction(##A, nth(x,env), nth(y,env), nth(z,env))" |
13348 | 1214 |
by (simp add: restriction_fm_def restriction_def) |
1215 |
||
1216 |
lemma restriction_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1217 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1218 |
i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1219 |
\<Longrightarrow> restriction(##A, x, y, z) \<longleftrightarrow> sats(A, restriction_fm(i,j,k), env)" |
13348 | 1220 |
by simp |
1221 |
||
1222 |
theorem restriction_reflection: |
|
13429 | 1223 |
"REFLECTS[\<lambda>x. restriction(L,f(x),g(x),h(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1224 |
\<lambda>i x. restriction(##Lset(i),f(x),g(x),h(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
1225 |
apply (simp only: restriction_def) |
13429 | 1226 |
apply (intro FOL_reflections pair_reflection) |
13348 | 1227 |
done |
1228 |
||
60770 | 1229 |
subsubsection\<open>Order-Isomorphisms, Internalized\<close> |
13309 | 1230 |
|
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
1231 |
(* order_isomorphism :: "[i\<Rightarrow>o,i,i,i,i,i] \<Rightarrow> o" |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1232 |
"order_isomorphism(M,A,r,B,s,f) \<equiv> |
76214 | 1233 |
bijection(M,A,B,f) \<and> |
46823 | 1234 |
(\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow> |
13309 | 1235 |
(\<forall>p[M]. \<forall>fx[M]. \<forall>fy[M]. \<forall>q[M]. |
46823 | 1236 |
pair(M,x,y,p) \<longrightarrow> fun_apply(M,f,x,fx) \<longrightarrow> fun_apply(M,f,y,fy) \<longrightarrow> |
1237 |
pair(M,fx,fy,q) \<longrightarrow> (p\<in>r \<longleftrightarrow> q\<in>s))))" |
|
13309 | 1238 |
*) |
1239 |
||
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
1240 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
1241 |
order_isomorphism_fm :: "[i,i,i,i,i]\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1242 |
"order_isomorphism_fm(A,r,B,s,f) \<equiv> |
13429 | 1243 |
And(bijection_fm(A,B,f), |
13309 | 1244 |
Forall(Implies(Member(0,succ(A)), |
1245 |
Forall(Implies(Member(0,succ(succ(A))), |
|
1246 |
Forall(Forall(Forall(Forall( |
|
1247 |
Implies(pair_fm(5,4,3), |
|
1248 |
Implies(fun_apply_fm(f#+6,5,2), |
|
1249 |
Implies(fun_apply_fm(f#+6,4,1), |
|
13429 | 1250 |
Implies(pair_fm(2,1,0), |
13309 | 1251 |
Iff(Member(3,r#+6), Member(0,s#+6)))))))))))))))" |
1252 |
||
1253 |
lemma order_isomorphism_type [TC]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1254 |
"\<lbrakk>A \<in> nat; r \<in> nat; B \<in> nat; s \<in> nat; f \<in> nat\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1255 |
\<Longrightarrow> order_isomorphism_fm(A,r,B,s,f) \<in> formula" |
13429 | 1256 |
by (simp add: order_isomorphism_fm_def) |
13309 | 1257 |
|
1258 |
lemma sats_order_isomorphism_fm [simp]: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1259 |
"\<lbrakk>U \<in> nat; r \<in> nat; B \<in> nat; s \<in> nat; f \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1260 |
\<Longrightarrow> sats(A, order_isomorphism_fm(U,r,B,s,f), env) \<longleftrightarrow> |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1261 |
order_isomorphism(##A, nth(U,env), nth(r,env), nth(B,env), |
13309 | 1262 |
nth(s,env), nth(f,env))" |
1263 |
by (simp add: order_isomorphism_fm_def order_isomorphism_def) |
|
1264 |
||
1265 |
lemma order_isomorphism_iff_sats: |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1266 |
"\<lbrakk>nth(i,env) = U; nth(j,env) = r; nth(k,env) = B; nth(j',env) = s; |
13429 | 1267 |
nth(k',env) = f; |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1268 |
i \<in> nat; j \<in> nat; k \<in> nat; j' \<in> nat; k' \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1269 |
\<Longrightarrow> order_isomorphism(##A,U,r,B,s,f) \<longleftrightarrow> |
13429 | 1270 |
sats(A, order_isomorphism_fm(i,j,k,j',k'), env)" |
13309 | 1271 |
by simp |
1272 |
||
13314 | 1273 |
theorem order_isomorphism_reflection: |
13429 | 1274 |
"REFLECTS[\<lambda>x. order_isomorphism(L,f(x),g(x),h(x),g'(x),h'(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1275 |
\<lambda>i x. order_isomorphism(##Lset(i),f(x),g(x),h(x),g'(x),h'(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
1276 |
apply (simp only: order_isomorphism_def) |
13429 | 1277 |
apply (intro FOL_reflections function_reflections bijection_reflection) |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1278 |
done |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1279 |
|
60770 | 1280 |
subsubsection\<open>Limit Ordinals, Internalized\<close> |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1281 |
|
60770 | 1282 |
text\<open>A limit ordinal is a non-empty, successor-closed ordinal\<close> |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1283 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1284 |
(* "limit_ordinal(M,a) \<equiv> |
76214 | 1285 |
ordinal(M,a) \<and> \<not> empty(M,a) \<and> |
1286 |
(\<forall>x[M]. x\<in>a \<longrightarrow> (\<exists>y[M]. y\<in>a \<and> successor(M,x,y)))" *) |
|
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1287 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
1288 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
1289 |
limit_ordinal_fm :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1290 |
"limit_ordinal_fm(x) \<equiv> |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1291 |
And(ordinal_fm(x), |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1292 |
And(Neg(empty_fm(x)), |
13429 | 1293 |
Forall(Implies(Member(0,succ(x)), |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1294 |
Exists(And(Member(0,succ(succ(x))), |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1295 |
succ_fm(1,0)))))))" |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1296 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1297 |
lemma limit_ordinal_type [TC]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1298 |
"x \<in> nat \<Longrightarrow> limit_ordinal_fm(x) \<in> formula" |
13429 | 1299 |
by (simp add: limit_ordinal_fm_def) |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1300 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1301 |
lemma sats_limit_ordinal_fm [simp]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1302 |
"\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1303 |
\<Longrightarrow> sats(A, limit_ordinal_fm(x), env) \<longleftrightarrow> limit_ordinal(##A, nth(x,env))" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1304 |
by (simp add: limit_ordinal_fm_def limit_ordinal_def sats_ordinal_fm') |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1305 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1306 |
lemma limit_ordinal_iff_sats: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1307 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1308 |
i \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1309 |
\<Longrightarrow> limit_ordinal(##A, x) \<longleftrightarrow> sats(A, limit_ordinal_fm(i), env)" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1310 |
by simp |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1311 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1312 |
theorem limit_ordinal_reflection: |
13429 | 1313 |
"REFLECTS[\<lambda>x. limit_ordinal(L,f(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1314 |
\<lambda>i x. limit_ordinal(##Lset(i),f(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
1315 |
apply (simp only: limit_ordinal_def) |
13429 | 1316 |
apply (intro FOL_reflections ordinal_reflection |
1317 |
empty_reflection successor_reflection) |
|
13314 | 1318 |
done |
13309 | 1319 |
|
60770 | 1320 |
subsubsection\<open>Finite Ordinals: The Predicate ``Is A Natural Number''\<close> |
13493
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1321 |
|
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1322 |
(* "finite_ordinal(M,a) \<equiv> |
76214 | 1323 |
ordinal(M,a) \<and> \<not> limit_ordinal(M,a) \<and> |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1324 |
(\<forall>x[M]. x\<in>a \<longrightarrow> \<not> limit_ordinal(M,x))" *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
1325 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
1326 |
finite_ordinal_fm :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1327 |
"finite_ordinal_fm(x) \<equiv> |
13493
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1328 |
And(ordinal_fm(x), |
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1329 |
And(Neg(limit_ordinal_fm(x)), |
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1330 |
Forall(Implies(Member(0,succ(x)), |
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1331 |
Neg(limit_ordinal_fm(0))))))" |
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1332 |
|
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1333 |
lemma finite_ordinal_type [TC]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1334 |
"x \<in> nat \<Longrightarrow> finite_ordinal_fm(x) \<in> formula" |
13493
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1335 |
by (simp add: finite_ordinal_fm_def) |
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1336 |
|
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1337 |
lemma sats_finite_ordinal_fm [simp]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1338 |
"\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1339 |
\<Longrightarrow> sats(A, finite_ordinal_fm(x), env) \<longleftrightarrow> finite_ordinal(##A, nth(x,env))" |
13493
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1340 |
by (simp add: finite_ordinal_fm_def sats_ordinal_fm' finite_ordinal_def) |
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1341 |
|
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1342 |
lemma finite_ordinal_iff_sats: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1343 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1344 |
i \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1345 |
\<Longrightarrow> finite_ordinal(##A, x) \<longleftrightarrow> sats(A, finite_ordinal_fm(i), env)" |
13493
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1346 |
by simp |
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1347 |
|
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1348 |
theorem finite_ordinal_reflection: |
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1349 |
"REFLECTS[\<lambda>x. finite_ordinal(L,f(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1350 |
\<lambda>i x. finite_ordinal(##Lset(i),f(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
1351 |
apply (simp only: finite_ordinal_def) |
13493
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1352 |
apply (intro FOL_reflections ordinal_reflection limit_ordinal_reflection) |
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1353 |
done |
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1354 |
|
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
paulson
parents:
13440
diff
changeset
|
1355 |
|
60770 | 1356 |
subsubsection\<open>Omega: The Set of Natural Numbers\<close> |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1357 |
|
76214 | 1358 |
(* omega(M,a) \<equiv> limit_ordinal(M,a) \<and> (\<forall>x[M]. x\<in>a \<longrightarrow> \<not> limit_ordinal(M,x)) *) |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
1359 |
definition |
76215
a642599ffdea
More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents:
76214
diff
changeset
|
1360 |
omega_fm :: "i\<Rightarrow>i" where |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1361 |
"omega_fm(x) \<equiv> |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1362 |
And(limit_ordinal_fm(x), |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1363 |
Forall(Implies(Member(0,succ(x)), |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1364 |
Neg(limit_ordinal_fm(0)))))" |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1365 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1366 |
lemma omega_type [TC]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1367 |
"x \<in> nat \<Longrightarrow> omega_fm(x) \<in> formula" |
13429 | 1368 |
by (simp add: omega_fm_def) |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1369 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1370 |
lemma sats_omega_fm [simp]: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1371 |
"\<lbrakk>x \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1372 |
\<Longrightarrow> sats(A, omega_fm(x), env) \<longleftrightarrow> omega(##A, nth(x,env))" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1373 |
by (simp add: omega_fm_def omega_def) |
13316 | 1374 |
|
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1375 |
lemma omega_iff_sats: |
76213
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1376 |
"\<lbrakk>nth(i,env) = x; nth(j,env) = y; |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1377 |
i \<in> nat; env \<in> list(A)\<rbrakk> |
e44d86131648
Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents:
71568
diff
changeset
|
1378 |
\<Longrightarrow> omega(##A, x) \<longleftrightarrow> sats(A, omega_fm(i), env)" |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1379 |
by simp |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1380 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1381 |
theorem omega_reflection: |
13429 | 1382 |
"REFLECTS[\<lambda>x. omega(L,f(x)), |
13807
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
paulson
parents:
13655
diff
changeset
|
1383 |
\<lambda>i x. omega(##Lset(i),f(x))]" |
13655
95b95cdb4704
Tidying up. New primitives is_iterates and is_iterates_fm.
paulson
parents:
13651
diff
changeset
|
1384 |
apply (simp only: omega_def) |
13429 | 1385 |
apply (intro FOL_reflections limit_ordinal_reflection) |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1386 |
done |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1387 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1388 |
|
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1389 |
lemmas fun_plus_reflections = |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1390 |
typed_function_reflection composition_reflection |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1391 |
injection_reflection surjection_reflection |
13348 | 1392 |
bijection_reflection restriction_reflection |
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13493
diff
changeset
|
1393 |
order_isomorphism_reflection finite_ordinal_reflection |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1394 |
ordinal_reflection limit_ordinal_reflection omega_reflection |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1395 |
|
13429 | 1396 |
lemmas fun_plus_iff_sats = |
1397 |
typed_function_iff_sats composition_iff_sats |
|
1398 |
injection_iff_sats surjection_iff_sats |
|
1399 |
bijection_iff_sats restriction_iff_sats |
|
13496
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
paulson
parents:
13493
diff
changeset
|
1400 |
order_isomorphism_iff_sats finite_ordinal_iff_sats |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13316
diff
changeset
|
1401 |
ordinal_iff_sats limit_ordinal_iff_sats omega_iff_sats |
13316 | 1402 |
|
13223 | 1403 |
end |