src/ZF/Constructible/Separation.thy
author paulson
Wed, 10 Jul 2002 16:54:07 +0200
changeset 13339 0f89104dd377
parent 13324 39d1b3a4c6f4
child 13350 626b79677dfa
permissions -rw-r--r--
Fixed quantified variable name preservation for ball and bex (bounded quants) Requires tweaking of other scripts. Also routine tidying.
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: 13324
diff changeset
     1
header{*Early Instances of Separation and Strong Replacement*}
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
     2
13324
39d1b3a4c6f4 more and simpler separation proofs
paulson
parents: 13323
diff changeset
     3
theory Separation = L_axioms + WF_absolute:
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
     4
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
     5
text{*This theory proves all instances needed for locale @{text "M_axioms"}*}
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
     6
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
     7
text{*Helps us solve for de Bruijn indices!*}
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
     8
lemma nth_ConsI: "[|nth(n,l) = x; n \<in> nat|] ==> nth(succ(n), Cons(a,l)) = x"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
     9
by simp
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    10
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
    11
lemmas nth_rules = nth_0 nth_ConsI nat_0I nat_succI
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
    12
lemmas sep_rules = nth_0 nth_ConsI FOL_iff_sats function_iff_sats 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
    13
                   fun_plus_iff_sats
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    14
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    15
lemma Collect_conj_in_DPow:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    16
     "[| {x\<in>A. P(x)} \<in> DPow(A);  {x\<in>A. Q(x)} \<in> DPow(A) |] 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    17
      ==> {x\<in>A. P(x) & Q(x)} \<in> DPow(A)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    18
by (simp add: Int_in_DPow Collect_Int_Collect_eq [symmetric]) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    19
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    20
lemma Collect_conj_in_DPow_Lset:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    21
     "[|z \<in> Lset(j); {x \<in> Lset(j). P(x)} \<in> DPow(Lset(j))|]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    22
      ==> {x \<in> Lset(j). x \<in> z & P(x)} \<in> DPow(Lset(j))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    23
apply (frule mem_Lset_imp_subset_Lset)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    24
apply (simp add: Collect_conj_in_DPow Collect_mem_eq 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    25
                 subset_Int_iff2 elem_subset_in_DPow)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    26
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    27
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    28
lemma separation_CollectI:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    29
     "(\<And>z. L(z) ==> L({x \<in> z . P(x)})) ==> separation(L, \<lambda>x. P(x))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    30
apply (unfold separation_def, clarify) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    31
apply (rule_tac x="{x\<in>z. P(x)}" in rexI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    32
apply simp_all
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    33
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    34
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    35
text{*Reduces the original comprehension to the reflected one*}
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    36
lemma reflection_imp_L_separation:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    37
      "[| \<forall>x\<in>Lset(j). P(x) <-> Q(x);
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    38
          {x \<in> Lset(j) . Q(x)} \<in> DPow(Lset(j)); 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    39
          Ord(j);  z \<in> Lset(j)|] ==> L({x \<in> z . P(x)})"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    40
apply (rule_tac i = "succ(j)" in L_I)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    41
 prefer 2 apply simp
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    42
apply (subgoal_tac "{x \<in> z. P(x)} = {x \<in> Lset(j). x \<in> z & (Q(x))}")
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    43
 prefer 2
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    44
 apply (blast dest: mem_Lset_imp_subset_Lset) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    45
apply (simp add: Lset_succ Collect_conj_in_DPow_Lset)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    46
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    47
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    48
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
    49
subsection{*Separation for Intersection*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    50
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    51
lemma Inter_Reflects:
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
    52
     "REFLECTS[\<lambda>x. \<forall>y[L]. y\<in>A --> x \<in> y, 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
    53
               \<lambda>i x. \<forall>y\<in>Lset(i). y\<in>A --> x \<in> y]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
    54
by (intro FOL_reflections)  
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    55
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    56
lemma Inter_separation:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    57
     "L(A) ==> separation(L, \<lambda>x. \<forall>y[L]. y\<in>A --> x\<in>y)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    58
apply (rule separation_CollectI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    59
apply (rule_tac A="{A,z}" in subset_LsetE, blast ) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    60
apply (rule ReflectsE [OF Inter_Reflects], assumption)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    61
apply (drule subset_Lset_ltD, assumption) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    62
apply (erule reflection_imp_L_separation)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    63
  apply (simp_all add: lt_Ord2, clarify)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    64
apply (rule DPowI2) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    65
apply (rule ball_iff_sats) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    66
apply (rule imp_iff_sats)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    67
apply (rule_tac [2] i=1 and j=0 and env="[y,x,A]" in mem_iff_sats)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    68
apply (rule_tac i=0 and j=2 in mem_iff_sats)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    69
apply (simp_all add: succ_Un_distrib [symmetric])
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    70
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    71
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
    72
subsection{*Separation for Cartesian Product*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    73
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
    74
lemma cartprod_Reflects:
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
    75
     "REFLECTS[\<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)),
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    76
                \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A & (\<exists>y\<in>Lset(i). y\<in>B & 
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
    77
                                   pair(**Lset(i),x,y,z))]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
    78
by (intro FOL_reflections function_reflections)
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    79
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    80
lemma cartprod_separation:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    81
     "[| L(A); L(B) |] 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    82
      ==> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. y\<in>B & pair(L,x,y,z)))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    83
