| author | kleing | 
| Wed, 25 May 2005 11:14:59 +0200 | |
| changeset 16076 | 03e8a88c0b54 | 
| parent 15766 | b08feb003f3c | 
| child 16417 | 9bc16273c2d4 | 
| permissions | -rw-r--r-- | 
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 1 | (* Title: ZF/Constructible/Rec_Separation.thy | 
| 13634 | 2 | ID: $Id$ | 
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | 
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 4 | *) | 
| 13429 | 5 | |
| 6 | header {*Separation for Facts About Recursion*}
 | |
| 13348 | 7 | |
| 13496 
6f0c57def6d5
In ZF/Constructible, moved many results from Satisfies_absolute, etc., to
 paulson parents: 
13493diff
changeset | 8 | theory Rec_Separation = Separation + Internalize: | 
| 13348 | 9 | |
| 10 | text{*This theory proves all instances needed for locales @{text
 | |
| 13634 | 11 | "M_trancl"} and @{text "M_datatypes"}*}
 | 
| 13348 | 12 | |
| 13363 | 13 | lemma eq_succ_imp_lt: "[|i = succ(j); Ord(i)|] ==> j<i" | 
| 13428 | 14 | by simp | 
| 13363 | 15 | |
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 16 | |
| 13348 | 17 | subsection{*The Locale @{text "M_trancl"}*}
 | 
| 18 | ||
| 19 | subsubsection{*Separation for Reflexive/Transitive Closure*}
 | |
| 20 | ||
| 21 | text{*First, The Defining Formula*}
 | |
