| author | Fabian Huch <huch@in.tum.de> | 
| Thu, 24 Oct 2024 14:07:13 +0200 | |
| changeset 81364 | 84e4388f8ab1 | 
| parent 76215 | a642599ffdea | 
| permissions | -rw-r--r-- | 
| 
13437
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
1  | 
(* Title: ZF/Constructible/Rec_Separation.thy  | 
| 
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
3  | 
*)  | 
| 13429 | 4  | 
|
| 60770 | 5  | 
section \<open>Separation for Facts About Recursion\<close>  | 
| 13348 | 6  | 
|
| 16417 | 7  | 
theory Rec_Separation imports Separation Internalize begin  | 
| 13348 | 8  | 
|
| 61798 | 9  | 
text\<open>This theory proves all instances needed for locales \<open>M_trancl\<close> and \<open>M_datatypes\<close>\<close>  | 
| 13348 | 10  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
11  | 
lemma eq_succ_imp_lt: "\<lbrakk>i = succ(j); Ord(i)\<rbrakk> \<Longrightarrow> j<i"  | 
| 13428 | 12  | 
by simp  | 
| 13363 | 13  | 
|
| 
13493
 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 
paulson 
parents: 
13441 
diff
changeset
 | 
14  | 
|
| 61798 | 15  | 
subsection\<open>The Locale \<open>M_trancl\<close>\<close>  | 
| 13348 | 16  | 
|
| 60770 | 17  | 
subsubsection\<open>Separation for Reflexive/Transitive Closure\<close>  | 
| 13348 | 18  | 
|
| 60770 | 19  | 
text\<open>First, The Defining Formula\<close>  | 
| 13348 | 20  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
21  | 
(* "rtran_closure_mem(M,A,r,p) \<equiv>  | 
| 13428 | 22  | 
\<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M].  | 
| 76214 | 23  | 
omega(M,nnat) \<and> n\<in>nnat \<and> successor(M,n,n') \<and>  | 
24  | 
(\<exists>f[M]. typed_function(M,n',A,f) \<and>  | 
|
25  | 
(\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) \<and> empty(M,zero) \<and>  | 
|
26  | 
fun_apply(M,f,zero,x) \<and> fun_apply(M,f,n,y)) \<and>  | 
|
| 46823 | 27  | 
(\<forall>j[M]. j\<in>n \<longrightarrow>  | 
| 13428 | 28  | 
(\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M].  | 
| 76214 | 29  | 
fun_apply(M,f,j,fj) \<and> successor(M,j,sj) \<and>  | 
30  | 
fun_apply(M,f,sj,fsj) \<and> pair(M,fj,fsj,ffp) \<and> ffp \<in> r)))"*)  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21233 
diff
changeset
 | 