apply (rule separation_CollectI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    84
apply (rule_tac A="{A,B,z}" in subset_LsetE, blast ) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    85
apply (rule ReflectsE [OF cartprod_Reflects], assumption)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    86
apply (drule subset_Lset_ltD, assumption) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    87
apply (erule reflection_imp_L_separation)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    88
  apply (simp_all add: lt_Ord2, clarify) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    89
apply (rule DPowI2)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    90
apply (rename_tac u)  
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    91
apply (rule bex_iff_sats) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    92
apply (rule conj_iff_sats)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    93
apply (rule_tac i=0 and j=2 and env="[x,u,A,B]" in mem_iff_sats, simp_all)
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
    94
apply (rule sep_rules | simp)+
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    95
apply (simp_all add: succ_Un_distrib [symmetric])
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    96
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    97
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
    98
subsection{*Separation for Image*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
    99
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   100
lemma image_Reflects:
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   101
     "REFLECTS[\<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>A & pair(L,x,y,p)),
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   102
           \<lambda>i y. \<exists>p\<in>Lset(i). p\<in>r & (\<exists>x\<in>Lset(i). x\<in>A & pair(**Lset(i),x,y,p))]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   103
by (intro FOL_reflections function_reflections)
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   104
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   105
lemma image_separation:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   106
     "[| L(A); L(r) |] 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   107
      ==> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & (\<exists>x[L]. x\<in>A & pair(L,x,y,p)))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   108
apply (rule separation_CollectI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   109
apply (rule_tac A="{A,r,z}" in subset_LsetE, blast ) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   110
apply (rule ReflectsE [OF image_Reflects], assumption)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   111
apply (drule subset_Lset_ltD, assumption) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   112
apply (erule reflection_imp_L_separation)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   113
  apply (simp_all add: lt_Ord2, clarify)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   114
apply (rule DPowI2)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   115
apply (rule bex_iff_sats) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   116
apply (rule conj_iff_sats)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   117
apply (rule_tac env="[p,y,A,r]" in mem_iff_sats)
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   118
apply (rule sep_rules | simp)+
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   119
apply (simp_all add: succ_Un_distrib [symmetric])
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   120
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   121
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   122
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   123
subsection{*Separation for Converse*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   124
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   125
lemma converse_Reflects:
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   126
  "REFLECTS[\<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) & pair(L,y,x,z)),
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   127
     \<lambda>i z. \<exists>p\<in>Lset(i). p\<in>r & (\<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i). 
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   128
                     pair(**Lset(i),x,y,p) & pair(**Lset(i),y,x,z))]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   129
by (intro FOL_reflections function_reflections)
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   130
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   131
lemma converse_separation:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   132
     "L(r) ==> separation(L, 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   133
         \<lambda>z. \<exists>p[L]. p\<in>r & (\<exists>x[L]. \<exists>y[L]. pair(L,x,y,p) & pair(L,y,x,z)))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   134
