| author | nipkow | 
| Sun, 03 Jun 2007 12:58:28 +0200 | |
| changeset 23209 | 098a23702aba | 
| parent 22710 | f44439cdce77 | 
| child 32960 | 69916a850301 | 
| permissions | -rw-r--r-- | 
| 13505 | 1 | (* Title: ZF/Constructible/Datatype_absolute.thy | 
| 2 | ID: $Id$ | |
| 3 | Author: Lawrence C Paulson, Cambridge University Computer Laboratory | |
| 4 | *) | |
| 5 | ||
| 13306 | 6 | header {*Absoluteness Properties for Recursive Datatypes*}
 | 
| 7 | ||
| 16417 | 8 | theory Datatype_absolute imports Formula WF_absolute begin | 
| 13268 | 9 | |
| 10 | ||
| 11 | subsection{*The lfp of a continuous function can be expressed as a union*}
 | |
| 12 | ||
| 21233 | 13 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 14 | directed :: "i=>o" where | 
| 13385 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 15 | "directed(A) == A\<noteq>0 & (\<forall>x\<in>A. \<forall>y\<in>A. x \<union> y \<in> A)" | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 16 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 17 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 18 | contin :: "(i=>i) => o" where | 
| 13385 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 19 | "contin(h) == (\<forall>A. directed(A) --> h(\<Union>A) = (\<Union>X\<in>A. h(X)))" | 
| 13268 | 20 | |
| 21 | lemma bnd_mono_iterates_subset: "[|bnd_mono(D, h); n \<in> nat|] ==> h^n (0) <= D" | |
| 22 | apply (induct_tac n) | |
| 23 | apply (simp_all add: bnd_mono_def, blast) | |
| 24 | done | |
| 25 | ||
| 13385 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 26 | lemma bnd_mono_increasing [rule_format]: | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 27 | "[|i \<in> nat; j \<in> nat; bnd_mono(D,h)|] ==> i \<le> j --> h^i(0) \<subseteq> h^j(0)" | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 28 | apply (rule_tac m=i and n=j in diff_induct, simp_all) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 29 | apply (blast del: subsetI | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 30 | intro: bnd_mono_iterates_subset bnd_monoD2 [of concl: h]) | 
| 13385 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 31 | done | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 32 | |
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 33 | lemma directed_iterates: "bnd_mono(D,h) ==> directed({h^n (0). n\<in>nat})"
 | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 34 | apply (simp add: directed_def, clarify) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 35 | apply (rename_tac i j) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 36 | apply (rule_tac x="i \<union> j" in bexI) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 37 | apply (rule_tac i = i and j = j in Ord_linear_le) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 38 | apply (simp_all add: subset_Un_iff [THEN iffD1] le_imp_subset | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 39 | subset_Un_iff2 [THEN iffD1]) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 40 | apply (simp_all add: subset_Un_iff [THEN iff_sym] bnd_mono_increasing | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 41 | subset_Un_iff2 [THEN iff_sym]) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 42 | done | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 43 | |
| 13268 | 44 | |
| 45 | lemma contin_iterates_eq: | |
| 13385 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 46 | "[|bnd_mono(D, h); contin(h)|] | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 47 | ==> h(\<Union>n\<in>nat. h^n (0)) = (\<Union>n\<in>nat. h^n (0))" | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 48 | apply (simp add: contin_def directed_iterates) | 
| 13268 | 49 | apply (rule trans) | 
| 50 | apply (rule equalityI) | |
| 13385 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 51 | apply (simp_all add: UN_subset_iff) | 
| 13268 | 52 | apply safe | 
| 53 | apply (erule_tac [2] natE) | |
| 54 | apply (rule_tac a="succ(x)" in UN_I) | |
| 55 | apply simp_all | |
| 56 | apply blast | |
| 57 | done | |
| 58 | ||
| 59 | lemma lfp_subset_Union: | |
| 60 | "[|bnd_mono(D, h); contin(h)|] ==> lfp(D,h) <= (\<Union>n\<in>nat. h^n(0))" | |
| 61 | apply (rule lfp_lowerbound) | |
| 62 | apply (simp add: contin_iterates_eq) | |
| 63 | apply (simp add: contin_def bnd_mono_iterates_subset UN_subset_iff) | |
| 64 | done | |
| 65 | ||
| 66 | lemma Union_subset_lfp: | |
| 67 | "bnd_mono(D,h) ==> (\<Union>n\<in>nat. h^n(0)) <= lfp(D,h)" | |
| 68 | apply (simp add: UN_subset_iff) | |
| 69 | apply (rule ballI) | |
| 13339 
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
 paulson parents: 
13306diff
changeset | 70 | apply (induct_tac n, simp_all) | 
| 13268 | 71 | apply (rule subset_trans [of _ "h(lfp(D,h))"]) | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 72 | apply (blast dest: bnd_monoD2 [OF _ _ lfp_subset]) | 
| 13268 | 73 | apply (erule lfp_lemma2) | 
| 74 | done | |
| 75 | ||
| 76 | lemma lfp_eq_Union: | |
| 77 | "[|bnd_mono(D, h); contin(h)|] ==> lfp(D,h) = (\<Union>n\<in>nat. h^n(0))" | |
| 78 | by (blast del: subsetI | |
| 79 | intro: lfp_subset_Union Union_subset_lfp) | |
| 80 | ||
| 81 | ||
| 13385 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 82 | subsubsection{*Some Standard Datatype Constructions Preserve Continuity*}
 | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 83 | |
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 84 | lemma contin_imp_mono: "[|X\<subseteq>Y; contin(F)|] ==> F(X) \<subseteq> F(Y)" | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 85 | apply (simp add: contin_def) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 86 | apply (drule_tac x="{X,Y}" in spec) 
 | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 87 | apply (simp add: directed_def subset_Un_iff2 Un_commute) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 88 | done | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 89 | |
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 90 | lemma sum_contin: "[|contin(F); contin(G)|] ==> contin(\<lambda>X. F(X) + G(X))" | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 91 | by (simp add: contin_def, blast) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 92 | |
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 93 | lemma prod_contin: "[|contin(F); contin(G)|] ==> contin(\<lambda>X. F(X) * G(X))" | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 94 | apply (subgoal_tac "\<forall>B C. F(B) \<subseteq> F(B \<union> C)") | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 95 | prefer 2 apply (simp add: Un_upper1 contin_imp_mono) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 96 | apply (subgoal_tac "\<forall>B C. G(C) \<subseteq> G(B \<union> C)") | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 97 | prefer 2 apply (simp add: Un_upper2 contin_imp_mono) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 98 | apply (simp add: contin_def, clarify) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 99 | apply (rule equalityI) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 100 | prefer 2 apply blast | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 101 | apply clarify | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 102 | apply (rename_tac B C) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 103 | apply (rule_tac a="B \<union> C" in UN_I) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 104 | apply (simp add: directed_def, blast) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 105 | done | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 106 | |
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 107 | lemma const_contin: "contin(\<lambda>X. A)" | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 108 | by (simp add: contin_def directed_def) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 109 | |
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 110 | lemma id_contin: "contin(\<lambda>X. X)" | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 111 | by (simp add: contin_def) | 
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 112 | |
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 113 | |
| 
31df66ca0780
Expressing Lset and L without using length and arity; simplifies Separation
 paulson parents: 
13382diff
changeset | 114 | |
| 13268 | 115 | subsection {*Absoluteness for "Iterates"*}
 | 
