src/ZF/Constructible/WF_absolute.thy
author paulson
Fri, 16 Aug 2002 16:41:48 +0200
changeset 13505 52a16cb7fefb
parent 13428 99e52e78eb65
child 13506 acc18a5d2b1a
permissions -rw-r--r--
Relativized right up to L satisfies V=L!
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
    ID:         $Id$
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     3
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     4
    Copyright   2002  University of Cambridge
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     5
*)
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     6
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
     7
header {*Absoluteness for Well-Founded Relations and Well-Founded Recursion*}
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
     8
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
     9
theory WF_absolute = WFrec:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    10
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    11
subsection{*Every well-founded relation is a subset of some inverse image of
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    12
      an ordinal*}
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    13
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    14
lemma wf_rvimage_Ord: "Ord(i) \<Longrightarrow> wf(rvimage(A, f, Memrel(i)))"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    15
by (blast intro: wf_rvimage wf_Memrel)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    16
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    17
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    18
constdefs
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    19
  wfrank :: "[i,i]=>i"
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    20
    "wfrank(r,a) == wfrec(r, a, %x f. \<Union>y \<in> r-``{x}. succ(f`y))"
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    21
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    22
constdefs
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    23
  wftype :: "i=>i"
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    24
    "wftype(r) == \<Union>y \<in> range(r). succ(wfrank(r,y))"
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    25
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    26
lemma wfrank: "wf(r) ==> wfrank(r,a) = (\<Union>y \<in> r-``{a}. succ(wfrank(r,y)))"
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    27
by (subst wfrank_def [THEN def_wfrec], simp_all)
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    28
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    29
lemma Ord_wfrank: "wf(r) ==> Ord(wfrank(r,a))"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
    30
apply (rule_tac a=a in wf_induct, assumption)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    31
apply (subst wfrank, assumption)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    32
apply (rule Ord_succ [THEN Ord_UN], blast)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    33
done
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    34
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    35
lemma wfrank_lt: "[|wf(r); <a,b> \<in> r|] ==> wfrank(r,a) < wfrank(r,b)"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
    36
apply (rule_tac a1 = b in wfrank [THEN ssubst], assumption)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    37
apply (rule UN_I [THEN ltI])
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    38
apply (simp add: Ord_wfrank vimage_iff)+
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    39
done
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    40
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    41
lemma Ord_wftype: "wf(r) ==> Ord(wftype(r))"
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    42
by (simp add: wftype_def Ord_wfrank)
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    43
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    44
lemma wftypeI: "\<lbrakk>wf(r);  x \<in> field(r)\<rbrakk> \<Longrightarrow> wfrank(r,x) \<in> wftype(r)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    45
apply (simp add: wftype_def)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    46
apply (blast intro: wfrank_lt [THEN ltD])
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    47
done
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    48
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    49
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    50
lemma wf_imp_subset_rvimage:
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    51
     "[|wf(r); r \<subseteq> A*A|] ==> \<exists>i f. Ord(i) & r <= rvimage(A, f, Memrel(i))"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    52