apply (rule separation_CollectI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   135
apply (rule_tac A="{r,z}" in subset_LsetE, blast ) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   136
apply (rule ReflectsE [OF converse_Reflects], assumption)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   137
apply (drule subset_Lset_ltD, assumption) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   138
apply (erule reflection_imp_L_separation)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   139
  apply (simp_all add: lt_Ord2, clarify)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   140
apply (rule DPowI2)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   141
apply (rename_tac u) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   142
apply (rule bex_iff_sats) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   143
apply (rule conj_iff_sats)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   144
apply (rule_tac i=0 and j=2 and env="[p,u,r]" in mem_iff_sats, simp_all)
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   145
apply (rule sep_rules | simp)+
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   146
apply (simp_all add: succ_Un_distrib [symmetric])
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   147
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   148
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   149
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   150
subsection{*Separation for Restriction*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   151
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   152
lemma restrict_Reflects:
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   153
     "REFLECTS[\<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. pair(L,x,y,z)),
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   154
        \<lambda>i z. \<exists>x\<in>Lset(i). x\<in>A & (\<exists>y\<in>Lset(i). pair(**Lset(i),x,y,z))]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   155
by (intro FOL_reflections function_reflections)
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   156
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   157
lemma restrict_separation:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   158
   "L(A) ==> separation(L, \<lambda>z. \<exists>x[L]. x\<in>A & (\<exists>y[L]. pair(L,x,y,z)))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   159
apply (rule separation_CollectI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   160
apply (rule_tac A="{A,z}" in subset_LsetE, blast ) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   161
apply (rule ReflectsE [OF restrict_Reflects], assumption)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   162
apply (drule subset_Lset_ltD, assumption) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   163
apply (erule reflection_imp_L_separation)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   164
  apply (simp_all add: lt_Ord2, clarify)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   165
apply (rule DPowI2)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   166
apply (rename_tac u) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   167
apply (rule bex_iff_sats) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   168
apply (rule conj_iff_sats)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   169
apply (rule_tac i=0 and j=2 and env="[x,u,A]" in mem_iff_sats, simp_all)
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   170
apply (rule sep_rules | simp)+
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   171
apply (simp_all add: succ_Un_distrib [symmetric])
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   172
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   173
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   174
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   175
subsection{*Separation for Composition*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   176
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   177
lemma comp_Reflects:
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   178
     "REFLECTS[\<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L]. 
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   179
		  pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   180
                  xy\<in>s & yz\<in>r,
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   181
        \<lambda>i xz. \<exists>x\<in>Lset(i). \<exists>y\<in>Lset(i). \<exists>z\<in>Lset(i). \<exists>xy\<in>Lset(i). \<exists>yz\<in>Lset(i). 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   182
		  pair(**Lset(i),x,z,xz) & pair(**Lset(i),x,y,xy) & 
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   183
                  pair(**Lset(i),y,z,yz) & xy\<in>s & yz\<in>r]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   184
by (intro FOL_reflections function_reflections)
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   185
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   186
lemma comp_separation:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   187
     "[| L(r); L(s) |]
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   188
      ==> separation(L, \<lambda>xz. \<exists>x[L]. \<exists>y[L]. \<exists>z[L]. \<exists>xy[L]. \<exists>yz[L]. 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   189
		  pair(L,x,z,xz) & pair(L,x,y,xy) & pair(L,y,z,yz) & 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   190
                  xy\<in>s & yz\<in>r)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   191
apply (rule separation_CollectI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   192
apply (rule_tac A="{r,s,z}" in subset_LsetE, blast ) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   193
apply (rule ReflectsE [OF comp_Reflects], assumption)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   194
apply (drule subset_Lset_ltD, assumption) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   195
apply (erule reflection_imp_L_separation)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   196
  apply (simp_all add: lt_Ord2, clarify)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   197
apply (rule DPowI2)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   198
apply (rename_tac u) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   199
apply (rule bex_iff_sats)+
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   200
apply (rename_tac x y z)  
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   201
apply (rule conj_iff_sats)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   202
apply (rule_tac env="[z,y,x,u,r,s]" in pair_iff_sats)
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   203
apply (rule sep_rules | simp)+
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   204
apply (simp_all add: succ_Un_distrib [symmetric])
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   205
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   206
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   207
subsection{*Separation for Predecessors in an Order*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   208
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   209
lemma pred_Reflects:
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   210
     "REFLECTS[\<lambda>y. \<exists>p[L]. p\<in>r & pair(L,y,x,p),
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   211
                    \<lambda>i y. \<exists>p \<in> Lset(i). p\<in>r & pair(**Lset(i),y,x,p)]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   212
by (intro FOL_reflections function_reflections)
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   213
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   214
lemma pred_separation:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   215
     "[| L(r); L(x) |] ==> separation(L, \<lambda>y. \<exists>p[L]. p\<in>r & pair(L,y,x,p))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   216
apply (rule separation_CollectI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   217
apply (rule_tac A="{r,x,z}" in subset_LsetE, blast ) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   218
apply (rule ReflectsE [OF pred_Reflects], assumption)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   219
apply (drule subset_Lset_ltD, assumption) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   220
apply (erule reflection_imp_L_separation)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   221
  apply (simp_all add: lt_Ord2, clarify)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   222
apply (rule DPowI2)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   223
apply (rename_tac u) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   224
apply (rule bex_iff_sats)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   225
apply (rule conj_iff_sats)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   226
apply (rule_tac env = "[p,u,r,x]" in mem_iff_sats) 
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   227
apply (rule sep_rules | simp)+
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   228
apply (simp_all add: succ_Un_distrib [symmetric])
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   229
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   230
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   231
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   232
subsection{*Separation for the Membership Relation*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   233
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   234
lemma Memrel_Reflects:
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   235
     "REFLECTS[\<lambda>z. \<exists>x[L]. \<exists>y[L]. pair(L,x,y,z) & x \<in> y,
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   236
            \<lambda>i z. \<exists>x \<in> Lset(i). \<exists>y \<in> Lset(i). pair(**Lset(i),x,y,z) & x \<in> y]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   237
by (intro FOL_reflections function_reflections)
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   238
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   239
lemma Memrel_separation:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   240
     "separation(L, \<lambda>z. \<exists>x[L]. \<exists>y[L]. pair(L,x,y,z) & x \<in> y)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   241
apply (rule separation_CollectI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   242
apply (rule_tac A="{z}" in subset_LsetE, blast ) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   243
apply (rule ReflectsE [OF Memrel_Reflects], assumption)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   244
apply (drule subset_Lset_ltD, assumption) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   245
apply (erule reflection_imp_L_separation)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   246
  apply (simp_all add: lt_Ord2)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   247
apply (rule DPowI2)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   248
apply (rename_tac u) 
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   249
apply (rule bex_iff_sats conj_iff_sats)+
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   250
apply (rule_tac env = "[y,x,u]" in pair_iff_sats) 
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   251
apply (rule sep_rules | simp)+
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   252
apply (simp_all add: succ_Un_distrib [symmetric])
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   253
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   254
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   255
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   256
subsection{*Replacement for FunSpace*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   257
		
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   258
lemma funspace_succ_Reflects:
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   259
 "REFLECTS[\<lambda>z. \<exists>p[L]. p\<in>A & (\<exists>f[L]. \<exists>b[L]. \<exists>nb[L]. \<exists>cnbf[L]. 
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   260
	    pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) &
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   261
	    upair(L,cnbf,cnbf,z)),
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   262
	\<lambda>i z. \<exists>p \<in> Lset(i). p\<in>A & (\<exists>f \<in> Lset(i). \<exists>b \<in> Lset(i). 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   263
	      \<exists>nb \<in> Lset(i). \<exists>cnbf \<in> Lset(i). 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   264
		pair(**Lset(i),f,b,p) & pair(**Lset(i),n,b,nb) & 
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   265
		is_cons(**Lset(i),nb,f,cnbf) & upair(**Lset(i),cnbf,cnbf,z))]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   266
by (intro FOL_reflections function_reflections)
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   267
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   268
lemma funspace_succ_replacement:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   269
     "L(n) ==> 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   270
      strong_replacement(L, \<lambda>p z. \<exists>f[L]. \<exists>b[L]. \<exists>nb[L]. \<exists>cnbf[L]. 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   271
                pair(L,f,b,p) & pair(L,n,b,nb) & is_cons(L,nb,f,cnbf) &
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   272
                upair(L,cnbf,cnbf,z))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   273
apply (rule strong_replacementI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   274
apply (rule rallI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   275
apply (rule separation_CollectI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   276
apply (rule_tac A="{n,A,z}" in subset_LsetE, blast ) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   277
apply (rule ReflectsE [OF funspace_succ_Reflects], assumption)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   278
apply (drule subset_Lset_ltD, assumption) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   279
apply (erule reflection_imp_L_separation)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   280
  apply (simp_all add: lt_Ord2)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   281
apply (rule DPowI2)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   282
apply (rename_tac u) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   283
apply (rule bex_iff_sats)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   284
apply (rule conj_iff_sats)
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   285
apply (rule_tac env = "[p,u,n,A]" in mem_iff_sats) 
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   286
apply (rule sep_rules | simp)+
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   287
apply (simp_all add: succ_Un_distrib [symmetric])
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   288
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   289
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   290
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   291
subsection{*Separation for Order-Isomorphisms*}
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   292
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   293
lemma well_ord_iso_Reflects:
13314
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   294
  "REFLECTS[\<lambda>x. x\<in>A --> 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   295
                (\<exists>y[L]. \<exists>p[L]. fun_apply(L,f,x,y) & pair(L,y,x,p) & p \<in> r),
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   296
        \<lambda>i x. x\<in>A --> (\<exists>y \<in> Lset(i). \<exists>p \<in> Lset(i). 
84b9de3cbc91 Defining a meta-existential quantifier.
paulson
parents: 13306
diff changeset
   297
                fun_apply(**Lset(i),f,x,y) & pair(**Lset(i),y,x,p) & p \<in> r)]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   298
by (intro FOL_reflections function_reflections)
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   299
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   300
lemma well_ord_iso_separation:
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   301
     "[| L(A); L(f); L(r) |] 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   302
      ==> separation (L, \<lambda>x. x\<in>A --> (\<exists>y[L]. (\<exists>p[L]. 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   303
		     fun_apply(L,f,x,y) & pair(L,y,x,p) & p \<in> r)))"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   304
apply (rule separation_CollectI) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   305
apply (rule_tac A="{A,f,r,z}" in subset_LsetE, blast ) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   306
apply (rule ReflectsE [OF well_ord_iso_Reflects], assumption)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   307
apply (drule subset_Lset_ltD, assumption) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   308
apply (erule reflection_imp_L_separation)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   309
  apply (simp_all add: lt_Ord2)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   310
apply (rule DPowI2)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   311
apply (rename_tac u) 
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   312
apply (rule imp_iff_sats)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   313
apply (rule_tac env = "[u,A,f,r]" in mem_iff_sats) 
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   314
apply (rule sep_rules | simp)+
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   315
apply (simp_all add: succ_Un_distrib [symmetric])
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   316
done
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   317
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   318
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   319
subsection{*Separation for @{term "obase"}*}
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   320
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   321
lemma obase_reflects:
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   322
  "REFLECTS[\<lambda>a. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L]. 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   323
	     ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) &
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   324
	     order_isomorphism(L,par,r,x,mx,g),
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   325
        \<lambda>i a. \<exists>x \<in> Lset(i). \<exists>g \<in> Lset(i). \<exists>mx \<in> Lset(i). \<exists>par \<in> Lset(i). 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   326
	     ordinal(**Lset(i),x) & membership(**Lset(i),x,mx) & pred_set(**Lset(i),A,a,r,par) &
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   327
	     order_isomorphism(**Lset(i),par,r,x,mx,g)]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   328
by (intro FOL_reflections function_reflections fun_plus_reflections)
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   329
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   330
lemma obase_separation:
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   331
     --{*part of the order type formalization*}
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   332
     "[| L(A); L(r) |] 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   333
      ==> separation(L, \<lambda>a. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L]. 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   334
	     ordinal(L,x) & membership(L,x,mx) & pred_set(L,A,a,r,par) &
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   335
	     order_isomorphism(L,par,r,x,mx,g))"
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   336
apply (rule separation_CollectI) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   337
apply (rule_tac A="{A,r,z}" in subset_LsetE, blast ) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   338
apply (rule ReflectsE [OF obase_reflects], assumption)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   339
apply (drule subset_Lset_ltD, assumption) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   340
apply (erule reflection_imp_L_separation)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   341
  apply (simp_all add: lt_Ord2)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   342
apply (rule DPowI2)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   343
apply (rename_tac u) 
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   344
apply (rule bex_iff_sats)
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   345
apply (rule conj_iff_sats)
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   346
apply (rule_tac env = "[x,u,A,r]" in ordinal_iff_sats) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   347
apply (rule sep_rules | simp)+
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   348
apply (simp_all add: succ_Un_distrib [symmetric])
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   349
done
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   350
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   351
13319
23de7b3af453 More Separation proofs
paulson
parents: 13316
diff changeset
   352
subsection{*Separation for a Theorem about @{term "obase"}*}
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   353
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   354
lemma obase_equals_reflects:
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   355
  "REFLECTS[\<lambda>x. x\<in>A --> ~(\<exists>y[L]. \<exists>g[L]. 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   356
		ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L]. 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   357
		membership(L,y,my) & pred_set(L,A,x,r,pxr) &
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   358
		order_isomorphism(L,pxr,r,y,my,g))),
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   359
	\<lambda>i x. x\<in>A --> ~(\<exists>y \<in> Lset(i). \<exists>g \<in> Lset(i). 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   360
		ordinal(**Lset(i),y) & (\<exists>my \<in> Lset(i). \<exists>pxr \<in> Lset(i). 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   361
		membership(**Lset(i),y,my) & pred_set(**Lset(i),A,x,r,pxr) &
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   362
		order_isomorphism(**Lset(i),pxr,r,y,my,g)))]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   363