| 116 | ||
| 21233 | 117 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 118 | iterates_MH :: "[i=>o, [i,i]=>o, i, i, i, i] => o" where | 
| 13353 | 119 | "iterates_MH(M,isF,v,n,g,z) == | 
| 120 | is_nat_case(M, v, \<lambda>m u. \<exists>gm[M]. fun_apply(M,g,m,gm) & isF(gm,u), | |
| 121 | n, z)" | |
| 122 | ||
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 123 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 124 | is_iterates :: "[i=>o, [i,i]=>o, i, i, i] => o" where | 
| 13687 | 125 | "is_iterates(M,isF,v,n,Z) == | 
| 126 | \<exists>sn[M]. \<exists>msn[M]. successor(M,n,sn) & membership(M,sn,msn) & | |
| 127 | is_wfrec(M, iterates_MH(M,isF,v), msn, n, Z)" | |
| 128 | ||
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 129 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 130 | iterates_replacement :: "[i=>o, [i,i]=>o, i] => o" where | 
| 13353 | 131 | "iterates_replacement(M,isF,v) == | 
| 13363 | 132 | \<forall>n[M]. n\<in>nat --> | 
| 13353 | 133 | wfrec_replacement(M, iterates_MH(M,isF,v), Memrel(succ(n)))" | 
| 134 | ||
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 135 | lemma (in M_basic) iterates_MH_abs: | 
| 13634 | 136 | "[| relation1(M,isF,F); M(n); M(g); M(z) |] | 
| 13353 | 137 | ==> iterates_MH(M,isF,v,n,g,z) <-> z = nat_case(v, \<lambda>m. F(g`m), n)" | 
| 13363 | 138 | by (simp add: nat_case_abs [of _ "\<lambda>m. F(g ` m)"] | 
| 13634 | 139 | relation1_def iterates_MH_def) | 
| 13353 | 140 | |
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 141 | lemma (in M_basic) iterates_imp_wfrec_replacement: | 
| 13634 | 142 | "[|relation1(M,isF,F); n \<in> nat; iterates_replacement(M,isF,v)|] | 
| 13353 | 143 | ==> wfrec_replacement(M, \<lambda>n f z. z = nat_case(v, \<lambda>m. F(f`m), n), | 
| 144 | Memrel(succ(n)))" | |
| 145 | by (simp add: iterates_replacement_def iterates_MH_abs) | |
| 146 | ||
| 147 | theorem (in M_trancl) iterates_abs: | |
| 13634 | 148 | "[| iterates_replacement(M,isF,v); relation1(M,isF,F); | 
| 13353 | 149 | n \<in> nat; M(v); M(z); \<forall>x[M]. M(F(x)) |] | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 150 | ==> is_iterates(M,isF,v,n,z) <-> z = iterates(F,n,v)" | 
| 13353 | 151 | apply (frule iterates_imp_wfrec_replacement, assumption+) | 
| 152 | apply (simp add: wf_Memrel trans_Memrel relation_Memrel nat_into_M | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 153 | is_iterates_def relation2_def iterates_MH_abs | 
| 13353 | 154 | iterates_nat_def recursor_def transrec_def | 
| 155 | eclose_sing_Ord_eq nat_into_M | |
| 156 | trans_wfrec_abs [of _ _ _ _ "\<lambda>n g. nat_case(v, \<lambda>m. F(g`m), n)"]) | |
| 157 | done | |
| 158 | ||
| 13268 | 159 | |
| 13634 | 160 | lemma (in M_trancl) iterates_closed [intro,simp]: | 
| 161 | "[| iterates_replacement(M,isF,v); relation1(M,isF,F); | |
| 13353 | 162 | n \<in> nat; M(v); \<forall>x[M]. M(F(x)) |] | 
| 13268 | 163 | ==> M(iterates(F,n,v))" | 
| 13353 | 164 | apply (frule iterates_imp_wfrec_replacement, assumption+) | 
| 165 | apply (simp add: wf_Memrel trans_Memrel relation_Memrel nat_into_M | |
| 13634 | 166 | relation2_def iterates_MH_abs | 
| 13353 | 167 | iterates_nat_def recursor_def transrec_def | 
| 168 | eclose_sing_Ord_eq nat_into_M | |
| 169 | trans_wfrec_closed [of _ _ _ "\<lambda>n g. nat_case(v, \<lambda>m. F(g`m), n)"]) | |
| 170 | done | |
| 13268 | 171 | |
| 172 | ||
| 13386 | 173 | subsection {*lists without univ*}
 | 
| 174 | ||
| 175 | lemmas datatype_univs = Inl_in_univ Inr_in_univ | |
| 176 | Pair_in_univ nat_into_univ A_into_univ | |
| 177 | ||
| 178 | lemma list_fun_bnd_mono: "bnd_mono(univ(A), \<lambda>X. {0} + A*X)"
 | |
| 179 | apply (rule bnd_monoI) | |
| 180 | apply (intro subset_refl zero_subset_univ A_subset_univ | |
| 181 | sum_subset_univ Sigma_subset_univ) | |
| 182 | apply (rule subset_refl sum_mono Sigma_mono | assumption)+ | |
| 183 | done | |
| 184 | ||
| 185 | lemma list_fun_contin: "contin(\<lambda>X. {0} + A*X)"
 | |
| 186 | by (intro sum_contin prod_contin id_contin const_contin) | |
| 187 | ||
| 188 | text{*Re-expresses lists using sum and product*}
 | |
| 189 | lemma list_eq_lfp2: "list(A) = lfp(univ(A), \<lambda>X. {0} + A*X)"
 | |
| 190 | apply (simp add: list_def) | |
| 191 | apply (rule equalityI) | |
| 192 | apply (rule lfp_lowerbound) | |
| 193 | prefer 2 apply (rule lfp_subset) | |
| 194 | apply (clarify, subst lfp_unfold [OF list_fun_bnd_mono]) | |
| 195 | apply (simp add: Nil_def Cons_def) | |
| 196 | apply blast | |
| 197 | txt{*Opposite inclusion*}
 | |
| 198 | apply (rule lfp_lowerbound) | |
| 199 | prefer 2 apply (rule lfp_subset) | |
| 200 | apply (clarify, subst lfp_unfold [OF list.bnd_mono]) | |
| 201 | apply (simp add: Nil_def Cons_def) | |
| 202 | apply (blast intro: datatype_univs | |
| 203 | dest: lfp_subset [THEN subsetD]) | |
| 204 | done | |
| 205 | ||
| 206 | text{*Re-expresses lists using "iterates", no univ.*}
 | |
| 207 | lemma list_eq_Union: | |
| 208 |      "list(A) = (\<Union>n\<in>nat. (\<lambda>X. {0} + A*X) ^ n (0))"
 | |
| 209 | by (simp add: list_eq_lfp2 lfp_eq_Union list_fun_bnd_mono list_fun_contin) | |
| 210 | ||
| 211 | ||
| 21233 | 212 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 213 | is_list_functor :: "[i=>o,i,i,i] => o" where | 
| 13350 | 214 | "is_list_functor(M,A,X,Z) == | 
| 215 | \<exists>n1[M]. \<exists>AX[M]. | |
| 216 | number1(M,n1) & cartprod(M,A,X,AX) & is_sum(M,n1,AX,Z)" | |
| 217 | ||
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 218 | lemma (in M_basic) list_functor_abs [simp]: | 
| 13350 | 219 |      "[| M(A); M(X); M(Z) |] ==> is_list_functor(M,A,X,Z) <-> (Z = {0} + A*X)"
 | 
| 220 | by (simp add: is_list_functor_def singleton_0 nat_into_M) | |
| 221 | ||
| 222 | ||
| 13386 | 223 | subsection {*formulas without univ*}
 | 
| 224 | ||
| 225 | lemma formula_fun_bnd_mono: | |
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 226 | "bnd_mono(univ(0), \<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X))" | 
| 13386 | 227 | apply (rule bnd_monoI) | 
| 228 | apply (intro subset_refl zero_subset_univ A_subset_univ | |
| 229 | sum_subset_univ Sigma_subset_univ nat_subset_univ) | |
| 230 | apply (rule subset_refl sum_mono Sigma_mono | assumption)+ | |
| 231 | done | |
| 232 | ||
| 233 | lemma formula_fun_contin: | |
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 234 | "contin(\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X))" | 
| 13386 | 235 | by (intro sum_contin prod_contin id_contin const_contin) | 
| 236 | ||
| 237 | ||
| 238 | text{*Re-expresses formulas using sum and product*}
 | |
| 239 | lemma formula_eq_lfp2: | |
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 240 | "formula = lfp(univ(0), \<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X))" | 
| 13386 | 241 | apply (simp add: formula_def) | 
| 242 | apply (rule equalityI) | |
| 243 | apply (rule lfp_lowerbound) | |
| 244 | prefer 2 apply (rule lfp_subset) | |
| 245 | apply (clarify, subst lfp_unfold [OF formula_fun_bnd_mono]) | |
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 246 | apply (simp add: Member_def Equal_def Nand_def Forall_def) | 
| 13386 | 247 | apply blast | 
| 248 | txt{*Opposite inclusion*}
 | |
