src/ZF/Constructible/WFrec.thy
author paulson
Fri, 16 Aug 2002 17:19:43 +0200
changeset 13506 acc18a5d2b1a
parent 13505 52a16cb7fefb
child 13564 1500a2e48d44
permissions -rw-r--r--
Various tweaks of the presentation
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/WFrec.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{*Relativized Well-Founded Recursion*}
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
     8
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     9
theory WFrec = Wellorderings:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    10
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    11
13506
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
    12
subsection{*General Lemmas*}
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
    13
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    14
(*Many of these might be useful in WF.thy*)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    15
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
    16
lemma apply_recfun2:
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
    17
    "[| is_recfun(r,a,H,f); <x,i>:f |] ==> i = H(x, restrict(f,r-``{x}))"
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
    18
apply (frule apply_recfun) 
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
    19
 apply (blast dest: is_recfun_type fun_is_rel) 
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
    20
apply (simp add: function_apply_equality [OF _ is_recfun_imp_function])
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    21
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    22
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    23
text{*Expresses @{text is_recfun} as a recursion equation*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    24
lemma is_recfun_iff_equation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    25
     "is_recfun(r,a,H,f) <->
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    26
	   f \<in> r -`` {a} \<rightarrow> range(f) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    27
	   (\<forall>x \<in> r-``{a}. f`x = H(x, restrict(f, r-``{x})))"  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    28
apply (rule iffI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    29
 apply (simp add: is_recfun_type apply_recfun Ball_def vimage_singleton_iff, 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    30
        clarify)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    31
apply (simp add: is_recfun_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    32
apply (rule fun_extension) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    33
  apply assumption
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    34
 apply (fast intro: lam_type, simp) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    35
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    36
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    37
lemma is_recfun_imp_in_r: "[|is_recfun(r,a,H,f); \<langle>x,i\<rangle> \<in> f|] ==> \<langle>x, a\<rangle> \<in> r"
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
    38
by (blast dest: is_recfun_type fun_is_rel)
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    39
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    40
lemma trans_Int_eq:
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    41
      "[| trans(r); <y,x> \<in> r |] ==> r -`` {x} \<inter> r -`` {y} = r -`` {y}"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    42
by (blast intro: transD) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    43
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    44
lemma is_recfun_restrict_idem:
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    45
     "is_recfun(r,a,H,f) ==> restrict(f, r -`` {a}) = f"
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    46
apply (drule is_recfun_type)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    47
apply (auto simp add: Pi_iff subset_Sigma_imp_relation restrict_idem)  
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    48
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    49
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    50
lemma is_recfun_cong_lemma:
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    51
  "[| is_recfun(r,a,H,f); r = r'; a = a'; f = f'; 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    52
      !!x g. [| <x,a'> \<in> r'; relation(g); domain(g) <= r' -``{x} |] 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    53
             ==> H(x,g) = H'(x,g) |]
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    54
   ==> is_recfun(r',a',H',f')"
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    55
apply (simp add: is_recfun_def) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    56
apply (erule trans) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    57
apply (rule lam_cong) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    58
apply (simp_all add: vimage_singleton_iff Int_lower2)  
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    59
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    60
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    61
text{*For @{text is_recfun} we need only pay attention to functions
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    62
      whose domains are initial segments of @{term r}.*}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    63
lemma is_recfun_cong:
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    64
  "[| r = r'; a = a'; f = f'; 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    65
      !!x g. [| <x,a'> \<in> r'; relation(g); domain(g) <= r' -``{x} |] 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    66
             ==> H(x,g) = H'(x,g) |]
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    67
   ==> is_recfun(r,a,H,f) <-> is_recfun(r',a',H',f')"
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    68
apply (rule iffI)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    69
txt{*Messy: fast and blast don't work for some reason*}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    70
apply (erule is_recfun_cong_lemma, auto) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    71
apply (erule is_recfun_cong_lemma)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    72
apply (blast intro: sym)+
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    73
done
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    74
13506
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
    75
subsection{*Reworking of the Recursion Theory Within @{term M}*}
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
    76
13319
23de7b3af453 More Separation proofs
paulson
parents: 13306
diff changeset
    77
lemma (in M_axioms) is_recfun_separation':
23de7b3af453 More Separation proofs
paulson
parents: 13306
diff changeset
    78
    "[| f \<in> r -`` {a} \<rightarrow> range(f); g \<in> r -`` {b} \<rightarrow> range(g);
23de7b3af453 More Separation proofs
paulson
parents: 13306
diff changeset
    79
        M(r); M(f); M(g); M(a); M(b) |] 
23de7b3af453 More Separation proofs
paulson
parents: 13306
diff changeset
    80
     ==> separation(M, \<lambda>x. \<not> (\<langle>x, a\<rangle> \<in> r \<longrightarrow> \<langle>x, b\<rangle> \<in> r \<longrightarrow> f ` x = g ` x))"
23de7b3af453 More Separation proofs
paulson
parents: 13306
diff changeset
    81
apply (insert is_recfun_separation [of r f g a b]) 
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
    82
apply (simp add: vimage_singleton_iff)
13319
23de7b3af453 More Separation proofs
paulson
parents: 13306
diff changeset
    83
done
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    84
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    85
text{*Stated using @{term "trans(r)"} rather than
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    86
      @{term "transitive_rel(M,A,r)"} because the latter rewrites to
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    87
      the former anyway, by @{text transitive_rel_abs}.
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    88
      As always, theorems should be expressed in simplified form.
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    89
      The last three M-premises are redundant because of @{term "M(r)"}, 
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    90
      but without them we'd have to undertake
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    91
      more work to set up the induction formula.*}
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    92
lemma (in M_axioms) is_recfun_equal [rule_format]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    93
    "[|is_recfun(r,a,H,f);  is_recfun(r,b,H,g);  
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    94
       wellfounded(M,r);  trans(r);
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    95
       M(f); M(g); M(r); M(x); M(a); M(b) |] 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    96
     ==> <x,a> \<in> r --> <x,b> \<in> r --> f`x=g`x"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13319
diff changeset
    97
apply (frule_tac f=f in is_recfun_type) 
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13319
diff changeset
    98
apply (frule_tac f=g in is_recfun_type) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    99
apply (simp add: is_recfun_def)
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   100
apply (erule_tac a=x in wellfounded_induct, assumption+)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   101
txt{*Separation to justify the induction*}
13319
23de7b3af453 More Separation proofs
paulson
parents: 13306
diff changeset
   102
 apply (blast intro: is_recfun_separation') 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   103
txt{*Now the inductive argument itself*}
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   104
apply clarify 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   105
apply (erule ssubst)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   106
apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   107
apply (rename_tac x1)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   108
apply (rule_tac t="%z. H(x1,z)" in subst_context) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   109
apply (subgoal_tac "ALL y : r-``{x1}. ALL z. <y,z>:f <-> <y,z>:g")
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   110
 apply (blast intro: transD) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   111
apply (simp add: apply_iff) 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   112
apply (blast intro: transD sym) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   113
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   114
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   115
lemma (in M_axioms) is_recfun_cut: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   116
    "[|is_recfun(r,a,H,f);  is_recfun(r,b,H,g);  
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   117
       wellfounded(M,r); trans(r); 
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   118
       M(f); M(g); M(r); <b,a> \<in> r |]   
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   119
      ==> restrict(f, r-``{b}) = g"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13319
diff changeset
   120
apply (frule_tac f=f in is_recfun_type) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   121
apply (rule fun_extension) 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   122
apply (blast intro: transD restrict_type2) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   123
apply (erule is_recfun_type, simp) 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   124
apply (blast intro: is_recfun_equal transD dest: transM) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   125
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   126
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   127
lemma (in M_axioms) is_recfun_functional:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   128
     "[|is_recfun(r,a,H,f);  is_recfun(r,a,H,g);  
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   129
       wellfounded(M,r); trans(r); M(f); M(g); M(r) |] ==> f=g"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   130
apply (rule fun_extension)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   131
apply (erule is_recfun_type)+
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   132
apply (blast intro!: is_recfun_equal dest: transM) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   133
done 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   134
13295
ca2e9b273472 document setup;
wenzelm
parents: 13293
diff changeset
   135
text{*Tells us that @{text is_recfun} can (in principle) be relativized.*}
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   136
lemma (in M_axioms) is_recfun_relativize:
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   137
  "[| M(r); M(f); \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   138
   ==> is_recfun(r,a,H,f) <->
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   139
       (\<forall>z[M]. z \<in> f <-> 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   140
        (\<exists>x[M]. <x,a> \<in> r & z = <x, H(x, restrict(f, r-``{x}))>))";
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   141
apply (simp add: is_recfun_def lam_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   142
apply (safe intro!: equalityI) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   143
   apply (drule equalityD1 [THEN subsetD], assumption) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   144
   apply (blast dest: pair_components_in_M) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   145
  apply (blast elim!: equalityE dest: pair_components_in_M)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   146
 apply (frule transM, assumption, rotate_tac -1) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   147
 apply simp  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   148
 apply blast
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   149
apply (subgoal_tac "is_function(M,f)")
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   150
 txt{*We use @{term "is_function"} rather than @{term "function"} because
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   151
      the subgoal's easier to prove with relativized quantifiers!*}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   152
 prefer 2 apply (simp add: is_function_def) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   153
apply (frule pair_components_in_M, assumption) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   154
apply (simp add: is_recfun_imp_function function_restrictI) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   155
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   156
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   157
lemma (in M_axioms) is_recfun_restrict:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   158
     "[| wellfounded(M,r); trans(r); is_recfun(r,x,H,f); \<langle>y,x\<rangle> \<in> r; 
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   159
       M(r); M(f); 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   160
       \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |]
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   161
       ==> is_recfun(r, y, H, restrict(f, r -`` {y}))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   162
apply (frule pair_components_in_M, assumption, clarify) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   163
apply (simp (no_asm_simp) add: is_recfun_relativize restrict_iff
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   164
           trans_Int_eq)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   165
apply safe
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   166
  apply (simp_all add: vimage_singleton_iff is_recfun_type [THEN apply_iff]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   167
  apply (frule_tac x=xa in pair_components_in_M, assumption)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   168
  apply (frule_tac x=xa in apply_recfun, blast intro: transD)  
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13245
diff changeset
   169
  apply (simp add: is_recfun_type [THEN apply_iff] 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   170
                   is_recfun_imp_function function_restrictI)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   171
apply (blast intro: apply_recfun dest: transD)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   172
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   173
 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   174
lemma (in M_axioms) restrict_Y_lemma:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   175
   "[| wellfounded(M,r); trans(r); M(r);
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   176
       \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g));  M(Y);
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   177
       \<forall>b[M]. 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   178
	   b \<in> Y <->
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   179
	   (\<exists>x[M]. <x,a1> \<in> r &
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   180
            (\<exists>y[M]. b = \<langle>x,y\<rangle> & (\<exists>g[M]. is_recfun(r,x,H,g) \<and> y = H(x,g))));
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   181
          \<langle>x,a1\<rangle> \<in> r; is_recfun(r,x,H,f); M(f) |]
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   182
       ==> restrict(Y, r -`` {x}) = f"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   183
apply (subgoal_tac "\<forall>y \<in> r-``{x}. \<forall>z. <y,z>:Y <-> <y,z>:f") 
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   184
 apply (simp (no_asm_simp) add: restrict_def) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   185
 apply (thin_tac "rall(M,?P)")+  --{*essential for efficiency*}
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   186
 apply (frule is_recfun_type [THEN fun_is_rel], blast)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   187
apply (frule pair_components_in_M, assumption, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   188
apply (rule iffI)
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
   189
 apply (frule_tac y="<y,z>" in transM, assumption)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   190
 apply (rotate_tac -1)   
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   191
 apply (clarsimp simp add: vimage_singleton_iff is_recfun_type [THEN apply_iff]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   192
			   apply_recfun is_recfun_cut) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   193
txt{*Opposite inclusion: something in f, show in Y*}
13293
paulson
parents: 13269
diff changeset
   194
apply (frule_tac y="<y,z>" in transM, assumption)  
paulson
parents: 13269
diff changeset
   195
apply (simp add: vimage_singleton_iff) 
paulson
parents: 13269
diff changeset
   196
apply (rule conjI) 
paulson
parents: 13269
diff changeset
   197
 apply (blast dest: transD) 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   198
apply (rule_tac x="restrict(f, r -`` {y})" in rexI) 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   199
apply (simp_all add: is_recfun_restrict
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   200
                     apply_recfun is_recfun_type [THEN apply_iff]) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   201
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   202
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   203
text{*For typical applications of Replacement for recursive definitions*}
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   204
lemma (in M_axioms) univalent_is_recfun:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   205
     "[|wellfounded(M,r); trans(r); M(r)|]
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   206
      ==> univalent (M, A, \<lambda>x p. 
13293
paulson
parents: 13269
diff changeset
   207
              \<exists>y[M]. p = \<langle>x,y\<rangle> & (\<exists>f[M]. is_recfun(r,x,H,f) & y = H(x,f)))"
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   208
apply (simp add: univalent_def) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   209
apply (blast dest: is_recfun_functional) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   210
done
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   211
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   212
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   213
text{*Proof of the inductive step for @{text exists_is_recfun}, since
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   214
      we must prove two versions.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   215
lemma (in M_axioms) exists_is_recfun_indstep:
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   216
    "[|\<forall>y. \<langle>y, a1\<rangle> \<in> r --> (\<exists>f[M]. is_recfun(r, y, H, f)); 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   217
       wellfounded(M,r); trans(r); M(r); M(a1);
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   218
       strong_replacement(M, \<lambda>x z. 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   219
              \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   220
       \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|]   
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   221
      ==> \<exists>f[M]. is_recfun(r,a1,H,f)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   222
apply (drule_tac A="r-``{a1}" in strong_replacementD)
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   223
  apply blast 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   224
 txt{*Discharge the "univalent" obligation of Replacement*}
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   225
 apply (simp add: univalent_is_recfun) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   226
txt{*Show that the constructed object satisfies @{text is_recfun}*} 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   227
apply clarify 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   228
apply (rule_tac x=Y in rexI)  
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   229
txt{*Unfold only the top-level occurrence of @{term is_recfun}*}
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   230
apply (simp (no_asm_simp) add: is_recfun_relativize [of concl: _ a1])
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   231
txt{*The big iff-formula defining @{term Y} is now redundant*}
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   232
apply safe 
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   233
 apply (simp add: vimage_singleton_iff restrict_Y_lemma [of r H _ a1]) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   234
txt{*one more case*}
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   235
apply (simp (no_asm_simp) add: Bex_def vimage_singleton_iff)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   236
apply (drule_tac x1=x in spec [THEN mp], assumption, clarify) 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   237
apply (rename_tac f) 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   238
apply (rule_tac x=f in rexI) 
13293
paulson
parents: 13269
diff changeset
   239
apply (simp_all add: restrict_Y_lemma [of r H])
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   240
txt{*FIXME: should not be needed!*}
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   241
apply (subst restrict_Y_lemma [of r H])
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   242
apply (simp add: vimage_singleton_iff)+
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   243
apply blast+
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   244
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   245
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   246
text{*Relativized version, when we have the (currently weaker) premise
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   247
      @{term "wellfounded(M,r)"}*}
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   248
lemma (in M_axioms) wellfounded_exists_is_recfun:
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   249
    "[|wellfounded(M,r);  trans(r);  
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   250
       separation(M, \<lambda>x. ~ (\<exists>f[M]. is_recfun(r, x, H, f)));
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   251
       strong_replacement(M, \<lambda>x z. 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   252
          \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   253
       M(r);  M(a);  
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   254
       \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |]   
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   255
      ==> \<exists>f[M]. is_recfun(r,a,H,f)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   256
apply (rule wellfounded_induct, assumption+, clarify)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   257
apply (rule exists_is_recfun_indstep, assumption+)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   258
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   259
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   260
lemma (in M_axioms) wf_exists_is_recfun [rule_format]:
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   261
    "[|wf(r);  trans(r);  M(r);  
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   262
       strong_replacement(M, \<lambda>x z. 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   263
         \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   264
       \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |]   
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   265
      ==> M(a) --> (\<exists>f[M]. is_recfun(r,a,H,f))"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   266
apply (rule wf_induct, assumption+)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   267
apply (frule wf_imp_relativized)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   268
apply (intro impI)
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   269
apply (rule exists_is_recfun_indstep) 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   270
      apply (blast dest: transM del: rev_rallE, assumption+)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   271
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   272
13506
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
   273
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
   274
subsection{*Relativization of the ZF Predicate @{term is_recfun}*}
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
   275
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   276
constdefs
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   277
  M_is_recfun :: "[i=>o, [i,i,i]=>o, i, i, i] => o"
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   278
   "M_is_recfun(M,MH,r,a,f) == 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   279
     \<forall>z[M]. z \<in> f <-> 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   280
            (\<exists>x[M]. \<exists>y[M]. \<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. \<exists>f_r_sx[M]. 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   281
	       pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) &
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   282
               pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) &
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   283
               xa \<in> r & MH(x, f_r_sx, y))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   284
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   285
  is_wfrec :: "[i=>o, [i,i,i]=>o, i, i, i] => o"
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   286
   "is_wfrec(M,MH,r,a,z) == 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   287
      \<exists>f[M]. M_is_recfun(M,MH,r,a,f) & MH(a,f,z)"
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   288
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   289
  wfrec_replacement :: "[i=>o, [i,i,i]=>o, i] => o"
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   290
   "wfrec_replacement(M,MH,r) ==
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   291
        strong_replacement(M, 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   292
             \<lambda>x z. \<exists>y[M]. pair(M,x,y,z) & is_wfrec(M,MH,r,x,y))"
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   293
13350
paulson
parents: 13348
diff changeset
   294
lemma (in M_axioms) is_recfun_abs:
paulson
parents: 13348
diff changeset
   295
     "[| \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g));  M(r); M(a); M(f); 
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   296
         relativize2(M,MH,H) |] 
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   297
      ==> M_is_recfun(M,MH,r,a,f) <-> is_recfun(r,a,H,f)"
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   298
apply (simp add: M_is_recfun_def relativize2_def is_recfun_relativize)
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   299
apply (rule rall_cong)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   300
apply (blast dest: transM)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   301
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   302
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   303
lemma M_is_recfun_cong [cong]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   304
     "[| r = r'; a = a'; f = f'; 
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   305
       !!x g y. [| M(x); M(g); M(y) |] ==> MH(x,g,y) <-> MH'(x,g,y) |]
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   306
      ==> M_is_recfun(M,MH,r,a,f) <-> M_is_recfun(M,MH',r',a',f')"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   307
by (simp add: M_is_recfun_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   308
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   309
lemma (in M_axioms) is_wfrec_abs:
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   310
     "[| \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)); 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   311
         relativize2(M,MH,H);  M(r); M(a); M(z) |]
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   312
      ==> is_wfrec(M,MH,r,a,z) <-> 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   313
          (\<exists>g[M]. is_recfun(r,a,H,g) & z = H(a,g))"
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   314
by (simp add: is_wfrec_def relativize2_def is_recfun_abs)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   315
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   316
text{*Relating @{term wfrec_replacement} to native constructs*}
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   317
lemma (in M_axioms) wfrec_replacement':
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   318
  "[|wfrec_replacement(M,MH,r);
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   319
     \<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)); 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   320
     relativize2(M,MH,H);  M(r)|] 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   321
   ==> strong_replacement(M, \<lambda>x z. \<exists>y[M]. 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   322
                pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g)))"
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   323
apply (rotate_tac 1) 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   324
apply (simp add: wfrec_replacement_def is_wfrec_abs) 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   325
done
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   326
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   327
lemma wfrec_replacement_cong [cong]:
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   328
     "[| !!x y z. [| M(x); M(y); M(z) |] ==> MH(x,y,z) <-> MH'(x,y,z);
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   329
         r=r' |] 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   330
      ==> wfrec_replacement(M, %x y. MH(x,y), r) <-> 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   331
          wfrec_replacement(M, %x y. MH'(x,y), r')" 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   332
by (simp add: is_wfrec_def wfrec_replacement_def) 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   333
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   334
13506
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
   335
subsection{*Ordinal Arithmetic: Two Examples of Recursion*}
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
   336
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
   337
subsubsection{*Ordinal Addition*}
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
   338
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   339
(*FIXME: update to use new techniques!!*)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   340
constdefs
13350
paulson
parents: 13348
diff changeset
   341
 (*This expresses ordinal addition in the language of ZF.  It also 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   342
   provides an abbreviation that can be used in the instance of strong
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   343
   replacement below.  Here j is used to define the relation, namely
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   344
   Memrel(succ(j)), while x determines the domain of f.*)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   345
 is_oadd_fun :: "[i=>o,i,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   346
    "is_oadd_fun(M,i,j,x,f) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   347
       (\<forall>sj msj. M(sj) --> M(msj) --> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   348
                 successor(M,j,sj) --> membership(M,sj,msj) --> 
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   349
	         M_is_recfun(M, 
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   350
		     %x g y. \<exists>gx[M]. image(M,g,x,gx) & union(M,i,gx,y),
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   351
		     msj, x, f))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   352
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   353
 is_oadd :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   354
    "is_oadd(M,i,j,k) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   355
        (~ ordinal(M,i) & ~ ordinal(M,j) & k=0) |
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   356
        (~ ordinal(M,i) & ordinal(M,j) & k=j) |
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   357
        (ordinal(M,i) & ~ ordinal(M,j) & k=i) |
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   358
        (ordinal(M,i) & ordinal(M,j) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   359
	 (\<exists>f fj sj. M(f) & M(fj) & M(sj) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   360
		    successor(M,j,sj) & is_oadd_fun(M,i,sj,sj,f) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   361
		    fun_apply(M,f,j,fj) & fj = k))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   362
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   363
 (*NEEDS RELATIVIZATION*)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   364
 omult_eqns :: "[i,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   365
    "omult_eqns(i,x,g,z) ==
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   366
            Ord(x) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   367
	    (x=0 --> z=0) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   368
            (\<forall>j. x = succ(j) --> z = g`j ++ i) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   369
            (Limit(x) --> z = \<Union>(g``x))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   370
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   371
 is_omult_fun :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   372
    "is_omult_fun(M,i,j,f) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   373
	    (\<exists>df. M(df) & is_function(M,f) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   374
                  is_domain(M,f,df) & subset(M, j, df)) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   375
            (\<forall>x\<in>j. omult_eqns(i,x,f,f`x))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   376
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   377
 is_omult :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   378
    "is_omult(M,i,j,k) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   379
	\<exists>f fj sj. M(f) & M(fj) & M(sj) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   380
                  successor(M,j,sj) & is_omult_fun(M,i,sj,f) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   381
                  fun_apply(M,f,j,fj) & fj = k"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   382
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   383
13428
99e52e78eb65 eliminate open locales and special ML code;
wenzelm
parents: 13382
diff changeset
   384
locale M_ord_arith = M_axioms +
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   385
  assumes oadd_strong_replacement:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   386
   "[| M(i); M(j) |] ==>
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   387
    strong_replacement(M, 
13293
paulson
parents: 13269
diff changeset
   388
         \<lambda>x z. \<exists>y[M]. pair(M,x,y,z) & 
paulson
parents: 13269
diff changeset
   389
                  (\<exists>f[M]. \<exists>fx[M]. is_oadd_fun(M,i,j,x,f) & 
paulson
parents: 13269
diff changeset
   390
		           image(M,f,x,fx) & y = i Un fx))"
paulson
parents: 13269
diff changeset
   391
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   392
 and omult_strong_replacement':
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   393
   "[| M(i); M(j) |] ==>
13293
paulson
parents: 13269
diff changeset
   394
    strong_replacement(M, 
paulson
parents: 13269
diff changeset
   395
         \<lambda>x z. \<exists>y[M]. z = <x,y> &
paulson
parents: 13269
diff changeset
   396
	     (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. THE z. omult_eqns(i,x,g,z),g) & 
paulson
parents: 13269
diff changeset
   397
	     y = (THE z. omult_eqns(i, x, g, z))))" 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   398
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   399
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   400
13295
ca2e9b273472 document setup;
wenzelm
parents: 13293
diff changeset
   401
text{*@{text is_oadd_fun}: Relating the pure "language of set theory" to Isabelle/ZF*}
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   402
lemma (in M_ord_arith) is_oadd_fun_iff:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   403
   "[| a\<le>j; M(i); M(j); M(a); M(f) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   404
    ==> is_oadd_fun(M,i,j,a,f) <->
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   405
	f \<in> a \<rightarrow> range(f) & (\<forall>x. M(x) --> x < a --> f`x = i Un f``x)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   406
apply (frule lt_Ord) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   407
apply (simp add: is_oadd_fun_def Memrel_closed Un_closed 
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   408
             relativize2_def is_recfun_abs [of "%x g. i Un g``x"]
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   409
             image_closed is_recfun_iff_equation  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   410
             Ball_def lt_trans [OF ltI, of _ a] lt_Memrel)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   411
apply (simp add: lt_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   412
apply (blast dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   413
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   414
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   415
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   416
lemma (in M_ord_arith) oadd_strong_replacement':
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   417
    "[| M(i); M(j) |] ==>
13293
paulson
parents: 13269
diff changeset
   418
     strong_replacement(M, 
paulson
parents: 13269
diff changeset
   419
            \<lambda>x z. \<exists>y[M]. z = <x,y> &
paulson
parents: 13269
diff changeset
   420
		  (\<exists>g[M]. is_recfun(Memrel(succ(j)),x,%x g. i Un g``x,g) & 
paulson
parents: 13269
diff changeset
   421
		  y = i Un g``x))" 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   422
apply (insert oadd_strong_replacement [of i j]) 
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   423
apply (simp add: is_oadd_fun_def relativize2_def is_recfun_abs [of "%x g. i Un g``x"])  
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   424
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   425
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   426
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   427
lemma (in M_ord_arith) exists_oadd:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   428
    "[| Ord(j);  M(i);  M(j) |]
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   429
     ==> \<exists>f[M]. is_recfun(Memrel(succ(j)), j, %x g. i Un g``x, f)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   430
apply (rule wf_exists_is_recfun [OF wf_Memrel trans_Memrel])
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   431
    apply (simp_all add: Memrel_type oadd_strong_replacement') 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   432
done 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   433
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   434
lemma (in M_ord_arith) exists_oadd_fun:
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   435
    "[| Ord(j);  M(i);  M(j) |] ==> \<exists>f[M]. is_oadd_fun(M,i,succ(j),succ(j),f)"
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   436
apply (rule exists_oadd [THEN rexE])
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   437
apply (erule Ord_succ, assumption, simp) 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   438
apply (rename_tac f) 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   439
apply (frule is_recfun_type)
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   440
apply (rule_tac x=f in rexI) 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   441
 apply (simp add: fun_is_function domain_of_fun lt_Memrel apply_recfun lt_def
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   442
                  is_oadd_fun_iff Ord_trans [OF _ succI1], assumption)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   443
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   444
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   445
lemma (in M_ord_arith) is_oadd_fun_apply:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   446
    "[| x < j; M(i); M(j); M(f); is_oadd_fun(M,i,j,j,f) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   447
     ==> f`x = i Un (\<Union>k\<in>x. {f ` k})"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   448
apply (simp add: is_oadd_fun_iff lt_Ord2, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   449
apply (frule lt_closed, simp)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   450
apply (frule leI [THEN le_imp_subset])  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   451
apply (simp add: image_fun, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   452
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   453
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   454
lemma (in M_ord_arith) is_oadd_fun_iff_oadd [rule_format]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   455
    "[| is_oadd_fun(M,i,J,J,f); M(i); M(J); M(f); Ord(i); Ord(j) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   456
     ==> j<J --> f`j = i++j"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   457
apply (erule_tac i=j in trans_induct, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   458
apply (subgoal_tac "\<forall>k\<in>x. k<J")
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   459
 apply (simp (no_asm_simp) add: is_oadd_def oadd_unfold is_oadd_fun_apply)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   460
apply (blast intro: lt_trans ltI lt_Ord) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   461
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   462
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   463
lemma (in M_ord_arith) Ord_oadd_abs:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   464
    "[| M(i); M(j); M(k); Ord(i); Ord(j) |] ==> is_oadd(M,i,j,k) <-> k = i++j"
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   465
apply (simp add: is_oadd_def is_oadd_fun_iff_oadd)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   466
apply (frule exists_oadd_fun [of j i], blast+)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   467
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   468
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   469
lemma (in M_ord_arith) oadd_abs:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   470
    "[| M(i); M(j); M(k) |] ==> is_oadd(M,i,j,k) <-> k = i++j"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   471
apply (case_tac "Ord(i) & Ord(j)")
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   472
 apply (simp add: Ord_oadd_abs)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   473
apply (auto simp add: is_oadd_def oadd_eq_if_raw_oadd)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   474
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   475
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   476
lemma (in M_ord_arith) oadd_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   477
    "[| M(i); M(j) |] ==> M(i++j)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   478
apply (simp add: oadd_eq_if_raw_oadd, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   479
apply (simp add: raw_oadd_eq_oadd) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   480
apply (frule exists_oadd_fun [of j i], auto)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   481
apply (simp add: apply_closed is_oadd_fun_iff_oadd [symmetric]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   482
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   483
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   484
13506
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
   485
subsubsection{*Ordinal Multiplication*}
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   486
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   487
lemma omult_eqns_unique:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   488
     "[| omult_eqns(i,x,g,z); omult_eqns(i,x,g,z') |] ==> z=z'";
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   489
apply (simp add: omult_eqns_def, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   490
apply (erule Ord_cases, simp_all) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   491
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   492
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   493
lemma omult_eqns_0: "omult_eqns(i,0,g,z) <-> z=0"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   494
by (simp add: omult_eqns_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   495
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   496
lemma the_omult_eqns_0: "(THE z. omult_eqns(i,0,g,z)) = 0"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   497
by (simp add: omult_eqns_0)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   498
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   499
lemma omult_eqns_succ: "omult_eqns(i,succ(j),g,z) <-> Ord(j) & z = g`j ++ i"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   500
by (simp add: omult_eqns_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   501
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   502
lemma the_omult_eqns_succ:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   503
     "Ord(j) ==> (THE z. omult_eqns(i,succ(j),g,z)) = g`j ++ i"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   504
by (simp add: omult_eqns_succ) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   505
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   506
lemma omult_eqns_Limit:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   507
     "Limit(x) ==> omult_eqns(i,x,g,z) <-> z = \<Union>(g``x)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   508
apply (simp add: omult_eqns_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   509
apply (blast intro: Limit_is_Ord) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   510
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   511
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   512
lemma the_omult_eqns_Limit:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   513
     "Limit(x) ==> (THE z. omult_eqns(i,x,g,z)) = \<Union>(g``x)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   514
by (simp add: omult_eqns_Limit)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   515
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   516
lemma omult_eqns_Not: "~ Ord(x) ==> ~ omult_eqns(i,x,g,z)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   517
by (simp add: omult_eqns_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   518
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   519
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   520
lemma (in M_ord_arith) the_omult_eqns_closed:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   521
    "[| M(i); M(x); M(g); function(g) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   522
     ==> M(THE z. omult_eqns(i, x, g, z))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   523
apply (case_tac "Ord(x)")
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   524
 prefer 2 apply (simp add: omult_eqns_Not) --{*trivial, non-Ord case*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   525
apply (erule Ord_cases) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   526
  apply (simp add: omult_eqns_0)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   527
 apply (simp add: omult_eqns_succ apply_closed oadd_closed) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   528
apply (simp add: omult_eqns_Limit) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   529
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   530
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   531
lemma (in M_ord_arith) exists_omult:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   532
    "[| Ord(j);  M(i);  M(j) |]
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   533
     ==> \<exists>f[M]. is_recfun(Memrel(succ(j)), j, %x g. THE z. omult_eqns(i,x,g,z), f)"
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   534
apply (rule wf_exists_is_recfun [OF wf_Memrel trans_Memrel])
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   535
    apply (simp_all add: Memrel_type omult_strong_replacement') 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   536
apply (blast intro: the_omult_eqns_closed) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   537
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   538
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   539
lemma (in M_ord_arith) exists_omult_fun:
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   540
    "[| Ord(j);  M(i);  M(j) |] ==> \<exists>f[M]. is_omult_fun(M,i,succ(j),f)"
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   541
apply (rule exists_omult [THEN rexE])
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   542
apply (erule Ord_succ, assumption, simp) 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   543
apply (rename_tac f) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   544
apply (frule is_recfun_type)
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   545
apply (rule_tac x=f in rexI) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   546
apply (simp add: fun_is_function domain_of_fun lt_Memrel apply_recfun lt_def
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   547
                 is_omult_fun_def Ord_trans [OF _ succI1])
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   548
 apply (force dest: Ord_in_Ord' 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   549
              simp add: omult_eqns_def the_omult_eqns_0 the_omult_eqns_succ
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   550
                        the_omult_eqns_Limit, assumption)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   551
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   552
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   553
lemma (in M_ord_arith) is_omult_fun_apply_0:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   554
    "[| 0 < j; is_omult_fun(M,i,j,f) |] ==> f`0 = 0"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   555
by (simp add: is_omult_fun_def omult_eqns_def lt_def ball_conj_distrib)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   556
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   557
lemma (in M_ord_arith) is_omult_fun_apply_succ:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   558
    "[| succ(x) < j; is_omult_fun(M,i,j,f) |] ==> f`succ(x) = f`x ++ i"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   559
by (simp add: is_omult_fun_def omult_eqns_def lt_def, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   560
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   561
lemma (in M_ord_arith) is_omult_fun_apply_Limit:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   562
    "[| x < j; Limit(x); M(j); M(f); is_omult_fun(M,i,j,f) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   563
     ==> f ` x = (\<Union>y\<in>x. f`y)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   564
apply (simp add: is_omult_fun_def omult_eqns_def domain_closed lt_def, clarify)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   565
apply (drule subset_trans [OF OrdmemD], assumption+)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   566
apply (simp add: ball_conj_distrib omult_Limit image_function)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   567
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   568
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   569
lemma (in M_ord_arith) is_omult_fun_eq_omult:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   570
    "[| is_omult_fun(M,i,J,f); M(J); M(f); Ord(i); Ord(j) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   571
     ==> j<J --> f`j = i**j"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   572
apply (erule_tac i=j in trans_induct3)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   573
apply (safe del: impCE)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   574
  apply (simp add: is_omult_fun_apply_0) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   575
 apply (subgoal_tac "x<J") 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   576
  apply (simp add: is_omult_fun_apply_succ omult_succ)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   577
 apply (blast intro: lt_trans) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   578
apply (subgoal_tac "\<forall>k\<in>x. k<J")
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   579
 apply (simp add: is_omult_fun_apply_Limit omult_Limit) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   580
apply (blast intro: lt_trans ltI lt_Ord) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   581
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   582
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   583
lemma (in M_ord_arith) omult_abs:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   584
    "[| M(i); M(j); M(k); Ord(i); Ord(j) |] ==> is_omult(M,i,j,k) <-> k = i**j"
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   585
apply (simp add: is_omult_def is_omult_fun_eq_omult)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   586
apply (frule exists_omult_fun [of j i], blast+)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   587
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   588
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   589
end
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   590