by (intro FOL_reflections function_reflections fun_plus_reflections)
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   364
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   365
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   366
lemma obase_equals_separation:
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   367
     "[| L(A); L(r) |] 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   368
      ==> separation (L, \<lambda>x. x\<in>A --> ~(\<exists>y[L]. \<exists>g[L]. 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   369
			      ordinal(L,y) & (\<exists>my[L]. \<exists>pxr[L]. 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   370
			      membership(L,y,my) & pred_set(L,A,x,r,pxr) &
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   371
			      order_isomorphism(L,pxr,r,y,my,g))))"
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   372
apply (rule separation_CollectI) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   373
apply (rule_tac A="{A,r,z}" in subset_LsetE, blast ) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   374
apply (rule ReflectsE [OF obase_equals_reflects], assumption)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   375
apply (drule subset_Lset_ltD, assumption) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   376
apply (erule reflection_imp_L_separation)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   377
  apply (simp_all add: lt_Ord2)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   378
apply (rule DPowI2)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   379
apply (rename_tac u) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   380
apply (rule imp_iff_sats ball_iff_sats disj_iff_sats not_iff_sats)+
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   381
apply (rule_tac env = "[u,A,r]" in mem_iff_sats) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   382
apply (rule sep_rules | simp)+
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   383
apply (simp_all add: succ_Un_distrib [symmetric])
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   384
done
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   385
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   386
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   387
subsection{*Replacement for @{term "omap"}*}
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   388
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   389
lemma omap_reflects:
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   390
 "REFLECTS[\<lambda>z. \<exists>a[L]. a\<in>B & (\<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L]. 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   391
     ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) & 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   392
     pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g)),
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   393
 \<lambda>i z. \<exists>a \<in> Lset(i). a\<in>B & (\<exists>x \<in> Lset(i). \<exists>g \<in> Lset(i). \<exists>mx \<in> Lset(i). 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   394
        \<exists>par \<in> Lset(i). 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   395
	 ordinal(**Lset(i),x) & pair(**Lset(i),a,x,z) & 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   396
         membership(**Lset(i),x,mx) & pred_set(**Lset(i),A,a,r,par) & 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   397
         order_isomorphism(**Lset(i),par,r,x,mx,g))]"
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   398
by (intro FOL_reflections function_reflections fun_plus_reflections)
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   399
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   400
lemma omap_replacement:
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   401
     "[| L(A); L(r) |] 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   402
      ==> strong_replacement(L,
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   403
             \<lambda>a z. \<exists>x[L]. \<exists>g[L]. \<exists>mx[L]. \<exists>par[L]. 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   404
	     ordinal(L,x) & pair(L,a,x,z) & membership(L,x,mx) & 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   405
	     pred_set(L,A,a,r,par) & order_isomorphism(L,par,r,x,mx,g))"
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   406
apply (rule strong_replacementI) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   407
apply (rule rallI)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   408
apply (rename_tac B)  
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   409
apply (rule separation_CollectI) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   410
apply (rule_tac A="{A,B,r,z}" in subset_LsetE, blast ) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   411
apply (rule ReflectsE [OF omap_reflects], assumption)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   412
apply (drule subset_Lset_ltD, assumption) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   413
apply (erule reflection_imp_L_separation)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   414
  apply (simp_all add: lt_Ord2)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   415
