src/ZF/Constructible/L_axioms.thy
author paulson
Tue, 16 Jul 2002 16:29:36 +0200
changeset 13363 c26eeb000470
parent 13352 3cd767f8d78b
child 13385 31df66ca0780
permissions -rw-r--r--
instantiation of locales M_trancl and M_wfrank; proofs of list_replacement{1,2}
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
     1
header {*The ZF Axioms (Except Separation) in L*}
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     2
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
     3
theory L_axioms = Formula + Relative + Reflection + MetaExists:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     4
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
     5
text {* The class L satisfies the premises of locale @{text M_triv_axioms} *}
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     6
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     7
lemma transL: "[| y\<in>x; L(x) |] ==> L(y)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     8
apply (insert Transset_Lset) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     9
apply (simp add: Transset_def L_def, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    10
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    11
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    12
lemma nonempty: "L(0)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    13
apply (simp add: L_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    14
apply (blast intro: zero_in_Lset) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    15
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    16
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    17
lemma upair_ax: "upair_ax(L)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    18
apply (simp add: upair_ax_def upair_def, clarify)
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
    19
apply (rule_tac x="{x,y}" in rexI)  
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
    20
apply (simp_all add: doubleton_in_L) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    21
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    22
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    23
lemma Union_ax: "Union_ax(L)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    24
apply (simp add: Union_ax_def big_union_def, clarify)
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
    25
apply (rule_tac x="Union(x)" in rexI)  
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
    26
apply (simp_all add: Union_in_L, auto) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    27
apply (blast intro: transL) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    28
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    29
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    30
lemma power_ax: "power_ax(L)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    31
apply (simp add: power_ax_def powerset_def Relative.subset_def, clarify)
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
    32
apply (rule_tac x="{y \<in> Pow(x). L(y)}" in rexI)  
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
    33
apply (simp_all add: LPow_in_L, auto)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    34
apply (blast intro: transL) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    35
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    36
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    37
subsubsection{*For L to satisfy Replacement *}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    38
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    39
(*Can't move these to Formula unless the definition of univalent is moved
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    40
there too!*)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    41
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    42
lemma LReplace_in_Lset:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    43
     "[|X \<in> Lset(i); univalent(L,X,Q); Ord(i)|] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    44
      ==> \<exists>j. Ord(j) & Replace(X, %x y. Q(x,y) & L(y)) \<subseteq> Lset(j)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    45
apply (rule_tac x="\<Union>y \<in> Replace(X, %x y. Q(x,y) & L(y)). succ(lrank(y))" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    46
       in exI)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    47
apply simp
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    48
apply clarify 
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
    49
apply (rule_tac a=x in UN_I)  
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    50
 apply (simp_all add: Replace_iff univalent_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    51
apply (blast dest: transL L_I) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    52
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    53
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    54
lemma LReplace_in_L: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    55
     "[|L(X); univalent(L,X,Q)|] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    56
      ==> \<exists>Y. L(Y) & Replace(X, %x y. Q(x,y) & L(y)) \<subseteq> Y"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    57
apply (drule L_D, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    58
apply (drule LReplace_in_Lset, assumption+)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    59
apply (blast intro: L_I Lset_in_Lset_succ)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    60
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    61
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    62
lemma replacement: "replacement(L,P)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    63
apply (simp add: replacement_def, clarify)
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13223
diff changeset
    64
apply (frule LReplace_in_L, assumption+, clarify) 
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
    65
apply (rule_tac x=Y in rexI)   
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
    66
apply (simp_all add: Replace_iff univalent_def, blast) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    67
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    68
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
    69
subsection{*Instantiating the locale @{text M_triv_axioms}*}
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
    70
text{*No instances of Separation yet.*}
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    71
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    72
lemma Lset_mono_le: "mono_le_subset(Lset)"
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    73
by (simp add: mono_le_subset_def le_imp_subset Lset_mono) 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    74
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    75
lemma Lset_cont: "cont_Ord(Lset)"
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    76
by (simp add: cont_Ord_def Limit_Lset_eq OUnion_def Limit_is_Ord) 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    77
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    78
lemmas Pair_in_Lset = Formula.Pair_in_LLimit;
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    79
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    80
lemmas L_nat = Ord_in_L [OF Ord_nat];
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    81
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    82
ML
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    83
{*
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    84
val transL = thm "transL";
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    85
val nonempty = thm "nonempty";
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    86
val upair_ax = thm "upair_ax";
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    87
val Union_ax = thm "Union_ax";
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    88
val power_ax = thm "power_ax";
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    89
val replacement = thm "replacement";
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    90
val L_nat = thm "L_nat";
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    91
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    92
fun kill_flex_triv_prems st = Seq.hd ((REPEAT_FIRST assume_tac) st);
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    93
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
    94
fun triv_axioms_L th =
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    95
    kill_flex_triv_prems 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    96
       ([transL, nonempty, upair_ax, Union_ax, power_ax, replacement, L_nat] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    97
        MRS (inst "M" "L" th));
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
    98
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
    99
bind_thm ("ball_abs", triv_axioms_L (thm "M_triv_axioms.ball_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   100
bind_thm ("rall_abs", triv_axioms_L (thm "M_triv_axioms.rall_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   101
bind_thm ("bex_abs", triv_axioms_L (thm "M_triv_axioms.bex_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   102
bind_thm ("rex_abs", triv_axioms_L (thm "M_triv_axioms.rex_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   103
bind_thm ("ball_iff_equiv", triv_axioms_L (thm "M_triv_axioms.ball_iff_equiv"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   104
bind_thm ("M_equalityI", triv_axioms_L (thm "M_triv_axioms.M_equalityI"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   105
bind_thm ("empty_abs", triv_axioms_L (thm "M_triv_axioms.empty_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   106
bind_thm ("subset_abs", triv_axioms_L (thm "M_triv_axioms.subset_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   107
bind_thm ("upair_abs", triv_axioms_L (thm "M_triv_axioms.upair_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   108
bind_thm ("upair_in_M_iff", triv_axioms_L (thm "M_triv_axioms.upair_in_M_iff"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   109
bind_thm ("singleton_in_M_iff", triv_axioms_L (thm "M_triv_axioms.singleton_in_M_iff"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   110
bind_thm ("pair_abs", triv_axioms_L (thm "M_triv_axioms.pair_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   111
bind_thm ("pair_in_M_iff", triv_axioms_L (thm "M_triv_axioms.pair_in_M_iff"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   112
bind_thm ("pair_components_in_M", triv_axioms_L (thm "M_triv_axioms.pair_components_in_M"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   113
bind_thm ("cartprod_abs", triv_axioms_L (thm "M_triv_axioms.cartprod_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   114
bind_thm ("union_abs", triv_axioms_L (thm "M_triv_axioms.union_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   115
bind_thm ("inter_abs", triv_axioms_L (thm "M_triv_axioms.inter_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   116
bind_thm ("setdiff_abs", triv_axioms_L (thm "M_triv_axioms.setdiff_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   117
bind_thm ("Union_abs", triv_axioms_L (thm "M_triv_axioms.Union_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   118
bind_thm ("Union_closed", triv_axioms_L (thm "M_triv_axioms.Union_closed"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   119
bind_thm ("Un_closed", triv_axioms_L (thm "M_triv_axioms.Un_closed"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   120
bind_thm ("cons_closed", triv_axioms_L (thm "M_triv_axioms.cons_closed"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   121
bind_thm ("successor_abs", triv_axioms_L (thm "M_triv_axioms.successor_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   122
bind_thm ("succ_in_M_iff", triv_axioms_L (thm "M_triv_axioms.succ_in_M_iff"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   123
bind_thm ("separation_closed", triv_axioms_L (thm "M_triv_axioms.separation_closed"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   124
bind_thm ("strong_replacementI", triv_axioms_L (thm "M_triv_axioms.strong_replacementI"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   125
bind_thm ("strong_replacement_closed", triv_axioms_L (thm "M_triv_axioms.strong_replacement_closed"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   126
bind_thm ("RepFun_closed", triv_axioms_L (thm "M_triv_axioms.RepFun_closed"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   127
bind_thm ("lam_closed", triv_axioms_L (thm "M_triv_axioms.lam_closed"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   128
bind_thm ("image_abs", triv_axioms_L (thm "M_triv_axioms.image_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   129
bind_thm ("powerset_Pow", triv_axioms_L (thm "M_triv_axioms.powerset_Pow"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   130
bind_thm ("powerset_imp_subset_Pow", triv_axioms_L (thm "M_triv_axioms.powerset_imp_subset_Pow"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   131
bind_thm ("nat_into_M", triv_axioms_L (thm "M_triv_axioms.nat_into_M"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   132
bind_thm ("nat_case_closed", triv_axioms_L (thm "M_triv_axioms.nat_case_closed"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   133
bind_thm ("Inl_in_M_iff", triv_axioms_L (thm "M_triv_axioms.Inl_in_M_iff"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   134
bind_thm ("Inr_in_M_iff", triv_axioms_L (thm "M_triv_axioms.Inr_in_M_iff"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   135
bind_thm ("lt_closed", triv_axioms_L (thm "M_triv_axioms.lt_closed"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   136
bind_thm ("transitive_set_abs", triv_axioms_L (thm "M_triv_axioms.transitive_set_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   137
bind_thm ("ordinal_abs", triv_axioms_L (thm "M_triv_axioms.ordinal_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   138
bind_thm ("limit_ordinal_abs", triv_axioms_L (thm "M_triv_axioms.limit_ordinal_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   139
bind_thm ("successor_ordinal_abs", triv_axioms_L (thm "M_triv_axioms.successor_ordinal_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   140
bind_thm ("finite_ordinal_abs", triv_axioms_L (thm "M_triv_axioms.finite_ordinal_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   141
bind_thm ("omega_abs", triv_axioms_L (thm "M_triv_axioms.omega_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   142
bind_thm ("number1_abs", triv_axioms_L (thm "M_triv_axioms.number1_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   143
bind_thm ("number1_abs", triv_axioms_L (thm "M_triv_axioms.number1_abs"));
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   144
bind_thm ("number3_abs", triv_axioms_L (thm "M_triv_axioms.number3_abs"));
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   145
*}
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   146
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   147
declare ball_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   148
declare rall_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   149
declare bex_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   150
declare rex_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   151
declare empty_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   152
declare subset_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   153
declare upair_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   154
declare upair_in_M_iff [iff]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   155
declare singleton_in_M_iff [iff]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   156
declare pair_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   157
declare pair_in_M_iff [iff]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   158
declare cartprod_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   159
declare union_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   160
declare inter_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   161
declare setdiff_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   162
declare Union_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   163
declare Union_closed [intro,simp]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   164
declare Un_closed [intro,simp]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   165
declare cons_closed [intro,simp]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   166
declare successor_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   167
declare succ_in_M_iff [iff]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   168
declare separation_closed [intro,simp]
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   169
declare strong_replacementI
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   170
declare strong_replacement_closed [intro,simp]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   171
declare RepFun_closed [intro,simp]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   172
declare lam_closed [intro,simp]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   173
declare image_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   174
declare nat_into_M [intro]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   175
declare Inl_in_M_iff [iff]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   176
declare Inr_in_M_iff [iff]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   177
declare transitive_set_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   178
declare ordinal_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   179
declare limit_ordinal_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   180
declare successor_ordinal_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   181
declare finite_ordinal_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   182
declare omega_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   183
declare number1_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   184
declare number1_abs [simp] 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   185
declare number3_abs [simp]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   186
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   187
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   188
subsection{*Instantiation of the locale @{text reflection}*}
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   189
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   190
text{*instances of locale constants*}
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   191
constdefs
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   192
  L_F0 :: "[i=>o,i] => i"
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   193
    "L_F0(P,y) == \<mu>b. (\<exists>z. L(z) \<and> P(<y,z>)) --> (\<exists>z\<in>Lset(b). P(<y,z>))"
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   194
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   195
  L_FF :: "[i=>o,i] => i"
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   196
    "L_FF(P)   == \<lambda>a. \<Union>y\<in>Lset(a). L_F0(P,y)"
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   197
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   198
  L_ClEx :: "[i=>o,i] => o"
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   199
    "L_ClEx(P) == \<lambda>a. Limit(a) \<and> normalize(L_FF(P),a) = a"
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   200
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   201
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   202
text{*We must use the meta-existential quantifier; otherwise the reflection
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   203
      terms become enormous!*} 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   204
constdefs
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   205
  L_Reflects :: "[i=>o,[i,i]=>o] => prop"      ("(3REFLECTS/ [_,/ _])")
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   206
    "REFLECTS[P,Q] == (??Cl. Closed_Unbounded(Cl) &
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   207
                           (\<forall>a. Cl(a) --> (\<forall>x \<in> Lset(a). P(x) <-> Q(a,x))))"
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   208
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   209
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   210
theorem Triv_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   211
     "REFLECTS[P, \<lambda>a x. P(x)]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   212
apply (simp add: L_Reflects_def) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   213
apply (rule meta_exI) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   214
apply (rule Closed_Unbounded_Ord) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   215
done
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   216
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   217
theorem Not_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   218
     "REFLECTS[P,Q] ==> REFLECTS[\<lambda>x. ~P(x), \<lambda>a x. ~Q(a,x)]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   219
apply (unfold L_Reflects_def) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   220
apply (erule meta_exE) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   221
apply (rule_tac x=Cl in meta_exI, simp) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   222
done
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   223
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   224
theorem And_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   225
     "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   226
      ==> REFLECTS[\<lambda>x. P(x) \<and> P'(x), \<lambda>a x. Q(a,x) \<and> Q'(a,x)]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   227
apply (unfold L_Reflects_def) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   228
apply (elim meta_exE) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   229
apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   230
apply (simp add: Closed_Unbounded_Int, blast) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   231
done
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   232
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   233
theorem Or_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   234
     "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   235
      ==> REFLECTS[\<lambda>x. P(x) \<or> P'(x), \<lambda>a x. Q(a,x) \<or> Q'(a,x)]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   236
apply (unfold L_Reflects_def) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   237
apply (elim meta_exE) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   238
apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   239
apply (simp add: Closed_Unbounded_Int, blast) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   240
done
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   241
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   242
theorem Imp_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   243
     "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   244
      ==> REFLECTS[\<lambda>x. P(x) --> P'(x), \<lambda>a x. Q(a,x) --> Q'(a,x)]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   245
apply (unfold L_Reflects_def) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   246
apply (elim meta_exE) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   247
apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   248
apply (simp add: Closed_Unbounded_Int, blast) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   249
done
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   250
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   251
theorem Iff_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   252
     "[| REFLECTS[P,Q]; REFLECTS[P',Q'] |] 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   253
      ==> REFLECTS[\<lambda>x. P(x) <-> P'(x), \<lambda>a x. Q(a,x) <-> Q'(a,x)]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   254
apply (unfold L_Reflects_def) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   255
apply (elim meta_exE) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   256
apply (rule_tac x="\<lambda>a. Cl(a) \<and> Cla(a)" in meta_exI) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   257
apply (simp add: Closed_Unbounded_Int, blast) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   258
done
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   259
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   260
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   261
theorem Ex_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   262
     "REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   263
      ==> REFLECTS[\<lambda>x. \<exists>z. L(z) \<and> P(x,z), \<lambda>a x. \<exists>z\<in>Lset(a). Q(a,x,z)]"
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   264
apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def) 
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   265
apply (elim meta_exE) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   266
apply (rule meta_exI)
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   267
apply (rule reflection.Ex_reflection [OF Lset_mono_le Lset_cont Pair_in_Lset],
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   268
       assumption+)
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   269
done
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   270
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   271
theorem All_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   272
     "REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   273
      ==> REFLECTS[\<lambda>x. \<forall>z. L(z) --> P(x,z), \<lambda>a x. \<forall>z\<in>Lset(a). Q(a,x,z)]" 
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   274
apply (unfold L_Reflects_def L_ClEx_def L_FF_def L_F0_def L_def) 
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   275
apply (elim meta_exE) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   276
apply (rule meta_exI)
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   277
apply (rule reflection.All_reflection [OF Lset_mono_le Lset_cont Pair_in_Lset],
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   278
       assumption+)
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   279
done
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   280
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   281
theorem Rex_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   282
     "REFLECTS[ \<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   283
      ==> REFLECTS[\<lambda>x. \<exists>z[L]. P(x,z), \<lambda>a x. \<exists>z\<in>Lset(a). Q(a,x,z)]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   284
apply (unfold rex_def) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   285
apply (intro And_reflection Ex_reflection, assumption)
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   286
done
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   287
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   288
theorem Rall_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   289
     "REFLECTS[\<lambda>x. P(fst(x),snd(x)), \<lambda>a x. Q(a,fst(x),snd(x))]
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   290
      ==> REFLECTS[\<lambda>x. \<forall>z[L]. P(x,z), \<lambda>a x. \<forall>z\<in>Lset(a). Q(a,x,z)]" 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   291
apply (unfold rall_def) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   292
apply (intro Imp_reflection All_reflection, assumption)
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   293
done
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   294
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   295
lemmas FOL_reflections = 
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   296
        Triv_reflection Not_reflection And_reflection Or_reflection
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   297
        Imp_reflection Iff_reflection Ex_reflection All_reflection
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   298
        Rex_reflection Rall_reflection
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   299
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   300
lemma ReflectsD:
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   301
     "[|REFLECTS[P,Q]; Ord(i)|] 
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   302
      ==> \<exists>j. i<j & (\<forall>x \<in> Lset(j). P(x) <-> Q(j,x))"
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   303
apply (unfold L_Reflects_def Closed_Unbounded_def) 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   304
apply (elim meta_exE, clarify) 
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   305
apply (blast dest!: UnboundedD) 
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   306
done
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   307
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   308
lemma ReflectsE:
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   309
     "[| REFLECTS[P,Q]; Ord(i);
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   310
         !!j. [|i<j;  \<forall>x \<in> Lset(j). P(x) <-> Q(j,x)|] ==> R |]
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   311
      ==> R"
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   312
apply (drule ReflectsD, assumption, blast) 
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   313
done
13291
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   314
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   315
lemma Collect_mem_eq: "{x\<in>A. x\<in>B} = A \<inter> B";
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   316
by blast
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   317
a73ab154f75c towards proving separation for L
paulson
parents: 13269
diff changeset
   318
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   319
subsection{*Internalized Formulas for some Set-Theoretic Concepts*}
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   320
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   321
lemmas setclass_simps = rall_setclass_is_ball rex_setclass_is_bex
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   322
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   323
subsubsection{*Some numbers to help write de Bruijn indices*}
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   324
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   325
syntax
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   326
    "3" :: i   ("3")
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   327
    "4" :: i   ("4")
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   328
    "5" :: i   ("5")
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   329
    "6" :: i   ("6")
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   330
    "7" :: i   ("7")
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   331
    "8" :: i   ("8")
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   332
    "9" :: i   ("9")
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   333
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   334
translations
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   335
   "3"  == "succ(2)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   336
   "4"  == "succ(3)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   337
   "5"  == "succ(4)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   338
   "6"  == "succ(5)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   339
   "7"  == "succ(6)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   340
   "8"  == "succ(7)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   341
   "9"  == "succ(8)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   342
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   343
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   344
subsubsection{*The Empty Set, Internalized*}
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   345
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   346
constdefs empty_fm :: "i=>i"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   347
    "empty_fm(x) == Forall(Neg(Member(0,succ(x))))"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   348
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   349
lemma empty_type [TC]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   350
     "x \<in> nat ==> empty_fm(x) \<in> formula"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   351
by (simp add: empty_fm_def) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   352
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   353
lemma arity_empty_fm [simp]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   354
     "x \<in> nat ==> arity(empty_fm(x)) = succ(x)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   355
by (simp add: empty_fm_def succ_Un_distrib [symmetric] Un_ac) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   356
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   357
lemma sats_empty_fm [simp]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   358
   "[| x \<in> nat; env \<in> list(A)|]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   359
    ==> sats(A, empty_fm(x), env) <-> empty(**A, nth(x,env))"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   360
by (simp add: empty_fm_def empty_def)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   361
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   362
lemma empty_iff_sats:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   363
      "[| nth(i,env) = x; nth(j,env) = y; 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   364
          i \<in> nat; env \<in> list(A)|]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   365
       ==> empty(**A, x) <-> sats(A, empty_fm(i), env)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   366
by simp
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   367
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   368
theorem empty_reflection:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   369
     "REFLECTS[\<lambda>x. empty(L,f(x)), 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   370
               \<lambda>i x. empty(**Lset(i),f(x))]"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   371
apply (simp only: empty_def setclass_simps)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   372
apply (intro FOL_reflections)  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   373
done
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   374
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   375
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   376
subsubsection{*Unordered Pairs, Internalized*}
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   377
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   378
constdefs upair_fm :: "[i,i,i]=>i"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   379
    "upair_fm(x,y,z) == 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   380
       And(Member(x,z), 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   381
           And(Member(y,z),
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   382
               Forall(Implies(Member(0,succ(z)), 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   383
                              Or(Equal(0,succ(x)), Equal(0,succ(y)))))))"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   384
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   385
lemma upair_type [TC]:
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   386
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> upair_fm(x,y,z) \<in> formula"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   387
by (simp add: upair_fm_def) 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   388
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   389
lemma arity_upair_fm [simp]:
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   390
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   391
      ==> arity(upair_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   392
by (simp add: upair_fm_def succ_Un_distrib [symmetric] Un_ac) 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   393
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   394
lemma sats_upair_fm [simp]:
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   395
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   396
    ==> sats(A, upair_fm(x,y,z), env) <-> 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   397
            upair(**A, nth(x,env), nth(y,env), nth(z,env))"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   398
by (simp add: upair_fm_def upair_def)
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   399
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   400
lemma upair_iff_sats:
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   401
      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   402
          i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   403
       ==> upair(**A, x, y, z) <-> sats(A, upair_fm(i,j,k), env)"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   404
by (simp add: sats_upair_fm)
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   405
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   406
text{*Useful? At least it refers to "real" unordered pairs*}
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   407
lemma sats_upair_fm2 [simp]:
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   408
   "[| x \<in> nat; y \<in> nat; z < length(env); env \<in> list(A); Transset(A)|]
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   409
    ==> sats(A, upair_fm(x,y,z), env) <-> 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   410
        nth(z,env) = {nth(x,env), nth(y,env)}"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   411
apply (frule lt_length_in_nat, assumption)  
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   412
apply (simp add: upair_fm_def Transset_def, auto) 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   413
apply (blast intro: nth_type) 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   414
done
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   415
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   416
theorem upair_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   417
     "REFLECTS[\<lambda>x. upair(L,f(x),g(x),h(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   418
               \<lambda>i x. upair(**Lset(i),f(x),g(x),h(x))]" 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   419
apply (simp add: upair_def)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   420
apply (intro FOL_reflections)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   421
done
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   422
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   423
subsubsection{*Ordered pairs, Internalized*}
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   424
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   425
constdefs pair_fm :: "[i,i,i]=>i"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   426
    "pair_fm(x,y,z) == 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   427
       Exists(And(upair_fm(succ(x),succ(x),0),
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   428
              Exists(And(upair_fm(succ(succ(x)),succ(succ(y)),0),
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   429
                         upair_fm(1,0,succ(succ(z)))))))"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   430
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   431
lemma pair_type [TC]:
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   432
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> pair_fm(x,y,z) \<in> formula"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   433
by (simp add: pair_fm_def) 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   434
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   435
lemma arity_pair_fm [simp]:
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   436
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   437
      ==> arity(pair_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   438
by (simp add: pair_fm_def succ_Un_distrib [symmetric] Un_ac) 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   439
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   440
lemma sats_pair_fm [simp]:
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   441
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   442
    ==> sats(A, pair_fm(x,y,z), env) <-> 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   443
        pair(**A, nth(x,env), nth(y,env), nth(z,env))"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   444
by (simp add: pair_fm_def pair_def)
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   445
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   446
lemma pair_iff_sats:
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   447
      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   448
          i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   449
       ==> pair(**A, x, y, z) <-> sats(A, pair_fm(i,j,k), env)"
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   450
by (simp add: sats_pair_fm)
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   451
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   452
theorem pair_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   453
     "REFLECTS[\<lambda>x. pair(L,f(x),g(x),h(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   454
               \<lambda>i x. pair(**Lset(i),f(x),g(x),h(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   455
apply (simp only: pair_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   456
apply (intro FOL_reflections upair_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   457
done
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   458
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   459
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   460
subsubsection{*Binary Unions, Internalized*}
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   461
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   462
constdefs union_fm :: "[i,i,i]=>i"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   463
    "union_fm(x,y,z) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   464
       Forall(Iff(Member(0,succ(z)),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   465
                  Or(Member(0,succ(x)),Member(0,succ(y)))))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   466
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   467
lemma union_type [TC]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   468
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> union_fm(x,y,z) \<in> formula"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   469
by (simp add: union_fm_def) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   470
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   471
lemma arity_union_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   472
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   473
      ==> arity(union_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   474
by (simp add: union_fm_def succ_Un_distrib [symmetric] Un_ac) 
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   475
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   476
lemma sats_union_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   477
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   478
    ==> sats(A, union_fm(x,y,z), env) <-> 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   479
        union(**A, nth(x,env), nth(y,env), nth(z,env))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   480
by (simp add: union_fm_def union_def)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   481
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   482
lemma union_iff_sats:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   483
      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   484
          i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   485
       ==> union(**A, x, y, z) <-> sats(A, union_fm(i,j,k), env)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   486
by (simp add: sats_union_fm)
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   487
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   488
theorem union_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   489
     "REFLECTS[\<lambda>x. union(L,f(x),g(x),h(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   490
               \<lambda>i x. union(**Lset(i),f(x),g(x),h(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   491
apply (simp only: union_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   492
apply (intro FOL_reflections)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   493
done
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   494
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   495
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   496
subsubsection{*Set ``Cons,'' Internalized*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   497
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   498
constdefs cons_fm :: "[i,i,i]=>i"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   499
    "cons_fm(x,y,z) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   500
       Exists(And(upair_fm(succ(x),succ(x),0),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   501
                  union_fm(0,succ(y),succ(z))))"
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   502
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   503
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   504
lemma cons_type [TC]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   505
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> cons_fm(x,y,z) \<in> formula"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   506
by (simp add: cons_fm_def) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   507
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   508
lemma arity_cons_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   509
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   510
      ==> arity(cons_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   511
by (simp add: cons_fm_def succ_Un_distrib [symmetric] Un_ac) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   512
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   513
lemma sats_cons_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   514
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   515
    ==> sats(A, cons_fm(x,y,z), env) <-> 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   516
        is_cons(**A, nth(x,env), nth(y,env), nth(z,env))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   517
by (simp add: cons_fm_def is_cons_def)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   518
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   519
lemma cons_iff_sats:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   520
      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   521
          i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   522
       ==> is_cons(**A, x, y, z) <-> sats(A, cons_fm(i,j,k), env)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   523
by simp
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   524
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   525
theorem cons_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   526
     "REFLECTS[\<lambda>x. is_cons(L,f(x),g(x),h(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   527
               \<lambda>i x. is_cons(**Lset(i),f(x),g(x),h(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   528
apply (simp only: is_cons_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   529
apply (intro FOL_reflections upair_reflection union_reflection)  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   530
done
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   531
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   532
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   533
subsubsection{*Successor Function, Internalized*}
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   534
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   535
constdefs succ_fm :: "[i,i]=>i"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   536
    "succ_fm(x,y) == cons_fm(x,x,y)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   537
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   538
lemma succ_type [TC]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   539
     "[| x \<in> nat; y \<in> nat |] ==> succ_fm(x,y) \<in> formula"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   540
by (simp add: succ_fm_def) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   541
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   542
lemma arity_succ_fm [simp]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   543
     "[| x \<in> nat; y \<in> nat |] 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   544
      ==> arity(succ_fm(x,y)) = succ(x) \<union> succ(y)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   545
by (simp add: succ_fm_def)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   546
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   547
lemma sats_succ_fm [simp]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   548
   "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   549
    ==> sats(A, succ_fm(x,y), env) <-> 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   550
        successor(**A, nth(x,env), nth(y,env))"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   551
by (simp add: succ_fm_def successor_def)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   552
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   553
lemma successor_iff_sats:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   554
      "[| nth(i,env) = x; nth(j,env) = y; 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   555
          i \<in> nat; j \<in> nat; env \<in> list(A)|]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   556
       ==> successor(**A, x, y) <-> sats(A, succ_fm(i,j), env)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   557
by simp
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   558
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   559
theorem successor_reflection:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   560
     "REFLECTS[\<lambda>x. successor(L,f(x),g(x)), 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   561
               \<lambda>i x. successor(**Lset(i),f(x),g(x))]"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   562
apply (simp only: successor_def setclass_simps)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   563
apply (intro cons_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   564
done
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   565
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   566
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   567
subsubsection{*The Number 1, Internalized*}
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   568
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   569
(* "number1(M,a) == (\<exists>x[M]. empty(M,x) & successor(M,x,a))" *)
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   570
constdefs number1_fm :: "i=>i"
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   571
    "number1_fm(a) == Exists(And(empty_fm(0), succ_fm(0,succ(a))))"
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   572
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   573
lemma number1_type [TC]:
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   574
     "x \<in> nat ==> number1_fm(x) \<in> formula"
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   575
by (simp add: number1_fm_def) 
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   576
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   577
lemma arity_number1_fm [simp]:
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   578
     "x \<in> nat ==> arity(number1_fm(x)) = succ(x)"
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   579
by (simp add: number1_fm_def succ_Un_distrib [symmetric] Un_ac) 
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   580
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   581
lemma sats_number1_fm [simp]:
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   582
   "[| x \<in> nat; env \<in> list(A)|]
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   583
    ==> sats(A, number1_fm(x), env) <-> number1(**A, nth(x,env))"
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   584
by (simp add: number1_fm_def number1_def)
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   585
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   586
lemma number1_iff_sats:
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   587
      "[| nth(i,env) = x; nth(j,env) = y; 
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   588
          i \<in> nat; env \<in> list(A)|]
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   589
       ==> number1(**A, x) <-> sats(A, number1_fm(i), env)"
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   590
by simp
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   591
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   592
theorem number1_reflection:
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   593
     "REFLECTS[\<lambda>x. number1(L,f(x)), 
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   594
               \<lambda>i x. number1(**Lset(i),f(x))]"
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   595
apply (simp only: number1_def setclass_simps)
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   596
apply (intro FOL_reflections empty_reflection successor_reflection)
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   597
done
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   598
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
   599
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   600
subsubsection{*Big Union, Internalized*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   601
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   602
(*  "big_union(M,A,z) == \<forall>x[M]. x \<in> z <-> (\<exists>y[M]. y\<in>A & x \<in> y)" *)
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   603
constdefs big_union_fm :: "[i,i]=>i"
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   604
    "big_union_fm(A,z) == 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   605
       Forall(Iff(Member(0,succ(z)),
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   606
                  Exists(And(Member(0,succ(succ(A))), Member(1,0)))))"
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   607
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   608
lemma big_union_type [TC]:
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   609
     "[| x \<in> nat; y \<in> nat |] ==> big_union_fm(x,y) \<in> formula"
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   610
by (simp add: big_union_fm_def) 
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   611
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   612
lemma arity_big_union_fm [simp]:
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   613
     "[| x \<in> nat; y \<in> nat |] 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   614
      ==> arity(big_union_fm(x,y)) = succ(x) \<union> succ(y)"
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   615
by (simp add: big_union_fm_def succ_Un_distrib [symmetric] Un_ac)
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   616
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   617
lemma sats_big_union_fm [simp]:
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   618
   "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   619
    ==> sats(A, big_union_fm(x,y), env) <-> 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   620
        big_union(**A, nth(x,env), nth(y,env))"
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   621
by (simp add: big_union_fm_def big_union_def)
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   622
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   623
lemma big_union_iff_sats:
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   624
      "[| nth(i,env) = x; nth(j,env) = y; 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   625
          i \<in> nat; j \<in> nat; env \<in> list(A)|]
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   626
       ==> big_union(**A, x, y) <-> sats(A, big_union_fm(i,j), env)"
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   627
by simp
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   628
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   629
theorem big_union_reflection:
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   630
     "REFLECTS[\<lambda>x. big_union(L,f(x),g(x)), 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   631
               \<lambda>i x. big_union(**Lset(i),f(x),g(x))]"
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   632
apply (simp only: big_union_def setclass_simps)
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   633
apply (intro FOL_reflections)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   634
done
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   635
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   636
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   637
subsubsection{*Variants of Satisfaction Definitions for Ordinals, etc.*}
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   638
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   639
text{*Differs from the one in Formula by using "ordinal" rather than "Ord"*}
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   640
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   641
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   642
lemma sats_subset_fm':
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   643
   "[|x \<in> nat; y \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   644
    ==> sats(A, subset_fm(x,y), env) <-> subset(**A, nth(x,env), nth(y,env))" 
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   645
by (simp add: subset_fm_def Relative.subset_def) 
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   646
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   647
theorem subset_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   648
     "REFLECTS[\<lambda>x. subset(L,f(x),g(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   649
               \<lambda>i x. subset(**Lset(i),f(x),g(x))]" 
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   650
apply (simp only: Relative.subset_def setclass_simps)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   651
apply (intro FOL_reflections)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   652
done
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   653
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   654
lemma sats_transset_fm':
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   655
   "[|x \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   656
    ==> sats(A, transset_fm(x), env) <-> transitive_set(**A, nth(x,env))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   657
by (simp add: sats_subset_fm' transset_fm_def transitive_set_def) 
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   658
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   659
theorem transitive_set_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   660
     "REFLECTS[\<lambda>x. transitive_set(L,f(x)),
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   661
               \<lambda>i x. transitive_set(**Lset(i),f(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   662
apply (simp only: transitive_set_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   663
apply (intro FOL_reflections subset_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   664
done
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   665
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   666
lemma sats_ordinal_fm':
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   667
   "[|x \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   668
    ==> sats(A, ordinal_fm(x), env) <-> ordinal(**A,nth(x,env))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   669
by (simp add: sats_transset_fm' ordinal_fm_def ordinal_def)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   670
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   671
lemma ordinal_iff_sats:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   672
      "[| nth(i,env) = x;  i \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   673
       ==> ordinal(**A, x) <-> sats(A, ordinal_fm(i), env)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   674
by (simp add: sats_ordinal_fm')
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   675
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   676
theorem ordinal_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   677
     "REFLECTS[\<lambda>x. ordinal(L,f(x)), \<lambda>i x. ordinal(**Lset(i),f(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   678
apply (simp only: ordinal_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   679
apply (intro FOL_reflections transitive_set_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   680
done
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   681
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   682
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   683
subsubsection{*Membership Relation, Internalized*}
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   684
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   685
constdefs Memrel_fm :: "[i,i]=>i"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   686
    "Memrel_fm(A,r) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   687
       Forall(Iff(Member(0,succ(r)),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   688
                  Exists(And(Member(0,succ(succ(A))),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   689
                             Exists(And(Member(0,succ(succ(succ(A)))),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   690
                                        And(Member(1,0),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   691
                                            pair_fm(1,0,2))))))))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   692
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   693
lemma Memrel_type [TC]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   694
     "[| x \<in> nat; y \<in> nat |] ==> Memrel_fm(x,y) \<in> formula"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   695
by (simp add: Memrel_fm_def) 
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   696
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   697
lemma arity_Memrel_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   698
     "[| x \<in> nat; y \<in> nat |] 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   699
      ==> arity(Memrel_fm(x,y)) = succ(x) \<union> succ(y)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   700
by (simp add: Memrel_fm_def succ_Un_distrib [symmetric] Un_ac) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   701
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   702
lemma sats_Memrel_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   703
   "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   704
    ==> sats(A, Memrel_fm(x,y), env) <-> 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   705
        membership(**A, nth(x,env), nth(y,env))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   706
by (simp add: Memrel_fm_def membership_def)
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   707
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   708
lemma Memrel_iff_sats:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   709
      "[| nth(i,env) = x; nth(j,env) = y; 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   710
          i \<in> nat; j \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   711
       ==> membership(**A, x, y) <-> sats(A, Memrel_fm(i,j), env)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   712
by simp
13304
43ef6c6dd906 more separation instances
paulson
parents: 13299
diff changeset
   713
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   714
theorem membership_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   715
     "REFLECTS[\<lambda>x. membership(L,f(x),g(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   716
               \<lambda>i x. membership(**Lset(i),f(x),g(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   717
apply (simp only: membership_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   718
apply (intro FOL_reflections pair_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   719
done
13304
43ef6c6dd906 more separation instances
paulson
parents: 13299
diff changeset
   720
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   721
subsubsection{*Predecessor Set, Internalized*}
13304
43ef6c6dd906 more separation instances
paulson
parents: 13299
diff changeset
   722
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   723
constdefs pred_set_fm :: "[i,i,i,i]=>i"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   724
    "pred_set_fm(A,x,r,B) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   725
       Forall(Iff(Member(0,succ(B)),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   726
                  Exists(And(Member(0,succ(succ(r))),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   727
                             And(Member(1,succ(succ(A))),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   728
                                 pair_fm(1,succ(succ(x)),0))))))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   729
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   730
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   731
lemma pred_set_type [TC]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   732
     "[| A \<in> nat; x \<in> nat; r \<in> nat; B \<in> nat |] 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   733
      ==> pred_set_fm(A,x,r,B) \<in> formula"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   734
by (simp add: pred_set_fm_def) 
13304
43ef6c6dd906 more separation instances
paulson
parents: 13299
diff changeset
   735
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   736
lemma arity_pred_set_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   737
   "[| A \<in> nat; x \<in> nat; r \<in> nat; B \<in> nat |] 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   738
    ==> arity(pred_set_fm(A,x,r,B)) = succ(A) \<union> succ(x) \<union> succ(r) \<union> succ(B)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   739
by (simp add: pred_set_fm_def succ_Un_distrib [symmetric] Un_ac) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   740
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   741
lemma sats_pred_set_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   742
   "[| U \<in> nat; x \<in> nat; r \<in> nat; B \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   743
    ==> sats(A, pred_set_fm(U,x,r,B), env) <-> 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   744
        pred_set(**A, nth(U,env), nth(x,env), nth(r,env), nth(B,env))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   745
by (simp add: pred_set_fm_def pred_set_def)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   746
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   747
lemma pred_set_iff_sats:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   748
      "[| nth(i,env) = U; nth(j,env) = x; nth(k,env) = r; nth(l,env) = B; 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   749
          i \<in> nat; j \<in> nat; k \<in> nat; l \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   750
       ==> pred_set(**A,U,x,r,B) <-> sats(A, pred_set_fm(i,j,k,l), env)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   751
by (simp add: sats_pred_set_fm)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   752
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   753
theorem pred_set_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   754
     "REFLECTS[\<lambda>x. pred_set(L,f(x),g(x),h(x),b(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   755
               \<lambda>i x. pred_set(**Lset(i),f(x),g(x),h(x),b(x))]" 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   756
apply (simp only: pred_set_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   757
apply (intro FOL_reflections pair_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   758
done
13304
43ef6c6dd906 more separation instances
paulson
parents: 13299
diff changeset
   759
43ef6c6dd906 more separation instances
paulson
parents: 13299
diff changeset
   760
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
   761
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   762
subsubsection{*Domain of a Relation, Internalized*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   763
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   764
(* "is_domain(M,r,z) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   765
	\<forall>x[M]. (x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w))))" *)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   766
constdefs domain_fm :: "[i,i]=>i"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   767
    "domain_fm(r,z) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   768
       Forall(Iff(Member(0,succ(z)),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   769
                  Exists(And(Member(0,succ(succ(r))),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   770
                             Exists(pair_fm(2,0,1))))))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   771
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   772
lemma domain_type [TC]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   773
     "[| x \<in> nat; y \<in> nat |] ==> domain_fm(x,y) \<in> formula"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   774
by (simp add: domain_fm_def) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   775
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   776
lemma arity_domain_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   777
     "[| x \<in> nat; y \<in> nat |] 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   778
      ==> arity(domain_fm(x,y)) = succ(x) \<union> succ(y)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   779
by (simp add: domain_fm_def succ_Un_distrib [symmetric] Un_ac) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   780
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   781
lemma sats_domain_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   782
   "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   783
    ==> sats(A, domain_fm(x,y), env) <-> 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   784
        is_domain(**A, nth(x,env), nth(y,env))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   785
by (simp add: domain_fm_def is_domain_def)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   786
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   787
lemma domain_iff_sats:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   788
      "[| nth(i,env) = x; nth(j,env) = y; 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   789
          i \<in> nat; j \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   790
       ==> is_domain(**A, x, y) <-> sats(A, domain_fm(i,j), env)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   791
by simp
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   792
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   793
theorem domain_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   794
     "REFLECTS[\<lambda>x. is_domain(L,f(x),g(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   795
               \<lambda>i x. is_domain(**Lset(i),f(x),g(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   796
apply (simp only: is_domain_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   797
apply (intro FOL_reflections pair_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   798
done
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   799
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   800
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   801
subsubsection{*Range of a Relation, Internalized*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   802
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   803
(* "is_range(M,r,z) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   804
	\<forall>y[M]. (y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w))))" *)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   805
constdefs range_fm :: "[i,i]=>i"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   806
    "range_fm(r,z) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   807
       Forall(Iff(Member(0,succ(z)),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   808
                  Exists(And(Member(0,succ(succ(r))),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   809
                             Exists(pair_fm(0,2,1))))))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   810
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   811
lemma range_type [TC]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   812
     "[| x \<in> nat; y \<in> nat |] ==> range_fm(x,y) \<in> formula"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   813
by (simp add: range_fm_def) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   814
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   815
lemma arity_range_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   816
     "[| x \<in> nat; y \<in> nat |] 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   817
      ==> arity(range_fm(x,y)) = succ(x) \<union> succ(y)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   818
by (simp add: range_fm_def succ_Un_distrib [symmetric] Un_ac) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   819
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   820
lemma sats_range_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   821
   "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   822
    ==> sats(A, range_fm(x,y), env) <-> 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   823
        is_range(**A, nth(x,env), nth(y,env))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   824
by (simp add: range_fm_def is_range_def)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   825
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   826
lemma range_iff_sats:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   827
      "[| nth(i,env) = x; nth(j,env) = y; 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   828
          i \<in> nat; j \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   829
       ==> is_range(**A, x, y) <-> sats(A, range_fm(i,j), env)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   830
by simp
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   831
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   832
theorem range_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   833
     "REFLECTS[\<lambda>x. is_range(L,f(x),g(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   834
               \<lambda>i x. is_range(**Lset(i),f(x),g(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   835
apply (simp only: is_range_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   836
apply (intro FOL_reflections pair_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   837
done
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   838
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   839
 
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   840
subsubsection{*Field of a Relation, Internalized*}
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   841
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   842
(* "is_field(M,r,z) == 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   843
	\<exists>dr[M]. is_domain(M,r,dr) & 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   844
            (\<exists>rr[M]. is_range(M,r,rr) & union(M,dr,rr,z))" *)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   845
constdefs field_fm :: "[i,i]=>i"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   846
    "field_fm(r,z) == 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   847
       Exists(And(domain_fm(succ(r),0), 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   848
              Exists(And(range_fm(succ(succ(r)),0), 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   849
                         union_fm(1,0,succ(succ(z)))))))"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   850
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   851
lemma field_type [TC]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   852
     "[| x \<in> nat; y \<in> nat |] ==> field_fm(x,y) \<in> formula"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   853
by (simp add: field_fm_def) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   854
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   855
lemma arity_field_fm [simp]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   856
     "[| x \<in> nat; y \<in> nat |] 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   857
      ==> arity(field_fm(x,y)) = succ(x) \<union> succ(y)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   858
by (simp add: field_fm_def succ_Un_distrib [symmetric] Un_ac) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   859
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   860
lemma sats_field_fm [simp]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   861
   "[| x \<in> nat; y \<in> nat; env \<in> list(A)|]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   862
    ==> sats(A, field_fm(x,y), env) <-> 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   863
        is_field(**A, nth(x,env), nth(y,env))"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   864
by (simp add: field_fm_def is_field_def)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   865
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   866
lemma field_iff_sats:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   867
      "[| nth(i,env) = x; nth(j,env) = y; 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   868
          i \<in> nat; j \<in> nat; env \<in> list(A)|]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   869
       ==> is_field(**A, x, y) <-> sats(A, field_fm(i,j), env)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   870
by simp
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   871
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   872
theorem field_reflection:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   873
     "REFLECTS[\<lambda>x. is_field(L,f(x),g(x)), 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   874
               \<lambda>i x. is_field(**Lset(i),f(x),g(x))]"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   875
apply (simp only: is_field_def setclass_simps)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   876
apply (intro FOL_reflections domain_reflection range_reflection
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   877
             union_reflection)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   878
done
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   879
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   880
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
   881
subsubsection{*Image under a Relation, Internalized*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   882
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   883
(* "image(M,r,A,z) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   884
        \<forall>y[M]. (y \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w))))" *)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   885
constdefs image_fm :: "[i,i,i]=>i"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   886
    "image_fm(r,A,z) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   887
       Forall(Iff(Member(0,succ(z)),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   888
                  Exists(And(Member(0,succ(succ(r))),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   889
                             Exists(And(Member(0,succ(succ(succ(A)))),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   890
	 			        pair_fm(0,2,1)))))))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   891
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   892
lemma image_type [TC]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   893
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> image_fm(x,y,z) \<in> formula"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   894
by (simp add: image_fm_def) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   895
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   896
lemma arity_image_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   897
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   898
      ==> arity(image_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   899
by (simp add: image_fm_def succ_Un_distrib [symmetric] Un_ac) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   900
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   901
lemma sats_image_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   902
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   903
    ==> sats(A, image_fm(x,y,z), env) <-> 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   904
        image(**A, nth(x,env), nth(y,env), nth(z,env))"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   905
by (simp add: image_fm_def Relative.image_def)
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   906
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   907
lemma image_iff_sats:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   908
      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   909
          i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   910
       ==> image(**A, x, y, z) <-> sats(A, image_fm(i,j,k), env)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   911
by (simp add: sats_image_fm)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   912
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   913
theorem image_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   914
     "REFLECTS[\<lambda>x. image(L,f(x),g(x),h(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   915
               \<lambda>i x. image(**Lset(i),f(x),g(x),h(x))]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   916
apply (simp only: Relative.image_def setclass_simps)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
   917
apply (intro FOL_reflections pair_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
   918
done
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   919
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
   920
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   921
subsubsection{*Pre-Image under a Relation, Internalized*}
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   922
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   923
(* "pre_image(M,r,A,z) == 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   924
	\<forall>x[M]. x \<in> z <-> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))" *)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   925
constdefs pre_image_fm :: "[i,i,i]=>i"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   926
    "pre_image_fm(r,A,z) == 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   927
       Forall(Iff(Member(0,succ(z)),
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   928
                  Exists(And(Member(0,succ(succ(r))),
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   929
                             Exists(And(Member(0,succ(succ(succ(A)))),
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   930
	 			        pair_fm(2,0,1)))))))"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   931
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   932
lemma pre_image_type [TC]:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   933
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> pre_image_fm(x,y,z) \<in> formula"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   934
by (simp add: pre_image_fm_def) 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   935
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   936
lemma arity_pre_image_fm [simp]:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   937
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   938
      ==> arity(pre_image_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   939
by (simp add: pre_image_fm_def succ_Un_distrib [symmetric] Un_ac) 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   940
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   941
lemma sats_pre_image_fm [simp]:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   942
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   943
    ==> sats(A, pre_image_fm(x,y,z), env) <-> 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   944
        pre_image(**A, nth(x,env), nth(y,env), nth(z,env))"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   945
by (simp add: pre_image_fm_def Relative.pre_image_def)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   946
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   947
lemma pre_image_iff_sats:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   948
      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   949
          i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   950
       ==> pre_image(**A, x, y, z) <-> sats(A, pre_image_fm(i,j,k), env)"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   951
by (simp add: sats_pre_image_fm)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   952
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   953
theorem pre_image_reflection:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   954
     "REFLECTS[\<lambda>x. pre_image(L,f(x),g(x),h(x)), 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   955
               \<lambda>i x. pre_image(**Lset(i),f(x),g(x),h(x))]"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   956
apply (simp only: Relative.pre_image_def setclass_simps)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   957
apply (intro FOL_reflections pair_reflection)  
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   958
done
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   959
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
   960
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   961
subsubsection{*Function Application, Internalized*}
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   962
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   963
(* "fun_apply(M,f,x,y) == 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   964
        (\<exists>xs[M]. \<exists>fxs[M]. 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   965
         upair(M,x,x,xs) & image(M,f,xs,fxs) & big_union(M,fxs,y))" *)
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   966
constdefs fun_apply_fm :: "[i,i,i]=>i"
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   967
    "fun_apply_fm(f,x,y) == 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   968
       Exists(Exists(And(upair_fm(succ(succ(x)), succ(succ(x)), 1),
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   969
                         And(image_fm(succ(succ(f)), 1, 0), 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   970
                             big_union_fm(0,succ(succ(y)))))))"
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   971
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   972
lemma fun_apply_type [TC]:
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   973
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> fun_apply_fm(x,y,z) \<in> formula"
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   974
by (simp add: fun_apply_fm_def) 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   975
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   976
lemma arity_fun_apply_fm [simp]:
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   977
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   978
      ==> arity(fun_apply_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   979
by (simp add: fun_apply_fm_def succ_Un_distrib [symmetric] Un_ac) 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   980
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   981
lemma sats_fun_apply_fm [simp]:
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   982
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   983
    ==> sats(A, fun_apply_fm(x,y,z), env) <-> 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   984
        fun_apply(**A, nth(x,env), nth(y,env), nth(z,env))"
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   985
by (simp add: fun_apply_fm_def fun_apply_def)
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   986
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   987
lemma fun_apply_iff_sats:
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   988
      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   989
          i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   990
       ==> fun_apply(**A, x, y, z) <-> sats(A, fun_apply_fm(i,j,k), env)"
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   991
by simp
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   992
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   993
theorem fun_apply_reflection:
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   994
     "REFLECTS[\<lambda>x. fun_apply(L,f(x),g(x),h(x)), 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   995
               \<lambda>i x. fun_apply(**Lset(i),f(x),g(x),h(x))]" 
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   996
apply (simp only: fun_apply_def setclass_simps)
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   997
apply (intro FOL_reflections upair_reflection image_reflection
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   998
             big_union_reflection)  
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
   999
done
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
  1000
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
  1001
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
  1002
subsubsection{*The Concept of Relation, Internalized*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1003
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1004
(* "is_relation(M,r) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1005
        (\<forall>z[M]. z\<in>r --> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))" *)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1006
constdefs relation_fm :: "i=>i"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1007
    "relation_fm(r) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1008
       Forall(Implies(Member(0,succ(r)), Exists(Exists(pair_fm(1,0,2)))))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1009
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1010
lemma relation_type [TC]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1011
     "[| x \<in> nat |] ==> relation_fm(x) \<in> formula"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1012
by (simp add: relation_fm_def) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1013
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1014
lemma arity_relation_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1015
     "x \<in> nat ==> arity(relation_fm(x)) = succ(x)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1016
by (simp add: relation_fm_def succ_Un_distrib [symmetric] Un_ac) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1017
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1018
lemma sats_relation_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1019
   "[| x \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1020
    ==> sats(A, relation_fm(x), env) <-> is_relation(**A, nth(x,env))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1021
by (simp add: relation_fm_def is_relation_def)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1022
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1023
lemma relation_iff_sats:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1024
      "[| nth(i,env) = x; nth(j,env) = y; 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1025
          i \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1026
       ==> is_relation(**A, x) <-> sats(A, relation_fm(i), env)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1027
by simp
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1028
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1029
theorem is_relation_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1030
     "REFLECTS[\<lambda>x. is_relation(L,f(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1031
               \<lambda>i x. is_relation(**Lset(i),f(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1032
apply (simp only: is_relation_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1033
apply (intro FOL_reflections pair_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1034
done
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1035
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1036
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
  1037
subsubsection{*The Concept of Function, Internalized*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1038
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1039
(* "is_function(M,r) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1040
	\<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M]. 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1041
           pair(M,x,y,p) --> pair(M,x,y',p') --> p\<in>r --> p'\<in>r --> y=y'" *)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1042
constdefs function_fm :: "i=>i"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1043
    "function_fm(r) == 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1044
       Forall(Forall(Forall(Forall(Forall(
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1045
         Implies(pair_fm(4,3,1),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1046
                 Implies(pair_fm(4,2,0),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1047
                         Implies(Member(1,r#+5),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1048
                                 Implies(Member(0,r#+5), Equal(3,2))))))))))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1049
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1050
lemma function_type [TC]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1051
     "[| x \<in> nat |] ==> function_fm(x) \<in> formula"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1052
by (simp add: function_fm_def) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1053
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1054
lemma arity_function_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1055
     "x \<in> nat ==> arity(function_fm(x)) = succ(x)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1056
by (simp add: function_fm_def succ_Un_distrib [symmetric] Un_ac) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1057
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1058
lemma sats_function_fm [simp]:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1059
   "[| x \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1060
    ==> sats(A, function_fm(x), env) <-> is_function(**A, nth(x,env))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1061
by (simp add: function_fm_def is_function_def)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1062
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1063
lemma function_iff_sats:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1064
      "[| nth(i,env) = x; nth(j,env) = y; 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1065
          i \<in> nat; env \<in> list(A)|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1066
       ==> is_function(**A, x) <-> sats(A, function_fm(i), env)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1067
by simp
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13304
diff changeset
  1068
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1069
theorem is_function_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1070
     "REFLECTS[\<lambda>x. is_function(L,f(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1071
               \<lambda>i x. is_function(**Lset(i),f(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1072
apply (simp only: is_function_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1073
apply (intro FOL_reflections pair_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1074
done
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
  1075
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13291
diff changeset
  1076
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
  1077
subsubsection{*Typed Functions, Internalized*}
13309
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1078
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1079
(* "typed_function(M,A,B,r) == 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1080
        is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) &
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1081
        (\<forall>u[M]. u\<in>r --> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) --> y\<in>B))" *)
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1082
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1083
constdefs typed_function_fm :: "[i,i,i]=>i"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1084
    "typed_function_fm(A,B,r) == 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1085
       And(function_fm(r),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1086
         And(relation_fm(r),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1087
           And(domain_fm(r,A),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1088
             Forall(Implies(Member(0,succ(r)),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1089
                  Forall(Forall(Implies(pair_fm(1,0,2),Member(0,B#+3)))))))))"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1090
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1091
lemma typed_function_type [TC]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1092
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> typed_function_fm(x,y,z) \<in> formula"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1093
by (simp add: typed_function_fm_def) 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1094
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1095
lemma arity_typed_function_fm [simp]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1096
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1097
      ==> arity(typed_function_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1098
by (simp add: typed_function_fm_def succ_Un_distrib [symmetric] Un_ac) 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1099
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1100
lemma sats_typed_function_fm [simp]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1101
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1102
    ==> sats(A, typed_function_fm(x,y,z), env) <-> 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1103
        typed_function(**A, nth(x,env), nth(y,env), nth(z,env))"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1104
by (simp add: typed_function_fm_def typed_function_def)
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1105
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1106
lemma typed_function_iff_sats:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1107
  "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1108
      i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1109
   ==> typed_function(**A, x, y, z) <-> sats(A, typed_function_fm(i,j,k), env)"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1110
by simp
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1111
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1112
lemmas function_reflections = 
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
  1113
        empty_reflection number1_reflection
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
  1114
	upair_reflection pair_reflection union_reflection
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13348
diff changeset
  1115
	big_union_reflection cons_reflection successor_reflection 
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1116
        fun_apply_reflection subset_reflection
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1117
	transitive_set_reflection membership_reflection
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1118
	pred_set_reflection domain_reflection range_reflection field_reflection
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1119
        image_reflection pre_image_reflection
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1120
	is_relation_reflection is_function_reflection
13309
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1121
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1122
lemmas function_iff_sats = 
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
  1123
        empty_iff_sats number1_iff_sats 
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13352
diff changeset
  1124
	upair_iff_sats pair_iff_sats union_iff_sats
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1125
	cons_iff_sats successor_iff_sats
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1126
        fun_apply_iff_sats  Memrel_iff_sats
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1127
	pred_set_iff_sats domain_iff_sats range_iff_sats field_iff_sats
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1128
        image_iff_sats pre_image_iff_sats 
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1129
	relation_iff_sats function_iff_sats
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1130
13309
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1131
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1132
theorem typed_function_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1133
     "REFLECTS[\<lambda>x. typed_function(L,f(x),g(x),h(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1134
               \<lambda>i x. typed_function(**Lset(i),f(x),g(x),h(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1135
apply (simp only: typed_function_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1136
apply (intro FOL_reflections function_reflections)  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1137
done
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1138
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1139
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
  1140
subsubsection{*Composition of Relations, Internalized*}
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1141
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1142
(* "composition(M,r,s,t) == 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1143
        \<forall>p[M]. p \<in> t <-> 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1144
               (\<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M]. 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1145
                pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) & 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1146
                xy \<in> s & yz \<in> r)" *)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1147
constdefs composition_fm :: "[i,i,i]=>i"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1148
  "composition_fm(r,s,t) == 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1149
     Forall(Iff(Member(0,succ(t)),
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1150
             Exists(Exists(Exists(Exists(Exists( 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1151
              And(pair_fm(4,2,5),
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1152
               And(pair_fm(4,3,1),
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1153
                And(pair_fm(3,2,0),
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1154
                 And(Member(1,s#+6), Member(0,r#+6))))))))))))"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1155
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1156
lemma composition_type [TC]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1157
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> composition_fm(x,y,z) \<in> formula"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1158
by (simp add: composition_fm_def) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1159
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1160
lemma arity_composition_fm [simp]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1161
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1162
      ==> arity(composition_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1163
by (simp add: composition_fm_def succ_Un_distrib [symmetric] Un_ac) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1164
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1165
lemma sats_composition_fm [simp]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1166
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1167
    ==> sats(A, composition_fm(x,y,z), env) <-> 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1168
        composition(**A, nth(x,env), nth(y,env), nth(z,env))"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1169
by (simp add: composition_fm_def composition_def)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1170
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1171
lemma composition_iff_sats:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1172
      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1173
          i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1174
       ==> composition(**A, x, y, z) <-> sats(A, composition_fm(i,j,k), env)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1175
by simp
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1176
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1177
theorem composition_reflection:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1178
     "REFLECTS[\<lambda>x. composition(L,f(x),g(x),h(x)), 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1179
               \<lambda>i x. composition(**Lset(i),f(x),g(x),h(x))]"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1180
apply (simp only: composition_def setclass_simps)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1181
apply (intro FOL_reflections pair_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1182
done
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1183
13309
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1184
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
  1185
subsubsection{*Injections, Internalized*}
13309
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1186
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1187
(* "injection(M,A,B,f) == 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1188
	typed_function(M,A,B,f) &
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1189
        (\<forall>x[M]. \<forall>x'[M]. \<forall>y[M]. \<forall>p[M]. \<forall>p'[M]. 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1190
          pair(M,x,y,p) --> pair(M,x',y,p') --> p\<in>f --> p'\<in>f --> x=x')" *)
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1191
constdefs injection_fm :: "[i,i,i]=>i"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1192
 "injection_fm(A,B,f) == 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1193
    And(typed_function_fm(A,B,f),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1194
       Forall(Forall(Forall(Forall(Forall(
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1195
         Implies(pair_fm(4,2,1),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1196
                 Implies(pair_fm(3,2,0),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1197
                         Implies(Member(1,f#+5),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1198
                                 Implies(Member(0,f#+5), Equal(4,3)))))))))))"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1199
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1200
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1201
lemma injection_type [TC]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1202
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> injection_fm(x,y,z) \<in> formula"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1203
by (simp add: injection_fm_def) 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1204
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1205
lemma arity_injection_fm [simp]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1206
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1207
      ==> arity(injection_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1208
by (simp add: injection_fm_def succ_Un_distrib [symmetric] Un_ac) 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1209
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1210
lemma sats_injection_fm [simp]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1211
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1212
    ==> sats(A, injection_fm(x,y,z), env) <-> 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1213
        injection(**A, nth(x,env), nth(y,env), nth(z,env))"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1214
by (simp add: injection_fm_def injection_def)
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1215
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1216
lemma injection_iff_sats:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1217
  "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1218
      i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1219
   ==> injection(**A, x, y, z) <-> sats(A, injection_fm(i,j,k), env)"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1220
by simp
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1221
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1222
theorem injection_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1223
     "REFLECTS[\<lambda>x. injection(L,f(x),g(x),h(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1224
               \<lambda>i x. injection(**Lset(i),f(x),g(x),h(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1225
apply (simp only: injection_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1226
apply (intro FOL_reflections function_reflections typed_function_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1227
done
13309
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1228
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1229
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
  1230
subsubsection{*Surjections, Internalized*}
13309
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1231
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1232
(*  surjection :: "[i=>o,i,i,i] => o"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1233
    "surjection(M,A,B,f) == 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1234
        typed_function(M,A,B,f) &
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1235
        (\<forall>y[M]. y\<in>B --> (\<exists>x[M]. x\<in>A & fun_apply(M,f,x,y)))" *)
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1236
constdefs surjection_fm :: "[i,i,i]=>i"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1237
 "surjection_fm(A,B,f) == 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1238
    And(typed_function_fm(A,B,f),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1239
       Forall(Implies(Member(0,succ(B)),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1240
                      Exists(And(Member(0,succ(succ(A))),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1241
                                 fun_apply_fm(succ(succ(f)),0,1))))))"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1242
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1243
lemma surjection_type [TC]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1244
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> surjection_fm(x,y,z) \<in> formula"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1245
by (simp add: surjection_fm_def) 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1246
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1247
lemma arity_surjection_fm [simp]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1248
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1249
      ==> arity(surjection_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1250
by (simp add: surjection_fm_def succ_Un_distrib [symmetric] Un_ac) 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1251
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1252
lemma sats_surjection_fm [simp]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1253
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1254
    ==> sats(A, surjection_fm(x,y,z), env) <-> 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1255
        surjection(**A, nth(x,env), nth(y,env), nth(z,env))"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1256
by (simp add: surjection_fm_def surjection_def)
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1257
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1258
lemma surjection_iff_sats:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1259
  "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1260
      i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1261
   ==> surjection(**A, x, y, z) <-> sats(A, surjection_fm(i,j,k), env)"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1262
by simp
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1263
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1264
theorem surjection_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1265
     "REFLECTS[\<lambda>x. surjection(L,f(x),g(x),h(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1266
               \<lambda>i x. surjection(**Lset(i),f(x),g(x),h(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1267
apply (simp only: surjection_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1268
apply (intro FOL_reflections function_reflections typed_function_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1269
done
13309
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1270
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1271
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1272
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
  1273
subsubsection{*Bijections, Internalized*}
13309
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1274
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1275
(*   bijection :: "[i=>o,i,i,i] => o"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1276
    "bijection(M,A,B,f) == injection(M,A,B,f) & surjection(M,A,B,f)" *)
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1277
constdefs bijection_fm :: "[i,i,i]=>i"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1278
 "bijection_fm(A,B,f) == And(injection_fm(A,B,f), surjection_fm(A,B,f))"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1279
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1280
lemma bijection_type [TC]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1281
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> bijection_fm(x,y,z) \<in> formula"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1282
by (simp add: bijection_fm_def) 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1283
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1284
lemma arity_bijection_fm [simp]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1285
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1286
      ==> arity(bijection_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1287
by (simp add: bijection_fm_def succ_Un_distrib [symmetric] Un_ac) 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1288
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1289
lemma sats_bijection_fm [simp]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1290
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1291
    ==> sats(A, bijection_fm(x,y,z), env) <-> 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1292
        bijection(**A, nth(x,env), nth(y,env), nth(z,env))"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1293
by (simp add: bijection_fm_def bijection_def)
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1294
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1295
lemma bijection_iff_sats:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1296
  "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1297
      i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1298
   ==> bijection(**A, x, y, z) <-> sats(A, bijection_fm(i,j,k), env)"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1299
by simp
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1300
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1301
theorem bijection_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1302
     "REFLECTS[\<lambda>x. bijection(L,f(x),g(x),h(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1303
               \<lambda>i x. bijection(**Lset(i),f(x),g(x),h(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1304
apply (simp only: bijection_def setclass_simps)
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1305
apply (intro And_reflection injection_reflection surjection_reflection)  
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1306
done
13309
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1307
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1308
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1309
subsubsection{*Restriction of a Relation, Internalized*}
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1310
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1311
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1312
(* "restriction(M,r,A,z) == 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1313
	\<forall>x[M]. x \<in> z <-> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))" *)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1314
constdefs restriction_fm :: "[i,i,i]=>i"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1315
    "restriction_fm(r,A,z) == 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1316
       Forall(Iff(Member(0,succ(z)),
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1317
                  And(Member(0,succ(r)),
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1318
                      Exists(And(Member(0,succ(succ(A))),
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1319
                                 Exists(pair_fm(1,0,2)))))))"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1320
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1321
lemma restriction_type [TC]:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1322
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] ==> restriction_fm(x,y,z) \<in> formula"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1323
by (simp add: restriction_fm_def) 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1324
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1325
lemma arity_restriction_fm [simp]:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1326
     "[| x \<in> nat; y \<in> nat; z \<in> nat |] 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1327
      ==> arity(restriction_fm(x,y,z)) = succ(x) \<union> succ(y) \<union> succ(z)"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1328
by (simp add: restriction_fm_def succ_Un_distrib [symmetric] Un_ac) 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1329
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1330
lemma sats_restriction_fm [simp]:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1331
   "[| x \<in> nat; y \<in> nat; z \<in> nat; env \<in> list(A)|]
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1332
    ==> sats(A, restriction_fm(x,y,z), env) <-> 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1333
        restriction(**A, nth(x,env), nth(y,env), nth(z,env))"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1334
by (simp add: restriction_fm_def restriction_def)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1335
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1336
lemma restriction_iff_sats:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1337
      "[| nth(i,env) = x; nth(j,env) = y; nth(k,env) = z; 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1338
          i \<in> nat; j \<in> nat; k \<in> nat; env \<in> list(A)|]
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1339
       ==> restriction(**A, x, y, z) <-> sats(A, restriction_fm(i,j,k), env)"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1340
by simp
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1341
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1342
theorem restriction_reflection:
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1343
     "REFLECTS[\<lambda>x. restriction(L,f(x),g(x),h(x)), 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1344
               \<lambda>i x. restriction(**Lset(i),f(x),g(x),h(x))]"
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1345
apply (simp only: restriction_def setclass_simps)
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1346
apply (intro FOL_reflections pair_reflection)  
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1347
done
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1348
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
  1349
subsubsection{*Order-Isomorphisms, Internalized*}
13309
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1350
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1351
(*  order_isomorphism :: "[i=>o,i,i,i,i,i] => o"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1352
   "order_isomorphism(M,A,r,B,s,f) == 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1353
        bijection(M,A,B,f) & 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1354
        (\<forall>x[M]. x\<in>A --> (\<forall>y[M]. y\<in>A -->
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1355
          (\<forall>p[M]. \<forall>fx[M]. \<forall>fy[M]. \<forall>q[M].
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1356
            pair(M,x,y,p) --> fun_apply(M,f,x,fx) --> fun_apply(M,f,y,fy) --> 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1357
            pair(M,fx,fy,q) --> (p\<in>r <-> q\<in>s))))"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1358
  *)
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1359
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1360
constdefs order_isomorphism_fm :: "[i,i,i,i,i]=>i"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1361
 "order_isomorphism_fm(A,r,B,s,f) == 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1362
   And(bijection_fm(A,B,f), 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1363
     Forall(Implies(Member(0,succ(A)),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1364
       Forall(Implies(Member(0,succ(succ(A))),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1365
         Forall(Forall(Forall(Forall(
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1366
           Implies(pair_fm(5,4,3),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1367
             Implies(fun_apply_fm(f#+6,5,2),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1368
               Implies(fun_apply_fm(f#+6,4,1),
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1369
                 Implies(pair_fm(2,1,0), 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1370
                   Iff(Member(3,r#+6), Member(0,s#+6)))))))))))))))"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1371
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1372
lemma order_isomorphism_type [TC]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1373
     "[| A \<in> nat; r \<in> nat; B \<in> nat; s \<in> nat; f \<in> nat |]  
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1374
      ==> order_isomorphism_fm(A,r,B,s,f) \<in> formula"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1375
by (simp add: order_isomorphism_fm_def) 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1376
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1377
lemma arity_order_isomorphism_fm [simp]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1378
     "[| A \<in> nat; r \<in> nat; B \<in> nat; s \<in> nat; f \<in> nat |] 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1379
      ==> arity(order_isomorphism_fm(A,r,B,s,f)) = 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1380
          succ(A) \<union> succ(r) \<union> succ(B) \<union> succ(s) \<union> succ(f)" 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1381
by (simp add: order_isomorphism_fm_def succ_Un_distrib [symmetric] Un_ac) 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1382
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1383
lemma sats_order_isomorphism_fm [simp]:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1384
   "[| U \<in> nat; r \<in> nat; B \<in> nat; s \<in> nat; f \<in> nat; env \<in> list(A)|]
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1385
    ==> sats(A, order_isomorphism_fm(U,r,B,s,f), env) <-> 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1386
        order_isomorphism(**A, nth(U,env), nth(r,env), nth(B,env), 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1387
                               nth(s,env), nth(f,env))"
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1388
by (simp add: order_isomorphism_fm_def order_isomorphism_def)
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1389
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1390
lemma order_isomorphism_iff_sats:
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1391
  "[| nth(i,env) = U; nth(j,env) = r; nth(k,env) = B; nth(j',env) = s; 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1392
      nth(k',env) = f; 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1393
      i \<in> nat; j \<in> nat; k \<in> nat; j' \<in> nat; k' \<in> nat; env \<in> list(A)|]
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1394
   ==> order_isomorphism(**A,U,r,B,s,f) <-> 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1395
       sats(A, order_isomorphism_fm(i,j,k,j',k'), env)" 
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1396
by simp
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1397
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1398
theorem order_isomorphism_reflection:
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1399
     "REFLECTS[\<lambda>x. order_isomorphism(L,f(x),g(x),h(x),g'(x),h'(x)), 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1400
               \<lambda>i x. order_isomorphism(**Lset(i),f(x),g(x),h(x),g'(x),h'(x))]"
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1401
apply (simp only: order_isomorphism_def setclass_simps)
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1402
apply (intro FOL_reflections function_reflections bijection_reflection)  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1403
done
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1404
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13323
diff changeset
  1405
subsubsection{*Limit Ordinals, Internalized*}
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1406
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1407
text{*A limit ordinal is a non-empty, successor-closed ordinal*}
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1408
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1409
(* "limit_ordinal(M,a) == 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1410
	ordinal(M,a) & ~ empty(M,a) & 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1411
        (\<forall>x[M]. x\<in>a --> (\<exists>y[M]. y\<in>a & successor(M,x,y)))" *)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1412
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1413
constdefs limit_ordinal_fm :: "i=>i"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1414
    "limit_ordinal_fm(x) == 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1415
        And(ordinal_fm(x),
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1416
            And(Neg(empty_fm(x)),
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1417
	        Forall(Implies(Member(0,succ(x)),
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1418
                               Exists(And(Member(0,succ(succ(x))),
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1419
                                          succ_fm(1,0)))))))"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1420
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1421
lemma limit_ordinal_type [TC]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1422
     "x \<in> nat ==> limit_ordinal_fm(x) \<in> formula"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1423
by (simp add: limit_ordinal_fm_def) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1424
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1425
lemma arity_limit_ordinal_fm [simp]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1426
     "x \<in> nat ==> arity(limit_ordinal_fm(x)) = succ(x)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1427
by (simp add: limit_ordinal_fm_def succ_Un_distrib [symmetric] Un_ac) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1428
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1429
lemma sats_limit_ordinal_fm [simp]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1430
   "[| x \<in> nat; env \<in> list(A)|]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1431
    ==> sats(A, limit_ordinal_fm(x), env) <-> limit_ordinal(**A, nth(x,env))"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1432
by (simp add: limit_ordinal_fm_def limit_ordinal_def sats_ordinal_fm')
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1433
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1434
lemma limit_ordinal_iff_sats:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1435
      "[| nth(i,env) = x; nth(j,env) = y; 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1436
          i \<in> nat; env \<in> list(A)|]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1437
       ==> limit_ordinal(**A, x) <-> sats(A, limit_ordinal_fm(i), env)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1438
by simp
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1439
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1440
theorem limit_ordinal_reflection:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1441
     "REFLECTS[\<lambda>x. limit_ordinal(L,f(x)), 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1442
               \<lambda>i x. limit_ordinal(**Lset(i),f(x))]"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1443
apply (simp only: limit_ordinal_def setclass_simps)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1444
apply (intro FOL_reflections ordinal_reflection 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1445
             empty_reflection successor_reflection)  
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13309
diff changeset
  1446
done
13309
a6adee8ea75a reflection for more internal formulas
paulson
parents: 13306
diff changeset
  1447
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1448
subsubsection{*Omega: The Set of Natural Numbers*}
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1449
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1450
(* omega(M,a) == limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a --> ~ limit_ordinal(M,x)) *)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1451
constdefs omega_fm :: "i=>i"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1452
    "omega_fm(x) == 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1453
       And(limit_ordinal_fm(x),
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1454
           Forall(Implies(Member(0,succ(x)),
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1455
                          Neg(limit_ordinal_fm(0)))))"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1456
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1457
lemma omega_type [TC]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1458
     "x \<in> nat ==> omega_fm(x) \<in> formula"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1459
by (simp add: omega_fm_def) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1460
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1461
lemma arity_omega_fm [simp]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1462
     "x \<in> nat ==> arity(omega_fm(x)) = succ(x)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1463
by (simp add: omega_fm_def succ_Un_distrib [symmetric] Un_ac) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1464
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1465
lemma sats_omega_fm [simp]:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1466
   "[| x \<in> nat; env \<in> list(A)|]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1467
    ==> sats(A, omega_fm(x), env) <-> omega(**A, nth(x,env))"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1468
by (simp add: omega_fm_def omega_def)
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
  1469
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1470
lemma omega_iff_sats:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1471
      "[| nth(i,env) = x; nth(j,env) = y; 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1472
          i \<in> nat; env \<in> list(A)|]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1473
       ==> omega(**A, x) <-> sats(A, omega_fm(i), env)"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1474
by simp
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1475
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1476
theorem omega_reflection:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1477
     "REFLECTS[\<lambda>x. omega(L,f(x)), 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1478
               \<lambda>i x. omega(**Lset(i),f(x))]"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1479
apply (simp only: omega_def setclass_simps)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1480
apply (intro FOL_reflections limit_ordinal_reflection)  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1481
done
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1482
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1483
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1484
lemmas fun_plus_reflections =
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1485
        typed_function_reflection composition_reflection
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1486
        injection_reflection surjection_reflection
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1487
        bijection_reflection restriction_reflection
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1488
        order_isomorphism_reflection
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1489
        ordinal_reflection limit_ordinal_reflection omega_reflection
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1490
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1491
lemmas fun_plus_iff_sats = 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1492
	typed_function_iff_sats composition_iff_sats
13348
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1493
        injection_iff_sats surjection_iff_sats 
374d05460db4 Separation/Replacement up to M_wfrank!
paulson
parents: 13339
diff changeset
  1494
        bijection_iff_sats restriction_iff_sats 
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
  1495
        order_isomorphism_iff_sats
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13316
diff changeset
  1496
        ordinal_iff_sats limit_ordinal_iff_sats omega_iff_sats
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
  1497
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1498
end