| 249 | apply (rule lfp_lowerbound) | |
| 250 | prefer 2 apply (rule lfp_subset, clarify) | |
| 251 | apply (subst lfp_unfold [OF formula.bnd_mono, simplified]) | |
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 252 | apply (simp add: Member_def Equal_def Nand_def Forall_def) | 
| 13386 | 253 | apply (elim sumE SigmaE, simp_all) | 
| 254 | apply (blast intro: datatype_univs dest: lfp_subset [THEN subsetD])+ | |
| 255 | done | |
| 256 | ||
| 257 | text{*Re-expresses formulas using "iterates", no univ.*}
 | |
| 258 | lemma formula_eq_Union: | |
| 259 | "formula = | |
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 260 | (\<Union>n\<in>nat. (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X)) ^ n (0))" | 
| 13386 | 261 | by (simp add: formula_eq_lfp2 lfp_eq_Union formula_fun_bnd_mono | 
| 262 | formula_fun_contin) | |
| 263 | ||
| 264 | ||
| 21233 | 265 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 266 | is_formula_functor :: "[i=>o,i,i] => o" where | 
| 13386 | 267 | "is_formula_functor(M,X,Z) == | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 268 | \<exists>nat'[M]. \<exists>natnat[M]. \<exists>natnatsum[M]. \<exists>XX[M]. \<exists>X3[M]. | 
| 13386 | 269 | omega(M,nat') & cartprod(M,nat',nat',natnat) & | 
| 270 | is_sum(M,natnat,natnat,natnatsum) & | |
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 271 | cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) & | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 272 | is_sum(M,natnatsum,X3,Z)" | 
| 13386 | 273 | |
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 274 | lemma (in M_basic) formula_functor_abs [simp]: | 
| 13386 | 275 | "[| M(X); M(Z) |] | 
| 276 | ==> is_formula_functor(M,X,Z) <-> | |
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 277 | Z = ((nat*nat) + (nat*nat)) + (X*X + X)" | 
| 13386 | 278 | by (simp add: is_formula_functor_def) | 
| 279 | ||
| 280 | ||
| 281 | subsection{*@{term M} Contains the List and Formula Datatypes*}
 | |
| 13395 | 282 | |
| 21233 | 283 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 284 | list_N :: "[i,i] => i" where | 
| 13397 | 285 |     "list_N(A,n) == (\<lambda>X. {0} + A * X)^n (0)"
 | 
| 286 | ||
| 287 | lemma Nil_in_list_N [simp]: "[] \<in> list_N(A,succ(n))" | |
| 288 | by (simp add: list_N_def Nil_def) | |
| 289 | ||
| 290 | lemma Cons_in_list_N [simp]: | |
| 291 | "Cons(a,l) \<in> list_N(A,succ(n)) <-> a\<in>A & l \<in> list_N(A,n)" | |
| 292 | by (simp add: list_N_def Cons_def) | |
| 293 | ||
| 294 | text{*These two aren't simprules because they reveal the underlying
 | |
| 295 | list representation.*} | |
| 296 | lemma list_N_0: "list_N(A,0) = 0" | |
| 297 | by (simp add: list_N_def) | |
| 298 | ||
| 299 | lemma list_N_succ: "list_N(A,succ(n)) = {0} + A * (list_N(A,n))"
 | |