apply (rule DPowI2)
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   416
apply (rename_tac u) 
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   417
apply (rule bex_iff_sats conj_iff_sats)+
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   418
apply (rule_tac env = "[a,u,A,B,r]" in mem_iff_sats) 
13316
d16629fd0f95 more and simpler separation proofs
paulson
parents: 13314
diff changeset
   419
apply (rule sep_rules | simp)+
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   420
apply (simp_all add: succ_Un_distrib [symmetric])
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   421
done
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   422
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   423
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   424
subsection{*Separation for a Theorem about @{term "obase"}*}
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   425
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   426
lemma is_recfun_reflects:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   427
  "REFLECTS[\<lambda>x. \<exists>xa[L]. \<exists>xb[L]. 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   428
                pair(L,x,a,xa) & xa \<in> r & pair(L,x,b,xb) & xb \<in> r & 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   429
                (\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) & 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   430
                                   fx \<noteq> gx),
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   431
   \<lambda>i x. \<exists>xa \<in> Lset(i). \<exists>xb \<in> Lset(i). 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   432
          pair(**Lset(i),x,a,xa) & xa \<in> r & pair(**Lset(i),x,b,xb) & xb \<in> r &
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   433
                (\<exists>fx \<in> Lset(i). \<exists>gx \<in> Lset(i). fun_apply(**Lset(i),f,x,fx) & 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   434
                  fun_apply(**Lset(i),g,x,gx) & fx \<noteq> gx)]"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   435
