src/ZF/Constructible/Rank_Separation.thy
author paulson <lp15@cam.ac.uk>
Tue, 04 Feb 2020 16:19:15 +0000
changeset 71417 89d05db6dd1f
parent 69593 3dda49e08b9d
child 76213 e44d86131648
permissions -rw-r--r--
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
     1
(*  Title:      ZF/Constructible/Rank_Separation.thy
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
     3
*)
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
     4
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
     5
section \<open>Separation for Facts About Order Types, Rank Functions and 
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
     6
      Well-Founded Relations\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
     7
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 13807
diff changeset
     8
theory Rank_Separation imports Rank Rec_Separation begin
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
     9
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    10
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
    11
text\<open>This theory proves all instances needed for locales
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
    12
 \<open>M_ordertype\<close> and  \<open>M_wfrank\<close>.  But the material is not
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
    13
 needed for proving the relative consistency of AC.\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    14
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
    15
subsection\<open>The Locale \<open>M_ordertype\<close>\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    16
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
    17
subsubsection\<open>Separation for Order-Isomorphisms\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    18
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    19
lemma well_ord_iso_Reflects:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    20
  "REFLECTS[\<lambda>x. x\<in>A \<longrightarrow>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    21
                (\<exists>y[L]. \<exists>p[L]. fun_apply(L,f,x,y) & pair(L,y,x,p) & p \<in> r),
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    22
        \<lambda>i x. x\<in>A \<longrightarrow> (\<exists>y \<in> Lset(i). \<exists>p \<in> Lset(i).
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
    23
                fun_apply(##Lset(i),f,x,y) & pair(##Lset(i),y,x,p) & p \<in> r)]"
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    24
by (intro FOL_reflections function_reflections)
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    25
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    26
lemma well_ord_iso_separation:
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    27
     "[| L(A); L(f); L(r) |]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    28
      ==> separation (L, \<lambda>x. x\<in>A \<longrightarrow> (\<exists>y[L]. (\<exists>p[L].
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    29
                     fun_apply(L,f,x,y) & pair(L,y,x,p) & p \<in> r)))"
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
    30
apply (rule gen_separation_multi [OF well_ord_iso_Reflects, of "{A,f,r}"], 
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
    31
       auto)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
    32
apply (rule_tac env="[A,f,r]" in DPow_LsetI)
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    33
apply (rule sep_rules | simp)+
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    34
done
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    35
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    36
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
    37
subsubsection\<open>Separation for \<^term>\<open>obase\<close>\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    38
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    39
lemma obase_reflects:
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    40
  "REFLECTS[\<lambda>a. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    41
             ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) &
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    42
             order_isomorphism(L,par,r,x,mx,g),
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    43
        \<lambda>i a. \<exists>x \<in> Lset(i). \<exists>g \<in> Lset(i). \<exists>mx \<in> Lset(i). \<exists>par \<in> Lset(i).
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
    44
             ordinal(##Lset(i),x) & membership(##Lset(i),x,mx) & pred_set(##Lset(i),A,a,r,par) &
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
    45
             order_isomorphism(##Lset(i),par,r,x,mx,g)]"
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    46
by (intro FOL_reflections function_reflections fun_plus_reflections)
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    47
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    48
lemma obase_separation:
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 67399
diff changeset
    49
     \<comment> \<open>part of the order type formalization\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    50
     "[| L(A); L(r) |]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    51
      ==> separation(L, \<lambda>a. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    52
             ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) &
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    53
             order_isomorphism(L,par,r,x,mx,g))"
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
    54
apply (rule gen_separation_multi [OF obase_reflects, of "{A,r}"], auto)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
    55
apply (rule_tac env="[A,r]" in DPow_LsetI)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
    56
apply (rule ordinal_iff_sats sep_rules | simp)+
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    57
done
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    58
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    59
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
    60
subsubsection\<open>Separation for a Theorem about \<^term>\<open>obase\<close>\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    61
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    62
lemma obase_equals_reflects:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    63
  "REFLECTS[\<lambda>x. x\<in>A \<longrightarrow> ~(\<exists>y[L]. \<exists>g[L].
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    64
                ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L].
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    65
                membership(L,y,my) & pred_set(L,A,x,r,pxr) &
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    66
                order_isomorphism(L,pxr,r,y,my,g))),
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    67
        \<lambda>i x. x\<in>A \<longrightarrow> ~(\<exists>y \<in> Lset(i). \<exists>g \<in> Lset(i).
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
    68
                ordinal(##Lset(i),y) & (\<exists>my \<in> Lset(i). \<exists>pxr \<in> Lset(i).
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
    69
                membership(##Lset(i),y,my) & pred_set(##Lset(i),A,x,r,pxr) &
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
    70
                order_isomorphism(##Lset(i),pxr,r,y,my,g)))]"
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    71
by (intro FOL_reflections function_reflections fun_plus_reflections)
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    72
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    73
lemma obase_equals_separation:
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    74
     "[| L(A); L(r) |]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    75
      ==> separation (L, \<lambda>x. x\<in>A \<longrightarrow> ~(\<exists>y[L]. \<exists>g[L].
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    76
                              ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L].
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    77
                              membership(L,y,my) & pred_set(L,A,x,r,pxr) &
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    78
                              order_isomorphism(L,pxr,r,y,my,g))))"
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
    79
apply (rule gen_separation_multi [OF obase_equals_reflects, of "{A,r}"], auto)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
    80
apply (rule_tac env="[A,r]" in DPow_LsetI)
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    81
apply (rule sep_rules | simp)+
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    82
done
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    83
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    84
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
    85
subsubsection\<open>Replacement for \<^term>\<open>omap\<close>\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    86
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    87
lemma omap_reflects:
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    88
 "REFLECTS[\<lambda>z. \<exists>a[L]. a\<in>B & (\<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    89
     ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) &
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    90
     pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g)),
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    91
 \<lambda>i z. \<exists>a \<in> Lset(i). a\<in>B & (\<exists>x \<in> Lset(i). \<exists>g \<in> Lset(i). \<exists>mx \<in> Lset(i).
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    92
        \<exists>par \<in> Lset(i).
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
    93
         ordinal(##Lset(i),x) & pair(##Lset(i),a,x,z) &
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
    94
         membership(##Lset(i),x,mx) & pred_set(##Lset(i),A,a,r,par) &
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
    95
         order_isomorphism(##Lset(i),par,r,x,mx,g))]"
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    96
by (intro FOL_reflections function_reflections fun_plus_reflections)
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    97
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    98
lemma omap_replacement:
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
    99
     "[| L(A); L(r) |]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   100
      ==> strong_replacement(L,
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   101
             \<lambda>a z. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L].
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   102
             ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) &
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   103
             pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g))"
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   104
apply (rule strong_replacementI)
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
   105
apply (rule_tac u="{A,r,B}" in gen_separation_multi [OF omap_reflects], auto)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
   106
apply (rule_tac env="[A,B,r]" in DPow_LsetI)
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   107
apply (rule sep_rules | simp)+
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   108
done
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   109
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   110
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   111
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   112
subsection\<open>Instantiating the locale \<open>M_ordertype\<close>\<close>
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   113
text\<open>Separation (and Strong Replacement) for basic set-theoretic constructions
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 58871
diff changeset
   114
such as intersection, Cartesian Product and image.\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   115
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   116
lemma M_ordertype_axioms_L: "M_ordertype_axioms(L)"
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   117
  apply (rule M_ordertype_axioms.intro)
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   118
       apply (assumption | rule well_ord_iso_separation
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 19931
diff changeset
   119
         obase_separation obase_equals_separation
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 19931
diff changeset
   120
         omap_replacement)+
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   121
  done
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   122
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   123
theorem M_ordertype_L: "M_ordertype(L)"
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   124
  apply (rule M_ordertype.intro)
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   125
   apply (rule M_basic_L)
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   126
  apply (rule M_ordertype_axioms_L) 
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   127
  done
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   128
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   129
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   130
subsection\<open>The Locale \<open>M_wfrank\<close>\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   131
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   132
subsubsection\<open>Separation for \<^term>\<open>wfrank\<close>\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   133
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   134
lemma wfrank_Reflects:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   135
 "REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   136
              ~ (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)),
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   137
      \<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) \<longrightarrow>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   138
         ~ (\<exists>f \<in> Lset(i).
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   139
            M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y),
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   140
                        rplus, x, f))]"
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   141
by (intro FOL_reflections function_reflections is_recfun_reflection tran_closure_reflection)
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   142
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   143
lemma wfrank_separation:
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   144
     "L(r) ==>
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   145
      separation (L, \<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   146
         ~ (\<exists>f[L]. M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f)))"
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   147
apply (rule gen_separation [OF wfrank_Reflects], simp)
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
   148