| 22 | ||
| 23 | (* "rtran_closure_mem(M,A,r,p) == | |
| 13428 | 24 | \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M]. | 
| 13348 | 25 | omega(M,nnat) & n\<in>nnat & successor(M,n,n') & | 
| 26 | (\<exists>f[M]. typed_function(M,n',A,f) & | |
| 13428 | 27 | (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) & | 
| 28 | fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) & | |
| 29 | (\<forall>j[M]. j\<in>n --> | |
| 30 | (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M]. | |
| 31 | fun_apply(M,f,j,fj) & successor(M,j,sj) & | |
| 32 | fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"*) | |
| 13348 | 33 | constdefs rtran_closure_mem_fm :: "[i,i,i]=>i" | 
| 13428 | 34 | "rtran_closure_mem_fm(A,r,p) == | 
| 13348 | 35 | Exists(Exists(Exists( | 
| 36 | And(omega_fm(2), | |
| 37 | And(Member(1,2), | |
| 38 | And(succ_fm(1,0), | |
| 39 | Exists(And(typed_function_fm(1, A#+4, 0), | |
| 13428 | 40 | And(Exists(Exists(Exists( | 
| 41 | And(pair_fm(2,1,p#+7), | |
| 42 | And(empty_fm(0), | |
| 43 | And(fun_apply_fm(3,0,2), fun_apply_fm(3,5,1))))))), | |
| 44 | Forall(Implies(Member(0,3), | |
| 45 | Exists(Exists(Exists(Exists( | |
| 46 | And(fun_apply_fm(5,4,3), | |
| 47 | And(succ_fm(4,2), | |
| 48 | And(fun_apply_fm(5,2,1), | |
| 49 | And(pair_fm(3,1,0), Member(0,r#+9))))))))))))))))))))" | |
| 13348 | 50 | |
| 51 | ||
| 52 | lemma rtran_closure_mem_type [TC]: | |
| 53 | "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> rtran_closure_mem_fm(x,y,z) \<in> formula" | |
| 13428 | 54 | by (simp add: rtran_closure_mem_fm_def) | 
| 13348 | 55 | |
| 56 | lemma sats_rtran_closure_mem_fm [simp]: | |
| 57 | "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|] | |
| 13428 | 58 | ==> sats(A, rtran_closure_mem_fm(x,y,z), env) <-> | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 59 | rtran_closure_mem(##A, nth(x,env), nth(y,env), nth(z,env))" | 
| 13348 | 60 | by (simp add: rtran_closure_mem_fm_def rtran_closure_mem_def) | 
| 61 | ||
| 62 | lemma rtran_closure_mem_iff_sats: | |
| 13428 | 63 | "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; | 
| 13348 | 64 | i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|] | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 65 | ==> rtran_closure_mem(##A, x, y, z) <-> sats(A, rtran_closure_mem_fm(i,j,k), env)" | 
| 13348 | 66 | by (simp add: sats_rtran_closure_mem_fm) | 
| 67 | ||
| 13566 | 68 | lemma rtran_closure_mem_reflection: | 
| 13428 | 69 | "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: 
13687diff
changeset | 70 | \<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: 
13651diff
changeset | 71 | apply (simp only: rtran_closure_mem_def) | 
| 13428 | 72 | apply (intro FOL_reflections function_reflections fun_plus_reflections) | 
| 13348 | 73 | done | 
| 74 | ||
| 75 | text{*Separation for @{term "rtrancl(r)"}.*}
 | |
| 76 | lemma rtrancl_separation: | |
| 77 | "[| L(r); L(A) |] ==> separation (L, rtran_closure_mem(L,A,r))" | |
| 13687 | 78 | apply (rule gen_separation_multi [OF rtran_closure_mem_reflection, of "{r,A}"],
 | 
| 79 | auto) | |
| 80 | apply (rule_tac env="[r,A]" in DPow_LsetI) | |
| 81 | apply (rule rtran_closure_mem_iff_sats sep_rules | simp)+ | |
| 13348 | 82 | done | 
| 83 | ||
| 84 | ||
| 85 | subsubsection{*Reflexive/Transitive Closure, Internalized*}
 | |
| 86 | ||
| 13428 | 87 | (* "rtran_closure(M,r,s) == | 
| 13348 | 88 | \<forall>A[M]. is_field(M,r,A) --> | 
| 13428 | 89 | (\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))" *) | 
| 13348 | 90 | constdefs rtran_closure_fm :: "[i,i]=>i" | 
| 13428 | 91 | "rtran_closure_fm(r,s) == | 
| 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]: | |
| 97 | "[| x \<in> nat; y \<in> nat |] ==> 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]: | |
| 101 | "[| x \<in> nat; y \<in> nat; env \<in> list(A)|] | |
| 13428 | 102 | ==> sats(A, rtran_closure_fm(x,y), env) <-> | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
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: | |
| 13428 | 107 | "[| nth(i,env) = x; nth(j,env) = y; | 
| 13348 | 108 | i \<in> nat; j \<in> nat; env \<in> list(A)|] | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 109 | ==> rtran_closure(##A, x, y) <-> 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: 
13687diff
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: 
13651diff
changeset | 115 | apply (simp only: rtran_closure_def) | 
| 13348 | 116 | apply (intro FOL_reflections function_reflections rtran_closure_mem_reflection) | 
| 117 | done | |
| 118 | ||
| 119 | ||
| 120 | subsubsection{*Transitive Closure of a Relation, Internalized*}
 | |
| 121 | ||
| 122 | (* "tran_closure(M,r,t) == | |
| 123 | \<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" *) | |
| 124 | constdefs tran_closure_fm :: "[i,i]=>i" | |
| 13428 | 125 | "tran_closure_fm(r,s) == | 
| 13348 | 126 | Exists(And(rtran_closure_fm(succ(r),0), composition_fm(succ(r),0,succ(s))))" | 
| 127 | ||
| 128 | lemma tran_closure_type [TC]: | |
| 129 | "[| x \<in> nat; y \<in> nat |] ==> tran_closure_fm(x,y) \<in> formula" | |
| 13428 | 130 | by (simp add: tran_closure_fm_def) | 
| 13348 | 131 | |
| 132 | lemma sats_tran_closure_fm [simp]: | |
| 133 | "[| x \<in> nat; y \<in> nat; env \<in> list(A)|] | |
| 13428 | 134 | ==> sats(A, tran_closure_fm(x,y), env) <-> | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 135 | tran_closure(##A, nth(x,env), nth(y,env))" | 
| 13348 | 136 | by (simp add: tran_closure_fm_def tran_closure_def) | 
| 137 | ||
| 138 | lemma tran_closure_iff_sats: | |
| 13428 | 139 | "[| nth(i,env) = x; nth(j,env) = y; | 
| 13348 | 140 | i \<in> nat; j \<in> nat; env \<in> list(A)|] | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 141 | ==> tran_closure(##A, x, y) <-> sats(A, tran_closure_fm(i,j), env)" | 
| 13348 | 142 | by simp | 
| 143 | ||
| 144 | theorem tran_closure_reflection: | |
| 13428 | 145 | "REFLECTS[\<lambda>x. tran_closure(L,f(x),g(x)), | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 146 | \<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: 
13651diff
changeset | 147 | apply (simp only: tran_closure_def) | 
| 13428 | 148 | apply (intro FOL_reflections function_reflections | 
| 13348 | 149 | rtran_closure_reflection composition_reflection) | 
| 150 | done | |
| 151 | ||
| 152 | ||
| 13506 | 153 | subsubsection{*Separation for the Proof of @{text "wellfounded_on_trancl"}*}
 | 
| 13348 | 154 | |
| 155 | lemma wellfounded_trancl_reflects: | |
| 13428 | 156 | "REFLECTS[\<lambda>x. \<exists>w[L]. \<exists>wx[L]. \<exists>rp[L]. | 
| 157 | w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp, | |
| 158 | \<lambda>i x. \<exists>w \<in> Lset(i). \<exists>wx \<in> Lset(i). \<exists>rp \<in> Lset(i). | |
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 159 | w \<in> Z & pair(##Lset(i),w,x,wx) & tran_closure(##Lset(i),r,rp) & | 
| 13348 | 160 | wx \<in> rp]" | 
| 13428 | 161 | by (intro FOL_reflections function_reflections fun_plus_reflections | 
| 13348 | 162 | tran_closure_reflection) | 
| 163 | ||
| 164 | lemma wellfounded_trancl_separation: | |
| 13428 | 165 | "[| L(r); L(Z) |] ==> | 
| 166 | separation (L, \<lambda>x. | |
| 167 | \<exists>w[L]. \<exists>wx[L]. \<exists>rp[L]. | |
| 168 | w \<in> Z & pair(L,w,x,wx) & tran_closure(L,r,rp) & wx \<in> rp)" | |
| 13687 | 169 | apply (rule gen_separation_multi [OF wellfounded_trancl_reflects, of "{r,Z}"],
 | 
| 170 | auto) | |
| 171 | apply (rule_tac env="[r,Z]" in DPow_LsetI) | |
| 13348 | 172 | apply (rule sep_rules tran_closure_iff_sats | simp)+ | 
| 173 | done | |
| 174 | ||
| 13363 | 175 | |
| 176 | subsubsection{*Instantiating the locale @{text M_trancl}*}
 | |
| 13428 | 177 | |
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 178 | lemma M_trancl_axioms_L: "M_trancl_axioms(L)" | 
| 13428 | 179 | apply (rule M_trancl_axioms.intro) | 
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 180 | apply (assumption | rule rtrancl_separation wellfounded_trancl_separation)+ | 
| 13428 | 181 | done | 
| 13363 | 182 | |
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 183 | theorem M_trancl_L: "PROP M_trancl(L)" | 
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 184 | by (rule M_trancl.intro | 
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13506diff
changeset | 185 | [OF M_trivial_L M_basic_axioms_L M_trancl_axioms_L]) | 
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 186 | |
| 15766 | 187 | interpretation M_trancl [L] by (rule M_trancl_axioms_L) | 
| 13363 | 188 | |
| 189 | ||
| 13428 | 190 | subsection{*@{term L} is Closed Under the Operator @{term list}*}
 | 
| 13363 | 191 | |
| 13386 | 192 | subsubsection{*Instances of Replacement for Lists*}
 | 
| 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: 
13687diff
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: 
13687diff
changeset | 199 | is_wfrec(##Lset(i), | 
| 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 200 | iterates_MH(##Lset(i), | 
| 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
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: | 
| 13363 | 207 | "L(A) ==> iterates_replacement(L, is_list_functor(L,A), 0)" | 
| 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], | 
| 212 | auto simp add: nonempty) | |
| 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 | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 221 | [\<lambda>x. \<exists>u[L]. u \<in> B & u \<in> nat & | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 222 | is_iterates(L, is_list_functor(L, A), 0, u, x), | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 223 | \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & u \<in> nat & | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
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: 
13651diff
changeset | 225 | by (intro FOL_reflections | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 226 | is_iterates_reflection list_functor_reflection) | 
| 13363 | 227 | |
| 13428 | 228 | lemma list_replacement2: | 
| 229 | "L(A) ==> strong_replacement(L, | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 230 | \<lambda>n y. n\<in>nat & 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], | 
| 234 | auto simp add: L_nat nonempty) | |
| 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: 
13651diff
changeset | 236 | apply (rule sep_rules list_functor_iff_sats is_iterates_iff_sats | simp)+ | 
| 13363 | 237 | done | 
| 238 | ||
| 13386 | 239 | |
| 13428 | 240 | subsection{*@{term L} is Closed Under the Operator @{term formula}*}
 | 
| 13386 | 241 | |
| 242 | subsubsection{*Instances of Replacement for Formulas*}
 | |
| 243 | ||
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
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: 
13651diff
changeset | 245 | need to expand iterates_replacement and wfrec_replacement*) | 
| 13386 | 246 | lemma formula_replacement1_Reflects: | 
| 247 | "REFLECTS | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 248 | [\<lambda>x. \<exists>u[L]. u \<in> B & (\<exists>y[L]. pair(L,u,y,x) & | 
| 13386 | 249 | is_wfrec(L, iterates_MH(L, is_formula_functor(L), 0), memsn, u, y)), | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 250 | \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) & | 
| 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 251 | is_wfrec(##Lset(i), | 
| 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 252 | iterates_MH(##Lset(i), | 
| 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
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], | 
| 263 | auto simp add: nonempty) | |
| 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 | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 271 | [\<lambda>x. \<exists>u[L]. u \<in> B & u \<in> nat & | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 272 | is_iterates(L, is_formula_functor(L), 0, u, x), | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 273 | \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & u \<in> nat & | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
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: 
13651diff
changeset | 275 | by (intro FOL_reflections | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 276 | is_iterates_reflection formula_functor_reflection) | 
| 13386 | 277 | |
| 13428 | 278 | lemma formula_replacement2: | 
| 279 | "strong_replacement(L, | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 280 | \<lambda>n y. n\<in>nat & 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], | 
| 284 | auto simp add: nonempty L_nat) | |
| 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: 
13651diff
changeset | 286 | apply (rule sep_rules formula_functor_iff_sats is_iterates_iff_sats | simp)+ | 
| 13386 | 287 | done | 
| 288 | ||
| 289 | text{*NB The proofs for type @{term formula} are virtually identical to those
 | |
| 290 | for @{term "list(A)"}.  It was a cut-and-paste job! *}
 | |
| 291 | ||
| 13387 | 292 | |
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 293 | subsubsection{*The Formula @{term is_nth}, Internalized*}
 | 
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 294 | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 295 | (* "is_nth(M,n,l,Z) == | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 296 | \<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,Z)" *) | 
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 297 | constdefs nth_fm :: "[i,i,i]=>i" | 
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 298 | "nth_fm(n,l,Z) == | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 299 | 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: 
13651diff
changeset | 300 | hd_fm(0,succ(Z))))" | 
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 301 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 302 | lemma nth_fm_type [TC]: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 303 | "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> nth_fm(x,y,z) \<in> formula" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 304 | by (simp add: nth_fm_def) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 305 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 306 | lemma sats_nth_fm [simp]: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 307 | "[| x < length(env); y \<in> nat; z \<in> nat; env \<in> list(A)|] | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 308 | ==> sats(A, nth_fm(x,y,z), env) <-> | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 309 | 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: 
13441diff
changeset | 310 | apply (frule lt_length_in_nat, assumption) | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 311 | 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: 
13441diff
changeset | 312 | done | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 313 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 314 | lemma nth_iff_sats: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 315 | "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 316 | i < length(env); j \<in> nat; k \<in> nat; env \<in> list(A)|] | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 317 | ==> is_nth(##A, x, y, z) <-> sats(A, nth_fm(i,j,k), env)" | 
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13441diff
changeset | 318 | by (simp add: sats_nth_fm) | 
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 319 | |
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 320 | theorem nth_reflection: | 
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 321 | "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: 
13687diff
changeset | 322 | \<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: 
13651diff
changeset | 323 | apply (simp only: is_nth_def) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 324 | apply (intro FOL_reflections is_iterates_reflection | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 325 | hd_reflection tl_reflection) | 
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 326 | done | 
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 327 | |
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 328 | |
| 13409 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 329 | subsubsection{*An Instance of Replacement for @{term nth}*}
 | 
| 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 330 | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 331 | (*FIXME: could prove a lemma iterates_replacementI to eliminate the | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 332 | need to expand iterates_replacement and wfrec_replacement*) | 
| 13409 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 333 | lemma nth_replacement_Reflects: | 
| 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 334 | "REFLECTS | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 335 | [\<lambda>x. \<exists>u[L]. u \<in> B & (\<exists>y[L]. pair(L,u,y,x) & | 
| 13409 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 336 | is_wfrec(L, iterates_MH(L, is_tl(L), z), memsn, u, y)), | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 337 | \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) & | 
| 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 338 | is_wfrec(##Lset(i), | 
| 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 339 | iterates_MH(##Lset(i), | 
| 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 340 | is_tl(##Lset(i)), z), memsn, u, y))]" | 
| 13428 | 341 | by (intro FOL_reflections function_reflections is_wfrec_reflection | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 342 | iterates_MH_reflection tl_reflection) | 
| 13409 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 343 | |
| 13428 | 344 | lemma nth_replacement: | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 345 | "L(w) ==> iterates_replacement(L, is_tl(L), w)" | 
| 13409 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 346 | apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) | 
| 13428 | 347 | apply (rule strong_replacementI) | 
| 13687 | 348 | apply (rule_tac u="{B,w,Memrel(succ(n))}" 
 | 
| 349 | in gen_separation_multi [OF nth_replacement_Reflects], | |
| 350 | auto) | |
| 351 | apply (rule_tac env="[B,w,Memrel(succ(n))]" in DPow_LsetI) | |
| 13434 | 352 | apply (rule sep_rules is_nat_case_iff_sats tl_iff_sats | 
| 13441 | 353 | is_wfrec_iff_sats iterates_MH_iff_sats quasinat_iff_sats | simp)+ | 
| 13409 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 354 | done | 
| 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 355 | |
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 356 | |
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 357 | subsubsection{*Instantiating the locale @{text M_datatypes}*}
 | 
| 13428 | 358 | |
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 359 | lemma M_datatypes_axioms_L: "M_datatypes_axioms(L)" | 
| 13428 | 360 | apply (rule M_datatypes_axioms.intro) | 
| 361 | apply (assumption | rule | |
| 362 | list_replacement1 list_replacement2 | |
| 363 | formula_replacement1 formula_replacement2 | |
| 364 | nth_replacement)+ | |
| 365 | done | |
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 366 | |
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 367 | theorem M_datatypes_L: "PROP M_datatypes(L)" | 
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 368 | apply (rule M_datatypes.intro) | 
| 13634 | 369 | apply (rule M_trancl.axioms [OF M_trancl_L])+ | 
| 13441 | 370 | apply (rule M_datatypes_axioms_L) | 
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 371 | done | 
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 372 | |
| 15766 | 373 | interpretation M_datatypes [L] by (rule M_datatypes_axioms_L) | 
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 374 | |
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 375 | |
| 13428 | 376 | subsection{*@{term L} is Closed Under the Operator @{term eclose}*}
 | 
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 377 | |
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 378 | subsubsection{*Instances of Replacement for @{term eclose}*}
 | 
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 379 | |
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 380 | lemma eclose_replacement1_Reflects: | 
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 381 | "REFLECTS | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 382 | [\<lambda>x. \<exists>u[L]. u \<in> B & (\<exists>y[L]. pair(L,u,y,x) & | 
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 383 | is_wfrec(L, iterates_MH(L, big_union(L), A), memsn, u, y)), | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 384 | \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & (\<exists>y \<in> Lset(i). pair(##Lset(i), u, y, x) & | 
| 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 385 | is_wfrec(##Lset(i), | 
| 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 386 | iterates_MH(##Lset(i), big_union(##Lset(i)), A), | 
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 387 | memsn, u, y))]" | 
| 13428 | 388 | by (intro FOL_reflections function_reflections is_wfrec_reflection | 
| 389 | iterates_MH_reflection) | |
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 390 | |
| 13428 | 391 | lemma eclose_replacement1: | 
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 392 | "L(A) ==> iterates_replacement(L, big_union(L), A)" | 
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 393 | apply (unfold iterates_replacement_def wfrec_replacement_def, clarify) | 
| 13428 | 394 | apply (rule strong_replacementI) | 
| 13566 | 395 | apply (rule_tac u="{B,A,n,Memrel(succ(n))}" 
 | 
| 13687 | 396 | in gen_separation_multi [OF eclose_replacement1_Reflects], auto) | 
| 397 | apply (rule_tac env="[B,A,n,Memrel(succ(n))]" in DPow_LsetI) | |
| 13434 | 398 | apply (rule sep_rules iterates_MH_iff_sats is_nat_case_iff_sats | 
| 13441 | 399 | is_wfrec_iff_sats big_union_iff_sats quasinat_iff_sats | simp)+ | 
| 13409 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 400 | done | 
| 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 401 | |
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 402 | |
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 403 | lemma eclose_replacement2_Reflects: | 
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 404 | "REFLECTS | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 405 | [\<lambda>x. \<exists>u[L]. u \<in> B & u \<in> nat & | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 406 | is_iterates(L, big_union(L), A, u, x), | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 407 | \<lambda>i x. \<exists>u \<in> Lset(i). u \<in> B & u \<in> nat & | 
| 13807 
a28a8fbc76d4
changed ** to ## to avoid conflict with new comment syntax
 paulson parents: 
13687diff
changeset | 408 | 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: 
13651diff
changeset | 409 | by (intro FOL_reflections function_reflections is_iterates_reflection) | 
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 410 | |
| 13428 | 411 | lemma eclose_replacement2: | 
| 412 | "L(A) ==> strong_replacement(L, | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 413 | \<lambda>n y. n\<in>nat & is_iterates(L, big_union(L), A, n, y))" | 
| 13428 | 414 | apply (rule strong_replacementI) | 
| 13566 | 415 | apply (rule_tac u="{A,B,nat}" 
 | 
| 13687 | 416 | in gen_separation_multi [OF eclose_replacement2_Reflects], | 
| 417 | auto simp add: L_nat) | |
| 418 | apply (rule_tac env="[A,B,nat]" in DPow_LsetI) | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13651diff
changeset | 419 | 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: 
13418diff
changeset | 420 | done | 
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 421 | |
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 422 | |
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 423 | subsubsection{*Instantiating the locale @{text M_eclose}*}
 | 
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 424 | |
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 425 | lemma M_eclose_axioms_L: "M_eclose_axioms(L)" | 
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 426 | apply (rule M_eclose_axioms.intro) | 
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 427 | apply (assumption | rule eclose_replacement1 eclose_replacement2)+ | 
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 428 | done | 
| 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 429 | |
| 13428 | 430 | theorem M_eclose_L: "PROP M_eclose(L)" | 
| 431 | apply (rule M_eclose.intro) | |
| 13429 | 432 | apply (rule M_datatypes.axioms [OF M_datatypes_L])+ | 
| 13437 
01b3fc0cc1b8
separate "axioms" proofs: more flexible for locale reasoning
 paulson parents: 
13434diff
changeset | 433 | apply (rule M_eclose_axioms_L) | 
| 13428 | 434 | done | 
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 435 | |
| 15766 | 436 | interpretation M_eclose [L] by (rule M_eclose_axioms_L) | 
| 437 | ||
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 438 | |
| 13348 | 439 | end |