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