src/ZF/Constructible/Relative.thy
author paulson
Mon, 24 Jun 2002 11:59:21 +0200
changeset 13245 714f7a423a15
parent 13223 45be08fbdcff
child 13247 e3c289f0724b
permissions -rw-r--r--
development and tweaks
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     1
header {*Relativization and Absoluteness*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     2
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     3
theory Relative = Main:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     4
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     5
subsection{* Relativized versions of standard set-theoretic concepts *}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     6
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     7
constdefs
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     8
  empty :: "[i=>o,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     9
    "empty(M,z) == \<forall>x. M(x) --> x \<notin> z"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    10
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    11
  subset :: "[i=>o,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    12
    "subset(M,A,B) == \<forall>x\<in>A. M(x) --> x \<in> B"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    13
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    14
  upair :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    15
    "upair(M,a,b,z) == a \<in> z & b \<in> z & (\<forall>x\<in>z. M(x) --> x = a | x = b)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    16
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    17
  pair :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    18
    "pair(M,a,b,z) == \<exists>x. M(x) & upair(M,a,a,x) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    19
                          (\<exists>y. M(y) & upair(M,a,b,y) & upair(M,x,y,z))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    20
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    21
  union :: "[i=>o,i,i,i] => o"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    22
    "union(M,a,b,z) == \<forall>x. M(x) --> (x \<in> z <-> x \<in> a | x \<in> b)"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    23
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    24
  successor :: "[i=>o,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    25
    "successor(M,a,z) == \<exists>x. M(x) & upair(M,a,a,x) & union(M,x,a,z)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    26
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    27
  powerset :: "[i=>o,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    28
    "powerset(M,A,z) == \<forall>x. M(x) --> (x \<in> z <-> subset(M,x,A))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    29
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    30
  inter :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    31
    "inter(M,a,b,z) == \<forall>x. M(x) --> (x \<in> z <-> x \<in> a & x \<in> b)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    32
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    33
  setdiff :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    34
    "setdiff(M,a,b,z) == \<forall>x. M(x) --> (x \<in> z <-> x \<in> a & x \<notin> b)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    35
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    36
  big_union :: "[i=>o,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    37
    "big_union(M,A,z) == \<forall>x. M(x) --> (x \<in> z <-> (\<exists>y\<in>A. M(y) & x \<in> y))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    38
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    39
  big_inter :: "[i=>o,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    40
    "big_inter(M,A,z) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    41
             (A=0 --> z=0) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    42
	     (A\<noteq>0 --> (\<forall>x. M(x) --> (x \<in> z <-> (\<forall>y\<in>A. M(y) --> x \<in> y))))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    43
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    44
  cartprod :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    45
    "cartprod(M,A,B,z) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    46
	\<forall>u. M(u) --> (u \<in> z <-> (\<exists>x\<in>A. M(x) & (\<exists>y\<in>B. M(y) & pair(M,x,y,u))))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    47
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    48
  is_converse :: "[i=>o,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    49
    "is_converse(M,r,z) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    50
	\<forall>x. M(x) --> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    51
            (x \<in> z <-> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    52
             (\<exists>w\<in>r. M(w) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    53
              (\<exists>u v. M(u) & M(v) & pair(M,u,v,w) & pair(M,v,u,x))))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    54
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    55
  pre_image :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    56
    "pre_image(M,r,A,z) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    57
	\<forall>x. M(x) --> (x \<in> z <-> (\<exists>w\<in>r. M(w) & (\<exists>y\<in>A. M(y) & pair(M,x,y,w))))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    58
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    59
  is_domain :: "[i=>o,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    60
    "is_domain(M,r,z) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    61
	\<forall>x. M(x) --> (x \<in> z <-> (\<exists>w\<in>r. M(w) & (\<exists>y. M(y) & pair(M,x,y,w))))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    62
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    63
  image :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    64
    "image(M,r,A,z) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    65
        \<forall>y. M(y) --> (y \<in> z <-> (\<exists>w\<in>r. M(w) & (\<exists>x\<in>A. M(x) & pair(M,x,y,w))))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    66
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    67
  is_range :: "[i=>o,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    68
    --{*the cleaner 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    69
      @{term "\<exists>r'. M(r') & is_converse(M,r,r') & is_domain(M,r',z)"}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    70
      unfortunately needs an instance of separation in order to prove 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    71
        @{term "M(converse(r))"}.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    72
    "is_range(M,r,z) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    73
	\<forall>y. M(y) --> (y \<in> z <-> (\<exists>w\<in>r. M(w) & (\<exists>x. M(x) & pair(M,x,y,w))))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    74
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    75
  is_field :: "[i=>o,i,i] => o"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    76
    "is_field(M,r,z) == 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    77
	\<exists>dr. M(dr) & is_domain(M,r,dr) & 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    78
            (\<exists>rr. M(rr) & is_range(M,r,rr) & union(M,dr,rr,z))"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    79
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    80
  is_relation :: "[i=>o,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    81
    "is_relation(M,r) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    82
        (\<forall>z\<in>r. M(z) --> (\<exists>x y. M(x) & M(y) & pair(M,x,y,z)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    83
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    84
  is_function :: "[i=>o,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    85
    "is_function(M,r) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    86
	(\<forall>x y y' p p'. M(x) --> M(y) --> M(y') --> M(p) --> M(p') --> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    87
                      pair(M,x,y,p) --> pair(M,x,y',p') --> p\<in>r --> p'\<in>r --> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    88
                      y=y')"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    89
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    90
  fun_apply :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    91
    "fun_apply(M,f,x,y) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    92
	(\<forall>y'. M(y') --> ((\<exists>u\<in>f. M(u) & pair(M,x,y',u)) <-> y=y'))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    93
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    94
  typed_function :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    95
    "typed_function(M,A,B,r) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    96
        is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    97
        (\<forall>u\<in>r. M(u) --> (\<forall>x y. M(x) & M(y) & pair(M,x,y,u) --> y\<in>B))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    98
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    99
  composition :: "[i=>o,i,i,i] => o"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   100
    "composition(M,r,s,t) == 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   101
        \<forall>p. M(p) --> (p \<in> t <-> 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   102
                      (\<exists>x. M(x) & (\<exists>y. M(y) & (\<exists>z. M(z) & 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   103
                           p = \<langle>x,z\<rangle> & \<langle>x,y\<rangle> \<in> s & \<langle>y,z\<rangle> \<in> r))))"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   104
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   105
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   106
  injection :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   107
    "injection(M,A,B,f) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   108
	typed_function(M,A,B,f) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   109
        (\<forall>x x' y p p'. M(x) --> M(x') --> M(y) --> M(p) --> M(p') --> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   110
                      pair(M,x,y,p) --> pair(M,x',y,p') --> p\<in>f --> p'\<in>f --> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   111
                      x=x')"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   112
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   113
  surjection :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   114
    "surjection(M,A,B,f) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   115
        typed_function(M,A,B,f) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   116
        (\<forall>y\<in>B. M(y) --> (\<exists>x\<in>A. M(x) & fun_apply(M,f,x,y)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   117
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   118
  bijection :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   119
    "bijection(M,A,B,f) == injection(M,A,B,f) & surjection(M,A,B,f)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   120
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   121
  restriction :: "[i=>o,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   122
    "restriction(M,r,A,z) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   123
	\<forall>x. M(x) --> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   124
            (x \<in> z <-> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   125
             (x \<in> r & (\<exists>u\<in>A. M(u) & (\<exists>v. M(v) & pair(M,u,v,x)))))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   126
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   127
  transitive_set :: "[i=>o,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   128
    "transitive_set(M,a) == \<forall>x\<in>a. M(x) --> subset(M,x,a)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   129
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   130
  ordinal :: "[i=>o,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   131
     --{*an ordinal is a transitive set of transitive sets*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   132
    "ordinal(M,a) == transitive_set(M,a) & (\<forall>x\<in>a. M(x) --> transitive_set(M,x))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   133
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   134
  limit_ordinal :: "[i=>o,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   135
    --{*a limit ordinal is a non-empty, successor-closed ordinal*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   136
    "limit_ordinal(M,a) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   137
	ordinal(M,a) & ~ empty(M,a) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   138
        (\<forall>x\<in>a. M(x) --> (\<exists>y\<in>a. M(y) & successor(M,x,y)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   139
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   140
  successor_ordinal :: "[i=>o,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   141
    --{*a successor ordinal is any ordinal that is neither empty nor limit*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   142
    "successor_ordinal(M,a) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   143
	ordinal(M,a) & ~ empty(M,a) & ~ limit_ordinal(M,a)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   144
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   145
  finite_ordinal :: "[i=>o,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   146
    --{*an ordinal is finite if neither it nor any of its elements are limit*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   147
    "finite_ordinal(M,a) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   148
	ordinal(M,a) & ~ limit_ordinal(M,a) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   149
        (\<forall>x\<in>a. M(x) --> ~ limit_ordinal(M,x))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   150
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   151
  omega :: "[i=>o,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   152
    --{*omega is a limit ordinal none of whose elements are limit*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   153
    "omega(M,a) == limit_ordinal(M,a) & (\<forall>x\<in>a. M(x) --> ~ limit_ordinal(M,x))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   154
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   155
  number1 :: "[i=>o,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   156
    "number1(M,a) == (\<exists>x. M(x) & empty(M,x) & successor(M,x,a))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   157
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   158
  number2 :: "[i=>o,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   159
    "number2(M,a) == (\<exists>x. M(x) & number1(M,x) & successor(M,x,a))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   160
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   161
  number3 :: "[i=>o,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   162
    "number3(M,a) == (\<exists>x. M(x) & number2(M,x) & successor(M,x,a))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   163
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   164
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   165
subsection {*The relativized ZF axioms*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   166
constdefs
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   167
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   168
  extensionality :: "(i=>o) => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   169
    "extensionality(M) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   170
	\<forall>x y. M(x) --> M(y) --> (\<forall>z. M(z) --> (z \<in> x <-> z \<in> y)) --> x=y"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   171
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   172
  separation :: "[i=>o, i=>o] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   173
    --{*Big problem: the formula @{text P} should only involve parameters
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   174
        belonging to @{text M}.  Don't see how to enforce that.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   175
    "separation(M,P) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   176
	\<forall>z. M(z) --> (\<exists>y. M(y) & (\<forall>x. M(x) --> (x \<in> y <-> x \<in> z & P(x))))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   177
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   178
  upair_ax :: "(i=>o) => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   179
    "upair_ax(M) == \<forall>x y. M(x) --> M(y) --> (\<exists>z. M(z) & upair(M,x,y,z))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   180
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   181
  Union_ax :: "(i=>o) => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   182
    "Union_ax(M) == \<forall>x. M(x) --> (\<exists>z. M(z) & big_union(M,x,z))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   183
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   184
  power_ax :: "(i=>o) => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   185
    "power_ax(M) == \<forall>x. M(x) --> (\<exists>z. M(z) & powerset(M,x,z))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   186
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   187
  univalent :: "[i=>o, i, [i,i]=>o] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   188
    "univalent(M,A,P) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   189
	(\<forall>x\<in>A. M(x) --> (\<forall>y z. M(y) --> M(z) --> P(x,y) & P(x,z) --> y=z))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   190
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   191
  replacement :: "[i=>o, [i,i]=>o] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   192
    "replacement(M,P) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   193
      \<forall>A. M(A) --> univalent(M,A,P) -->
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   194
      (\<exists>Y. M(Y) & (\<forall>b. M(b) --> ((\<exists>x\<in>A. M(x) & P(x,b)) --> b \<in> Y)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   195
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   196
  strong_replacement :: "[i=>o, [i,i]=>o] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   197
    "strong_replacement(M,P) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   198
      \<forall>A. M(A) --> univalent(M,A,P) -->
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   199
      (\<exists>Y. M(Y) & (\<forall>b. M(b) --> (b \<in> Y <-> (\<exists>x\<in>A. M(x) & P(x,b)))))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   200
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   201
  foundation_ax :: "(i=>o) => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   202
    "foundation_ax(M) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   203
	\<forall>x. M(x) --> (\<exists>y\<in>x. M(y))
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   204
                 --> (\<exists>y\<in>x. M(y) & ~(\<exists>z\<in>x. M(z) & z \<in> y))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   205
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   206
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   207
subsection{*A trivial consistency proof for $V_\omega$ *}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   208
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   209
text{*We prove that $V_\omega$ 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   210
      (or @{text univ} in Isabelle) satisfies some ZF axioms.
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   211
     Kunen, Theorem IV 3.13, page 123.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   212
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   213
lemma univ0_downwards_mem: "[| y \<in> x; x \<in> univ(0) |] ==> y \<in> univ(0)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   214
apply (insert Transset_univ [OF Transset_0])  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   215
apply (simp add: Transset_def, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   216
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   217
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   218
lemma univ0_Ball_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   219
     "A \<in> univ(0) ==> (\<forall>x\<in>A. x \<in> univ(0) --> P(x)) <-> (\<forall>x\<in>A. P(x))" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   220
by (blast intro: univ0_downwards_mem) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   221
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   222
lemma univ0_Bex_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   223
     "A \<in> univ(0) ==> (\<exists>x\<in>A. x \<in> univ(0) & P(x)) <-> (\<exists>x\<in>A. P(x))" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   224
by (blast intro: univ0_downwards_mem) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   225
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   226
text{*Congruence rule for separation: can assume the variable is in @{text M}*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   227
lemma [cong]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   228
     "(!!x. M(x) ==> P(x) <-> P'(x)) ==> separation(M,P) <-> separation(M,P')"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   229
by (simp add: separation_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   230
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   231
text{*Congruence rules for replacement*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   232
lemma [cong]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   233
     "[| A=A'; !!x y. [| x\<in>A; M(x); M(y) |] ==> P(x,y) <-> P'(x,y) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   234
      ==> univalent(M,A,P) <-> univalent(M,A',P')"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   235
by (simp add: univalent_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   236
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   237
lemma [cong]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   238
     "[| !!x y. [| M(x); M(y) |] ==> P(x,y) <-> P'(x,y) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   239
      ==> strong_replacement(M,P) <-> strong_replacement(M,P')" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   240
by (simp add: strong_replacement_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   241
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   242
text{*The extensionality axiom*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   243
lemma "extensionality(\<lambda>x. x \<in> univ(0))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   244
apply (simp add: extensionality_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   245
apply (blast intro: univ0_downwards_mem) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   246
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   247
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   248
text{*The separation axiom requires some lemmas*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   249
lemma Collect_in_Vfrom:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   250
     "[| X \<in> Vfrom(A,j);  Transset(A) |] ==> Collect(X,P) \<in> Vfrom(A, succ(j))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   251
apply (drule Transset_Vfrom)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   252
apply (rule subset_mem_Vfrom)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   253
apply (unfold Transset_def, blast)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   254
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   255
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   256
lemma Collect_in_VLimit:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   257
     "[| X \<in> Vfrom(A,i);  Limit(i);  Transset(A) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   258
      ==> Collect(X,P) \<in> Vfrom(A,i)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   259
apply (rule Limit_VfromE, assumption+)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   260
apply (blast intro: Limit_has_succ VfromI Collect_in_Vfrom)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   261
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   262
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   263
lemma Collect_in_univ:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   264
     "[| X \<in> univ(A);  Transset(A) |] ==> Collect(X,P) \<in> univ(A)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   265
by (simp add: univ_def Collect_in_VLimit Limit_nat)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   266
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   267
lemma "separation(\<lambda>x. x \<in> univ(0), P)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   268
apply (simp add: separation_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   269
apply (blast intro: Collect_in_univ Transset_0) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   270
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   271
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   272
text{*Unordered pairing axiom*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   273
lemma "upair_ax(\<lambda>x. x \<in> univ(0))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   274
apply (simp add: upair_ax_def upair_def)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   275
apply (blast intro: doubleton_in_univ) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   276
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   277
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   278
text{*Union axiom*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   279
lemma "Union_ax(\<lambda>x. x \<in> univ(0))"  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   280
apply (simp add: Union_ax_def big_union_def)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   281
apply (blast intro: Union_in_univ Transset_0 univ0_downwards_mem) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   282
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   283
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   284
text{*Powerset axiom*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   285
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   286
lemma Pow_in_univ:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   287
     "[| X \<in> univ(A);  Transset(A) |] ==> Pow(X) \<in> univ(A)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   288
apply (simp add: univ_def Pow_in_VLimit Limit_nat)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   289
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   290
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   291
lemma "power_ax(\<lambda>x. x \<in> univ(0))"  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   292
apply (simp add: power_ax_def powerset_def subset_def)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   293
apply (blast intro: Pow_in_univ Transset_0 univ0_downwards_mem) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   294
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   295
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   296
text{*Foundation axiom*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   297
lemma "foundation_ax(\<lambda>x. x \<in> univ(0))"  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   298
apply (simp add: foundation_ax_def, clarify)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   299
apply (cut_tac A=x in foundation, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   300
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   301
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   302
lemma "replacement(\<lambda>x. x \<in> univ(0), P)"  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   303
apply (simp add: replacement_def, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   304
oops
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   305
text{*no idea: maybe prove by induction on the rank of A?*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   306
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   307
text{*Still missing: Replacement, Choice*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   308
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   309
subsection{*lemmas needed to reduce some set constructions to instances
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   310
      of Separation*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   311
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   312
lemma image_iff_Collect: "r `` A = {y \<in> Union(Union(r)). \<exists>p\<in>r. \<exists>x\<in>A. p=<x,y>}"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   313
apply (rule equalityI, auto) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   314
apply (simp add: Pair_def, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   315
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   316
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   317
lemma vimage_iff_Collect:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   318
     "r -`` A = {x \<in> Union(Union(r)). \<exists>p\<in>r. \<exists>y\<in>A. p=<x,y>}"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   319
apply (rule equalityI, auto) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   320
apply (simp add: Pair_def, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   321
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   322
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   323
text{*These two lemmas lets us prove @{text domain_closed} and 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   324
      @{text range_closed} without new instances of separation*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   325
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   326
lemma domain_eq_vimage: "domain(r) = r -`` Union(Union(r))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   327
apply (rule equalityI, auto)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   328
apply (rule vimageI, assumption)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   329
apply (simp add: Pair_def, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   330
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   331
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   332
lemma range_eq_image: "range(r) = r `` Union(Union(r))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   333
apply (rule equalityI, auto)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   334
apply (rule imageI, assumption)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   335
apply (simp add: Pair_def, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   336
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   337
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   338
lemma replacementD:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   339
    "[| replacement(M,P); M(A);  univalent(M,A,P) |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   340
     ==> \<exists>Y. M(Y) & (\<forall>b. M(b) --> ((\<exists>x\<in>A. M(x) & P(x,b)) --> b \<in> Y))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   341
by (simp add: replacement_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   342
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   343
lemma strong_replacementD:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   344
    "[| strong_replacement(M,P); M(A);  univalent(M,A,P) |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   345
     ==> \<exists>Y. M(Y) & (\<forall>b. M(b) --> (b \<in> Y <-> (\<exists>x\<in>A. M(x) & P(x,b))))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   346
by (simp add: strong_replacement_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   347
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   348
lemma separationD:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   349
    "[| separation(M,P); M(z) |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   350
     ==> \<exists>y. M(y) & (\<forall>x. M(x) --> (x \<in> y <-> x \<in> z & P(x)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   351
by (simp add: separation_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   352
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   353
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   354
text{*More constants, for order types*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   355
constdefs
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   356
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   357
  order_isomorphism :: "[i=>o,i,i,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   358
    "order_isomorphism(M,A,r,B,s,f) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   359
        bijection(M,A,B,f) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   360
        (\<forall>x\<in>A. \<forall>y\<in>A. \<forall>p fx fy q. 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   361
            M(x) --> M(y) --> M(p) --> M(fx) --> M(fy) --> M(q) --> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   362
            pair(M,x,y,p) --> fun_apply(M,f,x,fx) --> fun_apply(M,f,y,fy) --> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   363
            pair(M,fx,fy,q) --> (p\<in>r <-> q\<in>s))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   364
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   365
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   366
  pred_set :: "[i=>o,i,i,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   367
    "pred_set(M,A,x,r,B) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   368
	\<forall>y. M(y) --> (y \<in> B <-> (\<exists>p\<in>r. M(p) & y \<in> A & pair(M,y,x,p)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   369
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   370
  membership :: "[i=>o,i,i] => o" --{*membership relation*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   371
    "membership(M,A,r) == 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   372
	\<forall>p. M(p) --> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   373
             (p \<in> r <-> (\<exists>x\<in>A. \<exists>y\<in>A. M(x) & M(y) & x\<in>y & pair(M,x,y,p)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   374
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   375
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   376
subsection{*Absoluteness for a transitive class model*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   377
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   378
text{*The class M is assumed to be transitive and to satisfy some
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   379
      relativized ZF axioms*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   380
locale M_axioms =
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   381
  fixes M
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   382
  assumes transM:           "[| y\<in>x; M(x) |] ==> M(y)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   383
      and nonempty [simp]:  "M(0)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   384
      and upair_ax:	    "upair_ax(M)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   385
      and Union_ax:	    "Union_ax(M)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   386
      and power_ax:         "power_ax(M)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   387
      and replacement:      "replacement(M,P)"
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   388
      and M_nat:            "M(nat)"   (*i.e. the axiom of infinity*)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   389
  and Inter_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   390
     "M(A) ==> separation(M, \<lambda>x. \<forall>y\<in>A. M(y) --> x\<in>y)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   391
  and cartprod_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   392
     "[| M(A); M(B) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   393
      ==> separation(M, \<lambda>z. \<exists>x\<in>A. \<exists>y\<in>B. M(x) & M(y) & pair(M,x,y,z))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   394
  and image_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   395
     "[| M(A); M(r) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   396
      ==> separation(M, \<lambda>y. \<exists>p\<in>r. M(p) & (\<exists>x\<in>A. M(x) & pair(M,x,y,p)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   397
  and vimage_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   398
     "[| M(A); M(r) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   399
      ==> separation(M, \<lambda>x. \<exists>p\<in>r. M(p) & (\<exists>y\<in>A. M(x) & pair(M,x,y,p)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   400
  and converse_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   401
     "M(r) ==> separation(M, \<lambda>z. \<exists>p\<in>r. M(p) & (\<exists>x y. M(x) & M(y) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   402
				     pair(M,x,y,p) & pair(M,y,x,z)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   403
  and restrict_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   404
     "M(A) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   405
      ==> separation(M, \<lambda>z. \<exists>x\<in>A. M(x) & (\<exists>y. M(y) & pair(M,x,y,z)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   406
  and comp_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   407
     "[| M(r); M(s) |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   408
      ==> separation(M, \<lambda>xz. \<exists>x y z. M(x) & M(y) & M(z) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   409
			   (\<exists>xy\<in>s. \<exists>yz\<in>r. M(xy) & M(yz) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   410
		  pair(M,x,z,xz) & pair(M,x,y,xy) & pair(M,y,z,yz)))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   411
  and pred_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   412
     "[| M(r); M(x) |] ==> separation(M, \<lambda>y. \<exists>p\<in>r. M(p) & pair(M,y,x,p))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   413
  and Memrel_separation:
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   414
     "separation(M, \<lambda>z. \<exists>x y. M(x) & M(y) & pair(M,x,y,z) & x \<in> y)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   415
  and obase_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   416
     --{*part of the order type formalization*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   417
     "[| M(A); M(r) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   418
      ==> separation(M, \<lambda>a. \<exists>x g mx par. M(x) & M(g) & M(mx) & M(par) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   419
	     ordinal(M,x) & membership(M,x,mx) & pred_set(M,A,a,r,par) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   420
	     order_isomorphism(M,par,r,x,mx,g))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   421
  and well_ord_iso_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   422
     "[| M(A); M(f); M(r) |] 
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   423
      ==> separation (M, \<lambda>x. x\<in>A --> (\<exists>y. M(y) & (\<exists>p. M(p) & 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   424
		     fun_apply(M,f,x,y) & pair(M,y,x,p) & p \<in> r)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   425
  and obase_equals_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   426
     "[| M(A); M(r) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   427
      ==> separation
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   428
      (M, \<lambda>x. x\<in>A --> ~(\<exists>y. M(y) & (\<exists>g. M(g) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   429
	      ordinal(M,y) & (\<exists>my pxr. M(my) & M(pxr) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   430
	      membership(M,y,my) & pred_set(M,A,x,r,pxr) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   431
	      order_isomorphism(M,pxr,r,y,my,g)))))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   432
  and is_recfun_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   433
     --{*for well-founded recursion.  NEEDS RELATIVIZATION*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   434
     "[| M(A); M(f); M(g); M(a); M(b) |] 
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   435
     ==> separation(M, \<lambda>x. x \<in> A --> \<langle>x,a\<rangle> \<in> r & \<langle>x,b\<rangle> \<in> r & f`x \<noteq> g`x)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   436
  and omap_replacement:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   437
     "[| M(A); M(r) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   438
      ==> strong_replacement(M,
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   439
             \<lambda>a z. \<exists>x g mx par. M(x) & M(g) & M(mx) & M(par) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   440
	     ordinal(M,x) & pair(M,a,x,z) & membership(M,x,mx) & 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   441
	     pred_set(M,A,a,r,par) & order_isomorphism(M,par,r,x,mx,g))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   442
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   443
lemma (in M_axioms) Ball_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   444
     "M(A) ==> (\<forall>x\<in>A. M(x) --> P(x)) <-> (\<forall>x\<in>A. P(x))" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   445
by (blast intro: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   446
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   447
lemma (in M_axioms) Bex_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   448
     "M(A) ==> (\<exists>x\<in>A. M(x) & P(x)) <-> (\<exists>x\<in>A. P(x))" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   449
by (blast intro: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   450
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   451
lemma (in M_axioms) Ball_iff_equiv: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   452
     "M(A) ==> (\<forall>x. M(x) --> (x\<in>A <-> P(x))) <-> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   453
               (\<forall>x\<in>A. P(x)) & (\<forall>x. P(x) --> M(x) --> x\<in>A)" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   454
by (blast intro: transM)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   455
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   456
text{*Simplifies proofs of equalities when there's an iff-equality
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   457
      available for rewriting, universally quantified over M. *}
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   458
lemma (in M_axioms) M_equalityI: 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   459
     "[| !!x. M(x) ==> x\<in>A <-> x\<in>B; M(A); M(B) |] ==> A=B"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   460
by (blast intro!: equalityI dest: transM) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   461
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   462
lemma (in M_axioms) empty_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   463
     "M(z) ==> empty(M,z) <-> z=0"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   464
apply (simp add: empty_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   465
apply (blast intro: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   466
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   467
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   468
lemma (in M_axioms) subset_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   469
     "M(A) ==> subset(M,A,B) <-> A \<subseteq> B"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   470
apply (simp add: subset_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   471
apply (blast intro: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   472
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   473
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   474
lemma (in M_axioms) upair_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   475
     "M(z) ==> upair(M,a,b,z) <-> z={a,b}"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   476
apply (simp add: upair_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   477
apply (blast intro: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   478
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   479
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   480
lemma (in M_axioms) upair_in_M_iff [iff]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   481
     "M({a,b}) <-> M(a) & M(b)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   482
apply (insert upair_ax, simp add: upair_ax_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   483
apply (blast intro: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   484
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   485
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   486
lemma (in M_axioms) singleton_in_M_iff [iff]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   487
     "M({a}) <-> M(a)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   488
by (insert upair_in_M_iff [of a a], simp) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   489
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   490
lemma (in M_axioms) pair_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   491
     "M(z) ==> pair(M,a,b,z) <-> z=<a,b>"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   492
apply (simp add: pair_def ZF.Pair_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   493
apply (blast intro: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   494
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   495
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   496
lemma (in M_axioms) pair_in_M_iff [iff]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   497
     "M(<a,b>) <-> M(a) & M(b)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   498
by (simp add: ZF.Pair_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   499
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   500
lemma (in M_axioms) pair_components_in_M:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   501
     "[| <x,y> \<in> A; M(A) |] ==> M(x) & M(y)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   502
apply (simp add: Pair_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   503
apply (blast dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   504
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   505
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   506
lemma (in M_axioms) cartprod_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   507
     "[| M(A); M(B); M(z) |] ==> cartprod(M,A,B,z) <-> z = A*B"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   508
apply (simp add: cartprod_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   509
apply (rule iffI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   510
apply (blast intro!: equalityI intro: transM dest!: spec) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   511
apply (blast dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   512
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   513
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   514
lemma (in M_axioms) union_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   515
     "[| M(a); M(b); M(z) |] ==> union(M,a,b,z) <-> z = a Un b"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   516
apply (simp add: union_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   517
apply (blast intro: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   518
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   519
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   520
lemma (in M_axioms) inter_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   521
     "[| M(a); M(b); M(z) |] ==> inter(M,a,b,z) <-> z = a Int b"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   522
apply (simp add: inter_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   523
apply (blast intro: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   524
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   525
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   526
lemma (in M_axioms) setdiff_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   527
     "[| M(a); M(b); M(z) |] ==> setdiff(M,a,b,z) <-> z = a-b"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   528
apply (simp add: setdiff_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   529
apply (blast intro: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   530
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   531
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   532
lemma (in M_axioms) Union_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   533
     "[| M(A); M(z) |] ==> big_union(M,A,z) <-> z = Union(A)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   534
apply (simp add: big_union_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   535
apply (blast intro!: equalityI dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   536
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   537
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   538
lemma (in M_axioms) Union_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   539
     "M(A) ==> M(Union(A))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   540
by (insert Union_ax, simp add: Union_ax_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   541
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   542
lemma (in M_axioms) Un_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   543
     "[| M(A); M(B) |] ==> M(A Un B)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   544
by (simp only: Un_eq_Union, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   545
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   546
lemma (in M_axioms) cons_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   547
     "[| M(a); M(A) |] ==> M(cons(a,A))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   548
by (subst cons_eq [symmetric], blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   549
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   550
lemma (in M_axioms) successor_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   551
     "[| M(a); M(z) |] ==> successor(M,a,z) <-> z=succ(a)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   552
by (simp add: successor_def, blast)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   553
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   554
lemma (in M_axioms) succ_in_M_iff [iff]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   555
     "M(succ(a)) <-> M(a)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   556
apply (simp add: succ_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   557
apply (blast intro: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   558
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   559
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   560
lemma (in M_axioms) separation_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   561
     "[| separation(M,P); M(A) |] ==> M(Collect(A,P))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   562
apply (insert separation, simp add: separation_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   563
apply (drule spec [THEN mp], assumption, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   564
apply (subgoal_tac "y = Collect(A,P)", blast)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   565
apply (blast dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   566
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   567
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   568
text{*Probably the premise and conclusion are equivalent*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   569
lemma (in M_axioms) strong_replacementI [rule_format]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   570
    "[| \<forall>A. M(A) --> separation(M, %u. \<exists>x\<in>A. P(x,u)) |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   571
     ==> strong_replacement(M,P)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   572
apply (simp add: strong_replacement_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   573
apply (clarify ); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   574
apply (frule replacementD [OF replacement]) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   575
apply assumption
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   576
apply (clarify ); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   577
apply (drule_tac x=A in spec)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   578
apply (clarify );  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   579
apply (drule_tac z=Y in separationD) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   580
apply assumption; 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   581
apply (clarify ); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   582
apply (blast dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   583
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   584
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   585
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   586
(*The last premise expresses that P takes M to M*)
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   587
lemma (in M_axioms) strong_replacement_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   588
     "[| strong_replacement(M,P); M(A); univalent(M,A,P); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   589
       !!x y. [| P(x,y); M(x) |] ==> M(y) |] ==> M(Replace(A,P))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   590
apply (simp add: strong_replacement_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   591
apply (drule spec [THEN mp], auto) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   592
apply (subgoal_tac "Replace(A,P) = Y")
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   593
 apply (simp add: ); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   594
apply (rule equality_iffI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   595
apply (simp add: Replace_iff) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   596
apply safe;
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   597
 apply (blast dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   598
apply (frule transM, assumption) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   599
 apply (simp add: univalent_def);
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   600
 apply (drule spec [THEN mp, THEN iffD1], assumption, assumption)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   601
 apply (blast dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   602
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   603
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   604
(*The first premise can't simply be assumed as a schema.
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   605
  It is essential to take care when asserting instances of Replacement.
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   606
  Let K be a nonconstructible subset of nat and define
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   607
  f(x) = x if x:K and f(x)=0 otherwise.  Then RepFun(nat,f) = cons(0,K), a 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   608
  nonconstructible set.  So we cannot assume that M(X) implies M(RepFun(X,f))
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   609
  even for f : M -> M.
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   610
*)
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   611
lemma (in M_axioms) RepFun_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   612
     "[| strong_replacement(M, \<lambda>x y. y = f(x)); M(A); \<forall>x. M(x) --> M(f(x)) |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   613
      ==> M(RepFun(A,f))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   614
apply (simp add: RepFun_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   615
apply (rule strong_replacement_closed) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   616
apply (auto dest: transM  simp add: univalent_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   617
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   618
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   619
lemma (in M_axioms) converse_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   620
     "[| M(r); M(z) |] ==> is_converse(M,r,z) <-> z = converse(r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   621
apply (simp add: is_converse_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   622
apply (rule iffI)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   623
 apply (rule equalityI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   624
  apply (blast dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   625
 apply (clarify, frule transM, assumption, simp, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   626
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   627
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   628
lemma (in M_axioms) image_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   629
     "[| M(r); M(A); M(z) |] ==> image(M,r,A,z) <-> z = r``A"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   630
apply (simp add: image_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   631
apply (rule iffI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   632
 apply (blast intro!: equalityI dest: transM, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   633
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   634
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   635
text{*What about @{text Pow_abs}?  Powerset is NOT absolute!
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   636
      This result is one direction of absoluteness.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   637
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   638
lemma (in M_axioms) powerset_Pow: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   639
     "powerset(M, x, Pow(x))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   640
by (simp add: powerset_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   641
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   642
text{*But we can't prove that the powerset in @{text M} includes the
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   643
      real powerset.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   644
lemma (in M_axioms) powerset_imp_subset_Pow: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   645
     "[| powerset(M,x,y); M(y) |] ==> y <= Pow(x)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   646
apply (simp add: powerset_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   647
apply (blast dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   648
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   649
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   650
lemma (in M_axioms) cartprod_iff_lemma:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   651
     "[| M(C); \<forall>u. M(u) --> u \<in> C <-> (\<exists>x\<in>A. \<exists>y\<in>B. u = {{x}, {x,y}}); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   652
       powerset(M, A \<union> B, p1); powerset(M, p1, p2); M(p2) |]
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   653
       ==> C = {u \<in> p2 . \<exists>x\<in>A. \<exists>y\<in>B. u = {{x}, {x,y}}}"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   654
apply (simp add: powerset_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   655
apply (rule equalityI, clarify, simp) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   656
 apply (frule transM, assumption, simp) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   657
 apply blast 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   658
apply clarify
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   659
apply (frule transM, assumption, force) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   660
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   661
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   662
lemma (in M_axioms) cartprod_iff:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   663
     "[| M(A); M(B); M(C) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   664
      ==> cartprod(M,A,B,C) <-> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   665
          (\<exists>p1 p2. M(p1) & M(p2) & powerset(M,A Un B,p1) & powerset(M,p1,p2) &
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   666
                   C = {z \<in> p2. \<exists>x\<in>A. \<exists>y\<in>B. z = <x,y>})"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   667
apply (simp add: Pair_def cartprod_def, safe)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   668
defer 1 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   669
  apply (simp add: powerset_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   670
 apply blast 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   671
txt{*Final, difficult case: the left-to-right direction of the theorem.*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   672
apply (insert power_ax, simp add: power_ax_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   673
apply (frule_tac x="A Un B" and P="\<lambda>x. M(x) --> Ex(?Q(x))" in spec) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   674
apply (erule impE, blast, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   675
apply (drule_tac x=z and P="\<lambda>x. M(x) --> Ex(?Q(x))" in spec) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   676
apply (blast intro: cartprod_iff_lemma) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   677
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   678
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   679
lemma (in M_axioms) cartprod_closed_lemma:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   680
     "[| M(A); M(B) |] ==> \<exists>C. M(C) & cartprod(M,A,B,C)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   681
apply (simp del: cartprod_abs add: cartprod_iff)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   682
apply (insert power_ax, simp add: power_ax_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   683
apply (frule_tac x="A Un B" and P="\<lambda>x. M(x) --> Ex(?Q(x))" in spec) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   684
apply (erule impE, blast, clarify) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   685
apply (drule_tac x=z and P="\<lambda>x. M(x) --> Ex(?Q(x))" in spec) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   686
apply (erule impE, blast, clarify)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   687
apply (intro exI conjI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   688
prefer 6 apply (rule refl) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   689
prefer 4 apply assumption
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   690
prefer 4 apply assumption
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   691
apply (insert cartprod_separation [of A B], auto)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   692
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   693
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   694
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   695
text{*All the lemmas above are necessary because Powerset is not absolute.
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   696
      I should have used Replacement instead!*}
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   697
lemma (in M_axioms) cartprod_closed [intro,simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   698
     "[| M(A); M(B) |] ==> M(A*B)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   699
by (frule cartprod_closed_lemma, assumption, force)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   700
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   701
lemma (in M_axioms) image_closed [intro,simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   702
     "[| M(A); M(r) |] ==> M(r``A)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   703
apply (simp add: image_iff_Collect)
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   704
apply (insert image_separation [of A r], simp) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   705
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   706
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   707
lemma (in M_axioms) vimage_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   708
     "[| M(r); M(A); M(z) |] ==> pre_image(M,r,A,z) <-> z = r-``A"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   709
apply (simp add: pre_image_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   710
apply (rule iffI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   711
 apply (blast intro!: equalityI dest: transM, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   712
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   713
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   714
lemma (in M_axioms) vimage_closed [intro,simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   715
     "[| M(A); M(r) |] ==> M(r-``A)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   716
apply (simp add: vimage_iff_Collect)
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   717
apply (insert vimage_separation [of A r], simp) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   718
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   719
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   720
lemma (in M_axioms) domain_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   721
     "[| M(r); M(z) |] ==> is_domain(M,r,z) <-> z = domain(r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   722
apply (simp add: is_domain_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   723
apply (blast intro!: equalityI dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   724
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   725
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   726
lemma (in M_axioms) domain_closed [intro,simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   727
     "M(r) ==> M(domain(r))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   728
apply (simp add: domain_eq_vimage)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   729
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   730
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   731
lemma (in M_axioms) range_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   732
     "[| M(r); M(z) |] ==> is_range(M,r,z) <-> z = range(r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   733
apply (simp add: is_range_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   734
apply (blast intro!: equalityI dest: transM)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   735
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   736
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   737
lemma (in M_axioms) range_closed [intro,simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   738
     "M(r) ==> M(range(r))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   739
apply (simp add: range_eq_image)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   740
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   741
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   742
lemma (in M_axioms) field_abs [simp]: 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   743
     "[| M(r); M(z) |] ==> is_field(M,r,z) <-> z = field(r)"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   744
by (simp add: domain_closed range_closed is_field_def field_def)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   745
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   746
lemma (in M_axioms) field_closed [intro,simp]: 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   747
     "M(r) ==> M(field(r))"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   748
by (simp add: domain_closed range_closed Un_closed field_def) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   749
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   750
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   751
lemma (in M_axioms) M_converse_iff:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   752
     "M(r) ==> 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   753
      converse(r) = 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   754
      {z \<in> range(r) * domain(r). 
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   755
        \<exists>p\<in>r. \<exists>x. M(x) & (\<exists>y. M(y) & p = \<langle>x,y\<rangle> & z = \<langle>y,x\<rangle>)}"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   756
by (blast dest: transM)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   757
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   758
lemma (in M_axioms) converse_closed [intro,simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   759
     "M(r) ==> M(converse(r))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   760
apply (simp add: M_converse_iff)
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   761
apply (insert converse_separation [of r], simp)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   762
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   763
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   764
lemma (in M_axioms) relation_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   765
     "M(r) ==> is_relation(M,r) <-> relation(r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   766
apply (simp add: is_relation_def relation_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   767
apply (blast dest!: bspec dest: pair_components_in_M)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   768
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   769
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   770
lemma (in M_axioms) function_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   771
     "M(r) ==> is_function(M,r) <-> function(r)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   772
apply (simp add: is_function_def function_def, safe) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   773
   apply (frule transM, assumption) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   774
  apply (blast dest: pair_components_in_M)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   775
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   776
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   777
lemma (in M_axioms) apply_closed [intro,simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   778
     "[|M(f); M(a)|] ==> M(f`a)"
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   779
apply (simp add: apply_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   780
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   781
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   782
lemma (in M_axioms) apply_abs: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   783
     "[| function(f); M(f); M(y) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   784
      ==> fun_apply(M,f,x,y) <-> x \<in> domain(f) & f`x = y"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   785
apply (simp add: fun_apply_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   786
apply (blast intro: function_apply_equality function_apply_Pair) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   787
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   788
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   789
lemma (in M_axioms) typed_apply_abs: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   790
     "[| f \<in> A -> B; M(f); M(y) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   791
      ==> fun_apply(M,f,x,y) <-> x \<in> A & f`x = y"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   792
by (simp add: apply_abs fun_is_function domain_of_fun) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   793
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   794
lemma (in M_axioms) typed_function_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   795
     "[| M(A); M(f) |] ==> typed_function(M,A,B,f) <-> f \<in> A -> B"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   796
apply (auto simp add: typed_function_def relation_def Pi_iff) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   797
apply (blast dest: pair_components_in_M)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   798
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   799
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   800
lemma (in M_axioms) injection_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   801
     "[| M(A); M(f) |] ==> injection(M,A,B,f) <-> f \<in> inj(A,B)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   802
apply (simp add: injection_def apply_iff inj_def apply_closed)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   803
apply (blast dest: transM [of _ A]); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   804
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   805
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   806
lemma (in M_axioms) surjection_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   807
     "[| M(A); M(B); M(f) |] ==> surjection(M,A,B,f) <-> f \<in> surj(A,B)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   808
by (simp add: typed_apply_abs surjection_def surj_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   809
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   810
lemma (in M_axioms) bijection_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   811
     "[| M(A); M(B); M(f) |] ==> bijection(M,A,B,f) <-> f \<in> bij(A,B)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   812
by (simp add: bijection_def bij_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   813
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   814
text{*no longer needed*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   815
lemma (in M_axioms) restriction_is_function: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   816
     "[| restriction(M,f,A,z); function(f); M(f); M(A); M(z) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   817
      ==> function(z)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   818
apply (rotate_tac 1)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   819
apply (simp add: restriction_def Ball_iff_equiv) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   820
apply (unfold function_def, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   821
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   822
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   823
lemma (in M_axioms) restriction_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   824
     "[| M(f); M(A); M(z) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   825
      ==> restriction(M,f,A,z) <-> z = restrict(f,A)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   826
apply (simp add: Ball_iff_equiv restriction_def restrict_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   827
apply (blast intro!: equalityI dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   828
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   829
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   830
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   831
lemma (in M_axioms) M_restrict_iff:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   832
     "M(r) ==> restrict(r,A) = {z \<in> r . \<exists>x\<in>A. \<exists>y. M(y) & z = \<langle>x, y\<rangle>}"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   833
by (simp add: restrict_def, blast dest: transM)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   834
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   835
lemma (in M_axioms) restrict_closed [intro,simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   836
     "[| M(A); M(r) |] ==> M(restrict(r,A))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   837
apply (simp add: M_restrict_iff)
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   838
apply (insert restrict_separation [of A], simp) 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   839
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   840
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   841
lemma (in M_axioms) M_comp_iff:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   842
     "[| M(r); M(s) |] 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   843
      ==> r O s = 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   844
          {xz \<in> domain(s) * range(r).  
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   845
            \<exists>x. M(x) & (\<exists>y. M(y) & (\<exists>z. M(z) & 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   846
                xz = \<langle>x,z\<rangle> & \<langle>x,y\<rangle> \<in> s & \<langle>y,z\<rangle> \<in> r))}"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   847
apply (simp add: comp_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   848
apply (rule equalityI) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   849
 apply (clarify ); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   850
 apply (simp add: ); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   851
 apply  (blast dest:  transM)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   852
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   853
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   854
lemma (in M_axioms) comp_closed [intro,simp]: 
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   855
     "[| M(r); M(s) |] ==> M(r O s)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   856
apply (simp add: M_comp_iff)
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   857
apply (insert comp_separation [of r s], simp) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   858
done
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   859
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   860
lemma (in M_axioms) composition_abs [simp]: 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   861
     "[| M(r); M(s); M(t) |] 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   862
      ==> composition(M,r,s,t) <-> t = r O s"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   863
apply safe;
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   864
 txt{*Proving @{term "composition(M, r, s, r O s)"}*}
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   865
 prefer 2 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   866
 apply (simp add: composition_def comp_def)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   867
 apply (blast dest: transM) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   868
txt{*Opposite implication*}
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   869
apply (rule M_equalityI)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   870
  apply (simp add: composition_def comp_def)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   871
  apply (blast del: allE dest: transM)+
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   872
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   873
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   874
lemma (in M_axioms) nat_into_M [intro]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   875
     "n \<in> nat ==> M(n)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   876
by (induct n rule: nat_induct, simp_all)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   877
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   878
lemma (in M_axioms) Inl_in_M_iff [iff]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   879
     "M(Inl(a)) <-> M(a)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   880
by (simp add: Inl_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   881
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   882
lemma (in M_axioms) Inr_in_M_iff [iff]:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   883
     "M(Inr(a)) <-> M(a)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   884
by (simp add: Inr_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   885
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   886
lemma (in M_axioms) Inter_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   887
     "[| M(A); M(z) |] ==> big_inter(M,A,z) <-> z = Inter(A)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   888
apply (simp add: big_inter_def Inter_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   889
apply (blast intro!: equalityI dest: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   890
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   891
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   892
lemma (in M_axioms) Inter_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   893
     "M(A) ==> M(Inter(A))"
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   894
by (insert Inter_separation, simp add: Inter_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   895
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   896
lemma (in M_axioms) Int_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   897
     "[| M(A); M(B) |] ==> M(A Int B)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   898
apply (subgoal_tac "M({A,B})")
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   899
apply (frule Inter_closed, force+); 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   900
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   901
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   902
text{*M contains all finite functions*}
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   903
lemma (in M_axioms) finite_fun_closed_lemma [rule_format]: 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   904
     "[| n \<in> nat; M(A) |] ==> \<forall>f \<in> n -> A. M(f)"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   905
apply (induct_tac n, simp)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   906
apply (rule ballI)  
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   907
apply (simp add: succ_def) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   908
apply (frule fun_cons_restrict_eq)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   909
apply (erule ssubst) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   910
apply (subgoal_tac "M(f`x) & restrict(f,x) \<in> x -> A") 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   911
 apply (simp add: cons_closed nat_into_M apply_closed) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   912
apply (blast intro: apply_funtype transM restrict_type2) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   913
done
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   914
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   915
lemma (in M_axioms) finite_fun_closed [rule_format]: 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   916
     "[| f \<in> n -> A; n \<in> nat; M(A) |] ==> M(f)"
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   917
by (blast intro: finite_fun_closed_lemma) 
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   918
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   919
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   920
subsection{*Absoluteness for ordinals*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   921
text{*These results constitute Theorem IV 5.1 of Kunen (page 126).*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   922
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   923
lemma (in M_axioms) lt_closed:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   924
     "[| j<i; M(i) |] ==> M(j)" 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   925
by (blast dest: ltD intro: transM) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   926
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   927
lemma (in M_axioms) transitive_set_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   928
     "M(a) ==> transitive_set(M,a) <-> Transset(a)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   929
by (simp add: transitive_set_def Transset_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   930
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   931
lemma (in M_axioms) ordinal_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   932
     "M(a) ==> ordinal(M,a) <-> Ord(a)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   933
by (simp add: ordinal_def Ord_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   934
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   935
lemma (in M_axioms) limit_ordinal_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   936
     "M(a) ==> limit_ordinal(M,a) <-> Limit(a)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   937
apply (simp add: limit_ordinal_def Ord_0_lt_iff Limit_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   938
apply (simp add: lt_def, blast) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   939
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   940
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   941
lemma (in M_axioms) successor_ordinal_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   942
     "M(a) ==> successor_ordinal(M,a) <-> Ord(a) & (\<exists>b. M(b) & a = succ(b))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   943
apply (simp add: successor_ordinal_def, safe)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   944
apply (drule Ord_cases_disj, auto) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   945
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   946
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   947
lemma finite_Ord_is_nat:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   948
      "[| Ord(a); ~ Limit(a); \<forall>x\<in>a. ~ Limit(x) |] ==> a \<in> nat"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   949
by (induct a rule: trans_induct3, simp_all)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   950
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   951
lemma naturals_not_limit: "a \<in> nat ==> ~ Limit(a)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   952
by (induct a rule: nat_induct, auto)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   953
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   954
lemma (in M_axioms) finite_ordinal_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   955
     "M(a) ==> finite_ordinal(M,a) <-> a \<in> nat"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   956
apply (simp add: finite_ordinal_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   957
apply (blast intro: finite_Ord_is_nat intro: nat_into_Ord 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   958
             dest: Ord_trans naturals_not_limit)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   959
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   960
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   961
lemma Limit_non_Limit_implies_nat: "[| Limit(a); \<forall>x\<in>a. ~ Limit(x) |] ==> a = nat"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   962
apply (rule le_anti_sym) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   963
apply (rule all_lt_imp_le, blast, blast intro: Limit_is_Ord)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   964
 apply (simp add: lt_def)  
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   965
 apply (blast intro: Ord_in_Ord Ord_trans finite_Ord_is_nat) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   966
apply (erule nat_le_Limit)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   967
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   968
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   969
lemma (in M_axioms) omega_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   970
     "M(a) ==> omega(M,a) <-> a = nat"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   971
apply (simp add: omega_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   972
apply (blast intro: Limit_non_Limit_implies_nat dest: naturals_not_limit)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   973
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   974
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   975
lemma (in M_axioms) number1_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   976
     "M(a) ==> number1(M,a) <-> a = 1"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   977
by (simp add: number1_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   978
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   979
lemma (in M_axioms) number1_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   980
     "M(a) ==> number2(M,a) <-> a = succ(1)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   981
by (simp add: number2_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   982
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   983
lemma (in M_axioms) number3_abs [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   984
     "M(a) ==> number3(M,a) <-> a = succ(succ(1))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   985
by (simp add: number3_def) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   986
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   987
text{*Kunen continued to 20...*}
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   988
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   989
(*Could not get this to work.  The \<lambda>x\<in>nat is essential because everything 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   990
  but the recursion variable must stay unchanged.  But then the recursion
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   991
  equations only hold for x\<in>nat (or in some other set) and not for the 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   992
  whole of the class M.
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   993
  consts
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   994
    natnumber_aux :: "[i=>o,i] => i"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   995
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   996
  primrec
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   997
      "natnumber_aux(M,0) = (\<lambda>x\<in>nat. if empty(M,x) then 1 else 0)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   998
      "natnumber_aux(M,succ(n)) = 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   999
	   (\<lambda>x\<in>nat. if (\<exists>y. M(y) & natnumber_aux(M,n)`y=1 & successor(M,y,x)) 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1000
		     then 1 else 0)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1001
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1002
  constdefs
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1003
    natnumber :: "[i=>o,i,i] => o"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1004
      "natnumber(M,n,x) == natnumber_aux(M,n)`x = 1"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1005
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1006
  lemma (in M_axioms) [simp]: 
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1007
       "natnumber(M,0,x) == x=0"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1008
*)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1009
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1010
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1011
end