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