| 300 | by (simp add: list_N_def) | |
| 301 | ||
| 302 | lemma list_N_imp_list: | |
| 303 | "[| l \<in> list_N(A,n); n \<in> nat |] ==> l \<in> list(A)" | |
| 304 | by (force simp add: list_eq_Union list_N_def) | |
| 305 | ||
| 306 | lemma list_N_imp_length_lt [rule_format]: | |
| 307 | "n \<in> nat ==> \<forall>l \<in> list_N(A,n). length(l) < n" | |
| 308 | apply (induct_tac n) | |
| 309 | apply (auto simp add: list_N_0 list_N_succ | |
| 310 | Nil_def [symmetric] Cons_def [symmetric]) | |
| 311 | done | |
| 312 | ||
| 313 | lemma list_imp_list_N [rule_format]: | |
| 314 | "l \<in> list(A) ==> \<forall>n\<in>nat. length(l) < n --> l \<in> list_N(A, n)" | |
| 315 | apply (induct_tac l) | |
| 316 | apply (force elim: natE)+ | |
| 317 | done | |
| 318 | ||
| 319 | lemma list_N_imp_eq_length: | |
| 320 | "[|n \<in> nat; l \<notin> list_N(A, n); l \<in> list_N(A, succ(n))|] | |
| 321 | ==> n = length(l)" | |
| 322 | apply (rule le_anti_sym) | |
| 323 | prefer 2 apply (simp add: list_N_imp_length_lt) | |
| 324 | apply (frule list_N_imp_list, simp) | |
| 325 | apply (simp add: not_lt_iff_le [symmetric]) | |
| 326 | apply (blast intro: list_imp_list_N) | |
| 327 | done | |
| 328 | ||
| 329 | text{*Express @{term list_rec} without using @{term rank} or @{term Vset},
 | |
| 330 | neither of which is absolute.*} | |
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 331 | lemma (in M_trivial) list_rec_eq: | 
| 13397 | 332 | "l \<in> list(A) ==> | 
| 333 | list_rec(a,g,l) = | |
| 334 | transrec (succ(length(l)), | |
| 13409 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 335 | \<lambda>x h. Lambda (list(A), | 
| 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 336 | list_case' (a, | 
| 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 337 | \<lambda>a l. g(a, l, h ` succ(length(l)) ` l)))) ` l" | 
| 13397 | 338 | apply (induct_tac l) | 
| 339 | apply (subst transrec, simp) | |
| 340 | apply (subst transrec) | |
| 341 | apply (simp add: list_imp_list_N) | |
| 342 | done | |
| 343 | ||
| 21233 | 344 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 345 | is_list_N :: "[i=>o,i,i,i] => o" where | 
| 13397 | 346 | "is_list_N(M,A,n,Z) == | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 347 | \<exists>zero[M]. empty(M,zero) & | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 348 | is_iterates(M, is_list_functor(M,A), zero, n, Z)" | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 349 | |
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 350 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 351 | mem_list :: "[i=>o,i,i] => o" where | 
| 13395 | 352 | "mem_list(M,A,l) == | 
| 353 | \<exists>n[M]. \<exists>listn[M]. | |
| 13397 | 354 | finite_ordinal(M,n) & is_list_N(M,A,n,listn) & l \<in> listn" | 
| 13395 | 355 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 356 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 357 | is_list :: "[i=>o,i,i] => o" where | 
| 13395 | 358 | "is_list(M,A,Z) == \<forall>l[M]. l \<in> Z <-> mem_list(M,A,l)" | 
| 359 | ||
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 360 | subsubsection{*Towards Absoluteness of @{term formula_rec}*}
 | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 361 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 362 | consts depth :: "i=>i" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 363 | primrec | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 364 | "depth(Member(x,y)) = 0" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 365 | "depth(Equal(x,y)) = 0" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 366 | "depth(Nand(p,q)) = succ(depth(p) \<union> depth(q))" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 367 | "depth(Forall(p)) = succ(depth(p))" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 368 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 369 | lemma depth_type [TC]: "p \<in> formula ==> depth(p) \<in> nat" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 370 | by (induct_tac p, simp_all) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 371 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 372 | |
| 21233 | 373 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 374 | formula_N :: "i => i" where | 
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 375 | "formula_N(n) == (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X)) ^ n (0)" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 376 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 377 | lemma Member_in_formula_N [simp]: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 378 | "Member(x,y) \<in> formula_N(succ(n)) <-> x \<in> nat & y \<in> nat" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 379 | by (simp add: formula_N_def Member_def) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 380 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 381 | lemma Equal_in_formula_N [simp]: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 382 | "Equal(x,y) \<in> formula_N(succ(n)) <-> x \<in> nat & y \<in> nat" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 383 | by (simp add: formula_N_def Equal_def) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 384 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 385 | lemma Nand_in_formula_N [simp]: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 386 | "Nand(x,y) \<in> formula_N(succ(n)) <-> x \<in> formula_N(n) & y \<in> formula_N(n)" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 387 | by (simp add: formula_N_def Nand_def) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 388 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 389 | lemma Forall_in_formula_N [simp]: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 390 | "Forall(x) \<in> formula_N(succ(n)) <-> x \<in> formula_N(n)" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 391 | by (simp add: formula_N_def Forall_def) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 392 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 393 | text{*These two aren't simprules because they reveal the underlying
 | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 394 | formula representation.*} | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 395 | lemma formula_N_0: "formula_N(0) = 0" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 396 | by (simp add: formula_N_def) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 397 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 398 | lemma formula_N_succ: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 399 | "formula_N(succ(n)) = | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 400 | ((nat*nat) + (nat*nat)) + (formula_N(n) * formula_N(n) + formula_N(n))" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 401 | by (simp add: formula_N_def) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 402 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 403 | lemma formula_N_imp_formula: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 404 | "[| p \<in> formula_N(n); n \<in> nat |] ==> p \<in> formula" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 405 | by (force simp add: formula_eq_Union formula_N_def) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 406 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 407 | lemma formula_N_imp_depth_lt [rule_format]: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 408 | "n \<in> nat ==> \<forall>p \<in> formula_N(n). depth(p) < n" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 409 | apply (induct_tac n) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 410 | apply (auto simp add: formula_N_0 formula_N_succ | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 411 | depth_type formula_N_imp_formula Un_least_lt_iff | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 412 | Member_def [symmetric] Equal_def [symmetric] | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 413 | Nand_def [symmetric] Forall_def [symmetric]) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 414 | done | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 415 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 416 | lemma formula_imp_formula_N [rule_format]: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 417 | "p \<in> formula ==> \<forall>n\<in>nat. depth(p) < n --> p \<in> formula_N(n)" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 418 | apply (induct_tac p) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 419 | apply (simp_all add: succ_Un_distrib Un_least_lt_iff) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 420 | apply (force elim: natE)+ | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 421 | done | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 422 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 423 | lemma formula_N_imp_eq_depth: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 424 | "[|n \<in> nat; p \<notin> formula_N(n); p \<in> formula_N(succ(n))|] | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 425 | ==> n = depth(p)" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 426 | apply (rule le_anti_sym) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 427 | prefer 2 apply (simp add: formula_N_imp_depth_lt) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 428 | apply (frule formula_N_imp_formula, simp) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 429 | apply (simp add: not_lt_iff_le [symmetric]) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 430 | apply (blast intro: formula_imp_formula_N) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 431 | done | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 432 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 433 | |
| 13647 | 434 | text{*This result and the next are unused.*}
 | 
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 435 | lemma formula_N_mono [rule_format]: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 436 | "[| m \<in> nat; n \<in> nat |] ==> m\<le>n --> formula_N(m) \<subseteq> formula_N(n)" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 437 | apply (rule_tac m = m and n = n in diff_induct) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 438 | apply (simp_all add: formula_N_0 formula_N_succ, blast) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 439 | done | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 440 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 441 | lemma formula_N_distrib: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 442 | "[| m \<in> nat; n \<in> nat |] ==> formula_N(m \<union> n) = formula_N(m) \<union> formula_N(n)" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 443 | apply (rule_tac i = m and j = n in Ord_linear_le, auto) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 444 | apply (simp_all add: subset_Un_iff [THEN iffD1] subset_Un_iff2 [THEN iffD1] | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 445 | le_imp_subset formula_N_mono) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 446 | done | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 447 | |
| 21233 | 448 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 449 | is_formula_N :: "[i=>o,i,i] => o" where | 
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 450 | "is_formula_N(M,n,Z) == | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 451 | \<exists>zero[M]. empty(M,zero) & | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 452 | is_iterates(M, is_formula_functor(M), zero, n, Z)" | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 453 | |
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 454 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 455 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 456 | mem_formula :: "[i=>o,i] => o" where | 
| 13395 | 457 | "mem_formula(M,p) == | 
| 458 | \<exists>n[M]. \<exists>formn[M]. | |
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 459 | finite_ordinal(M,n) & is_formula_N(M,n,formn) & p \<in> formn" | 
| 13395 | 460 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 461 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 462 | is_formula :: "[i=>o,i] => o" where | 
| 13395 | 463 | "is_formula(M,Z) == \<forall>p[M]. p \<in> Z <-> mem_formula(M,p)" | 
| 464 | ||
| 13634 | 465 | locale M_datatypes = M_trancl + | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 466 | assumes list_replacement1: | 
| 13363 | 467 | "M(A) ==> iterates_replacement(M, is_list_functor(M,A), 0)" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 468 | and list_replacement2: | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 469 | "M(A) ==> strong_replacement(M, | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 470 | \<lambda>n y. n\<in>nat & is_iterates(M, is_list_functor(M,A), 0, n, y))" | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 471 | and formula_replacement1: | 
| 13386 | 472 | "iterates_replacement(M, is_formula_functor(M), 0)" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 473 | and formula_replacement2: | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 474 | "strong_replacement(M, | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 475 | \<lambda>n y. n\<in>nat & is_iterates(M, is_formula_functor(M), 0, n, y))" | 
| 13422 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 476 | and nth_replacement: | 
| 
af9bc8d87a75
Added the assumption nth_replacement to locale M_datatypes.
 paulson parents: 
13418diff
changeset | 477 | "M(l) ==> iterates_replacement(M, %l t. is_tl(M,l,t), l)" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 478 | |
| 13395 | 479 | |
| 480 | subsubsection{*Absoluteness of the List Construction*}
 | |
| 481 | ||
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 482 | lemma (in M_datatypes) list_replacement2': | 
| 13353 | 483 |   "M(A) ==> strong_replacement(M, \<lambda>n y. n\<in>nat & y = (\<lambda>X. {0} + A * X)^n (0))"
 | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 484 | apply (insert list_replacement2 [of A]) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 485 | apply (rule strong_replacement_cong [THEN iffD1]) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 486 | apply (rule conj_cong [OF iff_refl iterates_abs [of "is_list_functor(M,A)"]]) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 487 | apply (simp_all add: list_replacement1 relation1_def) | 
| 13353 | 488 | done | 
| 13268 | 489 | |
| 490 | lemma (in M_datatypes) list_closed [intro,simp]: | |
| 491 | "M(A) ==> M(list(A))" | |
| 13353 | 492 | apply (insert list_replacement1) | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 493 | by (simp add: RepFun_closed2 list_eq_Union | 
| 13634 | 494 | list_replacement2' relation1_def | 
| 13353 | 495 | iterates_closed [of "is_list_functor(M,A)"]) | 
| 13397 | 496 | |
| 13423 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 497 | text{*WARNING: use only with @{text "dest:"} or with variables fixed!*}
 | 
| 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 498 | lemmas (in M_datatypes) list_into_M = transM [OF _ list_closed] | 
| 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 499 | |
| 13397 | 500 | lemma (in M_datatypes) list_N_abs [simp]: | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 501 | "[|M(A); n\<in>nat; M(Z)|] | 
| 13397 | 502 | ==> is_list_N(M,A,n,Z) <-> Z = list_N(A,n)" | 
| 13395 | 503 | apply (insert list_replacement1) | 
| 13634 | 504 | apply (simp add: is_list_N_def list_N_def relation1_def nat_into_M | 
| 13395 | 505 |                  iterates_abs [of "is_list_functor(M,A)" _ "\<lambda>X. {0} + A*X"])
 | 
| 506 | done | |
| 13268 | 507 | |
| 13397 | 508 | lemma (in M_datatypes) list_N_closed [intro,simp]: | 
| 509 | "[|M(A); n\<in>nat|] ==> M(list_N(A,n))" | |
| 510 | apply (insert list_replacement1) | |
| 13634 | 511 | apply (simp add: is_list_N_def list_N_def relation1_def nat_into_M | 
| 13397 | 512 | iterates_closed [of "is_list_functor(M,A)"]) | 
| 513 | done | |
| 514 | ||
| 13395 | 515 | lemma (in M_datatypes) mem_list_abs [simp]: | 
| 516 | "M(A) ==> mem_list(M,A,l) <-> l \<in> list(A)" | |
| 517 | apply (insert list_replacement1) | |
| 13634 | 518 | apply (simp add: mem_list_def list_N_def relation1_def list_eq_Union | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 519 | iterates_closed [of "is_list_functor(M,A)"]) | 
| 13395 | 520 | done | 
| 521 | ||
| 522 | lemma (in M_datatypes) list_abs [simp]: | |
| 523 | "[|M(A); M(Z)|] ==> is_list(M,A,Z) <-> Z = list(A)" | |
| 524 | apply (simp add: is_list_def, safe) | |
| 525 | apply (rule M_equalityI, simp_all) | |
| 526 | done | |
| 527 | ||
| 528 | subsubsection{*Absoluteness of Formulas*}
 | |
| 13293 | 529 | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 530 | lemma (in M_datatypes) formula_replacement2': | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 531 | "strong_replacement(M, \<lambda>n y. n\<in>nat & y = (\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X))^n (0))" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 532 | apply (insert formula_replacement2) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 533 | apply (rule strong_replacement_cong [THEN iffD1]) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 534 | apply (rule conj_cong [OF iff_refl iterates_abs [of "is_formula_functor(M)"]]) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 535 | apply (simp_all add: formula_replacement1 relation1_def) | 
| 13386 | 536 | done | 
| 537 | ||
| 538 | lemma (in M_datatypes) formula_closed [intro,simp]: | |
| 539 | "M(formula)" | |
| 540 | apply (insert formula_replacement1) | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 541 | apply (simp add: RepFun_closed2 formula_eq_Union | 
| 13634 | 542 | formula_replacement2' relation1_def | 
| 13386 | 543 | iterates_closed [of "is_formula_functor(M)"]) | 
| 544 | done | |
| 545 | ||
| 13423 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 546 | lemmas (in M_datatypes) formula_into_M = transM [OF _ formula_closed] | 
| 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 547 | |
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 548 | lemma (in M_datatypes) formula_N_abs [simp]: | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 549 | "[|n\<in>nat; M(Z)|] | 
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 550 | ==> is_formula_N(M,n,Z) <-> Z = formula_N(n)" | 
| 13395 | 551 | apply (insert formula_replacement1) | 
| 13634 | 552 | apply (simp add: is_formula_N_def formula_N_def relation1_def nat_into_M | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 553 | iterates_abs [of "is_formula_functor(M)" _ | 
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 554 | "\<lambda>X. ((nat*nat) + (nat*nat)) + (X*X + X)"]) | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 555 | done | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 556 | |
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 557 | lemma (in M_datatypes) formula_N_closed [intro,simp]: | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 558 | "n\<in>nat ==> M(formula_N(n))" | 
| 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 559 | apply (insert formula_replacement1) | 
| 13634 | 560 | apply (simp add: is_formula_N_def formula_N_def relation1_def nat_into_M | 
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 561 | iterates_closed [of "is_formula_functor(M)"]) | 
| 13395 | 562 | done | 
| 563 | ||
| 564 | lemma (in M_datatypes) mem_formula_abs [simp]: | |
| 565 | "mem_formula(M,l) <-> l \<in> formula" | |
| 566 | apply (insert formula_replacement1) | |
| 13634 | 567 | apply (simp add: mem_formula_def relation1_def formula_eq_Union formula_N_def | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 568 | iterates_closed [of "is_formula_functor(M)"]) | 
| 13395 | 569 | done | 
| 570 | ||
| 571 | lemma (in M_datatypes) formula_abs [simp]: | |
| 572 | "[|M(Z)|] ==> is_formula(M,Z) <-> Z = formula" | |
| 573 | apply (simp add: is_formula_def, safe) | |
| 574 | apply (rule M_equalityI, simp_all) | |
| 575 | done | |
| 576 | ||
| 577 | ||
| 578 | subsection{*Absoluteness for @{text \<epsilon>}-Closure: the @{term eclose} Operator*}
 | |
| 579 | ||
| 580 | text{*Re-expresses eclose using "iterates"*}
 | |
| 581 | lemma eclose_eq_Union: | |
| 582 | "eclose(A) = (\<Union>n\<in>nat. Union^n (A))" | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 583 | apply (simp add: eclose_def) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 584 | apply (rule UN_cong) | 
| 13395 | 585 | apply (rule refl) | 
| 586 | apply (induct_tac n) | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 587 | apply (simp add: nat_rec_0) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 588 | apply (simp add: nat_rec_succ) | 
| 13395 | 589 | done | 
| 590 | ||
| 21233 | 591 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 592 | is_eclose_n :: "[i=>o,i,i,i] => o" where | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 593 | "is_eclose_n(M,A,n,Z) == is_iterates(M, big_union(M), A, n, Z)" | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 594 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 595 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 596 | mem_eclose :: "[i=>o,i,i] => o" where | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 597 | "mem_eclose(M,A,l) == | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 598 | \<exists>n[M]. \<exists>eclosen[M]. | 
| 13395 | 599 | finite_ordinal(M,n) & is_eclose_n(M,A,n,eclosen) & l \<in> eclosen" | 
| 600 | ||
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 601 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 602 | is_eclose :: "[i=>o,i,i] => o" where | 
| 13395 | 603 | "is_eclose(M,A,Z) == \<forall>u[M]. u \<in> Z <-> mem_eclose(M,A,u)" | 
| 604 | ||
| 605 | ||
| 13428 | 606 | locale M_eclose = M_datatypes + | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 607 | assumes eclose_replacement1: | 
| 13395 | 608 | "M(A) ==> iterates_replacement(M, big_union(M), A)" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 609 | and eclose_replacement2: | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 610 | "M(A) ==> strong_replacement(M, | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 611 | \<lambda>n y. n\<in>nat & is_iterates(M, big_union(M), A, n, y))" | 
| 13395 | 612 | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 613 | lemma (in M_eclose) eclose_replacement2': | 
| 13395 | 614 | "M(A) ==> strong_replacement(M, \<lambda>n y. n\<in>nat & y = Union^n (A))" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 615 | apply (insert eclose_replacement2 [of A]) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 616 | apply (rule strong_replacement_cong [THEN iffD1]) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 617 | apply (rule conj_cong [OF iff_refl iterates_abs [of "big_union(M)"]]) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 618 | apply (simp_all add: eclose_replacement1 relation1_def) | 
| 13395 | 619 | done | 
| 620 | ||
| 621 | lemma (in M_eclose) eclose_closed [intro,simp]: | |
| 622 | "M(A) ==> M(eclose(A))" | |
| 623 | apply (insert eclose_replacement1) | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 624 | by (simp add: RepFun_closed2 eclose_eq_Union | 
| 13634 | 625 | eclose_replacement2' relation1_def | 
| 13395 | 626 | iterates_closed [of "big_union(M)"]) | 
| 627 | ||
| 628 | lemma (in M_eclose) is_eclose_n_abs [simp]: | |
| 629 | "[|M(A); n\<in>nat; M(Z)|] ==> is_eclose_n(M,A,n,Z) <-> Z = Union^n (A)" | |
| 630 | apply (insert eclose_replacement1) | |
| 13634 | 631 | apply (simp add: is_eclose_n_def relation1_def nat_into_M | 
| 13395 | 632 | iterates_abs [of "big_union(M)" _ "Union"]) | 
| 633 | done | |
| 634 | ||
| 635 | lemma (in M_eclose) mem_eclose_abs [simp]: | |
| 636 | "M(A) ==> mem_eclose(M,A,l) <-> l \<in> eclose(A)" | |
| 637 | apply (insert eclose_replacement1) | |
| 13634 | 638 | apply (simp add: mem_eclose_def relation1_def eclose_eq_Union | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 639 | iterates_closed [of "big_union(M)"]) | 
| 13395 | 640 | done | 
| 641 | ||
| 642 | lemma (in M_eclose) eclose_abs [simp]: | |
| 643 | "[|M(A); M(Z)|] ==> is_eclose(M,A,Z) <-> Z = eclose(A)" | |
| 644 | apply (simp add: is_eclose_def, safe) | |
| 645 | apply (rule M_equalityI, simp_all) | |
| 646 | done | |
| 647 | ||
| 648 | ||
| 649 | subsection {*Absoluteness for @{term transrec}*}
 | |
| 650 | ||
| 22710 | 651 | text{* @{prop "transrec(a,H) \<equiv> wfrec(Memrel(eclose({a})), a, H)"} *}
 | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 652 | |
| 21233 | 653 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 654 | is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o" where | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 655 | "is_transrec(M,MH,a,z) == | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 656 | \<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M]. | 
| 13395 | 657 | upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) & | 
| 658 | is_wfrec(M,MH,mesa,a,z)" | |
| 659 | ||
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 660 | definition | 
| 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 661 | transrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o" where | 
| 13395 | 662 | "transrec_replacement(M,MH,a) == | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 663 | \<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M]. | 
| 13395 | 664 | upair(M,a,a,sa) & is_eclose(M,sa,esa) & membership(M,esa,mesa) & | 
| 665 | wfrec_replacement(M,MH,mesa)" | |
| 666 | ||
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 667 | text{*The condition @{term "Ord(i)"} lets us use the simpler
 | 
| 13395 | 668 |   @{text "trans_wfrec_abs"} rather than @{text "trans_wfrec_abs"},
 | 
| 669 | which I haven't even proved yet. *} | |
| 670 | theorem (in M_eclose) transrec_abs: | |
| 13634 | 671 | "[|transrec_replacement(M,MH,i); relation2(M,MH,H); | 
| 13418 
7c0ba9dba978
tweaks, aiming towards relativization of "satisfies"
 paulson parents: 
13409diff
changeset | 672 | Ord(i); M(i); M(z); | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 673 | \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 674 | ==> is_transrec(M,MH,i,z) <-> z = transrec(i,H)" | 
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13564diff
changeset | 675 | by (simp add: trans_wfrec_abs transrec_replacement_def is_transrec_def | 
| 13395 | 676 | transrec_def eclose_sing_Ord_eq wf_Memrel trans_Memrel relation_Memrel) | 
| 677 | ||
| 678 | ||
| 679 | theorem (in M_eclose) transrec_closed: | |
| 13634 | 680 | "[|transrec_replacement(M,MH,i); relation2(M,MH,H); | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 681 | Ord(i); M(i); | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 682 | \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] | 
| 13395 | 683 | ==> M(transrec(i,H))" | 
| 13615 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13564diff
changeset | 684 | by (simp add: trans_wfrec_closed transrec_replacement_def is_transrec_def | 
| 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13564diff
changeset | 685 | transrec_def eclose_sing_Ord_eq wf_Memrel trans_Memrel relation_Memrel) | 
| 
449a70d88b38
Numerous cosmetic changes, prompted by the new simplifier
 paulson parents: 
13564diff
changeset | 686 | |
| 13395 | 687 | |
| 13440 | 688 | text{*Helps to prove instances of @{term transrec_replacement}*}
 | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 689 | lemma (in M_eclose) transrec_replacementI: | 
| 13440 | 690 | "[|M(a); | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 691 | strong_replacement (M, | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 692 | \<lambda>x z. \<exists>y[M]. pair(M, x, y, z) & | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 693 |                                is_wfrec(M,MH,Memrel(eclose({a})),x,y))|]
 | 
| 13440 | 694 | ==> transrec_replacement(M,MH,a)" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 695 | by (simp add: transrec_replacement_def wfrec_replacement_def) | 
| 13440 | 696 | |
| 13395 | 697 | |
| 13397 | 698 | subsection{*Absoluteness for the List Operator @{term length}*}
 | 
| 13647 | 699 | text{*But it is never used.*}
 | 
| 700 | ||
| 21233 | 701 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 702 | is_length :: "[i=>o,i,i,i] => o" where | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 703 | "is_length(M,A,l,n) == | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 704 | \<exists>sn[M]. \<exists>list_n[M]. \<exists>list_sn[M]. | 
| 13397 | 705 | is_list_N(M,A,n,list_n) & l \<notin> list_n & | 
| 706 | successor(M,n,sn) & is_list_N(M,A,sn,list_sn) & l \<in> list_sn" | |
| 707 | ||
| 708 | ||
| 709 | lemma (in M_datatypes) length_abs [simp]: | |
| 710 | "[|M(A); l \<in> list(A); n \<in> nat|] ==> is_length(M,A,l,n) <-> n = length(l)" | |
| 711 | apply (subgoal_tac "M(l) & M(n)") | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 712 | prefer 2 apply (blast dest: transM) | 
| 13397 | 713 | apply (simp add: is_length_def) | 
| 714 | apply (blast intro: list_imp_list_N nat_into_Ord list_N_imp_eq_length | |
| 715 | dest: list_N_imp_length_lt) | |
| 716 | done | |
| 717 | ||
| 718 | text{*Proof is trivial since @{term length} returns natural numbers.*}
 | |
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 719 | lemma (in M_trivial) length_closed [intro,simp]: | 
| 13397 | 720 | "l \<in> list(A) ==> M(length(l))" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 721 | by (simp add: nat_into_M) | 
| 13397 | 722 | |
| 723 | ||
| 13647 | 724 | subsection {*Absoluteness for the List Operator @{term nth}*}
 | 
| 13397 | 725 | |
| 726 | lemma nth_eq_hd_iterates_tl [rule_format]: | |
| 727 | "xs \<in> list(A) ==> \<forall>n \<in> nat. nth(n,xs) = hd' (tl'^n (xs))" | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 728 | apply (induct_tac xs) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 729 | apply (simp add: iterates_tl_Nil hd'_Nil, clarify) | 
| 13397 | 730 | apply (erule natE) | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 731 | apply (simp add: hd'_Cons) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 732 | apply (simp add: tl'_Cons iterates_commute) | 
| 13397 | 733 | done | 
| 734 | ||
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 735 | lemma (in M_basic) iterates_tl'_closed: | 
| 13397 | 736 | "[|n \<in> nat; M(x)|] ==> M(tl'^n (x))" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 737 | apply (induct_tac n, simp) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 738 | apply (simp add: tl'_Cons tl'_closed) | 
| 13397 | 739 | done | 
| 740 | ||
| 741 | text{*Immediate by type-checking*}
 | |
| 742 | lemma (in M_datatypes) nth_closed [intro,simp]: | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 743 | "[|xs \<in> list(A); n \<in> nat; M(A)|] ==> M(nth(n,xs))" | 
| 13397 | 744 | apply (case_tac "n < length(xs)") | 
| 745 | apply (blast intro: nth_type transM) | |
| 746 | apply (simp add: not_lt_iff_le nth_eq_0) | |
| 747 | done | |
| 748 | ||
| 21233 | 749 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 750 | is_nth :: "[i=>o,i,i,i] => o" where | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 751 | "is_nth(M,n,l,Z) == | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 752 | \<exists>X[M]. is_iterates(M, is_tl(M), l, n, X) & is_hd(M,X,Z)" | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 753 | |
| 13409 
d4ea094c650e
Relativization and Separation for the function "nth"
 paulson parents: 
13398diff
changeset | 754 | lemma (in M_datatypes) nth_abs [simp]: | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 755 | "[|M(A); n \<in> nat; l \<in> list(A); M(Z)|] | 
| 13397 | 756 | ==> is_nth(M,n,l,Z) <-> Z = nth(n,l)" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 757 | apply (subgoal_tac "M(l)") | 
| 13397 | 758 | prefer 2 apply (blast intro: transM) | 
| 759 | apply (simp add: is_nth_def nth_eq_hd_iterates_tl nat_into_M | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 760 | tl'_closed iterates_tl'_closed | 
| 13634 | 761 | iterates_abs [OF _ relation1_tl] nth_replacement) | 
| 13397 | 762 | done | 
| 763 | ||
| 13395 | 764 | |
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 765 | subsection{*Relativization and Absoluteness for the @{term formula} Constructors*}
 | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 766 | |
| 21233 | 767 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 768 | is_Member :: "[i=>o,i,i,i] => o" where | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 769 |      --{* because @{term "Member(x,y) \<equiv> Inl(Inl(\<langle>x,y\<rangle>))"}*}
 | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 770 | "is_Member(M,x,y,Z) == | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 771 | \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inl(M,u,Z)" | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 772 | |
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 773 | lemma (in M_trivial) Member_abs [simp]: | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 774 | "[|M(x); M(y); M(Z)|] ==> is_Member(M,x,y,Z) <-> (Z = Member(x,y))" | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 775 | by (simp add: is_Member_def Member_def) | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 776 | |
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 777 | lemma (in M_trivial) Member_in_M_iff [iff]: | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 778 | "M(Member(x,y)) <-> M(x) & M(y)" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 779 | by (simp add: Member_def) | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 780 | |
| 21233 | 781 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 782 | is_Equal :: "[i=>o,i,i,i] => o" where | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 783 |      --{* because @{term "Equal(x,y) \<equiv> Inl(Inr(\<langle>x,y\<rangle>))"}*}
 | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 784 | "is_Equal(M,x,y,Z) == | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 785 | \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inr(M,p,u) & is_Inl(M,u,Z)" | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 786 | |
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 787 | lemma (in M_trivial) Equal_abs [simp]: | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 788 | "[|M(x); M(y); M(Z)|] ==> is_Equal(M,x,y,Z) <-> (Z = Equal(x,y))" | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 789 | by (simp add: is_Equal_def Equal_def) | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 790 | |
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 791 | lemma (in M_trivial) Equal_in_M_iff [iff]: "M(Equal(x,y)) <-> M(x) & M(y)" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 792 | by (simp add: Equal_def) | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 793 | |
| 21233 | 794 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 795 | is_Nand :: "[i=>o,i,i,i] => o" where | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 796 |      --{* because @{term "Nand(x,y) \<equiv> Inr(Inl(\<langle>x,y\<rangle>))"}*}
 | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 797 | "is_Nand(M,x,y,Z) == | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 798 | \<exists>p[M]. \<exists>u[M]. pair(M,x,y,p) & is_Inl(M,p,u) & is_Inr(M,u,Z)" | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 799 | |
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 800 | lemma (in M_trivial) Nand_abs [simp]: | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 801 | "[|M(x); M(y); M(Z)|] ==> is_Nand(M,x,y,Z) <-> (Z = Nand(x,y))" | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 802 | by (simp add: is_Nand_def Nand_def) | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 803 | |
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 804 | lemma (in M_trivial) Nand_in_M_iff [iff]: "M(Nand(x,y)) <-> M(x) & M(y)" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 805 | by (simp add: Nand_def) | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 806 | |
| 21233 | 807 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 808 | is_Forall :: "[i=>o,i,i] => o" where | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 809 |      --{* because @{term "Forall(x) \<equiv> Inr(Inr(p))"}*}
 | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 810 | "is_Forall(M,p,Z) == \<exists>u[M]. is_Inr(M,p,u) & is_Inr(M,u,Z)" | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 811 | |
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 812 | lemma (in M_trivial) Forall_abs [simp]: | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 813 | "[|M(x); M(Z)|] ==> is_Forall(M,x,Z) <-> (Z = Forall(x))" | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 814 | by (simp add: is_Forall_def Forall_def) | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 815 | |
| 13564 
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
 paulson parents: 
13557diff
changeset | 816 | lemma (in M_trivial) Forall_in_M_iff [iff]: "M(Forall(x)) <-> M(x)" | 
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 817 | by (simp add: Forall_def) | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 818 | |
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 819 | |
| 13647 | 820 | |
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 821 | subsection {*Absoluteness for @{term formula_rec}*}
 | 
| 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 822 | |
| 21233 | 823 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 824 | formula_rec_case :: "[[i,i]=>i, [i,i]=>i, [i,i,i,i]=>i, [i,i]=>i, i, i] => i" where | 
| 13647 | 825 |     --{* the instance of @{term formula_case} in @{term formula_rec}*}
 | 
| 826 | "formula_rec_case(a,b,c,d,h) == | |
| 827 | formula_case (a, b, | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 828 | \<lambda>u v. c(u, v, h ` succ(depth(u)) ` u, | 
| 13647 | 829 | h ` succ(depth(v)) ` v), | 
| 830 | \<lambda>u. d(u, h ` succ(depth(u)) ` u))" | |
| 831 | ||
| 832 | text{*Unfold @{term formula_rec} to @{term formula_rec_case}.
 | |
| 833 |      Express @{term formula_rec} without using @{term rank} or @{term Vset},
 | |
| 834 | neither of which is absolute.*} | |
| 835 | lemma (in M_trivial) formula_rec_eq: | |
| 836 | "p \<in> formula ==> | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 837 | formula_rec(a,b,c,d,p) = | 
| 13647 | 838 | transrec (succ(depth(p)), | 
| 839 | \<lambda>x h. Lambda (formula, formula_rec_case(a,b,c,d,h))) ` p" | |
| 840 | apply (simp add: formula_rec_case_def) | |
| 841 | apply (induct_tac p) | |
| 842 |    txt{*Base case for @{term Member}*}
 | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 843 | apply (subst transrec, simp add: formula.intros) | 
