src/ZF/Constructible/WFrec.thy
author haftmann
Thu, 24 Jul 2025 17:46:29 +0200
changeset 82902 99a720d3ed8f
parent 76215 a642599ffdea
permissions -rw-r--r--
clarified code setup
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
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     3
*)
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13428
diff changeset
     4
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
     5
section\<open>Relativized Well-Founded Recursion\<close>
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
     6
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13721
diff changeset
     7
theory WFrec imports Wellorderings begin
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     8
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     9
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
    10
subsection\<open>General Lemmas\<close>
13506
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
    11
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    12
(*Many of these might be useful in WF.thy*)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    13
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
    14
lemma apply_recfun2:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    15
    "\<lbrakk>is_recfun(r,a,H,f); \<langle>x,i\<rangle>:f\<rbrakk> \<Longrightarrow> i = H(x, restrict(f,r-``{x}))"
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
    16
apply (frule apply_recfun) 
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
    17
 apply (blast dest: is_recfun_type fun_is_rel) 
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
    18
apply (simp add: function_apply_equality [OF _ is_recfun_imp_function])
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    19
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    20
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
    21
text\<open>Expresses \<open>is_recfun\<close> as a recursion equation\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    22
lemma is_recfun_iff_equation:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    23
     "is_recfun(r,a,H,f) \<longleftrightarrow>
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
    24
           f \<in> r -`` {a} \<rightarrow> range(f) \<and>
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
    25
           (\<forall>x \<in> r-``{a}. f`x = H(x, restrict(f, r-``{x})))"  
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    26
apply (rule iffI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    27
 apply (simp add: is_recfun_type apply_recfun Ball_def vimage_singleton_iff, 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    28
        clarify)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    29
apply (simp add: is_recfun_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    30
apply (rule fun_extension) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    31
  apply assumption
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    32
 apply (fast intro: lam_type, simp) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    33
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    34
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    35
lemma is_recfun_imp_in_r: "\<lbrakk>is_recfun(r,a,H,f); \<langle>x,i\<rangle> \<in> f\<rbrakk> \<Longrightarrow> \<langle>x, a\<rangle> \<in> r"
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
    36
by (blast dest: is_recfun_type fun_is_rel)
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    37
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    38
lemma trans_Int_eq:
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    39
      "\<lbrakk>trans(r); \<langle>y,x\<rangle> \<in> r\<rbrakk> \<Longrightarrow> 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
    40
by (blast intro: transD) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    41
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    42
lemma is_recfun_restrict_idem:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    43
     "is_recfun(r,a,H,f) \<Longrightarrow> restrict(f, r -`` {a}) = f"
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    44
apply (drule is_recfun_type)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    45
apply (auto simp add: Pi_iff subset_Sigma_imp_relation restrict_idem)  
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    46
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    47
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    48
lemma is_recfun_cong_lemma:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    49
  "\<lbrakk>is_recfun(r,a,H,f); r = r'; a = a'; f = f'; 
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    50
      \<And>x g. \<lbrakk><x,a'> \<in> r'; relation(g); domain(g) \<subseteq> r' -``{x}\<rbrakk> 
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    51
             \<Longrightarrow> H(x,g) = H'(x,g)\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    52
   \<Longrightarrow> is_recfun(r',a',H',f')"
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    53
apply (simp add: is_recfun_def) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    54
apply (erule trans) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    55
apply (rule lam_cong) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    56
apply (simp_all add: vimage_singleton_iff Int_lower2)  
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    57
done
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    58
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
    59
text\<open>For \<open>is_recfun\<close> we need only pay attention to functions
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
    60
      whose domains are initial segments of \<^term>\<open>r\<close>.\<close>
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    61
lemma is_recfun_cong:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    62
  "\<lbrakk>r = r'; a = a'; f = f'; 
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    63
      \<And>x g. \<lbrakk><x,a'> \<in> r'; relation(g); domain(g) \<subseteq> r' -``{x}\<rbrakk> 
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    64
             \<Longrightarrow> H(x,g) = H'(x,g)\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    65
   \<Longrightarrow> is_recfun(r,a,H,f) \<longleftrightarrow> is_recfun(r',a',H',f')"
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    66
apply (rule iffI)
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
    67
txt\<open>Messy: fast and blast don't work for some reason\<close>
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    68
apply (erule is_recfun_cong_lemma, auto) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    69
apply (erule is_recfun_cong_lemma)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    70
apply (blast intro: sym)+
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    71
done
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    72
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
    73
subsection\<open>Reworking of the Recursion Theory Within \<^term>\<open>M\<close>\<close>
13506
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
    74
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
    75
lemma (in M_basic) is_recfun_separation':
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    76
    "\<lbrakk>f \<in> r -`` {a} \<rightarrow> range(f); g \<in> r -`` {b} \<rightarrow> range(g);
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    77
        M(r); M(f); M(g); M(a); M(b)\<rbrakk> 
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    78
     \<Longrightarrow> separation(M, \<lambda>x. \<not> (\<langle>x, a\<rangle> \<in> r \<longrightarrow> \<langle>x, b\<rangle> \<in> r \<longrightarrow> f ` x = g ` x))"
13319
23de7b3af453 More Separation proofs
paulson
parents: 13306
diff changeset
    79
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
    80
apply (simp add: vimage_singleton_iff)
13319
23de7b3af453 More Separation proofs
paulson
parents: 13306
diff changeset
    81
done
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    82
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
    83
text\<open>Stated using \<^term>\<open>trans(r)\<close> rather than
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
    84
      \<^term>\<open>transitive_rel(M,A,r)\<close> because the latter rewrites to
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
    85
      the former anyway, by \<open>transitive_rel_abs\<close>.
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    86
      As always, theorems should be expressed in simplified form.
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
    87
      The last three M-premises are redundant because of \<^term>\<open>M(r)\<close>, 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
    88
      but without them we'd have to undertake
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
    89
      more work to set up the induction formula.\<close>
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
    90
lemma (in M_basic) is_recfun_equal [rule_format]: 
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    91
    "\<lbrakk>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
    92
       wellfounded(M,r);  trans(r);
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
    93
       M(f); M(g); M(r); M(x); M(a); M(b)\<rbrakk> 
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
    94
     \<Longrightarrow> \<langle>x,a\<rangle> \<in> r \<longrightarrow> \<langle>x,b\<rangle> \<in> r \<longrightarrow> f`x=g`x"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13319
diff changeset
    95
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
    96
apply (frule_tac f=g in is_recfun_type) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    97
apply (simp add: is_recfun_def)
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    98
apply (erule_tac a=x in wellfounded_induct, assumption+)
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
    99
txt\<open>Separation to justify the induction\<close>
13319
23de7b3af453 More Separation proofs
paulson
parents: 13306
diff changeset
   100
 apply (blast intro: is_recfun_separation') 
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   101
txt\<open>Now the inductive argument itself\<close>
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   102
apply clarify 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   103
apply (erule ssubst)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   104
apply (simp (no_asm_simp) add: vimage_singleton_iff restrict_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   105
apply (rename_tac x1)
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   106
apply (rule_tac t="\<lambda>z. H(x1,z)" in subst_context) 
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   107
apply (subgoal_tac "\<forall>y \<in> r-``{x1}. \<forall>z. \<langle>y,z\<rangle>\<in>f \<longleftrightarrow> \<langle>y,z\<rangle>\<in>g")
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   108
 apply (blast intro: transD) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   109
apply (simp add: apply_iff) 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   110
apply (blast intro: transD sym) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   111
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   112
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
   113
lemma (in M_basic) is_recfun_cut: 
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   114
    "\<lbrakk>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
   115
       wellfounded(M,r); trans(r); 
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   116
       M(f); M(g); M(r); \<langle>b,a\<rangle> \<in> r\<rbrakk>   
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   117
      \<Longrightarrow> restrict(f, r-``{b}) = g"
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13319
diff changeset
   118
apply (frule_tac f=f in is_recfun_type) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   119
apply (rule fun_extension) 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   120
apply (blast intro: transD restrict_type2) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   121
apply (erule is_recfun_type, simp) 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   122
apply (blast intro: is_recfun_equal transD dest: transM) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   123
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   124
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
   125
lemma (in M_basic) is_recfun_functional:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   126
     "\<lbrakk>is_recfun(r,a,H,f);  is_recfun(r,a,H,g);  
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   127
       wellfounded(M,r); trans(r); M(f); M(g); M(r)\<rbrakk> \<Longrightarrow> f=g"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   128
apply (rule fun_extension)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   129
apply (erule is_recfun_type)+
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   130
apply (blast intro!: is_recfun_equal dest: transM) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   131
done 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   132
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   133
text\<open>Tells us that \<open>is_recfun\<close> can (in principle) be relativized.\<close>
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
   134
lemma (in M_basic) is_recfun_relativize:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   135
  "\<lbrakk>M(r); M(f); \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk> 
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   136
   \<Longrightarrow> is_recfun(r,a,H,f) \<longleftrightarrow>
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   137
       (\<forall>z[M]. z \<in> f \<longleftrightarrow> 
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   138
        (\<exists>x[M]. \<langle>x,a\<rangle> \<in> r \<and> z = <x, H(x, restrict(f, r-``{x}))>))"
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   139
apply (simp add: is_recfun_def lam_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   140
apply (safe intro!: equalityI) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   141
   apply (drule equalityD1 [THEN subsetD], assumption) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   142
   apply (blast dest: pair_components_in_M) 
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   143
  apply (blast elim!: equalityE dest: pair_components_in_M)
13615
449a70d88b38 Numerous cosmetic changes, prompted by the new simplifier
paulson
parents: 13564
diff changeset
   144
 apply (frule transM, assumption) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   145
 apply simp  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   146
 apply blast
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   147
apply (subgoal_tac "is_function(M,f)")
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   148
 txt\<open>We use \<^term>\<open>is_function\<close> rather than \<^term>\<open>function\<close> because
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   149
      the subgoal's easier to prove with relativized quantifiers!\<close>
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   150
 prefer 2 apply (simp add: is_function_def) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   151
apply (frule pair_components_in_M, assumption) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   152
apply (simp add: is_recfun_imp_function function_restrictI) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   153
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   154
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
   155
lemma (in M_basic) is_recfun_restrict:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   156
     "\<lbrakk>wellfounded(M,r); trans(r); is_recfun(r,x,H,f); \<langle>y,x\<rangle> \<in> r; 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   157
       M(r); M(f); 
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   158
       \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   159
       \<Longrightarrow> is_recfun(r, y, H, restrict(f, r -`` {y}))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   160
apply (frule pair_components_in_M, assumption, clarify) 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   161
apply (simp (no_asm_simp) add: is_recfun_relativize restrict_iff
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   162
           trans_Int_eq)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   163
apply safe
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   164
  apply (simp_all add: vimage_singleton_iff is_recfun_type [THEN apply_iff]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   165
  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
   166
  apply (frule_tac x=xa in apply_recfun, blast intro: transD)  
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13245
diff changeset
   167
  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
   168
                   is_recfun_imp_function function_restrictI)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   169
apply (blast intro: apply_recfun dest: transD)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   170
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   171
 
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
   172
lemma (in M_basic) restrict_Y_lemma:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   173
   "\<lbrakk>wellfounded(M,r); trans(r); M(r);
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   174
       \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g));  M(Y);
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   175
       \<forall>b[M]. 
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   176
           b \<in> Y \<longleftrightarrow>
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   177
           (\<exists>x[M]. \<langle>x,a1\<rangle> \<in> r \<and>
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   178
            (\<exists>y[M]. b = \<langle>x,y\<rangle> \<and> (\<exists>g[M]. is_recfun(r,x,H,g) \<and> y = H(x,g))));
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   179
          \<langle>x,a1\<rangle> \<in> r; is_recfun(r,x,H,f); M(f)\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   180
       \<Longrightarrow> restrict(Y, r -`` {x}) = f"
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   181
apply (subgoal_tac "\<forall>y \<in> r-``{x}. \<forall>z. \<langle>y,z\<rangle>:Y \<longleftrightarrow> \<langle>y,z\<rangle>:f") 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   182
 apply (simp (no_asm_simp) add: restrict_def) 
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 61798
diff changeset
   183
 apply (thin_tac "rall(M,P)" for P)+  \<comment> \<open>essential for efficiency\<close>
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   184
 apply (frule is_recfun_type [THEN fun_is_rel], blast)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   185
apply (frule pair_components_in_M, assumption, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   186
apply (rule iffI)
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   187
 apply (frule_tac y="\<langle>y,z\<rangle>" in transM, assumption)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   188
 apply (clarsimp simp add: vimage_singleton_iff is_recfun_type [THEN apply_iff]
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 21404
diff changeset
   189
                           apply_recfun is_recfun_cut) 
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   190
txt\<open>Opposite inclusion: something in f, show in Y\<close>
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   191
apply (frule_tac y="\<langle>y,z\<rangle>" in transM, assumption)  
13293
paulson
parents: 13269
diff changeset
   192
apply (simp add: vimage_singleton_iff) 
paulson
parents: 13269
diff changeset
   193
apply (rule conjI) 
paulson
parents: 13269
diff changeset
   194
 apply (blast dest: transD) 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   195
apply (rule_tac x="restrict(f, r -`` {y})" in rexI) 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   196
apply (simp_all add: is_recfun_restrict
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   197
                     apply_recfun is_recfun_type [THEN apply_iff]) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   198
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   199
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   200
text\<open>For typical applications of Replacement for recursive definitions\<close>
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
   201
lemma (in M_basic) univalent_is_recfun:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   202
     "\<lbrakk>wellfounded(M,r); trans(r); M(r)\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   203
      \<Longrightarrow> univalent (M, A, \<lambda>x p. 
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   204
              \<exists>y[M]. p = \<langle>x,y\<rangle> \<and> (\<exists>f[M]. is_recfun(r,x,H,f) \<and> y = H(x,f)))"
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   205
apply (simp add: univalent_def) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   206
apply (blast dest: is_recfun_functional) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   207
done
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   208
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   209
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   210
text\<open>Proof of the inductive step for \<open>exists_is_recfun\<close>, since
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   211
      we must prove two versions.\<close>
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
   212
lemma (in M_basic) exists_is_recfun_indstep:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   213
    "\<lbrakk>\<forall>y. \<langle>y, a1\<rangle> \<in> r \<longrightarrow> (\<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
   214
       wellfounded(M,r); trans(r); M(r); M(a1);
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   215
       strong_replacement(M, \<lambda>x z. 
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   216
              \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) \<and> is_recfun(r,x,H,g) \<and> y = H(x,g)); 
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   217
       \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>   
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   218
      \<Longrightarrow> \<exists>f[M]. is_recfun(r,a1,H,f)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   219
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
   220
  apply blast 
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   221
 txt\<open>Discharge the "univalent" obligation of Replacement\<close>
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   222
 apply (simp add: univalent_is_recfun) 
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   223
txt\<open>Show that the constructed object satisfies \<open>is_recfun\<close>\<close> 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   224
apply clarify 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   225
apply (rule_tac x=Y in rexI)  
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   226
txt\<open>Unfold only the top-level occurrence of \<^term>\<open>is_recfun\<close>\<close>
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   227
apply (simp (no_asm_simp) add: is_recfun_relativize [of concl: _ a1])
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   228
txt\<open>The big iff-formula defining \<^term>\<open>Y\<close> is now redundant\<close>
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   229
apply safe 
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   230
 apply (simp add: vimage_singleton_iff restrict_Y_lemma [of r H _ a1]) 
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   231
txt\<open>one more case\<close>
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   232
apply (simp (no_asm_simp) add: Bex_def vimage_singleton_iff)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   233
apply (drule_tac x1=x in spec [THEN mp], assumption, clarify) 
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   234
apply (rename_tac f) 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   235
apply (rule_tac x=f in rexI) 
13293
paulson
parents: 13269
diff changeset
   236
apply (simp_all add: restrict_Y_lemma [of r H])
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   237
txt\<open>FIXME: should not be needed!\<close>
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   238
apply (subst restrict_Y_lemma [of r H])
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   239
apply (simp add: vimage_singleton_iff)+
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13295
diff changeset
   240
apply blast+
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   241
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   242
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   243
text\<open>Relativized version, when we have the (currently weaker) premise
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   244
      \<^term>\<open>wellfounded(M,r)\<close>\<close>
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
   245
lemma (in M_basic) wellfounded_exists_is_recfun:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   246
    "\<lbrakk>wellfounded(M,r);  trans(r);  
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   247
       separation(M, \<lambda>x. \<not> (\<exists>f[M]. is_recfun(r, x, H, f)));
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   248
       strong_replacement(M, \<lambda>x z. 
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   249
          \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) \<and> is_recfun(r,x,H,g) \<and> y = H(x,g)); 
13251
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   250
       M(r);  M(a);  
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   251
       \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>   
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   252
      \<Longrightarrow> \<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
   253
apply (rule wellfounded_induct, assumption+, clarify)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   254
apply (rule exists_is_recfun_indstep, assumption+)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   255
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   256
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
   257
lemma (in M_basic) wf_exists_is_recfun [rule_format]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   258
    "\<lbrakk>wf(r);  trans(r);  M(r);  
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   259
       strong_replacement(M, \<lambda>x z. 
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   260
         \<exists>y[M]. \<exists>g[M]. pair(M,x,y,z) \<and> is_recfun(r,x,H,g) \<and> y = H(x,g)); 
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   261
       \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))\<rbrakk>   
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   262
      \<Longrightarrow> M(a) \<longrightarrow> (\<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
   263
apply (rule wf_induct, assumption+)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   264
apply (frule wf_imp_relativized)
74cb2af8811e new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents: 13247
diff changeset
   265
apply (intro impI)
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   266
apply (rule exists_is_recfun_indstep) 
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   267
      apply (blast dest: transM del: rev_rallE, assumption+)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   268
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   269
13506
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
   270
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   271
subsection\<open>Relativization of the ZF Predicate \<^term>\<open>is_recfun\<close>\<close>
13506
acc18a5d2b1a Various tweaks of the presentation
paulson
parents: 13505
diff changeset
   272
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
   273
definition
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   274
  M_is_recfun :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, i, i, i] \<Rightarrow> o" where
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   275
   "M_is_recfun(M,MH,r,a,f) \<equiv> 
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   276
     \<forall>z[M]. z \<in> f \<longleftrightarrow> 
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   277
            (\<exists>x[M]. \<exists>y[M]. \<exists>xa[M]. \<exists>sx[M]. \<exists>r_sx[M]. \<exists>f_r_sx[M]. 
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   278
               pair(M,x,y,z) \<and> pair(M,x,a,xa) \<and> upair(M,x,x,sx) \<and>
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   279
               pre_image(M,r,sx,r_sx) \<and> restriction(M,f,r_sx,f_r_sx) \<and>
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   280
               xa \<in> r \<and> MH(x, f_r_sx, y))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   281
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   282
definition
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   283
  is_wfrec :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, i, i, i] \<Rightarrow> o" where
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   284
   "is_wfrec(M,MH,r,a,z) \<equiv> 
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   285
      \<exists>f[M]. M_is_recfun(M,MH,r,a,f) \<and> MH(a,f,z)"
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   286
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   287
definition
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   288
  wfrec_replacement :: "[i\<Rightarrow>o, [i,i,i]\<Rightarrow>o, i] \<Rightarrow> o" where
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   289
   "wfrec_replacement(M,MH,r) \<equiv>
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   290
        strong_replacement(M, 
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   291
             \<lambda>x z. \<exists>y[M]. pair(M,x,y,z) \<and> is_wfrec(M,MH,r,x,y))"
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   292
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
   293
lemma (in M_basic) is_recfun_abs:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   294
     "\<lbrakk>\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g));  M(r); M(a); M(f); 
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   295
         relation2(M,MH,H)\<rbrakk> 
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   296
      \<Longrightarrow> M_is_recfun(M,MH,r,a,f) \<longleftrightarrow> is_recfun(r,a,H,f)"
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   297
apply (simp add: M_is_recfun_def relation2_def is_recfun_relativize)
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   298
apply (rule rall_cong)
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   299
apply (blast dest: transM)
13223
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]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   303
     "\<lbrakk>r = r'; a = a'; f = f'; 
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   304
       \<And>x g y. \<lbrakk>M(x); M(g); M(y)\<rbrakk> \<Longrightarrow> MH(x,g,y) \<longleftrightarrow> MH'(x,g,y)\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   305
      \<Longrightarrow> M_is_recfun(M,MH,r,a,f) \<longleftrightarrow> M_is_recfun(M,MH',r',a',f')"
13223
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
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
   308
lemma (in M_basic) is_wfrec_abs:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   309
     "\<lbrakk>\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)); 
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   310
         relation2(M,MH,H);  M(r); M(a); M(z)\<rbrakk>
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   311
      \<Longrightarrow> is_wfrec(M,MH,r,a,z) \<longleftrightarrow> 
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   312
          (\<exists>g[M]. is_recfun(r,a,H,g) \<and> z = H(a,g))"
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13615
diff changeset
   313
by (simp add: is_wfrec_def relation2_def is_recfun_abs)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   314
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   315
text\<open>Relating \<^term>\<open>wfrec_replacement\<close> to native constructs\<close>
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13506
diff changeset
   316
lemma (in M_basic) wfrec_replacement':
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   317
  "\<lbrakk>wfrec_replacement(M,MH,r);
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   318
     \<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)); 
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   319
     relation2(M,MH,H);  M(r)\<rbrakk> 
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   320
   \<Longrightarrow> strong_replacement(M, \<lambda>x z. \<exists>y[M]. 
76214
0c18df79b1c8 more modernisation of syntax
paulson <lp15@cam.ac.uk>
parents: 76213
diff changeset
   321
                pair(M,x,y,z) \<and> (\<exists>g[M]. is_recfun(r,x,H,g) \<and> y = H(x,g)))"
13615
449a70d88b38 Numerous cosmetic changes, prompted by the new simplifier
paulson
parents: 13564
diff changeset
   322
by (simp add: wfrec_replacement_def is_wfrec_abs) 
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   323
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   324
lemma wfrec_replacement_cong [cong]:
76213
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   325
     "\<lbrakk>\<And>x y z. \<lbrakk>M(x); M(y); M(z)\<rbrakk> \<Longrightarrow> MH(x,y,z) \<longleftrightarrow> MH'(x,y,z);
e44d86131648 Removal of obsolete ASCII syntax
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   326
         r=r'\<rbrakk> 
76215
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   327
      \<Longrightarrow> wfrec_replacement(M, \<lambda>x y. MH(x,y), r) \<longleftrightarrow> 
a642599ffdea More syntactic cleanup. LaTeX markup working
paulson <lp15@cam.ac.uk>
parents: 76214
diff changeset
   328
          wfrec_replacement(M, \<lambda>x y. MH'(x,y), r')" 
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   329
by (simp add: is_wfrec_def wfrec_replacement_def) 
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   330
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   331
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   332
end
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   333