by (intro FOL_reflections function_reflections fun_plus_reflections)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   436
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   437
lemma is_recfun_separation:
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   438
     --{*for well-founded recursion*}
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   439
     "[| L(r); L(f); L(g); L(a); L(b) |] 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   440
     ==> separation(L, 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   441
            \<lambda>x. \<exists>xa[L]. \<exists>xb[L]. 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   442
                pair(L,x,a,xa) & xa \<in> r & pair(L,x,b,xb) & xb \<in> r & 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   443
                (\<exists>fx[L]. \<exists>gx[L]. fun_apply(L,f,x,fx) & fun_apply(L,g,x,gx) & 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   444
                                   fx \<noteq> gx))"
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   445
apply (rule separation_CollectI) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   446
apply (rule_tac A="{r,f,g,a,b,z}" in subset_LsetE, blast ) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   447
apply (rule ReflectsE [OF is_recfun_reflects], assumption)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   448
apply (drule subset_Lset_ltD, assumption) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   449
apply (erule reflection_imp_L_separation)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   450
  apply (simp_all add: lt_Ord2)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   451
apply (rule DPowI2)
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   452
apply (rename_tac u) 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   453
apply (rule bex_iff_sats conj_iff_sats)+
13339
0f89104dd377 Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents: 13324
diff changeset
   454