| 13647 | 844 |   txt{*Base case for @{term Equal}*}
 | 
| 845 | apply (subst transrec, simp add: formula.intros) | |
| 846 |  txt{*Inductive step for @{term Nand}*}
 | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 847 | apply (subst transrec) | 
| 13647 | 848 | apply (simp add: succ_Un_distrib formula.intros) | 
| 849 | txt{*Inductive step for @{term Forall}*}
 | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 850 | apply (subst transrec) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 851 | apply (simp add: formula_imp_formula_N formula.intros) | 
| 13647 | 852 | done | 
| 853 | ||
| 854 | ||
| 855 | subsubsection{*Absoluteness for the Formula Operator @{term depth}*}
 | |
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 856 | |
| 21233 | 857 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 858 | is_depth :: "[i=>o,i,i] => o" where | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 859 | "is_depth(M,p,n) == | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 860 | \<exists>sn[M]. \<exists>formula_n[M]. \<exists>formula_sn[M]. | 
| 13647 | 861 | is_formula_N(M,n,formula_n) & p \<notin> formula_n & | 
| 862 | successor(M,n,sn) & is_formula_N(M,sn,formula_sn) & p \<in> formula_sn" | |
| 863 | ||
| 864 | ||
| 865 | lemma (in M_datatypes) depth_abs [simp]: | |
| 866 | "[|p \<in> formula; n \<in> nat|] ==> is_depth(M,p,n) <-> n = depth(p)" | |
| 867 | apply (subgoal_tac "M(p) & M(n)") | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 868 | prefer 2 apply (blast dest: transM) | 
| 13647 | 869 | apply (simp add: is_depth_def) | 
| 870 | apply (blast intro: formula_imp_formula_N nat_into_Ord formula_N_imp_eq_depth | |
| 871 | dest: formula_N_imp_depth_lt) | |
| 872 | done | |
| 873 | ||
| 874 | text{*Proof is trivial since @{term depth} returns natural numbers.*}
 | |
