src/ZF/Constructible/WF_absolute.thy
author wenzelm
Sun, 02 Nov 2014 16:39:54 +0100
changeset 58871 c399ae4b836f
parent 46823 57bf0cecb366
child 60770 240563fbf41d
permissions -rw-r--r--
modernized header;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     1
(*  Title:      ZF/Constructible/WF_absolute.thy
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     3
*)
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     4
58871
c399ae4b836f modernized header;
wenzelm
parents: 46823
diff changeset
     5
section {*Absoluteness of Well-Founded Recursion*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
     6
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13647
diff changeset
     7
theory WF_absolute imports WFrec begin
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     8
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     9
subsection{*Transitive closure without fixedpoints*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    10
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    11
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    12
  rtrancl_alt :: "[i,i]=>i" where
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    13
    "rtrancl_alt(A,r) ==
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    14
       {p \<in> A*A. \<exists>n\<in>nat. \<exists>f \<in> succ(n) -> A.
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
    15
                 (\<exists>x y. p = <x,y> &  f`0 = x & f`n = y) &
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    16
                       (\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)}"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    17
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    18
lemma alt_rtrancl_lemma1 [rule_format]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    19
    "n \<in> nat
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    20
     ==> \<forall>f \<in> succ(n) -> field(r).
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    21
         (\<forall>i\<in>n. \<langle>f`i, f ` succ(i)\<rangle> \<in> r) \<longrightarrow> \<langle>f`0, f`n\<rangle> \<in> r^*"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    22
apply (induct_tac n)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    23
apply (simp_all add: apply_funtype rtrancl_refl, clarify)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    24
apply (rename_tac n f)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    25
apply (rule rtrancl_into_rtrancl)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    26
 prefer 2 apply assumption
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    27
apply (drule_tac x="restrict(f,succ(n))" in bspec)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    28
 apply (blast intro: restrict_type2)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    29
apply (simp add: Ord_succ_mem_iff nat_0_le [THEN ltD] leI [THEN ltD] ltI)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    30
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    31
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    32
lemma rtrancl_alt_subset_rtrancl: "rtrancl_alt(field(r),r) \<subseteq> r^*"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    33
apply (simp add: rtrancl_alt_def)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    34
apply (blast intro: alt_rtrancl_lemma1)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    35
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    36
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    37
lemma rtrancl_subset_rtrancl_alt: "r^* \<subseteq> rtrancl_alt(field(r),r)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    38
apply (simp add: rtrancl_alt_def, clarify)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    39
apply (frule rtrancl_type [THEN subsetD], clarify, simp)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    40
apply (erule rtrancl_induct)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    41
 txt{*Base case, trivial*}
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    42
 apply (rule_tac x=0 in bexI)
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    43
  apply (rule_tac x="\<lambda>x\<in>1. xa" in bexI)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    44
   apply simp_all
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    45
txt{*Inductive step*}
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    46
apply clarify
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    47
apply (rename_tac n f)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    48
apply (rule_tac x="succ(n)" in bexI)
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    49
 apply (rule_tac x="\<lambda>i\<in>succ(succ(n)). if i=succ(n) then z else f`i" in bexI)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    50
  apply (simp add: Ord_succ_mem_iff nat_0_le [THEN ltD] leI [THEN ltD] ltI)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    51
  apply (blast intro: mem_asym)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    52
 apply typecheck
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    53
 apply auto
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    54
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    55
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    56
lemma rtrancl_alt_eq_rtrancl: "rtrancl_alt(field(r),r) = r^*"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    57
by (blast del: subsetI
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
    58
          intro: rtrancl_alt_subset_rtrancl rtrancl_subset_rtrancl_alt)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    59
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    60
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    61
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    62
  rtran_closure_mem :: "[i=>o,i,i,i] => o" where
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
    63
    --{*The property of belonging to @{text "rtran_closure(r)"}*}
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
    64
    "rtran_closure_mem(M,A,r,p) ==
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
    65
              \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M]. 
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
    66
               omega(M,nnat) & n\<in>nnat & successor(M,n,n') &
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
    67
               (\<exists>f[M]. typed_function(M,n',A,f) &
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
    68
                (\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) &
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
    69
                  fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    70
                  (\<forall>j[M]. j\<in>n \<longrightarrow> 
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
    71
                    (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M]. 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
    72
                      fun_apply(M,f,j,fj) & successor(M,j,sj) &
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
    73
                      fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
    74
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    75
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    76
  rtran_closure :: "[i=>o,i,i] => o" where
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
    77
    "rtran_closure(M,r,s) == 
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    78
        \<forall>A[M]. is_field(M,r,A) \<longrightarrow>
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    79
         (\<forall>p[M]. p \<in> s \<longleftrightarrow> rtran_closure_mem(M,A,r,p))"
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
    80
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    81
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    82
  tran_closure :: "[i=>o,i,i] => o" where
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    83
    "tran_closure(M,r,t) ==
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
    84
         \<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)"
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
    85
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
    86
lemma (in M_basic) rtran_closure_mem_iff:
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
    87
     "[|M(A); M(r); M(p)|]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    88
      ==> rtran_closure_mem(M,A,r,p) \<longleftrightarrow>
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
    89
          (\<exists>n[M]. n\<in>nat & 
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
    90
           (\<exists>f[M]. f \<in> succ(n) -> A &
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
    91
            (\<exists>x[M]. \<exists>y[M]. p = <x,y> & f`0 = x & f`n = y) &
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
    92
                           (\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)))"
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
    93
by (simp add: rtran_closure_mem_def Ord_succ_mem_iff nat_0_le [THEN ltD]) 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
    94
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
    95
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
    96
locale M_trancl = M_basic +
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
    97
  assumes rtrancl_separation:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
    98
         "[| M(r); M(A) |] ==> separation (M, rtran_closure_mem(M,A,r))"
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
    99
      and wellfounded_trancl_separation:
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
   100
         "[| M(r); M(Z) |] ==> 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
   101
          separation (M, \<lambda>x. 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
   102
              \<exists>w[M]. \<exists>wx[M]. \<exists>rp[M]. 
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
   103
               w \<in> Z & pair(M,w,x,wx) & tran_closure(M,r,rp) & wx \<in> rp)"
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   104
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   105
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   106
lemma (in M_trancl) rtran_closure_rtrancl:
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   107
     "M(r) ==> rtran_closure(M,r,rtrancl(r))"
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   108
apply (simp add: rtran_closure_def rtran_closure_mem_iff 
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   109
                 rtrancl_alt_eq_rtrancl [symmetric] rtrancl_alt_def)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   110
apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) 
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   111
done
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   112
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   113
lemma (in M_trancl) rtrancl_closed [intro,simp]:
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   114
     "M(r) ==> M(rtrancl(r))"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   115
apply (insert rtrancl_separation [of r "field(r)"])
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   116
apply (simp add: rtrancl_alt_eq_rtrancl [symmetric]
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   117
                 rtrancl_alt_def rtran_closure_mem_iff)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   118
done
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   119
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   120
lemma (in M_trancl) rtrancl_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   121
     "[| M(r); M(z) |] ==> rtran_closure(M,r,z) \<longleftrightarrow> z = rtrancl(r)"
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   122
apply (rule iffI)
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   123
 txt{*Proving the right-to-left implication*}
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   124
 prefer 2 apply (blast intro: rtran_closure_rtrancl)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   125
apply (rule M_equalityI)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   126
apply (simp add: rtran_closure_def rtrancl_alt_eq_rtrancl [symmetric]
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   127
                 rtrancl_alt_def rtran_closure_mem_iff)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   128
apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) 
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   129
done
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   130
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   131
lemma (in M_trancl) trancl_closed [intro,simp]:
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   132
     "M(r) ==> M(trancl(r))"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   133
by (simp add: trancl_def comp_closed rtrancl_closed)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   134
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   135
lemma (in M_trancl) trancl_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   136
     "[| M(r); M(z) |] ==> tran_closure(M,r,z) \<longleftrightarrow> z = trancl(r)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   137
by (simp add: tran_closure_def trancl_def)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   138
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13306
diff changeset
   139
lemma (in M_trancl) wellfounded_trancl_separation':
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13306
diff changeset
   140
     "[| M(r); M(Z) |] ==> separation (M, \<lambda>x. \<exists>w[M]. w \<in> Z & <w,x> \<in> r^+)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13306
diff changeset
   141
by (insert wellfounded_trancl_separation [of r Z], simp) 
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   142
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   143
text{*Alternative proof of @{text wf_on_trancl}; inspiration for the
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   144
      relativized version.  Original version is on theory WF.*}
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   145
lemma "[| wf[A](r);  r-``A \<subseteq> A |] ==> wf[A](r^+)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   146
apply (simp add: wf_on_def wf_def)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   147
apply (safe intro!: equalityI)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   148
apply (drule_tac x = "{x\<in>A. \<exists>w. \<langle>w,x\<rangle> \<in> r^+ & w \<in> Z}" in spec)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   149
apply (blast elim: tranclE)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   150
done
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   151
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   152
lemma (in M_trancl) wellfounded_on_trancl:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   153
     "[| wellfounded_on(M,A,r);  r-``A \<subseteq> A; M(r); M(A) |]
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   154
      ==> wellfounded_on(M,A,r^+)"
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   155
apply (simp add: wellfounded_on_def)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   156
apply (safe intro!: equalityI)
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   157
apply (rename_tac Z x)
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   158
apply (subgoal_tac "M({x\<in>A. \<exists>w[M]. w \<in> Z & \<langle>w,x\<rangle> \<in> r^+})")
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   159
 prefer 2
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13306
diff changeset
   160
 apply (blast intro: wellfounded_trancl_separation') 
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13293
diff changeset
   161
apply (drule_tac x = "{x\<in>A. \<exists>w[M]. w \<in> Z & \<langle>w,x\<rangle> \<in> r^+}" in rspec, safe)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   162
apply (blast dest: transM, simp)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   163
apply (rename_tac y w)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   164
apply (drule_tac x=w in bspec, assumption, clarify)
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   165
apply (erule tranclE)
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   166
  apply (blast dest: transM)   (*transM is needed to prove M(xa)*)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   167
 apply blast
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   168
done
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   169
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   170
lemma (in M_trancl) wellfounded_trancl:
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   171
     "[|wellfounded(M,r); M(r)|] ==> wellfounded(M,r^+)"
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   172
apply (simp add: wellfounded_iff_wellfounded_on_field)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   173
apply (rule wellfounded_on_subset_A, erule wellfounded_on_trancl)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   174
   apply blast
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   175
  apply (simp_all add: trancl_type [THEN field_rel_subset])
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   176
done
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   177
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   178
13647
7f6f0ffc45c3 tidying and reorganization
paulson
parents: 13634
diff changeset
   179
text{*Absoluteness for wfrec-defined functions.*}
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   180
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   181
(*first use is_recfun, then M_is_recfun*)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   182
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   183
lemma (in M_trancl) wfrec_relativize:
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   184
  "[|wf(r); M(a); M(r);  
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   185
     strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   186
          pair(M,x,y,z) & 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   187
          is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   188
          y = H(x, restrict(g, r -`` {x}))); 
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   189
     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   190
   ==> wfrec(r,a,H) = z \<longleftrightarrow> 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   191
       (\<exists>f[M]. is_recfun(r^+, a, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   192
            z = H(a,restrict(f,r-``{a})))"
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   193
apply (frule wf_trancl) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   194
apply (simp add: wftrec_def wfrec_def, safe)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   195
 apply (frule wf_exists_is_recfun 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   196
              [of concl: "r^+" a "\<lambda>x f. H(x, restrict(f, r -`` {x}))"]) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   197
      apply (simp_all add: trans_trancl function_restrictI trancl_subset_times)
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   198
 apply (clarify, rule_tac x=x in rexI) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   199
 apply (simp_all add: the_recfun_eq trans_trancl trancl_subset_times)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   200
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   201
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   202
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   203
text{*Assuming @{term r} is transitive simplifies the occurrences of @{text H}.
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   204
      The premise @{term "relation(r)"} is necessary 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   205
      before we can replace @{term "r^+"} by @{term r}. *}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   206
theorem (in M_trancl) trans_wfrec_relativize:
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   207
  "[|wf(r);  trans(r);  relation(r);  M(r);  M(a);
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   208
     wfrec_replacement(M,MH,r);  relation2(M,MH,H);
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   209
     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   210
   ==> wfrec(r,a,H) = z \<longleftrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f) & z = H(a,f))" 
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   211
apply (frule wfrec_replacement', assumption+) 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   212
apply (simp cong: is_recfun_cong
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   213
           add: wfrec_relativize trancl_eq_r
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   214
                is_recfun_restrict_idem domain_restrict_idem)
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   215
done
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   216
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   217
theorem (in M_trancl) trans_wfrec_abs:
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   218
  "[|wf(r);  trans(r);  relation(r);  M(r);  M(a);  M(z);
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   219
     wfrec_replacement(M,MH,r);  relation2(M,MH,H);
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   220
     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   221
   ==> is_wfrec(M,MH,r,a,z) \<longleftrightarrow> z=wfrec(r,a,H)" 
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   222
by (simp add: trans_wfrec_relativize [THEN iff_sym] is_wfrec_abs, blast) 
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   223
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   224
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   225
lemma (in M_trancl) trans_eq_pair_wfrec_iff:
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   226
  "[|wf(r);  trans(r); relation(r); M(r);  M(y); 
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   227
     wfrec_replacement(M,MH,r);  relation2(M,MH,H);
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   228
     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   229
   ==> y = <x, wfrec(r, x, H)> \<longleftrightarrow> 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   230
       (\<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
13293
paulson
parents: 13269
diff changeset
   231
apply safe 
paulson
parents: 13269
diff changeset
   232
 apply (simp add: trans_wfrec_relativize [THEN iff_sym, of concl: _ x]) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   233
txt{*converse direction*}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   234
apply (rule sym)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   235
apply (simp add: trans_wfrec_relativize, blast) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   236
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   237
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   238
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   239
subsection{*M is closed under well-founded recursion*}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   240
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   241
text{*Lemma with the awkward premise mentioning @{text wfrec}.*}
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   242
lemma (in M_trancl) wfrec_closed_lemma [rule_format]:
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   243
     "[|wf(r); M(r); 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   244
        strong_replacement(M, \<lambda>x y. y = \<langle>x, wfrec(r, x, H)\<rangle>);
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   245
        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] 
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   246
      ==> M(a) \<longrightarrow> M(wfrec(r,a,H))"
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   247
apply (rule_tac a=a in wf_induct, assumption+)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   248
apply (subst wfrec, assumption, clarify)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   249
apply (drule_tac x1=x and x="\<lambda>x\<in>r -`` {x}. wfrec(r, x, H)" 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   250
       in rspec [THEN rspec]) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   251
apply (simp_all add: function_lam) 
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
   252
apply (blast intro: lam_closed dest: pair_components_in_M) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   253
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   255
text{*Eliminates one instance of replacement.*}
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   256
lemma (in M_trancl) wfrec_replacement_iff:
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   257
     "strong_replacement(M, \<lambda>x z. 
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   258
          \<exists>y[M]. pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g))) \<longleftrightarrow>
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   259
      strong_replacement(M, 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   260
           \<lambda>x y. \<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   261
apply simp 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   262
apply (rule strong_replacement_cong, blast) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   263
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   264
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   265
text{*Useful version for transitive relations*}
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   266
theorem (in M_trancl) trans_wfrec_closed:
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   267
     "[|wf(r); trans(r); relation(r); M(r); M(a);
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   268
       wfrec_replacement(M,MH,r);  relation2(M,MH,H);
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   269
        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   270
      ==> M(wfrec(r,a,H))"
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   271
apply (frule wfrec_replacement', assumption+) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   272
apply (frule wfrec_replacement_iff [THEN iffD1]) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   273
apply (rule wfrec_closed_lemma, assumption+) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   274
apply (simp_all add: wfrec_replacement_iff trans_eq_pair_wfrec_iff) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   275
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   276
13506
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
   277
subsection{*Absoluteness without assuming transitivity*}
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   278
lemma (in M_trancl) eq_pair_wfrec_iff:
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   279
  "[|wf(r);  M(r);  M(y); 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   280
     strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   281
          pair(M,x,y,z) & 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   282
          is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   283
          y = H(x, restrict(g, r -`` {x}))); 
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   284
     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] 
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   285
   ==> y = <x, wfrec(r, x, H)> \<longleftrightarrow> 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   286
       (\<exists>f[M]. is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   287
            y = <x, H(x,restrict(f,r-``{x}))>)"
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   288
apply safe  
13293
paulson
parents: 13269
diff changeset
   289
 apply (simp add: wfrec_relativize [THEN iff_sym, of concl: _ x]) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   290
txt{*converse direction*}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   291
apply (rule sym)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   292
apply (simp add: wfrec_relativize, blast) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   293
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   294
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   295
text{*Full version not assuming transitivity, but maybe not very useful.*}
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   296
theorem (in M_trancl) wfrec_closed:
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   297
     "[|wf(r); M(r); M(a);
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   298
        wfrec_replacement(M,MH,r^+);  
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   299
        relation2(M,MH, \<lambda>x f. H(x, restrict(f, r -`` {x})));
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   300
        \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   301
      ==> M(wfrec(r,a,H))"
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   302
apply (frule wfrec_replacement' 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   303
               [of MH "r^+" "\<lambda>x f. H(x, restrict(f, r -`` {x}))"])
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   304
   prefer 4
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   305
   apply (frule wfrec_replacement_iff [THEN iffD1]) 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   306
   apply (rule wfrec_closed_lemma, assumption+) 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   307
     apply (simp_all add: eq_pair_wfrec_iff func.function_restrictI) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   308
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   309
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   310
end