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