| 875 | lemma (in M_trivial) depth_closed [intro,simp]: | |
| 876 | "p \<in> formula ==> M(depth(p))" | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 877 | by (simp add: nat_into_M) | 
| 13647 | 878 | |
| 879 | ||
| 13423 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 880 | subsubsection{*@{term is_formula_case}: relativization of @{term formula_case}*}
 | 
| 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 881 | |
| 21233 | 882 | definition | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 883 | is_formula_case :: | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 884 | "[i=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i,i]=>o, [i,i]=>o, i, i] => o" where | 
| 13423 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 885 |   --{*no constraint on non-formulas*}
 | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 886 | "is_formula_case(M, is_a, is_b, is_c, is_d, p, z) == | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 887 | (\<forall>x[M]. \<forall>y[M]. finite_ordinal(M,x) --> finite_ordinal(M,y) --> | 
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 888 | is_Member(M,x,y,p) --> is_a(x,y,z)) & | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 889 | (\<forall>x[M]. \<forall>y[M]. finite_ordinal(M,x) --> finite_ordinal(M,y) --> | 
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 890 | is_Equal(M,x,y,p) --> is_b(x,y,z)) & | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 891 | (\<forall>x[M]. \<forall>y[M]. mem_formula(M,x) --> mem_formula(M,y) --> | 
| 13423 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 892 | is_Nand(M,x,y,p) --> is_c(x,y,z)) & | 
| 13493 
5aa68c051725
Lots of new results concerning recursive datatypes, towards absoluteness of
 paulson parents: 
