src/ZF/Constructible/WFrec.thy
author paulson
Mon, 24 Jun 2002 11:59:21 +0200
changeset 13245 714f7a423a15
parent 13223 45be08fbdcff
child 13247 e3c289f0724b
permissions -rw-r--r--
development and tweaks
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     1
theory WFrec = Wellorderings:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     2
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     3
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
     4
(*FIXME: could move these to WF.thy*)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     5
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     6
lemma is_recfunI:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     7
     "f = (lam x: r-``{a}. H(x, restrict(f, r-``{x}))) ==> is_recfun(r,a,H,f)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     8
by (simp add: is_recfun_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     9
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    10
lemma is_recfun_imp_function: "is_recfun(r,a,H,f) ==> function(f)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    11
apply (simp add: is_recfun_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    12
apply (erule ssubst)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    13
apply (rule function_lam) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    14
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    15
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    16
text{*Expresses @{text is_recfun} as a recursion equation*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    17
lemma is_recfun_iff_equation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    18
     "is_recfun(r,a,H,f) <->
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    19
	   f \<in> r -`` {a} \<rightarrow> range(f) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    20
	   (\<forall>x \<in> r-``{a}. f`x = H(x, restrict(f, r-``{x})))"  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    21
apply (rule iffI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    22
 apply (simp add: is_recfun_type apply_recfun Ball_def vimage_singleton_iff, 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    23
        clarify)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    24
apply (simp add: is_recfun_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    25
apply (rule fun_extension) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    26
  apply assumption
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    27
 apply (fast intro: lam_type, simp) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    28
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    29
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    30
lemma is_recfun_imp_in_r: "[|is_recfun(r,a,H,f); \<langle>x,i\<rangle> \<in> f|] ==> \<langle>x, a\<rangle> \<in> r"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    31
by (blast dest:  is_recfun_type fun_is_rel)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    32
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    33
lemma apply_recfun2:
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    34
    "[| is_recfun(r,a,H,f); <x,i>:f |] ==> i = H(x, restrict(f,r-``{x}))"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    35
apply (frule apply_recfun) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    36
 apply (blast dest: is_recfun_type fun_is_rel) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    37
apply (simp add: function_apply_equality [OF _ is_recfun_imp_function])
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    38
done
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    39
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    40
lemma trans_on_Int_eq [simp]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    41
      "[| trans[A](r); <y,x> \<in> r;  r \<subseteq> A*A |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    42
       ==> r -`` {y} \<inter> r -`` {x} = r -`` {y}"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    43
by (blast intro: trans_onD) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    44
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    45
lemma trans_on_Int_eq2 [simp]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    46
      "[| trans[A](r); <y,x> \<in> r;  r \<subseteq> A*A |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    47
       ==> r -`` {x} \<inter> r -`` {y} = r -`` {y}"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    48
by (blast intro: trans_onD) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    49
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    50
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    51
text{*Stated using @{term "trans[A](r)"} rather than
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    52
      @{term "transitive_rel(M,A,r)"} because the latter rewrites to
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    53
      the former anyway, by @{text transitive_rel_abs}.
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    54
      As always, theorems should be expressed in simplified form.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    55
lemma (in M_axioms) is_recfun_equal [rule_format]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    56
    "[|is_recfun(r,a,H,f);  is_recfun(r,b,H,g);  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    57
       wellfounded_on(M,A,r);  trans[A](r); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    58
       M(A); M(f); M(g); M(a); M(b); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    59
       r \<subseteq> A*A;  x\<in>A |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    60
     ==> <x,a> \<in> r --> <x,b> \<in> r --> f`x=g`x"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    61
apply (frule_tac f="f" in is_recfun_type) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    62
apply (frule_tac f="g" in is_recfun_type) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    63
apply (simp add: is_recfun_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    64
apply (erule wellfounded_on_induct2, assumption+) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    65
apply (force intro: is_recfun_separation, clarify)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    66
apply (erule ssubst)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    67
apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    68
apply (rename_tac x1)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    69
apply (rule_tac t="%z. H(x1,z)" in subst_context) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    70
apply (subgoal_tac "ALL y : r-``{x1}. ALL z. <y,z>:f <-> <y,z>:g")
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    71
 apply (blast intro: trans_onD) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    72
apply (simp add: apply_iff) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    73
apply (blast intro: trans_onD sym) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    74
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    75
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    76
lemma (in M_axioms) is_recfun_cut: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    77
    "[|is_recfun(r,a,H,f);  is_recfun(r,b,H,g);  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    78
       wellfounded_on(M,A,r); trans[A](r); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    79
       M(A); M(f); M(g); M(a); M(b); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    80
       r \<subseteq> A*A;  <b,a>\<in>r |]   
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    81
      ==> restrict(f, r-``{b}) = g"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    82
apply (frule_tac f="f" in is_recfun_type) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    83
apply (rule fun_extension) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    84
apply (blast intro: trans_onD restrict_type2) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    85
apply (erule is_recfun_type, simp) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    86
apply (blast intro: is_recfun_equal trans_onD) 
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 (in M_axioms) is_recfun_functional:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    90
     "[|is_recfun(r,a,H,f);  is_recfun(r,a,H,g);  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    91
       wellfounded_on(M,A,r); trans[A](r); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    92
       M(A); M(f); M(g); M(a); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    93
       r \<subseteq> A*A |]   
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    94
      ==> f=g"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    95
apply (rule fun_extension)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    96
apply (erule is_recfun_type)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    97
apply (blast intro!: is_recfun_equal) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    98
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    99
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   100
text{*Tells us that is_recfun can (in principle) be relativized.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   101
lemma (in M_axioms) is_recfun_relativize:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   102
     "[| M(r); M(a); M(f); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   103
       \<forall>x g. M(x) & M(g) & function(g) --> M(H(x,g)) |] ==>
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   104
       is_recfun(r,a,H,f) <->
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   105
       (\<forall>z. z \<in> f <-> (\<exists>x y. M(x) & M(y) & z=<x,y> & <x,a> \<in> r & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   106
                              y = H(x, restrict(f, r-``{x}))))";
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   107
apply (simp add: is_recfun_def vimage_closed restrict_closed lam_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   108
apply (safe intro!: equalityI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   109
  apply (drule equalityD1 [THEN subsetD], assumption) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   110
  apply clarify 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   111
  apply (rule_tac x=x in exI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   112
  apply (blast dest: pair_components_in_M) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   113
 apply (blast elim!: equalityE dest: pair_components_in_M)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   114
 apply simp  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   115
 apply blast
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   116
 apply simp 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   117
apply (subgoal_tac "function(f)")  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   118
 prefer 2
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   119
 apply (simp add: function_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   120
apply (frule pair_components_in_M, assumption) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   121
  apply (simp add: is_recfun_imp_function function_restrictI restrict_closed vimage_closed) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   122
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   123
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   124
(* ideas for further weaking the H-closure premise:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   125
apply (drule spec [THEN spec]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   126
apply (erule mp)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   127
apply (intro conjI)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   128
apply (blast dest!: pair_components_in_M)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   129
apply (blast intro!: function_restrictI dest!: pair_components_in_M)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   130
apply (blast intro!: function_restrictI dest!: pair_components_in_M)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   131
apply (simp only: subset_iff domain_iff restrict_iff vimage_iff) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   132
apply (simp add:  vimage_singleton_iff) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   133
apply (intro allI impI conjI)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   134
apply (blast intro: transM dest!: pair_components_in_M)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   135
prefer 4;apply blast 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   136
*)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   137
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   138
lemma (in M_axioms) is_recfun_restrict:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   139
     "[| wellfounded_on(M,A,r); trans[A](r); is_recfun(r,x,H,f); \<langle>y,x\<rangle> \<in> r; 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   140
       M(A); M(r); M(f); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   141
       \<forall>x g. M(x) & M(g) & function(g) --> M(H(x,g)); r \<subseteq> A * A |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   142
       ==> is_recfun(r, y, H, restrict(f, r -`` {y}))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   143
apply (frule pair_components_in_M, assumption, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   144
apply (simp (no_asm_simp) add: is_recfun_relativize vimage_closed restrict_closed
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   145
     restrict_iff)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   146
apply safe
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   147
  apply (simp_all add: vimage_singleton_iff is_recfun_type [THEN apply_iff]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   148
  apply (frule_tac x=xa in pair_components_in_M, assumption)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   149
  apply (frule_tac x=xa in apply_recfun, blast intro: trans_onD)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   150
  apply (simp add: is_recfun_type [THEN apply_iff])
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   151
  (*???COMBINE*) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   152
  apply (simp add: is_recfun_imp_function function_restrictI restrict_closed vimage_closed) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   153
apply (blast intro: apply_recfun dest: trans_onD)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   154
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   155
 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   156
lemma (in M_axioms) restrict_Y_lemma:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   157
     "[| wellfounded_on(M,A,r); trans[A](r); M(A); M(r); r \<subseteq> A \<times> A;
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   158
       \<forall>x g. M(x) \<and> M(g) & function(g) --> M(H(x,g));  M(Y);
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   159
       \<forall>b. M(b) -->
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   160
	   b \<in> Y <->
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   161
	   (\<exists>x\<in>r -`` {a1}.
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   162
	       \<exists>y. M(y) \<and>
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   163
		   (\<exists>g. M(g) \<and> b = \<langle>x,y\<rangle> \<and> is_recfun(r,x,H,g) \<and> y = H(x,g)));
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   164
          \<langle>x,a1\<rangle> \<in> r; M(f); is_recfun(r,x,H,f) |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   165
       ==> restrict(Y, r -`` {x}) = f"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   166
apply (subgoal_tac "ALL y : r-``{x}. ALL z. <y,z>:Y <-> <y,z>:f") 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   167
apply (simp (no_asm_simp) add: restrict_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   168
apply (thin_tac "All(?P)")+  --{*essential for efficiency*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   169
apply (frule is_recfun_type [THEN fun_is_rel], blast)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   170
apply (frule pair_components_in_M, assumption, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   171
apply (rule iffI)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   172
 apply (frule_tac y="<y,z>" in transM, assumption )
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   173
 apply (rotate_tac -1)   
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   174
 apply (clarsimp simp add: vimage_singleton_iff is_recfun_type [THEN apply_iff]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   175
			   apply_recfun is_recfun_cut) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   176
txt{*Opposite inclusion: something in f, show in Y*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   177
apply (frule_tac y="<y,z>" in transM, assumption, simp) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   178
apply (rule_tac x=y in bexI)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   179
prefer 2 apply (blast dest: trans_onD)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   180
apply (rule_tac x=z in exI, simp) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   181
apply (rule_tac x="restrict(f, r -`` {y})" in exI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   182
apply (simp add: vimage_closed restrict_closed is_recfun_restrict
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   183
                 apply_recfun is_recfun_type [THEN apply_iff]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   184
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   185
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   186
(*FIXME: use this lemma just below*)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   187
text{*For typical applications of Replacement for recursive definitions*}
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   188
lemma (in M_axioms) univalent_is_recfun:
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   189
     "[|wellfounded_on(M,A,r); trans[A](r); r \<subseteq> A*A; M(r); M(A)|]
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   190
      ==> univalent (M, A, \<lambda>x p. \<exists>y. M(y) &
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   191
                    (\<exists>f. M(f) & p = \<langle>x, y\<rangle> & is_recfun(r,x,H,f) & y = H(x,f)))"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   192
apply (simp add: univalent_def) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   193
apply (blast dest: is_recfun_functional) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   194
done
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   195
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   196
text{*Proof of the inductive step for @{text exists_is_recfun}, since
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   197
      we must prove two versions.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   198
lemma (in M_axioms) exists_is_recfun_indstep:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   199
    "[|a1 \<in> A;  \<forall>y. \<langle>y, a1\<rangle> \<in> r --> (\<exists>f. M(f) & is_recfun(r, y, H, f)); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   200
       wellfounded_on(M,A,r); trans[A](r); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   201
       strong_replacement(M, \<lambda>x z. \<exists>y g. M(y) & M(g) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   202
                   pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   203
       M(A); M(r); r \<subseteq> A * A;
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   204
       \<forall>x g. M(x) & M(g) & function(g) --> M(H(x,g))|]   
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   205
      ==> \<exists>f. M(f) & is_recfun(r,a1,H,f)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   206
apply (frule_tac y=a1 in transM, assumption)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   207
apply (drule_tac A="r-``{a1}" in strong_replacementD)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   208
  apply blast
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   209
 txt{*Discharge the "univalent" obligation of Replacement*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   210
 apply (clarsimp simp add: univalent_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   211
 apply (blast dest!: is_recfun_functional)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   212
txt{*Show that the constructed object satisfies @{text is_recfun}*} 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   213
apply clarify 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   214
apply (rule_tac x=Y in exI)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   215
apply (simp (no_asm_simp) add: is_recfun_relativize vimage_closed restrict_closed) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   216
(*Tried using is_recfun_iff2 here.  Much more simplification takes place
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   217
  because an assumption can kick in.  Not sure how to relate the new 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   218
  proof state to the current one.*)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   219
apply safe
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   220
txt{*Show that elements of @{term Y} are in the right relationship.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   221
apply (frule_tac x=z and P="%b. M(b) --> ?Q(b)" in spec)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   222
apply (erule impE, blast intro: transM)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   223
txt{*We have an element of  @{term Y}, so we have x, y, z*} 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   224
apply (frule_tac y=z in transM, assumption, clarify)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   225
apply (simp add: vimage_closed restrict_closed restrict_Y_lemma [of A r H]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   226
txt{*one more case*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   227
apply (simp add: vimage_closed restrict_closed )
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   228
apply (rule_tac x=x in bexI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   229
prefer 2 apply blast 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   230
apply (rule_tac x="H(x, restrict(Y, r -`` {x}))" in exI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   231
apply (simp add: vimage_closed restrict_closed )
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   232
apply (drule_tac x1=x in spec [THEN mp], assumption, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   233
apply (rule_tac x=f in exI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   234
apply (simp add: restrict_Y_lemma [of A r H]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   235
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   236
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   237
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   238
text{*Relativized version, when we have the (currently weaker) premise
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   239
      @{term "wellfounded_on(M,A,r)"}*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   240
lemma (in M_axioms) wellfounded_exists_is_recfun:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   241
    "[|wellfounded_on(M,A,r);  trans[A](r);  a\<in>A; 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   242
       separation(M, \<lambda>x. x \<in> A --> ~ (\<exists>f. M(f) \<and> is_recfun(r, x, H, f)));
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   243
       strong_replacement(M, \<lambda>x z. \<exists>y g. M(y) & M(g) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   244
                   pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   245
       M(A);  M(r);  r \<subseteq> A*A;  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   246
       \<forall>x g. M(x) & M(g) & function(g) --> M(H(x,g)) |]   
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   247
      ==> \<exists>f. M(f) & is_recfun(r,a,H,f)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   248
apply (rule wellfounded_on_induct2, assumption+, clarify)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   249
apply (rule exists_is_recfun_indstep, assumption+)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   250
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   251
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   252
lemma (in M_axioms) wf_exists_is_recfun:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   253
    "[|wf[A](r);  trans[A](r);  a\<in>A; 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   254
       strong_replacement(M, \<lambda>x z. \<exists>y g. M(y) & M(g) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   255
                   pair(M,x,y,z) & is_recfun(r,x,H,g) & y = H(x,g)); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   256
       M(A);  M(r);  r \<subseteq> A*A;  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   257
       \<forall>x g. M(x) & M(g) & function(g) --> M(H(x,g)) |]   
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   258
      ==> \<exists>f. M(f) & is_recfun(r,a,H,f)"        
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   259
apply (rule wf_on_induct2, assumption+)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   260
apply (frule wf_on_imp_relativized)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   261
apply (rule exists_is_recfun_indstep, assumption+)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   262
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   263
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   264
    (*????????????????NOT USED????????????????*)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   265
    constdefs
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   266
      M_the_recfun :: "[i=>o, i, i, [i,i]=>i] => i"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   267
      "M_the_recfun(M,r,a,H) == (THE f. M(f) & is_recfun(r,a,H,f))"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   268
    
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   269
    (*If some f satisfies is_recfun(r,a,H,-) then so does M_the_recfun(M,r,a,H) *)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   270
    lemma (in M_axioms) M_is_the_recfun: 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   271
      "[|is_recfun(r,a,H,f);  
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   272
      wellfounded_on(M,A,r); trans[A](r); 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   273
      M(A); M(f); M(a); r \<subseteq> A*A |]   
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   274
      ==> M(M_the_recfun(M,r,a,H)) & 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   275
      is_recfun(r, a, H, M_the_recfun(M,r,a,H))"    
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   276
    apply (unfold M_the_recfun_def)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   277
    apply (rule ex1I [THEN theI2], fast)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   278
    apply (blast intro: is_recfun_functional, blast) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   279
    done
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   280
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   281
constdefs
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   282
   M_is_recfun :: "[i=>o, i, i, [i=>o,i,i,i]=>o, i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   283
     "M_is_recfun(M,r,a,MH,f) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   284
      \<forall>z. M(z) -->
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   285
          (z \<in> f <-> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   286
           (\<exists>x y xa sx r_sx f_r_sx. 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   287
              M(x) & M(y) & M(xa) & M(sx) & M(r_sx) & M(f_r_sx) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   288
	      pair(M,x,y,z) & pair(M,x,a,xa) & upair(M,x,x,sx) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   289
              pre_image(M,r,sx,r_sx) & restriction(M,f,r_sx,f_r_sx) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   290
              xa \<in> r & MH(M, x, f_r_sx, y)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   291
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   292
lemma (in M_axioms) is_recfun_iff_M:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   293
     "[| M(r); M(a); M(f); \<forall>x g. M(x) & M(g) & function(g) --> M(H(x,g));
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   294
       \<forall>x g y. M(x) --> M(g) --> M(y) --> MH(M,x,g,y) <-> y = H(x,g) |] ==>
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   295
       is_recfun(r,a,H,f) <-> M_is_recfun(M,r,a,MH,f)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   296
apply (simp add: vimage_closed restrict_closed M_is_recfun_def is_recfun_relativize)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   297
apply (rule all_cong, safe)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   298
 apply (thin_tac "\<forall>x. ?P(x)")+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   299
 apply (blast dest: transM)  (*or del: allE*)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   300
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   301
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   302
lemma M_is_recfun_cong [cong]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   303
     "[| r = r'; a = a'; f = f'; 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   304
       !!x g y. [| M(x); M(g); M(y) |] ==> MH(M,x,g,y) <-> MH'(M,x,g,y) |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   305
      ==> M_is_recfun(M,r,a,MH,f) <-> M_is_recfun(M,r',a',MH',f')"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   306
by (simp add: M_is_recfun_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   307
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   308
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   309
constdefs
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   310
 (*This expresses ordinal addition as a formula in the LAST.  It also 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   311
   provides an abbreviation that can be used in the instance of strong
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   312
   replacement below.  Here j is used to define the relation, namely
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   313
   Memrel(succ(j)), while x determines the domain of f.*)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   314
 is_oadd_fun :: "[i=>o,i,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   315
    "is_oadd_fun(M,i,j,x,f) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   316
       (\<forall>sj msj. M(sj) --> M(msj) --> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   317
                 successor(M,j,sj) --> membership(M,sj,msj) --> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   318
	         M_is_recfun(M, msj, x, 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   319
		     %M x g y. \<exists>gx. M(gx) & image(M,g,x,gx) & union(M,i,gx,y),
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   320
		     f))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   321
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   322
 is_oadd :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   323
    "is_oadd(M,i,j,k) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   324
        (~ ordinal(M,i) & ~ ordinal(M,j) & k=0) |
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   325
        (~ ordinal(M,i) & ordinal(M,j) & k=j) |
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   326
        (ordinal(M,i) & ~ ordinal(M,j) & k=i) |
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   327
        (ordinal(M,i) & ordinal(M,j) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   328
	 (\<exists>f fj sj. M(f) & M(fj) & M(sj) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   329
		    successor(M,j,sj) & is_oadd_fun(M,i,sj,sj,f) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   330
		    fun_apply(M,f,j,fj) & fj = k))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   331
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   332
 (*NEEDS RELATIVIZATION*)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   333
 omult_eqns :: "[i,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   334
    "omult_eqns(i,x,g,z) ==
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   335
            Ord(x) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   336
	    (x=0 --> z=0) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   337
            (\<forall>j. x = succ(j) --> z = g`j ++ i) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   338
            (Limit(x) --> z = \<Union>(g``x))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   339
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   340
 is_omult_fun :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   341
    "is_omult_fun(M,i,j,f) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   342
	    (\<exists>df. M(df) & is_function(M,f) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   343
                  is_domain(M,f,df) & subset(M, j, df)) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   344
            (\<forall>x\<in>j. omult_eqns(i,x,f,f`x))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   345
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   346
 is_omult :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   347
    "is_omult(M,i,j,k) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   348
	\<exists>f fj sj. M(f) & M(fj) & M(sj) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   349
                  successor(M,j,sj) & is_omult_fun(M,i,sj,f) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   350
                  fun_apply(M,f,j,fj) & fj = k"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   351
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   352
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   353
locale M_recursion = M_axioms +
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   354
  assumes oadd_strong_replacement:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   355
   "[| M(i); M(j) |] ==>
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   356
    strong_replacement(M, 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   357
         \<lambda>x z. \<exists>y f fx. M(y) & M(f) & M(fx) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   358
		         pair(M,x,y,z) & is_oadd_fun(M,i,j,x,f) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   359
		         image(M,f,x,fx) & y = i Un fx)" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   360
 and omult_strong_replacement':
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   361
   "[| M(i); M(j) |] ==>
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   362
    strong_replacement(M, \<lambda>x z. \<exists>y g. M(y) & M(g) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   363
	     pair(M,x,y,z) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   364
	     is_recfun(Memrel(succ(j)),x,%x g. THE z. omult_eqns(i,x,g,z),g) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   365
	     y = (THE z. omult_eqns(i, x, g, z)))" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   366
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   367
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   368
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   369
text{*is_oadd_fun: Relating the pure "language of set theory" to Isabelle/ZF*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   370
lemma (in M_recursion) is_oadd_fun_iff:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   371
   "[| a\<le>j; M(i); M(j); M(a); M(f) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   372
    ==> is_oadd_fun(M,i,j,a,f) <->
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   373
	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
   374
apply (frule lt_Ord) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   375
apply (simp add: is_oadd_fun_def Memrel_closed Un_closed 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   376
             is_recfun_iff_M [of concl: _ _ "%x g. i Un g``x", THEN iff_sym]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   377
             image_closed is_recfun_iff_equation  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   378
             Ball_def lt_trans [OF ltI, of _ a] lt_Memrel)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   379
apply (simp add: lt_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   380
apply (blast dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   381
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   382
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   383
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   384
lemma (in M_recursion) oadd_strong_replacement':
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   385
    "[| M(i); M(j) |] ==>
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   386
     strong_replacement(M, \<lambda>x z. \<exists>y g. M(y) & M(g) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   387
		  pair(M,x,y,z) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   388
		  is_recfun(Memrel(succ(j)),x,%x g. i Un g``x,g) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   389
		  y = i Un g``x)" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   390
apply (insert oadd_strong_replacement [of i j]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   391
apply (simp add: Memrel_closed Un_closed image_closed is_oadd_fun_def
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   392
                 is_recfun_iff_M)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   393
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   394
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   395
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   396
lemma (in M_recursion) exists_oadd:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   397
    "[| Ord(j);  M(i);  M(j) |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   398
     ==> \<exists>f. M(f) & is_recfun(Memrel(succ(j)), j, %x g. i Un g``x, f)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   399
apply (rule wf_exists_is_recfun) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   400
apply (rule wf_Memrel [THEN wf_imp_wf_on]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   401
apply (rule trans_Memrel [THEN trans_imp_trans_on], simp)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   402
apply (rule succI1) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   403
apply (blast intro: oadd_strong_replacement') 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   404
apply (simp_all add: Memrel_type Memrel_closed Un_closed image_closed)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   405
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   406
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   407
lemma (in M_recursion) exists_oadd_fun:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   408
    "[| Ord(j);  M(i);  M(j) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   409
     ==> \<exists>f. M(f) & is_oadd_fun(M,i,succ(j),succ(j),f)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   410
apply (rule exists_oadd [THEN exE])
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   411
apply (erule Ord_succ, assumption, simp) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   412
apply (rename_tac f, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   413
apply (frule is_recfun_type)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   414
apply (rule_tac x=f in exI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   415
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
   416
                 is_oadd_fun_iff Ord_trans [OF _ succI1])
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   417
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   418
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   419
lemma (in M_recursion) is_oadd_fun_apply:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   420
    "[| 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
   421
     ==> f`x = i Un (\<Union>k\<in>x. {f ` k})"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   422
apply (simp add: is_oadd_fun_iff lt_Ord2, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   423
apply (frule lt_closed, simp)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   424
apply (frule leI [THEN le_imp_subset])  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   425
apply (simp add: image_fun, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   426
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   427
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   428
lemma (in M_recursion) is_oadd_fun_iff_oadd [rule_format]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   429
    "[| 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
   430
     ==> j<J --> f`j = i++j"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   431
apply (erule_tac i=j in trans_induct, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   432
apply (subgoal_tac "\<forall>k\<in>x. k<J")
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   433
 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
   434
apply (blast intro: lt_trans ltI lt_Ord) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   435
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   436
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   437
lemma (in M_recursion) oadd_abs_fun_apply_iff:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   438
    "[| M(i); M(J); M(f); M(k); j<J; is_oadd_fun(M,i,J,J,f) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   439
     ==> fun_apply(M,f,j,k) <-> f`j = k"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   440
by (force simp add: lt_def is_oadd_fun_iff subsetD typed_apply_abs) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   441
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   442
lemma (in M_recursion) Ord_oadd_abs:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   443
    "[| M(i); M(j); M(k); Ord(i); Ord(j) |] ==> is_oadd(M,i,j,k) <-> k = i++j"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   444
apply (simp add: is_oadd_def oadd_abs_fun_apply_iff is_oadd_fun_iff_oadd)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   445
apply (frule exists_oadd_fun [of j i], blast+)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   446
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   447
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   448
lemma (in M_recursion) oadd_abs:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   449
    "[| 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
   450
apply (case_tac "Ord(i) & Ord(j)")
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   451
 apply (simp add: Ord_oadd_abs)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   452
apply (auto simp add: is_oadd_def oadd_eq_if_raw_oadd)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   453
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   454
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   455
lemma (in M_recursion) oadd_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   456
    "[| M(i); M(j) |] ==> M(i++j)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   457
apply (simp add: oadd_eq_if_raw_oadd, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   458
apply (simp add: raw_oadd_eq_oadd) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   459
apply (frule exists_oadd_fun [of j i], auto)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   460
apply (simp add: apply_closed is_oadd_fun_iff_oadd [symmetric]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   461
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   462
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   463
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   464
text{*Ordinal Multiplication*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   465
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   466
lemma omult_eqns_unique:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   467
     "[| omult_eqns(i,x,g,z); omult_eqns(i,x,g,z') |] ==> z=z'";
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   468
apply (simp add: omult_eqns_def, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   469
apply (erule Ord_cases, simp_all) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   470
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   471
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   472
lemma omult_eqns_0: "omult_eqns(i,0,g,z) <-> z=0"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   473
by (simp add: omult_eqns_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   474
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   475
lemma the_omult_eqns_0: "(THE z. omult_eqns(i,0,g,z)) = 0"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   476
by (simp add: omult_eqns_0)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   477
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   478
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
   479
by (simp add: omult_eqns_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   480
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   481
lemma the_omult_eqns_succ:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   482
     "Ord(j) ==> (THE z. omult_eqns(i,succ(j),g,z)) = g`j ++ i"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   483
by (simp add: omult_eqns_succ) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   484
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   485
lemma omult_eqns_Limit:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   486
     "Limit(x) ==> omult_eqns(i,x,g,z) <-> z = \<Union>(g``x)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   487
apply (simp add: omult_eqns_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   488
apply (blast intro: Limit_is_Ord) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   489
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   490
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   491
lemma the_omult_eqns_Limit:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   492
     "Limit(x) ==> (THE z. omult_eqns(i,x,g,z)) = \<Union>(g``x)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   493
by (simp add: omult_eqns_Limit)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   494
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   495
lemma omult_eqns_Not: "~ Ord(x) ==> ~ omult_eqns(i,x,g,z)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   496
by (simp add: omult_eqns_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   497
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   498
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   499
lemma (in M_recursion) the_omult_eqns_closed:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   500
    "[| M(i); M(x); M(g); function(g) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   501
     ==> M(THE z. omult_eqns(i, x, g, z))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   502
apply (case_tac "Ord(x)")
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   503
 prefer 2 apply (simp add: omult_eqns_Not) --{*trivial, non-Ord case*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   504
apply (erule Ord_cases) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   505
  apply (simp add: omult_eqns_0)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   506
 apply (simp add: omult_eqns_succ apply_closed oadd_closed) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   507
apply (simp add: omult_eqns_Limit) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   508
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   509
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   510
lemma (in M_recursion) exists_omult:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   511
    "[| Ord(j);  M(i);  M(j) |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   512
     ==> \<exists>f. M(f) & is_recfun(Memrel(succ(j)), j, %x g. THE z. omult_eqns(i,x,g,z), f)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   513
apply (rule wf_exists_is_recfun) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   514
apply (rule wf_Memrel [THEN wf_imp_wf_on]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   515
apply (rule trans_Memrel [THEN trans_imp_trans_on], simp)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   516
apply (rule succI1) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   517
apply (blast intro: omult_strong_replacement') 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   518
apply (simp_all add: Memrel_type Memrel_closed Un_closed image_closed)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   519
apply (blast intro: the_omult_eqns_closed) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   520
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   521
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   522
lemma (in M_recursion) exists_omult_fun:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   523
    "[| Ord(j);  M(i);  M(j) |] ==> \<exists>f. M(f) & is_omult_fun(M,i,succ(j),f)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   524
apply (rule exists_omult [THEN exE])
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   525
apply (erule Ord_succ, assumption, simp) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   526
apply (rename_tac f, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   527
apply (frule is_recfun_type)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   528
apply (rule_tac x=f in exI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   529
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
   530
                 is_omult_fun_def Ord_trans [OF _ succI1])
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   531
apply (force dest: Ord_in_Ord' 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   532
             simp add: omult_eqns_def the_omult_eqns_0 the_omult_eqns_succ
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   533
                       the_omult_eqns_Limit) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   534
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   535
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   536
lemma (in M_recursion) is_omult_fun_apply_0:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   537
    "[| 0 < j; is_omult_fun(M,i,j,f) |] ==> f`0 = 0"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   538
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
   539
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   540
lemma (in M_recursion) is_omult_fun_apply_succ:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   541
    "[| 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
   542
by (simp add: is_omult_fun_def omult_eqns_def lt_def, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   543
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   544
lemma (in M_recursion) is_omult_fun_apply_Limit:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   545
    "[| 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
   546
     ==> f ` x = (\<Union>y\<in>x. f`y)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   547
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
   548
apply (drule subset_trans [OF OrdmemD], assumption+)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   549
apply (simp add: ball_conj_distrib omult_Limit image_function)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   550
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   551
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   552
lemma (in M_recursion) is_omult_fun_eq_omult:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   553
    "[| 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
   554
     ==> j<J --> f`j = i**j"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   555
apply (erule_tac i=j in trans_induct3)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   556
apply (safe del: impCE)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   557
  apply (simp add: is_omult_fun_apply_0) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   558
 apply (subgoal_tac "x<J") 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   559
  apply (simp add: is_omult_fun_apply_succ omult_succ)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   560
 apply (blast intro: lt_trans) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   561
apply (subgoal_tac "\<forall>k\<in>x. k<J")
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   562
 apply (simp add: is_omult_fun_apply_Limit omult_Limit) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   563
apply (blast intro: lt_trans ltI lt_Ord) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   564
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   565
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   566
lemma (in M_recursion) omult_abs_fun_apply_iff:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   567
    "[| M(i); M(J); M(f); M(k); j<J; is_omult_fun(M,i,J,f) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   568
     ==> fun_apply(M,f,j,k) <-> f`j = k"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   569
by (auto simp add: lt_def is_omult_fun_def subsetD apply_abs) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   570
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   571
lemma (in M_recursion) omult_abs:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   572
    "[| M(i); M(j); M(k); Ord(i); Ord(j) |] ==> is_omult(M,i,j,k) <-> k = i**j"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   573
apply (simp add: is_omult_def omult_abs_fun_apply_iff is_omult_fun_eq_omult)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   574
apply (frule exists_omult_fun [of j i], blast+)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   575
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   576
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   577
end
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   578