apply (rule_tac env="[r]" in DPow_LsetI)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
   149
apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   150
done
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   151
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   152
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   153
subsubsection\<open>Replacement for \<^term>\<open>wfrank\<close>\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   154
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   155
lemma wfrank_replacement_Reflects:
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   156
 "REFLECTS[\<lambda>z. \<exists>x[L]. x \<in> A &
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   157
        (\<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   158
         (\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z)  &
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   159
                        M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) &
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   160
                        is_range(L,f,y))),
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   161
 \<lambda>i z. \<exists>x \<in> Lset(i). x \<in> A &
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   162
      (\<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) \<longrightarrow>
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   163
       (\<exists>y \<in> Lset(i). \<exists>f \<in> Lset(i). pair(##Lset(i),x,y,z)  &
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   164
         M_is_recfun(##Lset(i), %x f y. is_range(##Lset(i),f,y), rplus, x, f) &
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   165
         is_range(##Lset(i),f,y)))]"
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   166
by (intro FOL_reflections function_reflections fun_plus_reflections
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   167
             is_recfun_reflection tran_closure_reflection)
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   168
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   169
lemma wfrank_strong_replacement:
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   170
     "L(r) ==>
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   171
      strong_replacement(L, \<lambda>x z.
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   172
         \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   173
         (\<exists>y[L]. \<exists>f[L]. pair(L,x,y,z)  &
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   174
                        M_is_recfun(L, %x f y. is_range(L,f,y), rplus, x, f) &
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   175
                        is_range(L,f,y)))"
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   176
apply (rule strong_replacementI)
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
   177
apply (rule_tac u="{r,B}" 
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
   178
         in gen_separation_multi [OF wfrank_replacement_Reflects], 
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
   179
       auto)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
   180
apply (rule_tac env="[r,B]" in DPow_LsetI)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
   181
apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   182
done
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   183
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   184
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   185
subsubsection\<open>Separation for Proving \<open>Ord_wfrank_range\<close>\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   186
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   187
lemma Ord_wfrank_Reflects:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   188
 "REFLECTS[\<lambda>x. \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   189
          ~ (\<forall>f[L]. \<forall>rangef[L].
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   190
             is_range(L,f,rangef) \<longrightarrow>
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   191
             M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) \<longrightarrow>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   192
             ordinal(L,rangef)),
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   193
      \<lambda>i x. \<forall>rplus \<in> Lset(i). tran_closure(##Lset(i),r,rplus) \<longrightarrow>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   194
          ~ (\<forall>f \<in> Lset(i). \<forall>rangef \<in> Lset(i).
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   195
             is_range(##Lset(i),f,rangef) \<longrightarrow>
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   196
             M_is_recfun(##Lset(i), \<lambda>x f y. is_range(##Lset(i),f,y),
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   197
                         rplus, x, f) \<longrightarrow>
13807
a28a8fbc76d4 changed ** to ## to avoid conflict with new comment syntax
paulson
parents: 13687
diff changeset
   198
             ordinal(##Lset(i),rangef))]"
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   199
by (intro FOL_reflections function_reflections is_recfun_reflection
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   200
          tran_closure_reflection ordinal_reflection)
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   201
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   202
lemma  Ord_wfrank_separation:
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   203
     "L(r) ==>
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   204
      separation (L, \<lambda>x.
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   205
         \<forall>rplus[L]. tran_closure(L,r,rplus) \<longrightarrow>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   206
          ~ (\<forall>f[L]. \<forall>rangef[L].
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   207
             is_range(L,f,rangef) \<longrightarrow>
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   208
             M_is_recfun(L, \<lambda>x f y. is_range(L,f,y), rplus, x, f) \<longrightarrow>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   209
             ordinal(L,rangef)))"
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   210
apply (rule gen_separation [OF Ord_wfrank_Reflects], simp)
13687
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
   211
apply (rule_tac env="[r]" in DPow_LsetI)
22dce9134953 simpler separation/replacement proofs
paulson
parents: 13634
diff changeset
   212
apply (rule sep_rules tran_closure_iff_sats is_recfun_iff_sats | simp)+
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   213
done
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   214
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   215
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   216
subsubsection\<open>Instantiating the locale \<open>M_wfrank\<close>\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   217
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   218
lemma M_wfrank_axioms_L: "M_wfrank_axioms(L)"
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   219
  apply (rule M_wfrank_axioms.intro)
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   220
   apply (assumption | rule
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   221
     wfrank_separation wfrank_strong_replacement Ord_wfrank_separation)+
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   222
  done
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   223
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   224
theorem M_wfrank_L: "M_wfrank(L)"
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   225
  apply (rule M_wfrank.intro)
19931
fb32b43e7f80 Restructured locales with predicates: import is now an interpretation.
ballarin
parents: 16417
diff changeset
   226
   apply (rule M_trancl_L)
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   227
  apply (rule M_wfrank_axioms_L) 
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   228
  done
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   229
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   230
lemmas exists_wfrank = M_wfrank.exists_wfrank [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   231
  and M_wellfoundedrank = M_wfrank.M_wellfoundedrank [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   232
  and Ord_wfrank_range = M_wfrank.Ord_wfrank_range [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   233
  and Ord_range_wellfoundedrank = M_wfrank.Ord_range_wellfoundedrank [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   234
  and function_wellfoundedrank = M_wfrank.function_wellfoundedrank [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   235
  and domain_wellfoundedrank = M_wfrank.domain_wellfoundedrank [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   236
  and wellfoundedrank_type = M_wfrank.wellfoundedrank_type [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   237
  and Ord_wellfoundedrank = M_wfrank.Ord_wellfoundedrank [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   238
  and wellfoundedrank_eq = M_wfrank.wellfoundedrank_eq [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   239
  and wellfoundedrank_lt = M_wfrank.wellfoundedrank_lt [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   240
  and wellfounded_imp_subset_rvimage = M_wfrank.wellfounded_imp_subset_rvimage [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   241
  and wellfounded_imp_wf = M_wfrank.wellfounded_imp_wf [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   242
  and wellfounded_on_imp_wf_on = M_wfrank.wellfounded_on_imp_wf_on [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   243
  and wf_abs = M_wfrank.wf_abs [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   244
  and wf_on_abs = M_wfrank.wf_on_abs [OF M_wfrank_L]
99a593b49b04 Re-organization of Constructible theories
paulson
parents:
diff changeset
   245
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 61798
diff changeset
   246
end