apply (rule_tac env = "[xa,u,r,f,g,a,b]" in pair_iff_sats) 
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   455
apply (rule sep_rules | simp)+
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   456
apply (simp_all add: succ_Un_distrib [symmetric])
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   457
done
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   458
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   459
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   460
ML
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   461
{*
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   462
val Inter_separation = thm "Inter_separation";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   463
val cartprod_separation = thm "cartprod_separation";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   464
val image_separation = thm "image_separation";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   465
val converse_separation = thm "converse_separation";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   466
val restrict_separation = thm "restrict_separation";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   467
val comp_separation = thm "comp_separation";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   468
val pred_separation = thm "pred_separation";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   469
val Memrel_separation = thm "Memrel_separation";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   470
val funspace_succ_replacement = thm "funspace_succ_replacement";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   471
val well_ord_iso_separation = thm "well_ord_iso_separation";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   472
val obase_separation = thm "obase_separation";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   473
val obase_equals_separation = thm "obase_equals_separation";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   474
val omap_replacement = thm "omap_replacement";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   475
val is_recfun_separation = thm "is_recfun_separation";
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   476
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   477
val m_axioms = 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   478
    [Inter_separation, cartprod_separation, image_separation, 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   479
     converse_separation, restrict_separation, comp_separation, 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   480
     pred_separation, Memrel_separation, funspace_succ_replacement, 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   481
     well_ord_iso_separation, obase_separation,
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   482
     obase_equals_separation, omap_replacement, is_recfun_separation]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   483
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   484
fun axiomsL th =
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   485
    kill_flex_triv_prems (m_axioms MRS (trivaxL th));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   486
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   487
bind_thm ("cartprod_iff", axiomsL (thm "M_axioms.cartprod_iff"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   488
bind_thm ("cartprod_closed", axiomsL (thm "M_axioms.cartprod_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   489
bind_thm ("sum_closed", axiomsL (thm "M_axioms.sum_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   490
bind_thm ("M_converse_iff", axiomsL (thm "M_axioms.M_converse_iff"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   491
bind_thm ("converse_closed", axiomsL (thm "M_axioms.converse_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   492
bind_thm ("converse_abs", axiomsL (thm "M_axioms.converse_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   493
bind_thm ("image_closed", axiomsL (thm "M_axioms.image_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   494
bind_thm ("vimage_abs", axiomsL (thm "M_axioms.vimage_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   495
bind_thm ("vimage_closed", axiomsL (thm "M_axioms.vimage_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   496
bind_thm ("domain_abs", axiomsL (thm "M_axioms.domain_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   497
bind_thm ("domain_closed", axiomsL (thm "M_axioms.domain_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   498
bind_thm ("range_abs", axiomsL (thm "M_axioms.range_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   499
bind_thm ("range_closed", axiomsL (thm "M_axioms.range_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   500
bind_thm ("field_abs", axiomsL (thm "M_axioms.field_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   501
bind_thm ("field_closed", axiomsL (thm "M_axioms.field_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   502
bind_thm ("relation_abs", axiomsL (thm "M_axioms.relation_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   503
bind_thm ("function_abs", axiomsL (thm "M_axioms.function_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   504
bind_thm ("apply_closed", axiomsL (thm "M_axioms.apply_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   505
bind_thm ("apply_abs", axiomsL (thm "M_axioms.apply_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   506
bind_thm ("typed_apply_abs", axiomsL (thm "M_axioms.typed_apply_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   507
bind_thm ("typed_function_abs", axiomsL (thm "M_axioms.typed_function_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   508
bind_thm ("injection_abs", axiomsL (thm "M_axioms.injection_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   509
bind_thm ("surjection_abs", axiomsL (thm "M_axioms.surjection_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   510
bind_thm ("bijection_abs", axiomsL (thm "M_axioms.bijection_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   511
bind_thm ("M_comp_iff", axiomsL (thm "M_axioms.M_comp_iff"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   512
bind_thm ("comp_closed", axiomsL (thm "M_axioms.comp_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   513
bind_thm ("composition_abs", axiomsL (thm "M_axioms.composition_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   514
bind_thm ("restriction_is_function", axiomsL (thm "M_axioms.restriction_is_function"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   515
bind_thm ("restriction_abs", axiomsL (thm "M_axioms.restriction_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   516
bind_thm ("M_restrict_iff", axiomsL (thm "M_axioms.M_restrict_iff"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   517
bind_thm ("restrict_closed", axiomsL (thm "M_axioms.restrict_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   518
bind_thm ("Inter_abs", axiomsL (thm "M_axioms.Inter_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   519
bind_thm ("Inter_closed", axiomsL (thm "M_axioms.Inter_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   520
bind_thm ("Int_closed", axiomsL (thm "M_axioms.Int_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   521
bind_thm ("finite_fun_closed", axiomsL (thm "M_axioms.finite_fun_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   522
bind_thm ("is_funspace_abs", axiomsL (thm "M_axioms.is_funspace_abs"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   523
bind_thm ("succ_fun_eq2", axiomsL (thm "M_axioms.succ_fun_eq2"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   524
bind_thm ("funspace_succ", axiomsL (thm "M_axioms.funspace_succ"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   525
bind_thm ("finite_funspace_closed", axiomsL (thm "M_axioms.finite_funspace_closed"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   526
*}
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   527
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   528
ML
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   529
{*
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   530
bind_thm ("is_recfun_equal", axiomsL (thm "M_axioms.is_recfun_equal"));  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   531
bind_thm ("is_recfun_cut", axiomsL (thm "M_axioms.is_recfun_cut")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   532
bind_thm ("is_recfun_functional", axiomsL (thm "M_axioms.is_recfun_functional"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   533
bind_thm ("is_recfun_relativize", axiomsL (thm "M_axioms.is_recfun_relativize"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   534
bind_thm ("is_recfun_restrict", axiomsL (thm "M_axioms.is_recfun_restrict"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   535
bind_thm ("univalent_is_recfun", axiomsL (thm "M_axioms.univalent_is_recfun"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   536
bind_thm ("exists_is_recfun_indstep", axiomsL (thm "M_axioms.exists_is_recfun_indstep"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   537
bind_thm ("wellfounded_exists_is_recfun", axiomsL (thm "M_axioms.wellfounded_exists_is_recfun"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   538
bind_thm ("wf_exists_is_recfun", axiomsL (thm "M_axioms.wf_exists_is_recfun")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   539
bind_thm ("is_recfun_iff_M", axiomsL (thm "M_axioms.is_recfun_iff_M"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   540
bind_thm ("irreflexive_abs", axiomsL (thm "M_axioms.irreflexive_abs"));  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   541
bind_thm ("transitive_rel_abs", axiomsL (thm "M_axioms.transitive_rel_abs"));  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   542
bind_thm ("linear_rel_abs", axiomsL (thm "M_axioms.linear_rel_abs"));  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   543
bind_thm ("wellordered_is_trans_on", axiomsL (thm "M_axioms.wellordered_is_trans_on")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   544
bind_thm ("wellordered_is_linear", axiomsL (thm "M_axioms.wellordered_is_linear")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   545
bind_thm ("wellordered_is_wellfounded_on", axiomsL (thm "M_axioms.wellordered_is_wellfounded_on")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   546
bind_thm ("wellfounded_imp_wellfounded_on", axiomsL (thm "M_axioms.wellfounded_imp_wellfounded_on")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   547
bind_thm ("wellfounded_on_subset_A", axiomsL (thm "M_axioms.wellfounded_on_subset_A"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   548
bind_thm ("wellfounded_on_iff_wellfounded", axiomsL (thm "M_axioms.wellfounded_on_iff_wellfounded"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   549
bind_thm ("wellfounded_on_imp_wellfounded", axiomsL (thm "M_axioms.wellfounded_on_imp_wellfounded"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   550
bind_thm ("wellfounded_on_field_imp_wellfounded", axiomsL (thm "M_axioms.wellfounded_on_field_imp_wellfounded"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   551
bind_thm ("wellfounded_iff_wellfounded_on_field", axiomsL (thm "M_axioms.wellfounded_iff_wellfounded_on_field"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   552
bind_thm ("wellfounded_induct", axiomsL (thm "M_axioms.wellfounded_induct")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   553
bind_thm ("wellfounded_on_induct", axiomsL (thm "M_axioms.wellfounded_on_induct")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   554
bind_thm ("wellfounded_on_induct2", axiomsL (thm "M_axioms.wellfounded_on_induct2")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   555
bind_thm ("linear_imp_relativized", axiomsL (thm "M_axioms.linear_imp_relativized")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   556
bind_thm ("trans_on_imp_relativized", axiomsL (thm "M_axioms.trans_on_imp_relativized")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   557
bind_thm ("wf_on_imp_relativized", axiomsL (thm "M_axioms.wf_on_imp_relativized")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   558
bind_thm ("wf_imp_relativized", axiomsL (thm "M_axioms.wf_imp_relativized")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   559
bind_thm ("well_ord_imp_relativized", axiomsL (thm "M_axioms.well_ord_imp_relativized")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   560
bind_thm ("order_isomorphism_abs", axiomsL (thm "M_axioms.order_isomorphism_abs"));  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   561
bind_thm ("pred_set_abs", axiomsL (thm "M_axioms.pred_set_abs"));  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   562
*}
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   563
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   564
ML
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   565
{*
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   566
bind_thm ("pred_closed", axiomsL (thm "M_axioms.pred_closed"));  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   567
bind_thm ("membership_abs", axiomsL (thm "M_axioms.membership_abs"));  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   568
bind_thm ("M_Memrel_iff", axiomsL (thm "M_axioms.M_Memrel_iff"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   569
bind_thm ("Memrel_closed", axiomsL (thm "M_axioms.Memrel_closed"));  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   570
bind_thm ("wellordered_iso_predD", axiomsL (thm "M_axioms.wellordered_iso_predD"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   571
bind_thm ("wellordered_iso_pred_eq", axiomsL (thm "M_axioms.wellordered_iso_pred_eq"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   572
bind_thm ("wellfounded_on_asym", axiomsL (thm "M_axioms.wellfounded_on_asym"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   573
bind_thm ("wellordered_asym", axiomsL (thm "M_axioms.wellordered_asym"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   574
bind_thm ("ord_iso_pred_imp_lt", axiomsL (thm "M_axioms.ord_iso_pred_imp_lt"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   575
bind_thm ("obase_iff", axiomsL (thm "M_axioms.obase_iff"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   576
bind_thm ("omap_iff", axiomsL (thm "M_axioms.omap_iff"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   577
bind_thm ("omap_unique", axiomsL (thm "M_axioms.omap_unique"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   578
bind_thm ("omap_yields_Ord", axiomsL (thm "M_axioms.omap_yields_Ord"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   579
bind_thm ("otype_iff", axiomsL (thm "M_axioms.otype_iff"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   580
bind_thm ("otype_eq_range", axiomsL (thm "M_axioms.otype_eq_range"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   581
bind_thm ("Ord_otype", axiomsL (thm "M_axioms.Ord_otype"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   582
bind_thm ("domain_omap", axiomsL (thm "M_axioms.domain_omap"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   583
bind_thm ("omap_subset", axiomsL (thm "M_axioms.omap_subset")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   584
bind_thm ("omap_funtype", axiomsL (thm "M_axioms.omap_funtype")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   585
bind_thm ("wellordered_omap_bij", axiomsL (thm "M_axioms.wellordered_omap_bij"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   586
bind_thm ("omap_ord_iso", axiomsL (thm "M_axioms.omap_ord_iso"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   587
bind_thm ("Ord_omap_image_pred", axiomsL (thm "M_axioms.Ord_omap_image_pred"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   588
bind_thm ("restrict_omap_ord_iso", axiomsL (thm "M_axioms.restrict_omap_ord_iso"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   589
bind_thm ("obase_equals", axiomsL (thm "M_axioms.obase_equals")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   590
bind_thm ("omap_ord_iso_otype", axiomsL (thm "M_axioms.omap_ord_iso_otype"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   591
bind_thm ("obase_exists", axiomsL (thm "M_axioms.obase_exists"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   592
bind_thm ("omap_exists", axiomsL (thm "M_axioms.omap_exists"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   593
bind_thm ("otype_exists", axiomsL (thm "M_axioms.otype_exists"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   594
bind_thm ("omap_ord_iso_otype", axiomsL (thm "M_axioms.omap_ord_iso_otype"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   595
bind_thm ("ordertype_exists", axiomsL (thm "M_axioms.ordertype_exists"));
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   596
bind_thm ("relativized_imp_well_ord", axiomsL (thm "M_axioms.relativized_imp_well_ord")); 
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   597
bind_thm ("well_ord_abs", axiomsL (thm "M_axioms.well_ord_abs"));  
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   598
*}
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   599
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   600
declare cartprod_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   601
declare sum_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   602
declare converse_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   603
declare converse_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   604
declare image_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   605
declare vimage_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   606
declare vimage_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   607
declare domain_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   608
declare domain_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   609
declare range_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   610
declare range_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   611
declare field_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   612
declare field_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   613
declare relation_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   614
declare function_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   615
declare apply_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   616
declare typed_function_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   617
declare injection_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   618
declare surjection_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   619
declare bijection_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   620
declare comp_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   621
declare composition_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   622
declare restriction_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   623
declare restrict_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   624
declare Inter_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   625
declare Inter_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   626
declare Int_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   627
declare finite_fun_closed [rule_format]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   628
declare is_funspace_abs [simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   629
declare finite_funspace_closed [intro,simp]
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   630
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents:
diff changeset
   631
end