src/ZF/Constructible/Datatype_absolute.thy
author paulson
Fri, 19 Jul 2002 13:29:22 +0200
changeset 13397 6e5f4d911435
parent 13395 4eb948d1eb4e
child 13398 1cadd412da48
permissions -rw-r--r--
Absoluteness of the function "nth"
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
31df66ca0780 Expressing Lset and L without using length and arity; simplifies Separation
paulson
parents: 13382
diff changeset
    24
	     intro: bnd_mono_iterates_subset bnd_monoD2 [of concl: h] ) 
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))"])
240509babf00 more use of relativized quantifiers
paulson
parents:
diff changeset
    66
 apply (blast dest: bnd_monoD2 [OF _ _ lfp_subset] )  
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:
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   215
     "bnd_mono(univ(0), \<lambda>X. ((nat*nat) + (nat*nat)) + (X + (X*X + X)))"
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:
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   223
     "contin(\<lambda>X. ((nat*nat) + (nat*nat)) + (X + (X*X + X)))"
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:
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   229
    "formula = lfp(univ(0), \<lambda>X. ((nat*nat) + (nat*nat)) + (X + (X*X + X)))"
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])
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   235
 apply (simp add: Member_def Equal_def Neg_def And_def Forall_def)
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]) 
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   241
apply (simp add: Member_def Equal_def Neg_def And_def Forall_def)  
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 = 
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   249
      (\<Union>n\<in>nat. (\<lambda>X. ((nat*nat) + (nat*nat)) + (X + (X*X + X))) ^ n (0))"
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) == 
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   257
        \<exists>nat'[M]. \<exists>natnat[M]. \<exists>natnatsum[M]. \<exists>XX[M]. \<exists>X3[M]. \<exists>X4[M]. 
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) &
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   260
          cartprod(M,X,X,XX) & is_sum(M,XX,X,X3) & is_sum(M,X,X3,X4) &
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   261
          is_sum(M,natnatsum,X4,Z)"
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) <-> 
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   266
          Z = ((nat*nat) + (nat*nat)) + (X + (X*X + X))"
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)),
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   324
      \<lambda>x h. Lambda (list_N(A,x),
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   325
             list_case' (a, 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   326
                \<lambda>a l. g(a, l, h ` succ(length(l)) ` l)))) ` l"
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)))"
13350
paulson
parents: 13348
diff changeset
   382
13395
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   383
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   384
subsubsection{*Absoluteness of the List Construction*}
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   385
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   386
lemma (in M_datatypes) list_replacement2': 
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   387
  "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
   388
apply (insert list_replacement2 [of A]) 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   389
apply (rule strong_replacement_cong [THEN iffD1])  
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   390
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
   391
apply (simp_all add: list_replacement1 relativize1_def) 
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   392
done
13268
240509babf00 more use of relativized quantifiers
paulson
parents:
diff changeset
   393
240509babf00 more use of relativized quantifiers
paulson
parents:
diff changeset
   394
lemma (in M_datatypes) list_closed [intro,simp]:
240509babf00 more use of relativized quantifiers
paulson
parents:
diff changeset
   395
     "M(A) ==> M(list(A))"
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   396
apply (insert list_replacement1)
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   397
by  (simp add: RepFun_closed2 list_eq_Union 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   398
               list_replacement2' relativize1_def
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   399
               iterates_closed [of "is_list_functor(M,A)"])
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   400
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   401
lemma (in M_datatypes) list_N_abs [simp]:
13395
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   402
     "[|M(A); n\<in>nat; M(Z)|] 
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   403
      ==> is_list_N(M,A,n,Z) <-> Z = list_N(A,n)"
13395
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   404
apply (insert list_replacement1)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   405
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
   406
                 iterates_abs [of "is_list_functor(M,A)" _ "\<lambda>X. {0} + A*X"])
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   407
done
13268
240509babf00 more use of relativized quantifiers
paulson
parents:
diff changeset
   408
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   409
lemma (in M_datatypes) list_N_closed [intro,simp]:
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   410
     "[|M(A); n\<in>nat|] ==> M(list_N(A,n))"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   411
apply (insert list_replacement1)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   412
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
   413
                 iterates_closed [of "is_list_functor(M,A)"])
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   414
done
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   415
13395
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   416
lemma (in M_datatypes) mem_list_abs [simp]:
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   417
     "M(A) ==> mem_list(M,A,l) <-> l \<in> list(A)"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   418
apply (insert list_replacement1)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   419
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
   420
                 iterates_closed [of "is_list_functor(M,A)"]) 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   421
done
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   422
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   423
lemma (in M_datatypes) list_abs [simp]:
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   424
     "[|M(A); M(Z)|] ==> is_list(M,A,Z) <-> Z = list(A)"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   425
apply (simp add: is_list_def, safe)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   426
apply (rule M_equalityI, simp_all)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   427
done
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   428
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   429
subsubsection{*Absoluteness of Formulas*}
13293
paulson
parents: 13269
diff changeset
   430
13386
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   431
lemma (in M_datatypes) formula_replacement2': 
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   432
  "strong_replacement(M, \<lambda>n y. n\<in>nat & y = (\<lambda>X. ((nat*nat) + (nat*nat)) + (X + (X*X + X)))^n (0))"
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   433
apply (insert formula_replacement2) 
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   434
apply (rule strong_replacement_cong [THEN iffD1])  
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   435
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
   436
apply (simp_all add: formula_replacement1 relativize1_def) 
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   437
done
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   438
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   439
lemma (in M_datatypes) formula_closed [intro,simp]:
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   440
     "M(formula)"
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   441
apply (insert formula_replacement1)
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   442
apply  (simp add: RepFun_closed2 formula_eq_Union 
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   443
                  formula_replacement2' relativize1_def
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   444
                  iterates_closed [of "is_formula_functor(M)"])
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   445
done
f3e9e8b21aba Formulas (and lists) in M (and L!)
paulson
parents: 13385
diff changeset
   446
13395
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   447
lemma (in M_datatypes) is_formula_n_abs [simp]:
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   448
     "[|n\<in>nat; M(Z)|] 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   449
      ==> is_formula_n(M,n,Z) <-> 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   450
          Z = (\<lambda>X. ((nat*nat) + (nat*nat)) + (X + (X*X + X)))^n (0)"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   451
apply (insert formula_replacement1)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   452
apply (simp add: is_formula_n_def relativize1_def nat_into_M
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   453
                 iterates_abs [of "is_formula_functor(M)" _ 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   454
                        "\<lambda>X. ((nat*nat) + (nat*nat)) + (X + (X*X + X))"])
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   455
done
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   456
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   457
lemma (in M_datatypes) mem_formula_abs [simp]:
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   458
     "mem_formula(M,l) <-> l \<in> formula"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   459
apply (insert formula_replacement1)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   460
apply (simp add: mem_formula_def relativize1_def formula_eq_Union
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   461
                 iterates_closed [of "is_formula_functor(M)"]) 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   462
done
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   463
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   464
lemma (in M_datatypes) formula_abs [simp]:
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   465
     "[|M(Z)|] ==> is_formula(M,Z) <-> Z = formula"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   466
apply (simp add: is_formula_def, safe)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   467
apply (rule M_equalityI, simp_all)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   468
done
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   469
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   470
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   471
subsection{*Absoluteness for Some List Operators*}
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   472
13395
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   473
subsection{*Absoluteness for @{text \<epsilon>}-Closure: the @{term eclose} Operator*}
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   474
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   475
text{*Re-expresses eclose using "iterates"*}
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   476
lemma eclose_eq_Union:
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   477
     "eclose(A) = (\<Union>n\<in>nat. Union^n (A))"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   478
apply (simp add: eclose_def) 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   479
apply (rule UN_cong) 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   480
apply (rule refl)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   481
apply (induct_tac n)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   482
apply (simp add: nat_rec_0)  
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   483
apply (simp add: nat_rec_succ) 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   484
done
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   485
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   486
constdefs
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   487
  is_eclose_n :: "[i=>o,i,i,i] => o"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   488
    "is_eclose_n(M,A,n,Z) == 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   489
      \<exists>sn[M]. \<exists>msn[M]. 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   490
       successor(M,n,sn) & membership(M,sn,msn) &
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   491
       is_wfrec(M, iterates_MH(M, big_union(M), A), msn, n, Z)"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   492
  
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   493
  mem_eclose :: "[i=>o,i,i] => o"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   494
    "mem_eclose(M,A,l) == 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   495
      \<exists>n[M]. \<exists>eclosen[M]. 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   496
       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
   497
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   498
  is_eclose :: "[i=>o,i,i] => o"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   499
    "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
   500
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   501
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   502
locale (open) M_eclose = M_wfrank +
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   503
 assumes eclose_replacement1: 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   504
   "M(A) ==> iterates_replacement(M, big_union(M), A)"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   505
  and eclose_replacement2: 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   506
   "M(A) ==> strong_replacement(M, 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   507
         \<lambda>n y. n\<in>nat & 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   508
               (\<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
   509
               is_wfrec(M, iterates_MH(M,big_union(M), A), 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   510
                        msn, n, y)))"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   511
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   512
lemma (in M_eclose) eclose_replacement2': 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   513
  "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
   514
apply (insert eclose_replacement2 [of A]) 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   515
apply (rule strong_replacement_cong [THEN iffD1])  
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   516
apply (rule conj_cong [OF iff_refl iterates_abs [of "big_union(M)"]]) 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   517
apply (simp_all add: eclose_replacement1 relativize1_def) 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   518
done
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   519
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   520
lemma (in M_eclose) eclose_closed [intro,simp]:
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   521
     "M(A) ==> M(eclose(A))"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   522
apply (insert eclose_replacement1)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   523
by  (simp add: RepFun_closed2 eclose_eq_Union 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   524
               eclose_replacement2' relativize1_def
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   525
               iterates_closed [of "big_union(M)"])
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   526
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   527
lemma (in M_eclose) is_eclose_n_abs [simp]:
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   528
     "[|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
   529
apply (insert eclose_replacement1)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   530
apply (simp add: is_eclose_n_def relativize1_def nat_into_M
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   531
                 iterates_abs [of "big_union(M)" _ "Union"])
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   532
done
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   533
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   534
lemma (in M_eclose) mem_eclose_abs [simp]:
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   535
     "M(A) ==> mem_eclose(M,A,l) <-> l \<in> eclose(A)"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   536
apply (insert eclose_replacement1)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   537
apply (simp add: mem_eclose_def relativize1_def eclose_eq_Union
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   538
                 iterates_closed [of "big_union(M)"]) 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   539
done
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   540
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   541
lemma (in M_eclose) eclose_abs [simp]:
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   542
     "[|M(A); M(Z)|] ==> is_eclose(M,A,Z) <-> Z = eclose(A)"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   543
apply (simp add: is_eclose_def, safe)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   544
apply (rule M_equalityI, simp_all)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   545
done
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   546
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   547
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
subsection {*Absoluteness for @{term transrec}*}
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   551
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   552
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   553
text{* @{term "transrec(a,H) \<equiv> wfrec(Memrel(eclose({a})), a, H)"} *}
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   554
constdefs
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   555
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   556
  is_transrec :: "[i=>o, [i,i,i]=>o, i, i] => o"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   557
   "is_transrec(M,MH,a,z) == 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   558
      \<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M]. 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   559
       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
   560
       is_wfrec(M,MH,mesa,a,z)"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   561
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   562
  transrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   563
   "transrec_replacement(M,MH,a) ==
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   564
      \<exists>sa[M]. \<exists>esa[M]. \<exists>mesa[M]. 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   565
       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
   566
       wfrec_replacement(M,MH,mesa)"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   567
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   568
text{*The condition @{term "Ord(i)"} lets us use the simpler 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   569
  @{text "trans_wfrec_abs"} rather than @{text "trans_wfrec_abs"},
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   570
  which I haven't even proved yet. *}
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   571
theorem (in M_eclose) transrec_abs:
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   572
  "[|Ord(i);  M(i);  M(z);
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   573
     transrec_replacement(M,MH,i);  relativize2(M,MH,H);
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   574
     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   575
   ==> is_transrec(M,MH,i,z) <-> z = transrec(i,H)" 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   576
by (simp add: trans_wfrec_abs transrec_replacement_def is_transrec_def
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   577
       transrec_def eclose_sing_Ord_eq wf_Memrel trans_Memrel relation_Memrel)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   578
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   579
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   580
theorem (in M_eclose) transrec_closed:
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   581
     "[|Ord(i);  M(i);  M(z);
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   582
	transrec_replacement(M,MH,i);  relativize2(M,MH,H);
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   583
	\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   584
      ==> M(transrec(i,H))"
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   585
by (simp add: trans_wfrec_closed transrec_replacement_def is_transrec_def
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   586
       transrec_def eclose_sing_Ord_eq wf_Memrel trans_Memrel relation_Memrel)
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   587
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   588
4eb948d1eb4e absoluteness for "formula" and "eclose"
paulson
parents: 13386
diff changeset
   589
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   590
subsection{*Absoluteness for the List Operator @{term length}*}
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   591
constdefs
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   592
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   593
  is_length :: "[i=>o,i,i,i] => o"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   594
    "is_length(M,A,l,n) == 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   595
       \<exists>sn[M]. \<exists>list_n[M]. \<exists>list_sn[M]. 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   596
        is_list_N(M,A,n,list_n) & l \<notin> list_n &
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   597
        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
   598
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   599
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   600
lemma (in M_datatypes) length_abs [simp]:
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   601
     "[|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
   602
apply (subgoal_tac "M(l) & M(n)")
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   603
 prefer 2 apply (blast dest: transM)  
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   604
apply (simp add: is_length_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   605
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
   606
             dest: list_N_imp_length_lt)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   607
done
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   608
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   609
text{*Proof is trivial since @{term length} returns natural numbers.*}
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   610
lemma (in M_triv_axioms) length_closed [intro,simp]:
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   611
     "l \<in> list(A) ==> M(length(l))"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   612
by (simp add: nat_into_M ) 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   613
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   614
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   615
subsection {*Absoluteness for @{term nth}*}
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   616
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   617
lemma nth_eq_hd_iterates_tl [rule_format]:
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   618
     "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
   619
apply (induct_tac xs) 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   620
apply (simp add: iterates_tl_Nil hd'_Nil, clarify) 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   621
apply (erule natE)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   622
apply (simp add: hd'_Cons) 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   623
apply (simp add: tl'_Cons iterates_commute) 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   624
done
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   625
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   626
lemma (in M_axioms) iterates_tl'_closed:
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   627
     "[|n \<in> nat; M(x)|] ==> M(tl'^n (x))"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   628
apply (induct_tac n, simp) 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   629
apply (simp add: tl'_Cons tl'_closed) 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   630
done
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   631
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   632
locale (open) M_nth = M_datatypes +
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   633
 assumes nth_replacement1: 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   634
   "M(xs) ==> iterates_replacement(M, %l t. is_tl(M,l,t), xs)"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   635
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   636
text{*Immediate by type-checking*}
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   637
lemma (in M_datatypes) nth_closed [intro,simp]:
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   638
     "[|xs \<in> list(A); n \<in> nat; M(A)|] ==> M(nth(n,xs))" 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   639
apply (case_tac "n < length(xs)")
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   640
 apply (blast intro: nth_type transM)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   641
apply (simp add: not_lt_iff_le nth_eq_0)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   642
done
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   643
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   644
constdefs
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   645
  is_nth :: "[i=>o,i,i,i] => o"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   646
    "is_nth(M,n,l,Z) == 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   647
      \<exists>X[M]. \<exists>sn[M]. \<exists>msn[M]. 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   648
       successor(M,n,sn) & membership(M,sn,msn) &
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   649
       is_wfrec(M, iterates_MH(M, is_tl(M), l), msn, n, X) &
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   650
       is_hd(M,X,Z)"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   651
 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   652
lemma (in M_nth) nth_abs [simp]:
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   653
     "[|M(A); n \<in> nat; l \<in> list(A); M(Z)|] 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   654
      ==> is_nth(M,n,l,Z) <-> Z = nth(n,l)"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   655
apply (subgoal_tac "M(l)") 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   656
 prefer 2 apply (blast intro: transM)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   657
apply (insert nth_replacement1)
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 
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13395
diff changeset
   660
                 iterates_abs [OF _ relativize1_tl])
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
13268
240509babf00 more use of relativized quantifiers
paulson
parents:
diff changeset
   664
end