13440diff
changeset | 893 | (\<forall>x[M]. mem_formula(M,x) --> is_Forall(M,x,p) --> is_d(x,z))" | 
| 13423 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 894 | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 895 | lemma (in M_datatypes) formula_case_abs [simp]: | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 896 | "[| Relation2(M,nat,nat,is_a,a); Relation2(M,nat,nat,is_b,b); | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 897 | Relation2(M,formula,formula,is_c,c); Relation1(M,formula,is_d,d); | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 898 | p \<in> formula; M(z) |] | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 899 | ==> is_formula_case(M,is_a,is_b,is_c,is_d,p,z) <-> | 
| 13423 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 900 | z = formula_case(a,b,c,d,p)" | 
| 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 901 | apply (simp add: formula_into_M is_formula_case_def) | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 902 | apply (erule formula.cases) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 903 | apply (simp_all add: Relation1_def Relation2_def) | 
| 13423 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 904 | done | 
| 
7ec771711c09
More lemmas, working towards relativization of "satisfies"
 paulson parents: 
13422diff
changeset | 905 | |
| 13418 
7c0ba9dba978
tweaks, aiming towards relativization of "satisfies"
 paulson parents: 
13409diff
changeset | 906 | lemma (in M_datatypes) formula_case_closed [intro,simp]: | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 907 | "[|p \<in> formula; | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 908 | \<forall>x[M]. \<forall>y[M]. x\<in>nat --> y\<in>nat --> M(a(x,y)); | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 909 | \<forall>x[M]. \<forall>y[M]. x\<in>nat --> y\<in>nat --> M(b(x,y)); | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 910 | \<forall>x[M]. \<forall>y[M]. x\<in>formula --> y\<in>formula --> M(c(x,y)); | 
| 13418 
7c0ba9dba978
tweaks, aiming towards relativization of "satisfies"
 paulson parents: 