apply (rule_tac x="wftype(r)" in exI)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    53
apply (rule_tac x="\<lambda>x\<in>A. wfrank(r,x)" in exI)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    54
apply (simp add: Ord_wftype, clarify)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    55
apply (frule subsetD, assumption, clarify)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    56
apply (simp add: rvimage_iff wfrank_lt [THEN ltD])
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    57
apply (blast intro: wftypeI)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    58
done
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    59
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    60
theorem wf_iff_subset_rvimage:
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    61
  "relation(r) ==> wf(r) <-> (\<exists>i f A. Ord(i) & r <= rvimage(A, f, Memrel(i)))"
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    62
by (blast dest!: relation_field_times_field wf_imp_subset_rvimage
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    63
          intro: wf_rvimage_Ord [THEN wf_subset])
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    64
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
    65
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    66
subsection{*Transitive closure without fixedpoints*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    67
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    68
constdefs
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    69
  rtrancl_alt :: "[i,i]=>i"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    70
    "rtrancl_alt(A,r) ==
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    71
       {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
    72
                 (\<exists>x y. p = <x,y> &  f`0 = x & f`n = y) &
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    73
                       (\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)}"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    74
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    75
lemma alt_rtrancl_lemma1 [rule_format]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    76
    "n \<in> nat
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    77
     ==> \<forall>f \<in> succ(n) -> field(r).
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    78
         (\<forall>i\<in>n. \<langle>f`i, f ` succ(i)\<rangle> \<in> r) --> \<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
    79
apply (induct_tac n)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    80
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
    81
apply (rename_tac n f)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    82
apply (rule rtrancl_into_rtrancl)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    83
 prefer 2 apply assumption
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    84
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
    85
 apply (blast intro: restrict_type2)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    86
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
    87
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    88
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    89
lemma rtrancl_alt_subset_rtrancl: "rtrancl_alt(field(r),r) <= r^*"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    90
apply (simp add: rtrancl_alt_def)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    91
apply (blast intro: alt_rtrancl_lemma1)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    92
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    93
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    94
lemma rtrancl_subset_rtrancl_alt: "r^* <= rtrancl_alt(field(r),r)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    95
apply (simp add: rtrancl_alt_def, clarify)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    96
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
    97
apply (erule rtrancl_induct)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    98
 txt{*Base case, trivial*}
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    99
 apply (rule_tac x=0 in bexI)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   100
  apply (rule_tac x="lam x:1. xa" in bexI)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   101
   apply simp_all
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   102
txt{*Inductive step*}
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   103
apply clarify
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   104
apply (rename_tac n f)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   105
apply (rule_tac x="succ(n)" in bexI)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   106
 apply (rule_tac x="lam i: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
   107
  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
   108
  apply (blast intro: mem_asym)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   109
 apply typecheck
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   110
 apply auto
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   111
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   112
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   113
lemma rtrancl_alt_eq_rtrancl: "rtrancl_alt(field(r),r) = r^*"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   114
by (blast del: subsetI
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   115
	  intro: rtrancl_alt_subset_rtrancl rtrancl_subset_rtrancl_alt)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   116
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   117
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   118
constdefs
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   119
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   120
  rtran_closure_mem :: "[i=>o,i,i,i] => o"
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   121
    --{*The property of belonging to @{text "rtran_closure(r)"}*}
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   122
    "rtran_closure_mem(M,A,r,p) ==
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   123
	      \<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M]. 
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   124
               omega(M,nnat) & n\<in>nnat & successor(M,n,n') &
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   125
	       (\<exists>f[M]. typed_function(M,n',A,f) &
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   126
		(\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) &
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   127
		  fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) &
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   128
		  (\<forall>j[M]. j\<in>n --> 
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   129
		    (\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M]. 
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   130
		      fun_apply(M,f,j,fj) & successor(M,j,sj) &
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   131
		      fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))"
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   132
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   133
  rtran_closure :: "[i=>o,i,i] => o"
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   134
    "rtran_closure(M,r,s) == 
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   135
        \<forall>A[M]. is_field(M,r,A) -->
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   136
 	 (\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))"
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   137
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   138
  tran_closure :: "[i=>o,i,i] => o"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   139
    "tran_closure(M,r,t) ==
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   140
         \<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)"
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   141
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   142
lemma (in M_axioms) rtran_closure_mem_iff:
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   143
     "[|M(A); M(r); M(p)|]
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   144
      ==> rtran_closure_mem(M,A,r,p) <->
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   145
          (\<exists>n[M]. n\<in>nat & 
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   146
           (\<exists>f[M]. f \<in> succ(n) -> A &
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   147
            (\<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
   148
                           (\<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
   149
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
   150
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   151
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13418
diff changeset
   152
locale M_trancl = M_axioms +
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   153
  assumes rtrancl_separation:
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   154
	 "[| M(r); M(A) |] ==> separation (M, rtran_closure_mem(M,A,r))"
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   155
      and wellfounded_trancl_separation:
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13306
diff changeset
   156
	 "[| M(r); M(Z) |] ==> 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13306
diff changeset
   157
	  separation (M, \<lambda>x. 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13306
diff changeset
   158
	      \<exists>w[M]. \<exists>wx[M]. \<exists>rp[M]. 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13306
diff changeset
   159
	       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
   160
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   161
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   162
lemma (in M_trancl) rtran_closure_rtrancl:
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   163
     "M(r) ==> rtran_closure(M,r,rtrancl(r))"
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   164
apply (simp add: rtran_closure_def rtran_closure_mem_iff 
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   165
                 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
   166
apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) 
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   167
done
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   168
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   169
lemma (in M_trancl) rtrancl_closed [intro,simp]:
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   170
     "M(r) ==> M(rtrancl(r))"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   171
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
   172
apply (simp add: rtrancl_alt_eq_rtrancl [symmetric]
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   173
                 rtrancl_alt_def rtran_closure_mem_iff)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   174
done
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   175
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   176
lemma (in M_trancl) rtrancl_abs [simp]:
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   177
     "[| M(r); M(z) |] ==> rtran_closure(M,r,z) <-> z = rtrancl(r)"
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   178
apply (rule iffI)
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   179
 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
   180
 prefer 2 apply (blast intro: rtran_closure_rtrancl)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   181
apply (rule M_equalityI)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   182
apply (simp add: rtran_closure_def rtrancl_alt_eq_rtrancl [symmetric]
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
   183
                 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
   184
apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) 
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   185
done
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   186
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   187
lemma (in M_trancl) trancl_closed [intro,simp]:
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   188
     "M(r) ==> M(trancl(r))"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   189
by (simp add: trancl_def comp_closed rtrancl_closed)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   190
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   191
lemma (in M_trancl) trancl_abs [simp]:
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   192
     "[| M(r); M(z) |] ==> tran_closure(M,r,z) <-> z = trancl(r)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   193
by (simp add: tran_closure_def trancl_def)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   194
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13306
diff changeset
   195
lemma (in M_trancl) wellfounded_trancl_separation':
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13306
diff changeset
   196
     "[| 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
   197
by (insert wellfounded_trancl_separation [of r Z], simp) 
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   198
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   199
text{*Alternative proof of @{text wf_on_trancl}; inspiration for the
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   200
      relativized version.  Original version is on theory WF.*}
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   201
lemma "[| wf[A](r);  r-``A <= A |] ==> wf[A](r^+)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   202
apply (simp add: wf_on_def wf_def)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   203
apply (safe intro!: equalityI)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   204
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
   205
apply (blast elim: tranclE)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   206
done
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   207
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   208
lemma (in M_trancl) wellfounded_on_trancl:
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   209
     "[| wellfounded_on(M,A,r);  r-``A <= A; M(r); M(A) |]
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   210
      ==> wellfounded_on(M,A,r^+)"
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   211
apply (simp add: wellfounded_on_def)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   212
apply (safe intro!: equalityI)
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   213
apply (rename_tac Z x)
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   214
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
   215
 prefer 2
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13306
diff changeset
   216
 apply (blast intro: wellfounded_trancl_separation') 
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13293
diff changeset
   217
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
   218
apply (blast dest: transM, simp)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   219
apply (rename_tac y w)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   220
apply (drule_tac x=w in bspec, assumption, clarify)
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   221
apply (erule tranclE)
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   222
  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
   223
 apply blast
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   224
done
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   225
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   226
lemma (in M_trancl) wellfounded_trancl:
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   227
     "[|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
   228
apply (rotate_tac -1)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   229
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
   230
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
   231
   apply blast
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   232
  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
   233
done
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   234
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   235
text{*Relativized to M: Every well-founded relation is a subset of some
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   236
inverse image of an ordinal.  Key step is the construction (in M) of a
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   237
rank function.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   238
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   239
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13418
diff changeset
   240
locale M_wfrank = M_trancl +
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   241
  assumes wfrank_separation:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   242
     "M(r) ==>
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   243
      separation (M, \<lambda>x. 
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   244
         \<forall>rplus[M]. tran_closure(M,r,rplus) -->
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   245
         ~ (\<exists>f[M]. M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f)))"
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   246
 and wfrank_strong_replacement:
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   247
     "M(r) ==>
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   248
      strong_replacement(M, \<lambda>x z. 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   249
         \<forall>rplus[M]. tran_closure(M,r,rplus) -->
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   250
         (\<exists>y[M]. \<exists>f[M]. pair(M,x,y,z)  & 
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   251
                        M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f) &
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   252
                        is_range(M,f,y)))"
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   253
 and Ord_wfrank_separation:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   254
     "M(r) ==>
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   255
      separation (M, \<lambda>x.
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   256
         \<forall>rplus[M]. tran_closure(M,r,rplus) --> 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   257
          ~ (\<forall>f[M]. \<forall>rangef[M]. 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   258
             is_range(M,f,rangef) -->
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   259
             M_is_recfun(M, \<lambda>x f y. is_range(M,f,y), rplus, x, f) -->
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   260
             ordinal(M,rangef)))" 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   261
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   262
text{*Proving that the relativized instances of Separation or Replacement
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   263
agree with the "real" ones.*}
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   264
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   265
lemma (in M_wfrank) wfrank_separation':
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   266
     "M(r) ==>
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   267
      separation
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   268
	   (M, \<lambda>x. ~ (\<exists>f[M]. is_recfun(r^+, x, %x f. range(f), f)))"
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   269
apply (insert wfrank_separation [of r])
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   270
apply (simp add: relativize2_def is_recfun_abs [of "%x. range"])
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   271
done
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   272
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   273
lemma (in M_wfrank) wfrank_strong_replacement':
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   274
     "M(r) ==>
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   275
      strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>f[M]. 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   276
		  pair(M,x,y,z) & is_recfun(r^+, x, %x f. range(f), f) &
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   277
		  y = range(f))"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   278
apply (insert wfrank_strong_replacement [of r])
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   279
apply (simp add: relativize2_def is_recfun_abs [of "%x. range"])
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   280
done
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   281
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   282
lemma (in M_wfrank) Ord_wfrank_separation':
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   283
     "M(r) ==>
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   284
      separation (M, \<lambda>x. 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   285
         ~ (\<forall>f[M]. is_recfun(r^+, x, \<lambda>x. range, f) --> Ord(range(f))))" 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   286
apply (insert Ord_wfrank_separation [of r])
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   287
apply (simp add: relativize2_def is_recfun_abs [of "%x. range"])
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   288
done
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   289
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   290
text{*This function, defined using replacement, is a rank function for
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   291
well-founded relations within the class M.*}
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   292
constdefs
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   293
 wellfoundedrank :: "[i=>o,i,i] => i"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   294
    "wellfoundedrank(M,r,A) ==
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   295
        {p. x\<in>A, \<exists>y[M]. \<exists>f[M]. 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   296
                       p = <x,y> & is_recfun(r^+, x, %x f. range(f), f) &
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   297
                       y = range(f)}"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   298
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   299
lemma (in M_wfrank) exists_wfrank:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   300
    "[| wellfounded(M,r); M(a); M(r) |]
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   301
     ==> \<exists>f[M]. is_recfun(r^+, a, %x f. range(f), f)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   302
apply (rule wellfounded_exists_is_recfun)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   303
      apply (blast intro: wellfounded_trancl)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   304
     apply (rule trans_trancl)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   305
    apply (erule wfrank_separation')
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   306
   apply (erule wfrank_strong_replacement')
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   307
apply (simp_all add: trancl_subset_times)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   308
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   309
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   310
lemma (in M_wfrank) M_wellfoundedrank:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   311
    "[| wellfounded(M,r); M(r); M(A) |] ==> M(wellfoundedrank(M,r,A))"
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   312
apply (insert wfrank_strong_replacement' [of r])
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   313
apply (simp add: wellfoundedrank_def)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   314
apply (rule strong_replacement_closed)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   315
   apply assumption+
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   316
 apply (rule univalent_is_recfun)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   317
   apply (blast intro: wellfounded_trancl)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   318
  apply (rule trans_trancl)
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
   319
 apply (simp add: trancl_subset_times) 
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
   320
apply (blast dest: transM) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   321
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   322
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   323
lemma (in M_wfrank) Ord_wfrank_range [rule_format]:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   324
    "[| wellfounded(M,r); a\<in>A; M(r); M(A) |]
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   325
     ==> \<forall>f[M]. is_recfun(r^+, a, %x f. range(f), f) --> Ord(range(f))"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   326
apply (drule wellfounded_trancl, assumption)
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13418
diff changeset
   327
apply (rule wellfounded_induct, assumption, erule (1) transM)
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   328
  apply simp
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   329
 apply (blast intro: Ord_wfrank_separation', clarify)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   330
txt{*The reasoning in both cases is that we get @{term y} such that
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   331
   @{term "\<langle>y, x\<rangle> \<in> r^+"}.  We find that
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   332
   @{term "f`y = restrict(f, r^+ -`` {y})"}. *}
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   333
apply (rule OrdI [OF _ Ord_is_Transset])
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   334
 txt{*An ordinal is a transitive set...*}
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   335
 apply (simp add: Transset_def)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   336
 apply clarify
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   337
 apply (frule apply_recfun2, assumption)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   338
 apply (force simp add: restrict_iff)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   339
txt{*...of ordinals.  This second case requires the induction hyp.*}
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   340
apply clarify
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   341
apply (rename_tac i y)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   342
apply (frule apply_recfun2, assumption)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   343
apply (frule is_recfun_imp_in_r, assumption)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   344
apply (frule is_recfun_restrict)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   345
    (*simp_all won't work*)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   346
    apply (simp add: trans_trancl trancl_subset_times)+
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   347
apply (drule spec [THEN mp], assumption)
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   348
apply (subgoal_tac "M(restrict(f, r^+ -`` {y}))")
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   349
 apply (drule_tac x="restrict(f, r^+ -`` {y})" in rspec)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   350
apply assumption
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   351
 apply (simp add: function_apply_equality [OF _ is_recfun_imp_function])
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   352
apply (blast dest: pair_components_in_M)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   353
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   354
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   355
lemma (in M_wfrank) Ord_range_wellfoundedrank:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   356
    "[| wellfounded(M,r); r \<subseteq> A*A;  M(r); M(A) |]
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   357
     ==> Ord (range(wellfoundedrank(M,r,A)))"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   358
apply (frule wellfounded_trancl, assumption)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   359
apply (frule trancl_subset_times)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   360
apply (simp add: wellfoundedrank_def)
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   361
apply (rule OrdI [OF _ Ord_is_Transset])
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   362
 prefer 2
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   363
 txt{*by our previous result the range consists of ordinals.*}
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   364
 apply (blast intro: Ord_wfrank_range)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   365
txt{*We still must show that the range is a transitive set.*}
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   366
apply (simp add: Transset_def, clarify, simp)
13293
paulson
parents: 13269
diff changeset
   367
apply (rename_tac x i f u)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   368
apply (frule is_recfun_imp_in_r, assumption)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   369
apply (subgoal_tac "M(u) & M(i) & M(x)")
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   370
 prefer 2 apply (blast dest: transM, clarify)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   371
apply (rule_tac a=u in rangeI)
13293
paulson
parents: 13269
diff changeset
   372
apply (rule_tac x=u in ReplaceI)
paulson
parents: 13269
diff changeset
   373
  apply simp 
paulson
parents: 13269
diff changeset
   374
  apply (rule_tac x="restrict(f, r^+ -`` {u})" in rexI)
paulson
parents: 13269
diff changeset
   375
   apply (blast intro: is_recfun_restrict trans_trancl dest: apply_recfun2)
paulson
parents: 13269
diff changeset
   376
  apply simp 
paulson
parents: 13269
diff changeset
   377
apply blast 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   378
txt{*Unicity requirement of Replacement*}
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   379
apply clarify
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   380
apply (frule apply_recfun2, assumption)
13293
paulson
parents: 13269
diff changeset
   381
apply (simp add: trans_trancl is_recfun_cut)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   382
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   383
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   384
lemma (in M_wfrank) function_wellfoundedrank:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   385
    "[| wellfounded(M,r); M(r); M(A)|]
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   386
     ==> function(wellfoundedrank(M,r,A))"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   387
apply (simp add: wellfoundedrank_def function_def, clarify)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   388
txt{*Uniqueness: repeated below!*}
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   389
apply (drule is_recfun_functional, assumption)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   390
     apply (blast intro: wellfounded_trancl)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   391
    apply (simp_all add: trancl_subset_times trans_trancl)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   392
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   393
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   394
lemma (in M_wfrank) domain_wellfoundedrank:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   395
    "[| wellfounded(M,r); M(r); M(A)|]
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   396
     ==> domain(wellfoundedrank(M,r,A)) = A"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   397
apply (simp add: wellfoundedrank_def function_def)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   398
apply (rule equalityI, auto)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   399
apply (frule transM, assumption)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   400
apply (frule_tac a=x in exists_wfrank, assumption+, clarify)
13293
paulson
parents: 13269
diff changeset
   401
apply (rule_tac b="range(f)" in domainI)
paulson
parents: 13269
diff changeset
   402
apply (rule_tac x=x in ReplaceI)
paulson
parents: 13269
diff changeset
   403
  apply simp 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   404
  apply (rule_tac x=f in rexI, blast, simp_all)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   405
txt{*Uniqueness (for Replacement): repeated above!*}
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   406
apply clarify
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   407
apply (drule is_recfun_functional, assumption)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   408
    apply (blast intro: wellfounded_trancl)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   409
    apply (simp_all add: trancl_subset_times trans_trancl)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   410
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   411
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   412
lemma (in M_wfrank) wellfoundedrank_type:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   413
    "[| wellfounded(M,r);  M(r); M(A)|]
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   414
     ==> wellfoundedrank(M,r,A) \<in> A -> range(wellfoundedrank(M,r,A))"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   415
apply (frule function_wellfoundedrank [of r A], assumption+)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   416
apply (frule function_imp_Pi)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   417
 apply (simp add: wellfoundedrank_def relation_def)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   418
 apply blast
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   419
apply (simp add: domain_wellfoundedrank)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   420
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   421
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   422
lemma (in M_wfrank) Ord_wellfoundedrank:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   423
    "[| wellfounded(M,r); a \<in> A; r \<subseteq> A*A;  M(r); M(A) |]
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   424
     ==> Ord(wellfoundedrank(M,r,A) ` a)"
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   425
by (blast intro: apply_funtype [OF wellfoundedrank_type]
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   426
                 Ord_in_Ord [OF Ord_range_wellfoundedrank])
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   427
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   428
lemma (in M_wfrank) wellfoundedrank_eq:
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   429
     "[| is_recfun(r^+, a, %x. range, f);
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   430
         wellfounded(M,r);  a \<in> A; M(f); M(r); M(A)|]
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   431
      ==> wellfoundedrank(M,r,A) ` a = range(f)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   432
apply (rule apply_equality)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   433
 prefer 2 apply (blast intro: wellfoundedrank_type)
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   434
apply (simp add: wellfoundedrank_def)
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   435
apply (rule ReplaceI)
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   436
  apply (rule_tac x="range(f)" in rexI) 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   437
  apply blast
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   438
 apply simp_all
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   439
txt{*Unicity requirement of Replacement*}
13242
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   440
apply clarify
f96bd927dd37 towards absoluteness of wf
paulson
parents: 13223
diff changeset
   441
apply (drule is_recfun_functional, assumption)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   442
    apply (blast intro: wellfounded_trancl)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   443
    apply (simp_all add: trancl_subset_times trans_trancl)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   444
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   445
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   446
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   447
lemma (in M_wfrank) wellfoundedrank_lt:
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   448
     "[| <a,b> \<in> r;
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   449
         wellfounded(M,r); r \<subseteq> A*A;  M(r); M(A)|]
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   450
      ==> wellfoundedrank(M,r,A) ` a < wellfoundedrank(M,r,A) ` b"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   451
apply (frule wellfounded_trancl, assumption)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   452
apply (subgoal_tac "a\<in>A & b\<in>A")
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   453
 prefer 2 apply blast
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   454
apply (simp add: lt_def Ord_wellfoundedrank, clarify)
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13418
diff changeset
   455
apply (frule exists_wfrank [of concl: _ b], erule (1) transM, assumption)
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13418
diff changeset
   456
apply clarify
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   457
apply (rename_tac fb)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   458
apply (frule is_recfun_restrict [of concl: "r^+" a])
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   459
    apply (rule trans_trancl, assumption)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   460
   apply (simp_all add: r_into_trancl trancl_subset_times)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   461
txt{*Still the same goal, but with new @{text is_recfun} assumptions.*}
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   462
apply (simp add: wellfoundedrank_eq)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   463
apply (frule_tac a=a in wellfoundedrank_eq, assumption+)
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   464
   apply (simp_all add: transM [of a])
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   465
txt{*We have used equations for wellfoundedrank and now must use some
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   466
    for  @{text is_recfun}. *}
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   467
apply (rule_tac a=a in rangeI)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   468
apply (simp add: is_recfun_type [THEN apply_iff] vimage_singleton_iff
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   469
                 r_into_trancl apply_recfun r_into_trancl)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   470
done
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   471
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   472
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   473
lemma (in M_wfrank) wellfounded_imp_subset_rvimage:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   474
     "[|wellfounded(M,r); r \<subseteq> A*A; M(r); M(A)|]
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   475
      ==> \<exists>i f. Ord(i) & r <= rvimage(A, f, Memrel(i))"
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   476
apply (rule_tac x="range(wellfoundedrank(M,r,A))" in exI)
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   477
apply (rule_tac x="wellfoundedrank(M,r,A)" in exI)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   478
apply (simp add: Ord_range_wellfoundedrank, clarify)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   479
apply (frule subsetD, assumption, clarify)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   480
apply (simp add: rvimage_iff wellfoundedrank_lt [THEN ltD])
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   481
apply (blast intro: apply_rangeI wellfoundedrank_type)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   482
done
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   483
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   484
lemma (in M_wfrank) wellfounded_imp_wf:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   485
     "[|wellfounded(M,r); relation(r); M(r)|] ==> wf(r)"
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   486
by (blast dest!: relation_field_times_field wellfounded_imp_subset_rvimage
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   487
          intro: wf_rvimage_Ord [THEN wf_subset])
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   488
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   489
lemma (in M_wfrank) wellfounded_on_imp_wf_on:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   490
     "[|wellfounded_on(M,A,r); relation(r); M(r); M(A)|] ==> wf[A](r)"
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   491
apply (simp add: wellfounded_on_iff_wellfounded wf_on_def)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   492
apply (rule wellfounded_imp_wf)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   493
apply (simp_all add: relation_def)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   494
done
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   495
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   496
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   497
theorem (in M_wfrank) wf_abs [simp]:
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   498
     "[|relation(r); M(r)|] ==> wellfounded(M,r) <-> wf(r)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   499
by (blast intro: wellfounded_imp_wf wf_imp_relativized)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   500
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   501
theorem (in M_wfrank) wf_on_abs [simp]:
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   502
     "[|relation(r); M(r); M(A)|] ==> wellfounded_on(M,A,r) <-> wf[A](r)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   503
by (blast intro: wellfounded_on_imp_wf_on wf_on_imp_relativized)
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13242
diff changeset
   504
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   505
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   506
text{*absoluteness for wfrec-defined functions.*}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   507
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   508
(*first use is_recfun, then M_is_recfun*)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   509
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   510
lemma (in M_trancl) wfrec_relativize:
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   511
  "[|wf(r); M(a); M(r);  
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   512
     strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   513
          pair(M,x,y,z) & 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   514
          is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   515
          y = H(x, restrict(g, r -`` {x}))); 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   516
     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   517
   ==> wfrec(r,a,H) = z <-> 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   518
       (\<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
   519
            z = H(a,restrict(f,r-``{a})))"
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   520
apply (frule wf_trancl) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   521
apply (simp add: wftrec_def wfrec_def, safe)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   522
 apply (frule wf_exists_is_recfun 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   523
              [of concl: "r^+" a "\<lambda>x f. H(x, restrict(f, r -`` {x}))"]) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   524
      apply (simp_all add: trans_trancl function_restrictI trancl_subset_times)
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   525
 apply (clarify, rule_tac x=x in rexI) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   526
 apply (simp_all add: the_recfun_eq trans_trancl trancl_subset_times)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   527
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   528
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   529
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   530
text{*Assuming @{term r} is transitive simplifies the occurrences of @{text H}.
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   531
      The premise @{term "relation(r)"} is necessary 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   532
      before we can replace @{term "r^+"} by @{term r}. *}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   533
theorem (in M_trancl) trans_wfrec_relativize:
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   534
  "[|wf(r);  trans(r);  relation(r);  M(r);  M(a);
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   535
     wfrec_replacement(M,MH,r);  relativize2(M,MH,H);
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   536
     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   537
   ==> wfrec(r,a,H) = z <-> (\<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
   538
apply (frule wfrec_replacement', assumption+) 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   539
apply (simp cong: is_recfun_cong
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   540
           add: wfrec_relativize trancl_eq_r
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   541
                is_recfun_restrict_idem domain_restrict_idem)
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   542
done
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   543
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   544
theorem (in M_trancl) trans_wfrec_abs:
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   545
  "[|wf(r);  trans(r);  relation(r);  M(r);  M(a);  M(z);
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   546
     wfrec_replacement(M,MH,r);  relativize2(M,MH,H);
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   547
     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   548
   ==> is_wfrec(M,MH,r,a,z) <-> z=wfrec(r,a,H)" 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   549
apply (simp add: trans_wfrec_relativize [THEN iff_sym] is_wfrec_abs, blast) 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   550
done
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   551
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   552
lemma (in M_trancl) trans_eq_pair_wfrec_iff:
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   553
  "[|wf(r);  trans(r); relation(r); M(r);  M(y); 
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   554
     wfrec_replacement(M,MH,r);  relativize2(M,MH,H);
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   555
     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   556
   ==> y = <x, wfrec(r, x, H)> <-> 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   557
       (\<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)"
13293
paulson
parents: 13269
diff changeset
   558
apply safe 
paulson
parents: 13269
diff changeset
   559
 apply (simp add: trans_wfrec_relativize [THEN iff_sym, of concl: _ x]) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   560
txt{*converse direction*}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   561
apply (rule sym)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   562
apply (simp add: trans_wfrec_relativize, blast) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   563
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   564
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   565
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   566
subsection{*M is closed under well-founded recursion*}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   567
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   568
text{*Lemma with the awkward premise mentioning @{text wfrec}.*}
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   569
lemma (in M_wfrank) wfrec_closed_lemma [rule_format]:
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   570
     "[|wf(r); M(r); 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   571
        strong_replacement(M, \<lambda>x y. y = \<langle>x, wfrec(r, x, H)\<rangle>);
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   572
        \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   573
      ==> M(a) --> M(wfrec(r,a,H))"
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   574
apply (rule_tac a=a in wf_induct, assumption+)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   575
apply (subst wfrec, assumption, clarify)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   576
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
   577
       in rspec [THEN rspec]) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   578
apply (simp_all add: function_lam) 
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
   579
apply (blast intro: lam_closed dest: pair_components_in_M) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   580
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   581
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   582
text{*Eliminates one instance of replacement.*}
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   583
lemma (in M_wfrank) wfrec_replacement_iff:
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   584
     "strong_replacement(M, \<lambda>x z. 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   585
          \<exists>y[M]. pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g))) <->
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   586
      strong_replacement(M, 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   587
           \<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
   588
apply simp 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   589
apply (rule strong_replacement_cong, blast) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   590
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   591
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   592
text{*Useful version for transitive relations*}
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   593
theorem (in M_wfrank) trans_wfrec_closed:
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   594
     "[|wf(r); trans(r); relation(r); M(r); M(a);
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   595
       wfrec_replacement(M,MH,r);  relativize2(M,MH,H);
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   596
        \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   597
      ==> M(wfrec(r,a,H))"
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   598
apply (frule wfrec_replacement', assumption+) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   599
apply (frule wfrec_replacement_iff [THEN iffD1]) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   600
apply (rule wfrec_closed_lemma, assumption+) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   601
apply (simp_all add: wfrec_replacement_iff trans_eq_pair_wfrec_iff) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   602
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   603
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   604
section{*Absoluteness without assuming transitivity*}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   605
lemma (in M_trancl) eq_pair_wfrec_iff:
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   606
  "[|wf(r);  M(r);  M(y); 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   607
     strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M].
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   608
          pair(M,x,y,z) & 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   609
          is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   610
          y = H(x, restrict(g, r -`` {x}))); 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   611
     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   612
   ==> y = <x, wfrec(r, x, H)> <-> 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   613
       (\<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
   614
            y = <x, H(x,restrict(f,r-``{x}))>)"
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   615
apply safe  
13293
paulson
parents: 13269
diff changeset
   616
 apply (simp add: wfrec_relativize [THEN iff_sym, of concl: _ x]) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   617
txt{*converse direction*}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   618
apply (rule sym)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   619
apply (simp add: wfrec_relativize, blast) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   620
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   621
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   622
text{*Full version not assuming transitivity, but maybe not very useful.*}
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   623
theorem (in M_wfrank) wfrec_closed:
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   624
     "[|wf(r); M(r); M(a);
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   625
        wfrec_replacement(M,MH,r^+);  
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   626
        relativize2(M,MH, \<lambda>x f. H(x, restrict(f, r -`` {x})));
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   627
        \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   628
      ==> M(wfrec(r,a,H))"
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   629
apply (frule wfrec_replacement' 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   630
               [of MH "r^+" "\<lambda>x f. H(x, restrict(f, r -`` {x}))"])
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   631
   prefer 4
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   632
   apply (frule wfrec_replacement_iff [THEN iffD1]) 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   633
   apply (rule wfrec_closed_lemma, assumption+) 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   634
     apply (simp_all add: eq_pair_wfrec_iff func.function_restrictI) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   635
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   636
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   637
end