31  | 
definition  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
32  | 
rtran_closure_mem_fm :: "[i,i,i]\<Rightarrow>i" where  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
33  | 
"rtran_closure_mem_fm(A,r,p) \<equiv>  | 
| 13348 | 34  | 
Exists(Exists(Exists(  | 
35  | 
And(omega_fm(2),  | 
|
36  | 
And(Member(1,2),  | 
|
37  | 
And(succ_fm(1,0),  | 
|
38  | 
Exists(And(typed_function_fm(1, A#+4, 0),  | 
|
| 13428 | 39  | 
And(Exists(Exists(Exists(  | 
40  | 
And(pair_fm(2,1,p#+7),  | 
|
41  | 
And(empty_fm(0),  | 
|
42  | 
And(fun_apply_fm(3,0,2), fun_apply_fm(3,5,1))))))),  | 
|
43  | 
Forall(Implies(Member(0,3),  | 
|
44  | 
Exists(Exists(Exists(Exists(  | 
|
45  | 
And(fun_apply_fm(5,4,3),  | 
|
46  | 
And(succ_fm(4,2),  | 
|
47  | 
And(fun_apply_fm(5,2,1),  | 
|
48  | 
And(pair_fm(3,1,0), Member(0,r#+9))))))))))))))))))))"  | 
|
| 13348 | 49  | 
|
50  | 
||
51  | 
lemma rtran_closure_mem_type [TC]:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
52  | 
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> rtran_closure_mem_fm(x,y,z) \<in> formula"  | 
| 13428 | 53  | 
by (simp add: rtran_closure_mem_fm_def)  | 
| 13348 | 54  | 
|
55  | 
lemma sats_rtran_closure_mem_fm [simp]:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
56  | 
"\<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
 | 
57  | 
\<Longrightarrow> sats(A, rtran_closure_mem_fm(x,y,z), env) \<longleftrightarrow>  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
58  | 
rtran_closure_mem(##A, nth(x,env), nth(y,env), nth(z,env))"  | 
| 13348 | 59  | 
by (simp add: rtran_closure_mem_fm_def rtran_closure_mem_def)  | 
60  | 
||
61  | 
lemma rtran_closure_mem_iff_sats:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
62  | 
"\<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
 | 
63  | 
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
 | 
64  | 
\<Longrightarrow> rtran_closure_mem(##A, x, y, z) \<longleftrightarrow> sats(A, rtran_closure_mem_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
 | 
65  | 
by (simp)  | 
| 13348 | 66  | 
|
| 13566 | 67  | 
lemma rtran_closure_mem_reflection:  | 
| 13428 | 68  | 
"REFLECTS[\<lambda>x. rtran_closure_mem(L,f(x),g(x),h(x)),  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
69  | 
\<lambda>i x. rtran_closure_mem(##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
 | 
70  | 
apply (simp only: rtran_closure_mem_def)  | 
| 13428 | 71  | 
apply (intro FOL_reflections function_reflections fun_plus_reflections)  | 
| 13348 | 72  | 
done  | 
73  | 
||
| 69593 | 74  | 
text\<open>Separation for \<^term>\<open>rtrancl(r)\<close>.\<close>  | 
| 13348 | 75  | 
lemma rtrancl_separation:  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
76  | 
"\<lbrakk>L(r); L(A)\<rbrakk> \<Longrightarrow> separation (L, rtran_closure_mem(L,A,r))"  | 
| 13687 | 77  | 
apply (rule gen_separation_multi [OF rtran_closure_mem_reflection, of "{r,A}"],
 | 
78  | 
auto)  | 
|
79  | 
apply (rule_tac env="[r,A]" in DPow_LsetI)  | 
|
80  | 
apply (rule rtran_closure_mem_iff_sats sep_rules | simp)+  | 
|
| 13348 | 81  | 
done  | 
82  | 
||
83  | 
||
| 60770 | 84  | 
subsubsection\<open>Reflexive/Transitive Closure, Internalized\<close>  | 
| 13348 | 85  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
86  | 
(* "rtran_closure(M,r,s) \<equiv>  | 
| 46823 | 87  | 
\<forall>A[M]. is_field(M,r,A) \<longrightarrow>  | 
88  | 
(\<forall>p[M]. p \<in> s \<longleftrightarrow> rtran_closure_mem(M,A,r,p))" *)  | 
|
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21233 
diff
changeset
 | 
89  | 
definition  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
90  | 
rtran_closure_fm :: "[i,i]\<Rightarrow>i" where  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
91  | 
"rtran_closure_fm(r,s) \<equiv>  | 
| 13348 | 92  | 
Forall(Implies(field_fm(succ(r),0),  | 
93  | 
Forall(Iff(Member(0,succ(succ(s))),  | 
|
94  | 
rtran_closure_mem_fm(1,succ(succ(r)),0)))))"  | 
|
95  | 
||
96  | 
lemma rtran_closure_type [TC]:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
97  | 
"\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> rtran_closure_fm(x,y) \<in> formula"  | 
| 13428 | 98  | 
by (simp add: rtran_closure_fm_def)  | 
| 13348 | 99  | 
|
100  | 
lemma sats_rtran_closure_fm [simp]:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
101  | 
"\<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
 | 
102  | 
\<Longrightarrow> sats(A, rtran_closure_fm(x,y), env) \<longleftrightarrow>  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
103  | 
rtran_closure(##A, nth(x,env), nth(y,env))"  | 
| 13348 | 104  | 
by (simp add: rtran_closure_fm_def rtran_closure_def)  | 
105  | 
||
106  | 
lemma rtran_closure_iff_sats:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
107  | 
"\<lbrakk>nth(i,env) = x; nth(j,env) = y;  | 
| 
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
108  | 
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
 | 
109  | 
\<Longrightarrow> rtran_closure(##A, x, y) \<longleftrightarrow> sats(A, rtran_closure_fm(i,j), env)"  | 
| 13348 | 110  | 
by simp  | 
111  | 
||
112  | 
theorem rtran_closure_reflection:  | 
|
| 13428 | 113  | 
"REFLECTS[\<lambda>x. rtran_closure(L,f(x),g(x)),  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
114  | 
\<lambda>i x. rtran_closure(##Lset(i),f(x),g(x))]"  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
115  | 
apply (simp only: rtran_closure_def)  | 
| 13348 | 116  | 
apply (intro FOL_reflections function_reflections rtran_closure_mem_reflection)  | 
117  | 
done  | 
|
118  | 
||
119  | 
||
| 60770 | 120  | 
subsubsection\<open>Transitive Closure of a Relation, Internalized\<close>  | 
| 13348 | 121  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
122  | 
(* "tran_closure(M,r,t) \<equiv>  | 
| 76214 | 123  | 
\<exists>s[M]. rtran_closure(M,r,s) \<and> composition(M,r,s,t)" *)  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21233 
diff
changeset
 | 
124  | 
definition  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
125  | 
tran_closure_fm :: "[i,i]\<Rightarrow>i" where  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
126  | 
"tran_closure_fm(r,s) \<equiv>  | 
| 13348 | 127  | 
Exists(And(rtran_closure_fm(succ(r),0), composition_fm(succ(r),0,succ(s))))"  | 
128  | 
||
129  | 
lemma tran_closure_type [TC]:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
130  | 
"\<lbrakk>x \<in> nat; y \<in> nat\<rbrakk> \<Longrightarrow> tran_closure_fm(x,y) \<in> formula"  | 
| 13428 | 131  | 
by (simp add: tran_closure_fm_def)  | 
| 13348 | 132  | 
|
133  | 
lemma sats_tran_closure_fm [simp]:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
134  | 
"\<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
 | 
135  | 
\<Longrightarrow> sats(A, tran_closure_fm(x,y), env) \<longleftrightarrow>  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
136  | 
tran_closure(##A, nth(x,env), nth(y,env))"  | 
| 13348 | 137  | 
by (simp add: tran_closure_fm_def tran_closure_def)  | 
138  | 
||
139  | 
lemma tran_closure_iff_sats:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
140  | 
"\<lbrakk>nth(i,env) = x; nth(j,env) = y;  | 
| 
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
141  | 
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
 | 
142  | 
\<Longrightarrow> tran_closure(##A, x, y) \<longleftrightarrow> sats(A, tran_closure_fm(i,j), env)"  | 
| 13348 | 143  | 
by simp  | 
144  | 
||
145  | 
theorem tran_closure_reflection:  | 
|
| 13428 | 146  | 
"REFLECTS[\<lambda>x. tran_closure(L,f(x),g(x)),  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
147  | 
\<lambda>i x. tran_closure(##Lset(i),f(x),g(x))]"  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
148  | 
apply (simp only: tran_closure_def)  | 
| 13428 | 149  | 
apply (intro FOL_reflections function_reflections  | 
| 13348 | 150  | 
rtran_closure_reflection composition_reflection)  | 
151  | 
done  | 
|
152  | 
||
153  | 
||
| 61798 | 154  | 
subsubsection\<open>Separation for the Proof of \<open>wellfounded_on_trancl\<close>\<close>  | 
| 13348 | 155  | 
|
156  | 
lemma wellfounded_trancl_reflects:  | 
|
| 13428 | 157  | 
"REFLECTS[\<lambda>x. \<exists>w[L]. \<exists>wx[L]. \<exists>rp[L].  | 
| 76214 | 158  | 
w \<in> Z \<and> pair(L,w,x,wx) \<and> tran_closure(L,r,rp) \<and> wx \<in> rp,  | 
| 13428 | 159  | 
\<lambda>i x. \<exists>w \<in> Lset(i). \<exists>wx \<in> Lset(i). \<exists>rp \<in> Lset(i).  | 
| 76214 | 160  | 
w \<in> Z \<and> pair(##Lset(i),w,x,wx) \<and> tran_closure(##Lset(i),r,rp) \<and>  | 
| 13348 | 161  | 
wx \<in> rp]"  | 
| 13428 | 162  | 
by (intro FOL_reflections function_reflections fun_plus_reflections  | 
| 13348 | 163  | 
tran_closure_reflection)  | 
164  | 
||
165  | 
lemma wellfounded_trancl_separation:  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
166  | 
"\<lbrakk>L(r); L(Z)\<rbrakk> \<Longrightarrow>  | 
| 13428 | 167  | 
separation (L, \<lambda>x.  | 
168  | 
\<exists>w[L]. \<exists>wx[L]. \<exists>rp[L].  | 
|
| 76214 | 169  | 
w \<in> Z \<and> pair(L,w,x,wx) \<and> tran_closure(L,r,rp) \<and> wx \<in> rp)"  | 
| 13687 | 170  | 
apply (rule gen_separation_multi [OF wellfounded_trancl_reflects, of "{r,Z}"],
 | 
171  | 
auto)  | 
|
172  | 
apply (rule_tac env="[r,Z]" in DPow_LsetI)  | 
|
| 13348 | 173  | 
apply (rule sep_rules tran_closure_iff_sats | simp)+  | 
174  | 
done  | 
|
175  | 
||
| 13363 | 176  | 
|
| 61798 | 177  | 
subsubsection\<open>Instantiating the locale \<open>M_trancl\<close>\<close>  | 
| 13428 | 178  | 
|
| 
13437
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
179  | 
lemma M_trancl_axioms_L: "M_trancl_axioms(L)"  | 
| 13428 | 180  | 
apply (rule M_trancl_axioms.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
 | 
181  | 
apply (assumption | rule rtrancl_separation wellfounded_trancl_separation L_nat)+  | 
| 13428 | 182  | 
done  | 
| 13363 | 183  | 
|
| 
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
 | 
184  | 
theorem M_trancl_L: "M_trancl(L)"  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
16417 
diff
changeset
 | 
185  | 
by (rule M_trancl.intro [OF M_basic_L M_trancl_axioms_L])  | 
| 
13437
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
186  | 
|
| 
71568
 
1005c50b2750
prefer strict qualification (default for 'interpretation', see 461ee3e49ad3) as proposed by Pedro Sánchez Terraf;
 
wenzelm 
parents: 
71417 
diff
changeset
 | 
187  | 
interpretation L: M_trancl L by (rule M_trancl_L)  | 
| 13363 | 188  | 
|
189  | 
||
| 69593 | 190  | 
subsection\<open>\<^term>\<open>L\<close> is Closed Under the Operator \<^term>\<open>list\<close>\<close>  | 
| 13363 | 191  | 
|
| 60770 | 192  | 
subsubsection\<open>Instances of Replacement for Lists\<close>  | 
| 13386 | 193  | 
|
| 13363 | 194  | 
lemma list_replacement1_Reflects:  | 
195  | 
"REFLECTS  | 
|
196  | 
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and>  | 
|
197  | 
is_wfrec(L, iterates_MH(L, is_list_functor(L,A), 0), memsn, u, y)),  | 
|
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
198  | 
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) \<and>  | 
| 
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
199  | 
is_wfrec(##Lset(i),  | 
| 
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
200  | 
iterates_MH(##Lset(i),  | 
| 
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
201  | 
is_list_functor(##Lset(i), A), 0), memsn, u, y))]"  | 
| 13428 | 202  | 
by (intro FOL_reflections function_reflections is_wfrec_reflection  | 
203  | 
iterates_MH_reflection list_functor_reflection)  | 
|
| 13363 | 204  | 
|
| 13441 | 205  | 
|
| 13428 | 206  | 
lemma list_replacement1:  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
207  | 
"L(A) \<Longrightarrow> iterates_replacement(L, is_list_functor(L,A), 0)"  | 
| 13363 | 208  | 
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)  | 
| 13428 | 209  | 
apply (rule strong_replacementI)  | 
| 13566 | 210  | 
apply (rule_tac u="{B,A,n,0,Memrel(succ(n))}" 
 | 
| 13687 | 211  | 
in gen_separation_multi [OF list_replacement1_Reflects],  | 
| 
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
 | 
212  | 
auto)  | 
| 13687 | 213  | 
apply (rule_tac env="[B,A,n,0,Memrel(succ(n))]" in DPow_LsetI)  | 
| 13434 | 214  | 
apply (rule sep_rules is_nat_case_iff_sats list_functor_iff_sats  | 
| 13441 | 215  | 
is_wfrec_iff_sats iterates_MH_iff_sats quasinat_iff_sats | simp)+  | 
| 13363 | 216  | 
done  | 
217  | 
||
| 13441 | 218  | 
|
| 13363 | 219  | 
lemma list_replacement2_Reflects:  | 
220  | 
"REFLECTS  | 
|
| 76214 | 221  | 
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> u \<in> nat \<and>  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
222  | 
is_iterates(L, is_list_functor(L, A), 0, u, x),  | 
| 76214 | 223  | 
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> u \<in> nat \<and>  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
224  | 
is_iterates(##Lset(i), is_list_functor(##Lset(i), A), 0, u, x)]"  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
225  | 
by (intro FOL_reflections  | 
| 
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
226  | 
is_iterates_reflection list_functor_reflection)  | 
| 13363 | 227  | 
|
| 13428 | 228  | 
lemma list_replacement2:  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
229  | 
"L(A) \<Longrightarrow> strong_replacement(L,  | 
| 76214 | 230  | 
\<lambda>n y. n\<in>nat \<and> is_iterates(L, is_list_functor(L,A), 0, n, y))"  | 
| 13428 | 231  | 
apply (rule strong_replacementI)  | 
| 13566 | 232  | 
apply (rule_tac u="{A,B,0,nat}" 
 | 
| 13687 | 233  | 
in gen_separation_multi [OF list_replacement2_Reflects],  | 
| 
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
 | 
234  | 
auto)  | 
| 13687 | 235  | 
apply (rule_tac env="[A,B,0,nat]" in DPow_LsetI)  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
236  | 
apply (rule sep_rules list_functor_iff_sats is_iterates_iff_sats | simp)+  | 
| 13363 | 237  | 
done  | 
238  | 
||
| 13386 | 239  | 
|
| 69593 | 240  | 
subsection\<open>\<^term>\<open>L\<close> is Closed Under the Operator \<^term>\<open>formula\<close>\<close>  | 
| 13386 | 241  | 
|
| 60770 | 242  | 
subsubsection\<open>Instances of Replacement for Formulas\<close>  | 
| 13386 | 243  | 
|
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
244  | 
(*FIXME: could prove a lemma iterates_replacementI to eliminate the  | 
| 
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
245  | 
need to expand iterates_replacement and wfrec_replacement*)  | 
| 13386 | 246  | 
lemma formula_replacement1_Reflects:  | 
247  | 
"REFLECTS  | 
|
| 76214 | 248  | 
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and>  | 
| 13386 | 249  | 
is_wfrec(L, iterates_MH(L, is_formula_functor(L), 0), memsn, u, y)),  | 
| 76214 | 250  | 
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) \<and>  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
251  | 
is_wfrec(##Lset(i),  | 
| 
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
252  | 
iterates_MH(##Lset(i),  | 
| 
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
253  | 
is_formula_functor(##Lset(i)), 0), memsn, u, y))]"  | 
| 13428 | 254  | 
by (intro FOL_reflections function_reflections is_wfrec_reflection  | 
255  | 
iterates_MH_reflection formula_functor_reflection)  | 
|
| 13386 | 256  | 
|
| 13428 | 257  | 
lemma formula_replacement1:  | 
| 13386 | 258  | 
"iterates_replacement(L, is_formula_functor(L), 0)"  | 
259  | 
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)  | 
|
| 13428 | 260  | 
apply (rule strong_replacementI)  | 
| 13566 | 261  | 
apply (rule_tac u="{B,n,0,Memrel(succ(n))}" 
 | 
| 13687 | 262  | 
in gen_separation_multi [OF formula_replacement1_Reflects],  | 
| 
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
 | 
263  | 
auto)  | 
| 13687 | 264  | 
apply (rule_tac env="[n,B,0,Memrel(succ(n))]" in DPow_LsetI)  | 
| 13434 | 265  | 
apply (rule sep_rules is_nat_case_iff_sats formula_functor_iff_sats  | 
| 13441 | 266  | 
is_wfrec_iff_sats iterates_MH_iff_sats quasinat_iff_sats | simp)+  | 
| 13386 | 267  | 
done  | 
268  | 
||
269  | 
lemma formula_replacement2_Reflects:  | 
|
270  | 
"REFLECTS  | 
|
| 76214 | 271  | 
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> u \<in> nat \<and>  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
272  | 
is_iterates(L, is_formula_functor(L), 0, u, x),  | 
| 76214 | 273  | 
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> u \<in> nat \<and>  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
274  | 
is_iterates(##Lset(i), is_formula_functor(##Lset(i)), 0, u, x)]"  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
275  | 
by (intro FOL_reflections  | 
| 
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
276  | 
is_iterates_reflection formula_functor_reflection)  | 
| 13386 | 277  | 
|
| 13428 | 278  | 
lemma formula_replacement2:  | 
279  | 
"strong_replacement(L,  | 
|
| 76214 | 280  | 
\<lambda>n y. n\<in>nat \<and> is_iterates(L, is_formula_functor(L), 0, n, y))"  | 
| 13428 | 281  | 
apply (rule strong_replacementI)  | 
| 13566 | 282  | 
apply (rule_tac u="{B,0,nat}" 
 | 
| 13687 | 283  | 
in gen_separation_multi [OF formula_replacement2_Reflects],  | 
| 
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
 | 
284  | 
auto)  | 
| 13687 | 285  | 
apply (rule_tac env="[B,0,nat]" in DPow_LsetI)  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
286  | 
apply (rule sep_rules formula_functor_iff_sats is_iterates_iff_sats | simp)+  | 
| 13386 | 287  | 
done  | 
288  | 
||
| 69593 | 289  | 
text\<open>NB The proofs for type \<^term>\<open>formula\<close> are virtually identical to those  | 
290  | 
for \<^term>\<open>list(A)\<close>. It was a cut-and-paste job!\<close>  | 
|
| 13386 | 291  | 
|
| 13387 | 292  | 
|
| 69593 | 293  | 
subsubsection\<open>The Formula \<^term>\<open>is_nth\<close>, Internalized\<close>  | 
| 
13437
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
294  | 
|
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
295  | 
(* "is_nth(M,n,l,Z) \<equiv>  | 
| 76214 | 296  | 
\<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) \<and> is_hd(M,X,Z)" *)  | 
| 
21404
 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 
wenzelm 
parents: 
21233 
diff
changeset
 | 
297  | 
definition  | 
| 
76215
 
a642599ffdea
More syntactic cleanup. LaTeX markup working
 
paulson <lp15@cam.ac.uk> 
parents: 
76214 
diff
changeset
 | 
298  | 
nth_fm :: "[i,i,i]\<Rightarrow>i" where  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
299  | 
"nth_fm(n,l,Z) \<equiv>  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
300  | 
Exists(And(is_iterates_fm(tl_fm(1,0), succ(l), succ(n), 0),  | 
| 
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
301  | 
hd_fm(0,succ(Z))))"  | 
| 
13493
 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 
paulson 
parents: 
13441 
diff
changeset
 | 
302  | 
|
| 
 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 
paulson 
parents: 
13441 
diff
changeset
 | 
303  | 
lemma nth_fm_type [TC]:  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
304  | 
"\<lbrakk>x \<in> nat; y \<in> nat; z \<in> nat\<rbrakk> \<Longrightarrow> nth_fm(x,y,z) \<in> formula"  | 
| 
13493
 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 
paulson 
parents: 
13441 
diff
changeset
 | 
305  | 
by (simp add: nth_fm_def)  | 
| 
 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 
paulson 
parents: 
13441 
diff
changeset
 | 
306  | 
|
| 
 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 
paulson 
parents: 
13441 
diff
changeset
 | 
307  | 
lemma sats_nth_fm [simp]:  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
308  | 
"\<lbrakk>x < length(env); 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
 | 
309  | 
\<Longrightarrow> sats(A, nth_fm(x,y,z), env) \<longleftrightarrow>  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
310  | 
is_nth(##A, nth(x,env), nth(y,env), nth(z,env))"  | 
| 
13493
 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 
paulson 
parents: 
13441 
diff
changeset
 | 
311  | 
apply (frule lt_length_in_nat, assumption)  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
312  | 
apply (simp add: nth_fm_def is_nth_def sats_is_iterates_fm)  | 
| 
13493
 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 
paulson 
parents: 
13441 
diff
changeset
 | 
313  | 
done  | 
| 
 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 
paulson 
parents: 
13441 
diff
changeset
 | 
314  | 
|
| 
 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 
paulson 
parents: 
13441 
diff
changeset
 | 
315  | 
lemma nth_iff_sats:  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
316  | 
"\<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
 | 
317  | 
i < length(env); 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
 | 
318  | 
\<Longrightarrow> is_nth(##A, x, y, z) \<longleftrightarrow> sats(A, nth_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
 | 
319  | 
by (simp)  | 
| 
13437
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
320  | 
|
| 
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
321  | 
theorem nth_reflection:  | 
| 
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
322  | 
"REFLECTS[\<lambda>x. is_nth(L, f(x), g(x), h(x)),  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
323  | 
\<lambda>i x. is_nth(##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
 | 
324  | 
apply (simp only: is_nth_def)  | 
| 
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
325  | 
apply (intro FOL_reflections is_iterates_reflection  | 
| 
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
326  | 
hd_reflection tl_reflection)  | 
| 
13437
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
327  | 
done  | 
| 
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
328  | 
|
| 
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
329  | 
|
| 69593 | 330  | 
subsubsection\<open>An Instance of Replacement for \<^term>\<open>nth\<close>\<close>  | 
| 
13409
 
d4ea094c650e
Relativization and Separation for the function "nth"
 
paulson 
parents: 
13398 
diff
changeset
 | 
331  | 
|
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
332  | 
(*FIXME: could prove a lemma iterates_replacementI to eliminate the  | 
| 
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
333  | 
need to expand iterates_replacement and wfrec_replacement*)  | 
| 
13409
 
d4ea094c650e
Relativization and Separation for the function "nth"
 
paulson 
parents: 
13398 
diff
changeset
 | 
334  | 
lemma nth_replacement_Reflects:  | 
| 
 
d4ea094c650e
Relativization and Separation for the function "nth"
 
paulson 
parents: 
13398 
diff
changeset
 | 
335  | 
"REFLECTS  | 
| 76214 | 336  | 
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and>  | 
| 
13409
 
d4ea094c650e
Relativization and Separation for the function "nth"
 
paulson 
parents: 
13398 
diff
changeset
 | 
337  | 
is_wfrec(L, iterates_MH(L, is_tl(L), z), memsn, u, y)),  | 
| 76214 | 338  | 
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) \<and>  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
339  | 
is_wfrec(##Lset(i),  | 
| 
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
340  | 
iterates_MH(##Lset(i),  | 
| 
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
341  | 
is_tl(##Lset(i)), z), memsn, u, y))]"  | 
| 13428 | 342  | 
by (intro FOL_reflections function_reflections is_wfrec_reflection  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
343  | 
iterates_MH_reflection tl_reflection)  | 
| 
13409
 
d4ea094c650e
Relativization and Separation for the function "nth"
 
paulson 
parents: 
13398 
diff
changeset
 | 
344  | 
|
| 13428 | 345  | 
lemma nth_replacement:  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
346  | 
"L(w) \<Longrightarrow> iterates_replacement(L, is_tl(L), w)"  | 
| 
13409
 
d4ea094c650e
Relativization and Separation for the function "nth"
 
paulson 
parents: 
13398 
diff
changeset
 | 
347  | 
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)  | 
| 13428 | 348  | 
apply (rule strong_replacementI)  | 
| 13687 | 349  | 
apply (rule_tac u="{B,w,Memrel(succ(n))}" 
 | 
350  | 
in gen_separation_multi [OF nth_replacement_Reflects],  | 
|
351  | 
auto)  | 
|
352  | 
apply (rule_tac env="[B,w,Memrel(succ(n))]" in DPow_LsetI)  | 
|
| 13434 | 353  | 
apply (rule sep_rules is_nat_case_iff_sats tl_iff_sats  | 
| 13441 | 354  | 
is_wfrec_iff_sats iterates_MH_iff_sats quasinat_iff_sats | simp)+  | 
| 
13409
 
d4ea094c650e
Relativization and Separation for the function "nth"
 
paulson 
parents: 
13398 
diff
changeset
 | 
355  | 
done  | 
| 
 
d4ea094c650e
Relativization and Separation for the function "nth"
 
paulson 
parents: 
13398 
diff
changeset
 | 
356  | 
|
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
357  | 
|
| 61798 | 358  | 
subsubsection\<open>Instantiating the locale \<open>M_datatypes\<close>\<close>  | 
| 13428 | 359  | 
|
| 
13437
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
360  | 
lemma M_datatypes_axioms_L: "M_datatypes_axioms(L)"  | 
| 13428 | 361  | 
apply (rule M_datatypes_axioms.intro)  | 
362  | 
apply (assumption | rule  | 
|
363  | 
list_replacement1 list_replacement2  | 
|
364  | 
formula_replacement1 formula_replacement2  | 
|
365  | 
nth_replacement)+  | 
|
366  | 
done  | 
|
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
367  | 
|
| 
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
 | 
368  | 
theorem M_datatypes_L: "M_datatypes(L)"  | 
| 
13437
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
369  | 
apply (rule M_datatypes.intro)  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
16417 
diff
changeset
 | 
370  | 
apply (rule M_trancl_L)  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
16417 
diff
changeset
 | 
371  | 
apply (rule M_datatypes_axioms_L)  | 
| 
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
16417 
diff
changeset
 | 
372  | 
done  | 
| 
13437
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
373  | 
|
| 
71568
 
1005c50b2750
prefer strict qualification (default for 'interpretation', see 461ee3e49ad3) as proposed by Pedro Sánchez Terraf;
 
wenzelm 
parents: 
71417 
diff
changeset
 | 
374  | 
interpretation L: M_datatypes L by (rule M_datatypes_L)  | 
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
375  | 
|
| 
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
376  | 
|
| 69593 | 377  | 
subsection\<open>\<^term>\<open>L\<close> is Closed Under the Operator \<^term>\<open>eclose\<close>\<close>  | 
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
378  | 
|
| 69593 | 379  | 
subsubsection\<open>Instances of Replacement for \<^term>\<open>eclose\<close>\<close>  | 
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
380  | 
|
| 
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
381  | 
lemma eclose_replacement1_Reflects:  | 
| 
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
382  | 
"REFLECTS  | 
| 76214 | 383  | 
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> (\<exists>y[L]. pair(L,u,y,x) \<and>  | 
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
384  | 
is_wfrec(L, iterates_MH(L, big_union(L), A), memsn, u, y)),  | 
| 76214 | 385  | 
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) \<and>  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
386  | 
is_wfrec(##Lset(i),  | 
| 
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
387  | 
iterates_MH(##Lset(i), big_union(##Lset(i)), A),  | 
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
388  | 
memsn, u, y))]"  | 
| 13428 | 389  | 
by (intro FOL_reflections function_reflections is_wfrec_reflection  | 
390  | 
iterates_MH_reflection)  | 
|
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
391  | 
|
| 13428 | 392  | 
lemma eclose_replacement1:  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
393  | 
"L(A) \<Longrightarrow> iterates_replacement(L, big_union(L), A)"  | 
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
394  | 
apply (unfold iterates_replacement_def wfrec_replacement_def, clarify)  | 
| 13428 | 395  | 
apply (rule strong_replacementI)  | 
| 13566 | 396  | 
apply (rule_tac u="{B,A,n,Memrel(succ(n))}" 
 | 
| 13687 | 397  | 
in gen_separation_multi [OF eclose_replacement1_Reflects], auto)  | 
398  | 
apply (rule_tac env="[B,A,n,Memrel(succ(n))]" in DPow_LsetI)  | 
|
| 13434 | 399  | 
apply (rule sep_rules iterates_MH_iff_sats is_nat_case_iff_sats  | 
| 13441 | 400  | 
is_wfrec_iff_sats big_union_iff_sats quasinat_iff_sats | simp)+  | 
| 
13409
 
d4ea094c650e
Relativization and Separation for the function "nth"
 
paulson 
parents: 
13398 
diff
changeset
 | 
401  | 
done  | 
| 
 
d4ea094c650e
Relativization and Separation for the function "nth"
 
paulson 
parents: 
13398 
diff
changeset
 | 
402  | 
|
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
403  | 
|
| 
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
404  | 
lemma eclose_replacement2_Reflects:  | 
| 
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
405  | 
"REFLECTS  | 
| 76214 | 406  | 
[\<lambda>x. \<exists>u[L]. u \<in> B \<and> u \<in> nat \<and>  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
407  | 
is_iterates(L, big_union(L), A, u, x),  | 
| 76214 | 408  | 
\<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B \<and> u \<in> nat \<and>  | 
| 
13807
 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 
paulson 
parents: 
13687 
diff
changeset
 | 
409  | 
is_iterates(##Lset(i), big_union(##Lset(i)), A, u, x)]"  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
410  | 
by (intro FOL_reflections function_reflections is_iterates_reflection)  | 
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
411  | 
|
| 13428 | 412  | 
lemma eclose_replacement2:  | 
| 
76213
 
e44d86131648
Removal of obsolete ASCII syntax
 
paulson <lp15@cam.ac.uk> 
parents: 
71568 
diff
changeset
 | 
413  | 
"L(A) \<Longrightarrow> strong_replacement(L,  | 
| 76214 | 414  | 
\<lambda>n y. n\<in>nat \<and> is_iterates(L, big_union(L), A, n, y))"  | 
| 13428 | 415  | 
apply (rule strong_replacementI)  | 
| 13566 | 416  | 
apply (rule_tac u="{A,B,nat}" 
 | 
| 13687 | 417  | 
in gen_separation_multi [OF eclose_replacement2_Reflects],  | 
| 
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
 | 
418  | 
auto)  | 
| 13687 | 419  | 
apply (rule_tac env="[A,B,nat]" in DPow_LsetI)  | 
| 
13655
 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 
paulson 
parents: 
13651 
diff
changeset
 | 
420  | 
apply (rule sep_rules is_iterates_iff_sats big_union_iff_sats | simp)+  | 
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
421  | 
done  | 
| 
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
422  | 
|
| 
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
423  | 
|
| 61798 | 424  | 
subsubsection\<open>Instantiating the locale \<open>M_eclose\<close>\<close>  | 
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
425  | 
|
| 
13437
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
426  | 
lemma M_eclose_axioms_L: "M_eclose_axioms(L)"  | 
| 
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
427  | 
apply (rule M_eclose_axioms.intro)  | 
| 
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
428  | 
apply (assumption | rule eclose_replacement1 eclose_replacement2)+  | 
| 
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
429  | 
done  | 
| 
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
430  | 
|
| 
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
 | 
431  | 
theorem M_eclose_L: "M_eclose(L)"  | 
| 13428 | 432  | 
apply (rule M_eclose.intro)  | 
| 
19931
 
fb32b43e7f80
Restructured locales with predicates: import is now an interpretation.
 
ballarin 
parents: 
16417 
diff
changeset
 | 
433  | 
apply (rule M_datatypes_L)  | 
| 
13437
 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 
paulson 
parents: 
13434 
diff
changeset
 | 
434  | 
apply (rule M_eclose_axioms_L)  | 
| 13428 | 435  | 
done  | 
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
436  | 
|
| 
71568
 
1005c50b2750
prefer strict qualification (default for 'interpretation', see 461ee3e49ad3) as proposed by Pedro Sánchez Terraf;
 
wenzelm 
parents: 
71417 
diff
changeset
 | 
437  | 
interpretation L: M_eclose L by (rule M_eclose_L)  | 
| 15766 | 438  | 
|
| 
13422
 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 
paulson 
parents: 
13418 
diff
changeset
 | 
439  | 
|
| 13348 | 440  | 
end  |