13409diff
changeset | 911 | \<forall>x[M]. x\<in>formula --> M(d(x))|] ==> M(formula_case(a,b,c,d,p))" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 912 | by (erule formula.cases, simp_all) | 
| 13418 
7c0ba9dba978
tweaks, aiming towards relativization of "satisfies"
 paulson parents: 
13409diff
changeset | 913 | |
| 13398 
1cadd412da48
Towards relativization and absoluteness of formula_rec
 paulson parents: 
13397diff
changeset | 914 | |
| 13647 | 915 | subsubsection {*Absoluteness for @{term formula_rec}: Final Results*}
 | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 916 | |
| 21233 | 917 | definition | 
| 21404 
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
 wenzelm parents: 
21233diff
changeset | 918 | is_formula_rec :: "[i=>o, [i,i,i]=>o, i, i] => o" where | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 919 |     --{* predicate to relativize the functional @{term formula_rec}*}
 | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 920 | "is_formula_rec(M,MH,p,z) == | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 921 | \<exists>dp[M]. \<exists>i[M]. \<exists>f[M]. finite_ordinal(M,dp) & is_depth(M,p,dp) & | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 922 | successor(M,dp,i) & fun_apply(M,f,p,z) & is_transrec(M,MH,i,f)" | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 923 | |
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 924 | |
| 13647 | 925 | text{*Sufficient conditions to relativize the instance of @{term formula_case}
 | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 926 |       in @{term formula_rec}*}
 | 
| 13634 | 927 | lemma (in M_datatypes) Relation1_formula_rec_case: | 
| 928 | "[|Relation2(M, nat, nat, is_a, a); | |
| 929 | Relation2(M, nat, nat, is_b, b); | |
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 930 | Relation2 (M, formula, formula, | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 931 | is_c, \<lambda>u v. c(u, v, h`succ(depth(u))`u, h`succ(depth(v))`v)); | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 932 | Relation1(M, formula, | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 933 | is_d, \<lambda>u. d(u, h ` succ(depth(u)) ` u)); | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 934 | M(h) |] | 
| 13634 | 935 | ==> Relation1(M, formula, | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 936 | is_formula_case (M, is_a, is_b, is_c, is_d), | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 937 | formula_rec_case(a, b, c, d, h))" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 938 | apply (simp (no_asm) add: formula_rec_case_def Relation1_def) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 939 | apply (simp add: formula_case_abs) | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 940 | done | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 941 | |
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 942 | |
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 943 | text{*This locale packages the premises of the following theorems,
 | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 944 | which is the normal purpose of locales. It doesn't accumulate | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 945 |       constraints on the class @{term M}, as in most of this deveopment.*}
 | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 946 | locale Formula_Rec = M_eclose + | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 947 | fixes a and is_a and b and is_b and c and is_c and d and is_d and MH | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 948 | defines | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 949 | "MH(u::i,f,z) == | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 950 | \<forall>fml[M]. is_formula(M,fml) --> | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 951 | is_lambda | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 952 | (M, fml, is_formula_case (M, is_a, is_b, is_c(f), is_d(f)), z)" | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 953 | |
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 954 | assumes a_closed: "[|x\<in>nat; y\<in>nat|] ==> M(a(x,y))" | 
| 13634 | 955 | and a_rel: "Relation2(M, nat, nat, is_a, a)" | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 956 | and b_closed: "[|x\<in>nat; y\<in>nat|] ==> M(b(x,y))" | 
| 13634 | 957 | and b_rel: "Relation2(M, nat, nat, is_b, b)" | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 958 | and c_closed: "[|x \<in> formula; y \<in> formula; M(gx); M(gy)|] | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 959 | ==> M(c(x, y, gx, gy))" | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 960 | and c_rel: | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 961 | "M(f) ==> | 
| 13634 | 962 | Relation2 (M, formula, formula, is_c(f), | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 963 | \<lambda>u v. c(u, v, f ` succ(depth(u)) ` u, f ` succ(depth(v)) ` v))" | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 964 | and d_closed: "[|x \<in> formula; M(gx)|] ==> M(d(x, gx))" | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 965 | and d_rel: | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 966 | "M(f) ==> | 
| 13634 | 967 | Relation1(M, formula, is_d(f), \<lambda>u. d(u, f ` succ(depth(u)) ` u))" | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 968 | and fr_replace: "n \<in> nat ==> transrec_replacement(M,MH,n)" | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 969 | and fr_lam_replace: | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 970 | "M(g) ==> | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 971 | strong_replacement | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 972 | (M, \<lambda>x y. x \<in> formula & | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 973 | y = \<langle>x, formula_rec_case(a,b,c,d,g,x)\<rangle>)"; | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 974 | |
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 975 | lemma (in Formula_Rec) formula_rec_case_closed: | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 976 | "[|M(g); p \<in> formula|] ==> M(formula_rec_case(a, b, c, d, g, p))" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 977 | by (simp add: formula_rec_case_def a_closed b_closed c_closed d_closed) | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 978 | |
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 979 | lemma (in Formula_Rec) formula_rec_lam_closed: | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 980 | "M(g) ==> M(Lambda (formula, formula_rec_case(a,b,c,d,g)))" | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 981 | by (simp add: lam_closed2 fr_lam_replace formula_rec_case_closed) | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 982 | |
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 983 | lemma (in Formula_Rec) MH_rel2: | 
| 13634 | 984 | "relation2 (M, MH, | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 985 | \<lambda>x h. Lambda (formula, formula_rec_case(a,b,c,d,h)))" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 986 | apply (simp add: relation2_def MH_def, clarify) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 987 | apply (rule lambda_abs2) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 988 | apply (rule Relation1_formula_rec_case) | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 989 | apply (simp_all add: a_rel b_rel c_rel d_rel formula_rec_case_closed) | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 990 | done | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 991 | |
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 992 | lemma (in Formula_Rec) fr_transrec_closed: | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 993 | "n \<in> nat | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 994 | ==> M(transrec | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 995 | (n, \<lambda>x h. Lambda(formula, formula_rec_case(a, b, c, d, h))))" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 996 | by (simp add: transrec_closed [OF fr_replace MH_rel2] | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 997 | nat_into_M formula_rec_lam_closed) | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 998 | |
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 999 | text{*The main two results: @{term formula_rec} is absolute for @{term M}.*}
 | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 1000 | theorem (in Formula_Rec) formula_rec_closed: | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 1001 | "p \<in> formula ==> M(formula_rec(a,b,c,d,p))" | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 1002 | by (simp add: formula_rec_eq fr_transrec_closed | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 1003 | transM [OF _ formula_closed]) | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 1004 | |
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 1005 | theorem (in Formula_Rec) formula_rec_abs: | 
| 13655 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 1006 | "[| p \<in> formula; M(z)|] | 
| 
95b95cdb4704
Tidying up.  New primitives is_iterates and is_iterates_fm.
 paulson parents: 
13647diff
changeset | 1007 | ==> is_formula_rec(M,MH,p,z) <-> z = formula_rec(a,b,c,d,p)" | 
| 13557 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 1008 | by (simp add: is_formula_rec_def formula_rec_eq transM [OF _ formula_closed] | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 1009 | transrec_abs [OF fr_replace MH_rel2] depth_type | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 1010 | fr_transrec_closed formula_rec_lam_closed eq_commute) | 
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 1011 | |
| 
6061d0045409
deleted redundant material (quasiformula, ...) and rationalized
 paulson parents: 
13505diff
changeset | 1012 | |
| 13268 | 1013 | end |