src/ZF/Constructible/Relative.thy
author wenzelm
Mon, 13 Apr 2020 22:08:14 +0200
changeset 71751 abf3e80bd815
parent 71417 89d05db6dd1f
child 76213 e44d86131648
permissions -rw-r--r--
tuned NEWS;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
     1
(*  Title:      ZF/Constructible/Relative.thy
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
     2
    Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
     3
                With modifications by E. Gunther, M. Pagano, 
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
     4
                and P. Sánchez Terraf
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
     5
*)
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
     6
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
     7
section \<open>Relativization and Absoluteness\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
     8
65449
c82e63b11b8b clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 61798
diff changeset
     9
theory Relative imports ZF begin
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    10
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
    11
subsection\<open>Relativized versions of standard set-theoretic concepts\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    12
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
    13
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    14
  empty :: "[i=>o,i] => o" where
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
    15
    "empty(M,z) == \<forall>x[M]. x \<notin> z"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    16
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    17
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    18
  subset :: "[i=>o,i,i] => o" where
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    19
    "subset(M,A,B) == \<forall>x[M]. x\<in>A \<longrightarrow> x \<in> B"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    20
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    21
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    22
  upair :: "[i=>o,i,i,i] => o" where
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    23
    "upair(M,a,b,z) == a \<in> z & b \<in> z & (\<forall>x[M]. x\<in>z \<longrightarrow> x = a | x = b)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    24
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    25
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    26
  pair :: "[i=>o,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
    27
    "pair(M,a,b,z) == \<exists>x[M]. upair(M,a,a,x) &
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    28
                     (\<exists>y[M]. upair(M,a,b,y) & upair(M,x,y,z))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    29
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
    30
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    31
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    32
  union :: "[i=>o,i,i,i] => o" where
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    33
    "union(M,a,b,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a | x \<in> b"
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
    34
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    35
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    36
  is_cons :: "[i=>o,i,i,i] => o" where
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
    37
    "is_cons(M,a,b,z) == \<exists>x[M]. upair(M,a,a,x) & union(M,x,b,z)"
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
    38
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    39
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    40
  successor :: "[i=>o,i,i] => o" where
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
    41
    "successor(M,a,z) == is_cons(M,a,a,z)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    42
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    43
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    44
  number1 :: "[i=>o,i] => o" where
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
    45
    "number1(M,a) == \<exists>x[M]. empty(M,x) & successor(M,x,a)"
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13353
diff changeset
    46
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    47
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    48
  number2 :: "[i=>o,i] => o" where
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
    49
    "number2(M,a) == \<exists>x[M]. number1(M,x) & successor(M,x,a)"
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13353
diff changeset
    50
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    51
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    52
  number3 :: "[i=>o,i] => o" where
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
    53
    "number3(M,a) == \<exists>x[M]. number2(M,x) & successor(M,x,a)"
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13353
diff changeset
    54
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    55
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    56
  powerset :: "[i=>o,i,i] => o" where
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    57
    "powerset(M,A,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> subset(M,x,A)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    58
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    59
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    60
  is_Collect :: "[i=>o,i,i=>o,i] => o" where
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    61
    "is_Collect(M,A,P,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> A & P(x)"
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
    62
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    63
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    64
  is_Replace :: "[i=>o,i,[i,i]=>o,i] => o" where
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    65
    "is_Replace(M,A,P,z) == \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,u))"
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
    66
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    67
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    68
  inter :: "[i=>o,i,i,i] => o" where
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    69
    "inter(M,a,b,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a & x \<in> b"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    70
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    71
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    72
  setdiff :: "[i=>o,i,i,i] => o" where
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    73
    "setdiff(M,a,b,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> x \<in> a & x \<notin> b"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    74
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    75
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    76
  big_union :: "[i=>o,i,i] => o" where
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    77
    "big_union(M,A,z) == \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>y[M]. y\<in>A & x \<in> y)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    78
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    79
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    80
  big_inter :: "[i=>o,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
    81
    "big_inter(M,A,z) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    82
             (A=0 \<longrightarrow> z=0) &
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    83
             (A\<noteq>0 \<longrightarrow> (\<forall>x[M]. x \<in> z \<longleftrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow> x \<in> y)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    84
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    85
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    86
  cartprod :: "[i=>o,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
    87
    "cartprod(M,A,B,z) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
    88
        \<forall>u[M]. u \<in> z \<longleftrightarrow> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,u)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
    89
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    90
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    91
  is_sum :: "[i=>o,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
    92
    "is_sum(M,A,B,Z) ==
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
    93
       \<exists>A0[M]. \<exists>n1[M]. \<exists>s1[M]. \<exists>B1[M].
13350
paulson
parents: 13348
diff changeset
    94
       number1(M,n1) & cartprod(M,n1,A,A0) & upair(M,n1,n1,s1) &
paulson
parents: 13348
diff changeset
    95
       cartprod(M,s1,B,B1) & union(M,A0,B1,Z)"
paulson
parents: 13348
diff changeset
    96
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    97
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
    98
  is_Inl :: "[i=>o,i,i] => o" where
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
    99
    "is_Inl(M,a,z) == \<exists>zero[M]. empty(M,zero) & pair(M,zero,a,z)"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
   100
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   101
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   102
  is_Inr :: "[i=>o,i,i] => o" where
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
   103
    "is_Inr(M,a,z) == \<exists>n1[M]. number1(M,n1) & pair(M,n1,a,z)"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
   104
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   105
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   106
  is_converse :: "[i=>o,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   107
    "is_converse(M,r,z) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   108
        \<forall>x[M]. x \<in> z \<longleftrightarrow>
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
   109
             (\<exists>w[M]. w\<in>r & (\<exists>u[M]. \<exists>v[M]. pair(M,u,v,w) & pair(M,v,u,x)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   110
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   111
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   112
  pre_image :: "[i=>o,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   113
    "pre_image(M,r,A,z) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   114
        \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. y\<in>A & pair(M,x,y,w)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   115
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   116
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   117
  is_domain :: "[i=>o,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   118
    "is_domain(M,r,z) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   119
        \<forall>x[M]. x \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>y[M]. pair(M,x,y,w)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   120
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   121
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   122
  image :: "[i=>o,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   123
    "image(M,r,A,z) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   124
        \<forall>y[M]. y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,w)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   125
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   126
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   127
  is_range :: "[i=>o,i,i] => o" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 65449
diff changeset
   128
    \<comment> \<open>the cleaner
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   129
      \<^term>\<open>\<exists>r'[M]. is_converse(M,r,r') & is_domain(M,r',z)\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   130
      unfortunately needs an instance of separation in order to prove
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   131
        \<^term>\<open>M(converse(r))\<close>.\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   132
    "is_range(M,r,z) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   133
        \<forall>y[M]. y \<in> z \<longleftrightarrow> (\<exists>w[M]. w\<in>r & (\<exists>x[M]. pair(M,x,y,w)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   134
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   135
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   136
  is_field :: "[i=>o,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   137
    "is_field(M,r,z) ==
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22710
diff changeset
   138
        \<exists>dr[M]. \<exists>rr[M]. is_domain(M,r,dr) & is_range(M,r,rr) &
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
   139
                        union(M,dr,rr,z)"
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   140
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   141
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   142
  is_relation :: "[i=>o,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   143
    "is_relation(M,r) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   144
        (\<forall>z[M]. z\<in>r \<longrightarrow> (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,z)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   145
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   146
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   147
  is_function :: "[i=>o,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   148
    "is_function(M,r) ==
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22710
diff changeset
   149
        \<forall>x[M]. \<forall>y[M]. \<forall>y'[M]. \<forall>p[M]. \<forall>p'[M].
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   150
           pair(M,x,y,p) \<longrightarrow> pair(M,x,y',p') \<longrightarrow> p\<in>r \<longrightarrow> p'\<in>r \<longrightarrow> y=y'"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   151
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   152
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   153
  fun_apply :: "[i=>o,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   154
    "fun_apply(M,f,x,y) ==
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   155
        (\<exists>xs[M]. \<exists>fxs[M].
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   156
         upair(M,x,x,xs) & image(M,f,xs,fxs) & big_union(M,fxs,y))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   157
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   158
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   159
  typed_function :: "[i=>o,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   160
    "typed_function(M,A,B,r) ==
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   161
        is_function(M,r) & is_relation(M,r) & is_domain(M,r,A) &
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   162
        (\<forall>u[M]. u\<in>r \<longrightarrow> (\<forall>x[M]. \<forall>y[M]. pair(M,x,y,u) \<longrightarrow> y\<in>B))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   163
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   164
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   165
  is_funspace :: "[i=>o,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   166
    "is_funspace(M,A,B,F) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   167
        \<forall>f[M]. f \<in> F \<longleftrightarrow> typed_function(M,A,B,f)"
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
   168
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   169
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   170
  composition :: "[i=>o,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   171
    "composition(M,r,s,t) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   172
        \<forall>p[M]. p \<in> t \<longleftrightarrow>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   173
               (\<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   174
                pair(M,x,z,p) & pair(M,x,y,xy) & pair(M,y,z,yz) &
13323
2c287f50c9f3 More relativization, reflection and proofs of separation
paulson
parents: 13319
diff changeset
   175
                xy \<in> s & yz \<in> r)"
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
   176
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   177
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   178
  injection :: "[i=>o,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   179
    "injection(M,A,B,f) ==
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22710
diff changeset
   180
        typed_function(M,A,B,f) &
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   181
        (\<forall>x[M]. \<forall>x'[M]. \<forall>y[M]. \<forall>p[M]. \<forall>p'[M].
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   182
          pair(M,x,y,p) \<longrightarrow> pair(M,x',y,p') \<longrightarrow> p\<in>f \<longrightarrow> p'\<in>f \<longrightarrow> x=x')"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   183
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   184
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   185
  surjection :: "[i=>o,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   186
    "surjection(M,A,B,f) ==
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   187
        typed_function(M,A,B,f) &
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   188
        (\<forall>y[M]. y\<in>B \<longrightarrow> (\<exists>x[M]. x\<in>A & fun_apply(M,f,x,y)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   189
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   190
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   191
  bijection :: "[i=>o,i,i,i] => o" where
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   192
    "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
   193
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   194
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   195
  restriction :: "[i=>o,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   196
    "restriction(M,r,A,z) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   197
        \<forall>x[M]. x \<in> z \<longleftrightarrow> (x \<in> r & (\<exists>u[M]. u\<in>A & (\<exists>v[M]. pair(M,u,v,x))))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   198
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   199
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   200
  transitive_set :: "[i=>o,i] => o" where
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   201
    "transitive_set(M,a) == \<forall>x[M]. x\<in>a \<longrightarrow> subset(M,x,a)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   202
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   203
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   204
  ordinal :: "[i=>o,i] => o" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 65449
diff changeset
   205
     \<comment> \<open>an ordinal is a transitive set of transitive sets\<close>
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   206
    "ordinal(M,a) == transitive_set(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> transitive_set(M,x))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   207
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   208
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   209
  limit_ordinal :: "[i=>o,i] => o" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 65449
diff changeset
   210
    \<comment> \<open>a limit ordinal is a non-empty, successor-closed ordinal\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   211
    "limit_ordinal(M,a) ==
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22710
diff changeset
   212
        ordinal(M,a) & ~ empty(M,a) &
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   213
        (\<forall>x[M]. x\<in>a \<longrightarrow> (\<exists>y[M]. y\<in>a & successor(M,x,y)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   214
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   215
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   216
  successor_ordinal :: "[i=>o,i] => o" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 65449
diff changeset
   217
    \<comment> \<open>a successor ordinal is any ordinal that is neither empty nor limit\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   218
    "successor_ordinal(M,a) ==
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22710
diff changeset
   219
        ordinal(M,a) & ~ empty(M,a) & ~ limit_ordinal(M,a)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   220
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   221
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   222
  finite_ordinal :: "[i=>o,i] => o" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 65449
diff changeset
   223
    \<comment> \<open>an ordinal is finite if neither it nor any of its elements are limit\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   224
    "finite_ordinal(M,a) ==
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22710
diff changeset
   225
        ordinal(M,a) & ~ limit_ordinal(M,a) &
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   226
        (\<forall>x[M]. x\<in>a \<longrightarrow> ~ limit_ordinal(M,x))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   227
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   228
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   229
  omega :: "[i=>o,i] => o" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 65449
diff changeset
   230
    \<comment> \<open>omega is a limit ordinal none of whose elements are limit\<close>
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   231
    "omega(M,a) == limit_ordinal(M,a) & (\<forall>x[M]. x\<in>a \<longrightarrow> ~ limit_ordinal(M,x))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   232
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   233
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   234
  is_quasinat :: "[i=>o,i] => o" where
13350
paulson
parents: 13348
diff changeset
   235
    "is_quasinat(M,z) == empty(M,z) | (\<exists>m[M]. successor(M,m,z))"
paulson
parents: 13348
diff changeset
   236
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   237
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   238
  is_nat_case :: "[i=>o, i, [i,i]=>o, i, i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   239
    "is_nat_case(M, a, is_b, k, z) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   240
       (empty(M,k) \<longrightarrow> z=a) &
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   241
       (\<forall>m[M]. successor(M,m,k) \<longrightarrow> is_b(m,z)) &
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13353
diff changeset
   242
       (is_quasinat(M,k) | empty(M,z))"
13350
paulson
parents: 13348
diff changeset
   243
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   244
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   245
  relation1 :: "[i=>o, [i,i]=>o, i=>i] => o" where
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   246
    "relation1(M,is_f,f) == \<forall>x[M]. \<forall>y[M]. is_f(x,y) \<longleftrightarrow> y = f(x)"
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   247
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   248
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   249
  Relation1 :: "[i=>o, i, [i,i]=>o, i=>i] => o" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 65449
diff changeset
   250
    \<comment> \<open>as above, but typed\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   251
    "Relation1(M,A,is_f,f) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   252
        \<forall>x[M]. \<forall>y[M]. x\<in>A \<longrightarrow> is_f(x,y) \<longleftrightarrow> y = f(x)"
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
   253
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   254
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   255
  relation2 :: "[i=>o, [i,i,i]=>o, [i,i]=>i] => o" where
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   256
    "relation2(M,is_f,f) == \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. is_f(x,y,z) \<longleftrightarrow> z = f(x,y)"
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   257
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   258
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   259
  Relation2 :: "[i=>o, i, i, [i,i,i]=>o, [i,i]=>i] => o" where
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   260
    "Relation2(M,A,B,is_f,f) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   261
        \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. x\<in>A \<longrightarrow> y\<in>B \<longrightarrow> is_f(x,y,z) \<longleftrightarrow> z = f(x,y)"
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
   262
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   263
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   264
  relation3 :: "[i=>o, [i,i,i,i]=>o, [i,i,i]=>i] => o" where
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   265
    "relation3(M,is_f,f) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   266
       \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>u[M]. is_f(x,y,z,u) \<longleftrightarrow> u = f(x,y,z)"
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   267
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   268
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   269
  Relation3 :: "[i=>o, i, i, i, [i,i,i,i]=>o, [i,i,i]=>i] => o" where
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   270
    "Relation3(M,A,B,C,is_f,f) ==
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   271
       \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>u[M].
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   272
         x\<in>A \<longrightarrow> y\<in>B \<longrightarrow> z\<in>C \<longrightarrow> is_f(x,y,z,u) \<longleftrightarrow> u = f(x,y,z)"
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
   273
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   274
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   275
  relation4 :: "[i=>o, [i,i,i,i,i]=>o, [i,i,i,i]=>i] => o" where
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   276
    "relation4(M,is_f,f) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   277
       \<forall>u[M]. \<forall>x[M]. \<forall>y[M]. \<forall>z[M]. \<forall>a[M]. is_f(u,x,y,z,a) \<longleftrightarrow> a = f(u,x,y,z)"
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
   278
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
   279
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   280
text\<open>Useful when absoluteness reasoning has replaced the predicates by terms\<close>
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   281
lemma triv_Relation1:
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   282
     "Relation1(M, A, \<lambda>x y. y = f(x), f)"
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   283
by (simp add: Relation1_def)
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
   284
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   285
lemma triv_Relation2:
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   286
     "Relation2(M, A, B, \<lambda>x y a. a = f(x,y), f)"
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   287
by (simp add: Relation2_def)
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
   288
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   289
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   290
subsection \<open>The relativized ZF axioms\<close>
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   291
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
   292
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   293
  extensionality :: "(i=>o) => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   294
    "extensionality(M) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   295
        \<forall>x[M]. \<forall>y[M]. (\<forall>z[M]. z \<in> x \<longleftrightarrow> z \<in> y) \<longrightarrow> x=y"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   296
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   297
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   298
  separation :: "[i=>o, i=>o] => o" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 65449
diff changeset
   299
    \<comment> \<open>The formula \<open>P\<close> should only involve parameters
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   300
        belonging to \<open>M\<close> and all its quantifiers must be relativized
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   301
        to \<open>M\<close>.  We do not have separation as a scheme; every instance
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   302
        that we need must be assumed (and later proved) separately.\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   303
    "separation(M,P) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   304
        \<forall>z[M]. \<exists>y[M]. \<forall>x[M]. x \<in> y \<longleftrightarrow> x \<in> z & P(x)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   305
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   306
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   307
  upair_ax :: "(i=>o) => o" where
13563
paulson
parents: 13543
diff changeset
   308
    "upair_ax(M) == \<forall>x[M]. \<forall>y[M]. \<exists>z[M]. upair(M,x,y,z)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   309
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   310
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   311
  Union_ax :: "(i=>o) => o" where
13514
cc3bbaf1b8d3 tweaked
paulson
parents: 13505
diff changeset
   312
    "Union_ax(M) == \<forall>x[M]. \<exists>z[M]. big_union(M,x,z)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   313
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   314
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   315
  power_ax :: "(i=>o) => o" where
13514
cc3bbaf1b8d3 tweaked
paulson
parents: 13505
diff changeset
   316
    "power_ax(M) == \<forall>x[M]. \<exists>z[M]. powerset(M,x,z)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   317
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   318
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   319
  univalent :: "[i=>o, i, [i,i]=>o] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   320
    "univalent(M,A,P) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   321
        \<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. \<forall>z[M]. P(x,y) & P(x,z) \<longrightarrow> y=z)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   322
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   323
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   324
  replacement :: "[i=>o, [i,i]=>o] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   325
    "replacement(M,P) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   326
      \<forall>A[M]. univalent(M,A,P) \<longrightarrow>
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   327
      (\<exists>Y[M]. \<forall>b[M]. (\<exists>x[M]. x\<in>A & P(x,b)) \<longrightarrow> b \<in> Y)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   328
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   329
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   330
  strong_replacement :: "[i=>o, [i,i]=>o] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   331
    "strong_replacement(M,P) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   332
      \<forall>A[M]. univalent(M,A,P) \<longrightarrow>
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   333
      (\<exists>Y[M]. \<forall>b[M]. b \<in> Y \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,b)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   334
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   335
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   336
  foundation_ax :: "(i=>o) => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   337
    "foundation_ax(M) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   338
        \<forall>x[M]. (\<exists>y[M]. y\<in>x) \<longrightarrow> (\<exists>y[M]. y\<in>x & ~(\<exists>z[M]. z\<in>x & z \<in> y))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   339
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   340
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   341
subsection\<open>A trivial consistency proof for $V_\omega$\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   342
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   343
text\<open>We prove that $V_\omega$
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   344
      (or \<open>univ\<close> in Isabelle) satisfies some ZF axioms.
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   345
     Kunen, Theorem IV 3.13, page 123.\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   346
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   347
lemma univ0_downwards_mem: "[| y \<in> x; x \<in> univ(0) |] ==> y \<in> univ(0)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   348
apply (insert Transset_univ [OF Transset_0])
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   349
apply (simp add: Transset_def, blast)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   350
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   351
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   352
lemma univ0_Ball_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   353
     "A \<in> univ(0) ==> (\<forall>x\<in>A. x \<in> univ(0) \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(x))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   354
by (blast intro: univ0_downwards_mem)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   355
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   356
lemma univ0_Bex_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   357
     "A \<in> univ(0) ==> (\<exists>x\<in>A. x \<in> univ(0) & P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(x))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   358
by (blast intro: univ0_downwards_mem)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   359
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   360
text\<open>Congruence rule for separation: can assume the variable is in \<open>M\<close>\<close>
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   361
lemma separation_cong [cong]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   362
     "(!!x. M(x) ==> P(x) \<longleftrightarrow> P'(x))
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   363
      ==> separation(M, %x. P(x)) \<longleftrightarrow> separation(M, %x. P'(x))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   364
by (simp add: separation_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   365
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   366
lemma univalent_cong [cong]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   367
     "[| A=A'; !!x y. [| x\<in>A; M(x); M(y) |] ==> P(x,y) \<longleftrightarrow> P'(x,y) |]
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   368
      ==> univalent(M, A, %x y. P(x,y)) \<longleftrightarrow> univalent(M, A', %x y. P'(x,y))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   369
by (simp add: univalent_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   370
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   371
lemma univalent_triv [intro,simp]:
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   372
     "univalent(M, A, \<lambda>x y. y = f(x))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   373
by (simp add: univalent_def)
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   374
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   375
lemma univalent_conjI2 [intro,simp]:
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   376
     "univalent(M,A,Q) ==> univalent(M, A, \<lambda>x y. P(x,y) & Q(x,y))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   377
by (simp add: univalent_def, blast)
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   378
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   379
text\<open>Congruence rule for replacement\<close>
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
   380
lemma strong_replacement_cong [cong]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   381
     "[| !!x y. [| M(x); M(y) |] ==> P(x,y) \<longleftrightarrow> P'(x,y) |]
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   382
      ==> strong_replacement(M, %x y. P(x,y)) \<longleftrightarrow>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   383
          strong_replacement(M, %x y. P'(x,y))"
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   384
by (simp add: strong_replacement_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   385
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   386
text\<open>The extensionality axiom\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   387
lemma "extensionality(\<lambda>x. x \<in> univ(0))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   388
apply (simp add: extensionality_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   389
apply (blast intro: univ0_downwards_mem)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   390
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   391
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   392
text\<open>The separation axiom requires some lemmas\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   393
lemma Collect_in_Vfrom:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   394
     "[| 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
   395
apply (drule Transset_Vfrom)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   396
apply (rule subset_mem_Vfrom)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   397
apply (unfold Transset_def, blast)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   398
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   399
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   400
lemma Collect_in_VLimit:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   401
     "[| X \<in> Vfrom(A,i);  Limit(i);  Transset(A) |]
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   402
      ==> Collect(X,P) \<in> Vfrom(A,i)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   403
apply (rule Limit_VfromE, assumption+)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   404
apply (blast intro: Limit_has_succ VfromI Collect_in_Vfrom)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   405
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   406
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   407
lemma Collect_in_univ:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   408
     "[| X \<in> univ(A);  Transset(A) |] ==> Collect(X,P) \<in> univ(A)"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   409
by (simp add: univ_def Collect_in_VLimit)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   410
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   411
lemma "separation(\<lambda>x. x \<in> univ(0), P)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   412
apply (simp add: separation_def, clarify)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   413
apply (rule_tac x = "Collect(z,P)" in bexI)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   414
apply (blast intro: Collect_in_univ Transset_0)+
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   415
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   416
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   417
text\<open>Unordered pairing axiom\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   418
lemma "upair_ax(\<lambda>x. x \<in> univ(0))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   419
apply (simp add: upair_ax_def upair_def)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   420
apply (blast intro: doubleton_in_univ)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   421
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   422
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   423
text\<open>Union axiom\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   424
lemma "Union_ax(\<lambda>x. x \<in> univ(0))"
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   425
apply (simp add: Union_ax_def big_union_def, clarify)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   426
apply (rule_tac x="\<Union>x" in bexI)
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
   427
 apply (blast intro: univ0_downwards_mem)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   428
apply (blast intro: Union_in_univ Transset_0)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   429
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   430
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   431
text\<open>Powerset axiom\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   432
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   433
lemma Pow_in_univ:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   434
     "[| X \<in> univ(A);  Transset(A) |] ==> Pow(X) \<in> univ(A)"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   435
apply (simp add: univ_def Pow_in_VLimit)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   436
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   437
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   438
lemma "power_ax(\<lambda>x. x \<in> univ(0))"
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   439
apply (simp add: power_ax_def powerset_def subset_def, clarify)
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
   440
apply (rule_tac x="Pow(x)" in bexI)
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
   441
 apply (blast intro: univ0_downwards_mem)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   442
apply (blast intro: Pow_in_univ Transset_0)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   443
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   444
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   445
text\<open>Foundation axiom\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   446
lemma "foundation_ax(\<lambda>x. x \<in> univ(0))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   447
apply (simp add: foundation_ax_def, clarify)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   448
apply (cut_tac A=x in foundation)
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
   449
apply (blast intro: univ0_downwards_mem)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   450
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   451
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   452
lemma "replacement(\<lambda>x. x \<in> univ(0), P)"
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   453
apply (simp add: replacement_def, clarify)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   454
oops
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   455
text\<open>no idea: maybe prove by induction on the rank of A?\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   456
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   457
text\<open>Still missing: Replacement, Choice\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   458
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   459
subsection\<open>Lemmas Needed to Reduce Some Set Constructions to Instances
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   460
      of Separation\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   461
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   462
lemma image_iff_Collect: "r `` A = {y \<in> \<Union>(\<Union>(r)). \<exists>p\<in>r. \<exists>x\<in>A. p=<x,y>}"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   463
apply (rule equalityI, auto)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   464
apply (simp add: Pair_def, blast)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   465
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   466
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   467
lemma vimage_iff_Collect:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   468
     "r -`` A = {x \<in> \<Union>(\<Union>(r)). \<exists>p\<in>r. \<exists>y\<in>A. p=<x,y>}"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   469
apply (rule equalityI, auto)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   470
apply (simp add: Pair_def, blast)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   471
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   472
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   473
text\<open>These two lemmas lets us prove \<open>domain_closed\<close> and
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   474
      \<open>range_closed\<close> without new instances of separation\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   475
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   476
lemma domain_eq_vimage: "domain(r) = r -`` Union(Union(r))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   477
apply (rule equalityI, auto)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   478
apply (rule vimageI, assumption)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   479
apply (simp add: Pair_def, blast)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   480
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   481
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   482
lemma range_eq_image: "range(r) = r `` Union(Union(r))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   483
apply (rule equalityI, auto)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   484
apply (rule imageI, assumption)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   485
apply (simp add: Pair_def, blast)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   486
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   487
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   488
lemma replacementD:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   489
    "[| replacement(M,P); M(A);  univalent(M,A,P) |]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   490
     ==> \<exists>Y[M]. (\<forall>b[M]. ((\<exists>x[M]. x\<in>A & P(x,b)) \<longrightarrow> b \<in> Y))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   491
by (simp add: replacement_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   492
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   493
lemma strong_replacementD:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   494
    "[| strong_replacement(M,P); M(A);  univalent(M,A,P) |]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   495
     ==> \<exists>Y[M]. (\<forall>b[M]. (b \<in> Y \<longleftrightarrow> (\<exists>x[M]. x\<in>A & P(x,b))))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   496
by (simp add: strong_replacement_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   497
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   498
lemma separationD:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   499
    "[| separation(M,P); M(z) |] ==> \<exists>y[M]. \<forall>x[M]. x \<in> y \<longleftrightarrow> x \<in> z & P(x)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   500
by (simp add: separation_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   501
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   502
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   503
text\<open>More constants, for order types\<close>
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   504
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
   505
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   506
  order_isomorphism :: "[i=>o,i,i,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   507
    "order_isomorphism(M,A,r,B,s,f) ==
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   508
        bijection(M,A,B,f) &
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   509
        (\<forall>x[M]. x\<in>A \<longrightarrow> (\<forall>y[M]. y\<in>A \<longrightarrow>
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
   510
          (\<forall>p[M]. \<forall>fx[M]. \<forall>fy[M]. \<forall>q[M].
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   511
            pair(M,x,y,p) \<longrightarrow> fun_apply(M,f,x,fx) \<longrightarrow> fun_apply(M,f,y,fy) \<longrightarrow>
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   512
            pair(M,fx,fy,q) \<longrightarrow> (p\<in>r \<longleftrightarrow> q\<in>s))))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   513
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   514
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   515
  pred_set :: "[i=>o,i,i,i,i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   516
    "pred_set(M,A,x,r,B) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   517
        \<forall>y[M]. y \<in> B \<longleftrightarrow> (\<exists>p[M]. p\<in>r & y \<in> A & pair(M,y,x,p))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   518
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   519
definition
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 65449
diff changeset
   520
  membership :: "[i=>o,i,i] => o" where \<comment> \<open>membership relation\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   521
    "membership(M,A,r) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   522
        \<forall>p[M]. p \<in> r \<longleftrightarrow> (\<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>A & x\<in>y & pair(M,x,y,p)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   523
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   524
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   525
subsection\<open>Introducing a Transitive Class Model\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
   526
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   527
text\<open>The class M is assumed to be transitive and inhabited\<close>
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   528
locale M_trans =
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   529
  fixes M
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   530
  assumes transM:   "[| y\<in>x; M(x) |] ==> M(y)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   531
    and M_inhabited: "\<exists>x . M(x)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   532
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   533
lemma (in M_trans) nonempty [simp]:  "M(0)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   534
proof -
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   535
  have "M(x) \<longrightarrow> M(0)" for x
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   536
  proof (rule_tac P="\<lambda>w. M(w) \<longrightarrow> M(0)" in eps_induct)
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   537
    {
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   538
    fix x
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   539
    assume "\<forall>y\<in>x. M(y) \<longrightarrow> M(0)" "M(x)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   540
    consider (a) "\<exists>y. y\<in>x" | (b) "x=0" by auto
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   541
    then 
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   542
    have "M(x) \<longrightarrow> M(0)" 
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   543
    proof cases
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   544
      case a
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   545
      then show ?thesis using \<open>\<forall>y\<in>x._\<close> \<open>M(x)\<close> transM by auto
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   546
    next
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   547
      case b
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   548
      then show ?thesis by simp
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   549
    qed
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   550
  }
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   551
  then show "M(x) \<longrightarrow> M(0)" if "\<forall>y\<in>x. M(y) \<longrightarrow> M(0)" for x
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   552
    using that by auto
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   553
  qed
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   554
  with M_inhabited
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   555
  show "M(0)" using M_inhabited by blast
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   556
qed
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   557
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   558
text\<open>The class M is assumed to be transitive and to satisfy some
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   559
      relativized ZF axioms\<close>
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   560
locale M_trivial = M_trans +
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   561
  assumes upair_ax:         "upair_ax(M)"
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22710
diff changeset
   562
      and Union_ax:         "Union_ax(M)"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   563
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   564
lemma (in M_trans) rall_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   565
     "M(A) ==> (\<forall>x[M]. x\<in>A \<longrightarrow> P(x)) \<longleftrightarrow> (\<forall>x\<in>A. P(x))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   566
by (blast intro: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   567
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   568
lemma (in M_trans) rex_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   569
     "M(A) ==> (\<exists>x[M]. x\<in>A & P(x)) \<longleftrightarrow> (\<exists>x\<in>A. P(x))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   570
by (blast intro: transM)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   571
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   572
lemma (in M_trans) ball_iff_equiv:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   573
     "M(A) ==> (\<forall>x[M]. (x\<in>A \<longleftrightarrow> P(x))) \<longleftrightarrow>
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   574
               (\<forall>x\<in>A. P(x)) & (\<forall>x. P(x) \<longrightarrow> M(x) \<longrightarrow> x\<in>A)"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   575
by (blast intro: transM)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   576
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   577
text\<open>Simplifies proofs of equalities when there's an iff-equality
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   578
      available for rewriting, universally quantified over M.
13702
c7cf8fa66534 Polishing.
paulson
parents: 13687
diff changeset
   579
      But it's not the only way to prove such equalities: its
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   580
      premises \<^term>\<open>M(A)\<close> and  \<^term>\<open>M(B)\<close> can be too strong.\<close>
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   581
lemma (in M_trans) M_equalityI:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   582
     "[| !!x. M(x) ==> x\<in>A \<longleftrightarrow> x\<in>B; M(A); M(B) |] ==> A=B"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   583
by (blast dest: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   584
13418
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   585
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   586
subsubsection\<open>Trivial Absoluteness Proofs: Empty Set, Pairs, etc.\<close>
13418
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   587
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   588
lemma (in M_trans) empty_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   589
     "M(z) ==> empty(M,z) \<longleftrightarrow> z=0"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   590
apply (simp add: empty_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   591
apply (blast intro: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   592
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   593
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   594
lemma (in M_trans) subset_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   595
     "M(A) ==> subset(M,A,B) \<longleftrightarrow> A \<subseteq> B"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   596
apply (simp add: subset_def)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   597
apply (blast intro: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   598
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   599
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   600
lemma (in M_trans) upair_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   601
     "M(z) ==> upair(M,a,b,z) \<longleftrightarrow> z={a,b}"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   602
apply (simp add: upair_def)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   603
apply (blast intro: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   604
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   605
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   606
lemma (in M_trivial) upair_in_MI [intro!]:
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   607
     " M(a) & M(b) \<Longrightarrow> M({a,b})"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   608
by (insert upair_ax, simp add: upair_ax_def)
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   609
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   610
lemma (in M_trans) upair_in_MD [dest!]:
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   611
     "M({a,b}) \<Longrightarrow> M(a) & M(b)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   612
  by (blast intro: transM)
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   613
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   614
lemma (in M_trivial) upair_in_M_iff [simp]:
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   615
  "M({a,b}) \<longleftrightarrow> M(a) & M(b)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   616
  by blast
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   617
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   618
lemma (in M_trivial) singleton_in_MI [intro!]:
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   619
     "M(a) \<Longrightarrow> M({a})"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   620
  by (insert upair_in_M_iff [of a a], simp)
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   621
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   622
lemma (in M_trans) singleton_in_MD [dest!]:
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   623
     "M({a}) \<Longrightarrow> M(a)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   624
  by (insert upair_in_MD [of a a], simp)
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   625
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   626
lemma (in M_trivial) singleton_in_M_iff [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   627
     "M({a}) \<longleftrightarrow> M(a)"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   628
  by blast
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   629
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   630
lemma (in M_trans) pair_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   631
     "M(z) ==> pair(M,a,b,z) \<longleftrightarrow> z=<a,b>"
65449
c82e63b11b8b clarified main ZF.thy / ZFC.thy, and avoid name clash with global HOL/Main.thy;
wenzelm
parents: 61798
diff changeset
   632
apply (simp add: pair_def Pair_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   633
apply (blast intro: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   634
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   635
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   636
lemma (in M_trans) pair_in_MD [dest!]:
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   637
     "M(<a,b>) \<Longrightarrow> M(a) & M(b)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   638
  by (simp add: Pair_def, blast intro: transM)
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   639
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   640
lemma (in M_trivial) pair_in_MI [intro!]:
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   641
     "M(a) & M(b) \<Longrightarrow> M(<a,b>)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   642
  by (simp add: Pair_def)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   643
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   644
lemma (in M_trivial) pair_in_M_iff [simp]:
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   645
     "M(<a,b>) \<longleftrightarrow> M(a) & M(b)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   646
  by blast
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   647
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   648
lemma (in M_trans) pair_components_in_M:
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   649
     "[| <x,y> \<in> A; M(A) |] ==> M(x) & M(y)"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   650
  by (blast dest: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   651
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   652
lemma (in M_trivial) cartprod_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   653
     "[| M(A); M(B); M(z) |] ==> cartprod(M,A,B,z) \<longleftrightarrow> z = A*B"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   654
apply (simp add: cartprod_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   655
apply (rule iffI)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   656
 apply (blast intro!: equalityI intro: transM dest!: rspec)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   657
apply (blast dest: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   658
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   659
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   660
subsubsection\<open>Absoluteness for Unions and Intersections\<close>
13418
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   661
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   662
lemma (in M_trans) union_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   663
     "[| M(a); M(b); M(z) |] ==> union(M,a,b,z) \<longleftrightarrow> z = a \<union> b"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   664
  unfolding union_def
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   665
  by (blast intro: transM )
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   666
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   667
lemma (in M_trans) inter_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   668
     "[| M(a); M(b); M(z) |] ==> inter(M,a,b,z) \<longleftrightarrow> z = a \<inter> b"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   669
  unfolding inter_def
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   670
  by (blast intro: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   671
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   672
lemma (in M_trans) setdiff_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   673
     "[| M(a); M(b); M(z) |] ==> setdiff(M,a,b,z) \<longleftrightarrow> z = a-b"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   674
  unfolding setdiff_def
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   675
  by (blast intro: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   676
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   677
lemma (in M_trans) Union_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   678
     "[| M(A); M(z) |] ==> big_union(M,A,z) \<longleftrightarrow> z = \<Union>(A)"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   679
  unfolding big_union_def
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   680
  by (blast  dest: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   681
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
   682
lemma (in M_trivial) Union_closed [intro,simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   683
     "M(A) ==> M(\<Union>(A))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   684
by (insert Union_ax, simp add: Union_ax_def)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   685
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
   686
lemma (in M_trivial) Un_closed [intro,simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   687
     "[| M(A); M(B) |] ==> M(A \<union> B)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   688
by (simp only: Un_eq_Union, blast)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   689
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
   690
lemma (in M_trivial) cons_closed [intro,simp]:
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   691
     "[| M(a); M(A) |] ==> M(cons(a,A))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   692
by (subst cons_eq [symmetric], blast)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   693
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   694
lemma (in M_trivial) cons_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   695
     "[| M(b); M(z) |] ==> is_cons(M,a,b,z) \<longleftrightarrow> z = cons(a,b)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   696
by (simp add: is_cons_def, blast intro: transM)
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
   697
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   698
lemma (in M_trivial) successor_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   699
     "[| M(a); M(z) |] ==> successor(M,a,z) \<longleftrightarrow> z = succ(a)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   700
by (simp add: successor_def, blast)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   701
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   702
lemma (in M_trans) succ_in_MD [dest!]:
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   703
     "M(succ(a)) \<Longrightarrow> M(a)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   704
  unfolding succ_def
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   705
  by (blast intro: transM)
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   706
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   707
lemma (in M_trivial) succ_in_MI [intro!]:
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   708
     "M(a) \<Longrightarrow> M(succ(a))"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   709
  unfolding succ_def
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   710
  by (blast intro: transM)
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   711
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   712
lemma (in M_trivial) succ_in_M_iff [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   713
     "M(succ(a)) \<longleftrightarrow> M(a)"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   714
  by blast
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   715
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   716
subsubsection\<open>Absoluteness for Separation and Replacement\<close>
13418
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   717
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   718
lemma (in M_trans) separation_closed [intro,simp]:
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   719
     "[| separation(M,P); M(A) |] ==> M(Collect(A,P))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   720
apply (insert separation, simp add: separation_def)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   721
apply (drule rspec, assumption, clarify)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   722
apply (subgoal_tac "y = Collect(A,P)", blast)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   723
apply (blast dest: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   724
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   725
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
   726
lemma separation_iff:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   727
     "separation(M,P) \<longleftrightarrow> (\<forall>z[M]. \<exists>y[M]. is_Collect(M,z,P,y))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   728
by (simp add: separation_def is_Collect_def)
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
   729
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   730
lemma (in M_trans) Collect_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   731
     "[| M(A); M(z) |] ==> is_Collect(M,A,P,z) \<longleftrightarrow> z = Collect(A,P)"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   732
  unfolding is_Collect_def
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   733
  by (blast dest: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   734
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   735
subsubsection\<open>The Operator \<^term>\<open>is_Replace\<close>\<close>
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   736
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   737
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   738
lemma is_Replace_cong [cong]:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   739
     "[| A=A';
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   740
         !!x y. [| M(x); M(y) |] ==> P(x,y) \<longleftrightarrow> P'(x,y);
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   741
         z=z' |]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   742
      ==> is_Replace(M, A, %x y. P(x,y), z) \<longleftrightarrow>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   743
          is_Replace(M, A', %x y. P'(x,y), z')"
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   744
by (simp add: is_Replace_def)
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   745
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   746
lemma (in M_trans) univalent_Replace_iff:
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   747
     "[| M(A); univalent(M,A,P);
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   748
         !!x y. [| x\<in>A; P(x,y) |] ==> M(y) |]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   749
      ==> u \<in> Replace(A,P) \<longleftrightarrow> (\<exists>x. x\<in>A & P(x,u))"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   750
  unfolding Replace_iff univalent_def
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   751
  by (blast dest: transM)
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   752
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   753
(*The last premise expresses that P takes M to M*)
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   754
lemma (in M_trans) strong_replacement_closed [intro,simp]:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   755
     "[| strong_replacement(M,P); M(A); univalent(M,A,P);
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   756
         !!x y. [| x\<in>A; P(x,y) |] ==> M(y) |] ==> M(Replace(A,P))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   757
apply (simp add: strong_replacement_def)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   758
apply (drule_tac x=A in rspec, safe)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   759
apply (subgoal_tac "Replace(A,P) = Y")
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   760
 apply simp
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   761
apply (rule equality_iffI)
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   762
apply (simp add: univalent_Replace_iff)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   763
apply (blast dest: transM)
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   764
done
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   765
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   766
lemma (in M_trans) Replace_abs:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   767
     "[| M(A); M(z); univalent(M,A,P);
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   768
         !!x y. [| x\<in>A; P(x,y) |] ==> M(y)  |]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   769
      ==> is_Replace(M,A,P,z) \<longleftrightarrow> z = Replace(A,P)"
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
   770
apply (simp add: is_Replace_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   771
apply (rule iffI)
13702
c7cf8fa66534 Polishing.
paulson
parents: 13687
diff changeset
   772
 apply (rule equality_iffI)
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   773
 apply (simp_all add: univalent_Replace_iff)
13702
c7cf8fa66534 Polishing.
paulson
parents: 13687
diff changeset
   774
 apply (blast dest: transM)+
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   775
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   776
13702
c7cf8fa66534 Polishing.
paulson
parents: 13687
diff changeset
   777
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   778
(*The first premise can't simply be assumed as a schema.
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   779
  It is essential to take care when asserting instances of Replacement.
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   780
  Let K be a nonconstructible subset of nat and define
46953
2b6e55924af3 replacing ":" by "\<in>"
paulson
parents: 46823
diff changeset
   781
  f(x) = x if x \<in> K and f(x)=0 otherwise.  Then RepFun(nat,f) = cons(0,K), a
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   782
  nonconstructible set.  So we cannot assume that M(X) implies M(RepFun(X,f))
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   783
  even for f \<in> M -> M.
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   784
*)
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   785
lemma (in M_trans) RepFun_closed:
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   786
     "[| strong_replacement(M, \<lambda>x y. y = f(x)); M(A); \<forall>x\<in>A. M(f(x)) |]
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   787
      ==> M(RepFun(A,f))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   788
apply (simp add: RepFun_def)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   789
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   790
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   791
lemma Replace_conj_eq: "{y . x \<in> A, x\<in>A & y=f(x)} = {y . x\<in>A, y=f(x)}"
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   792
by simp
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   793
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   794
text\<open>Better than \<open>RepFun_closed\<close> when having the formula \<^term>\<open>x\<in>A\<close>
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   795
      makes relativization easier.\<close>
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   796
lemma (in M_trans) RepFun_closed2:
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   797
     "[| strong_replacement(M, \<lambda>x y. x\<in>A & y = f(x)); M(A); \<forall>x\<in>A. M(f(x)) |]
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   798
      ==> M(RepFun(A, %x. f(x)))"
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   799
apply (simp add: RepFun_def)
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   800
apply (frule strong_replacement_closed, assumption)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   801
apply (auto dest: transM  simp add: Replace_conj_eq univalent_def)
13353
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   802
done
1800e7134d2e towards relativization of "iterates" and "wfrec"
paulson
parents: 13352
diff changeset
   803
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   804
subsubsection \<open>Absoluteness for \<^term>\<open>Lambda\<close>\<close>
13418
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   805
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
   806
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
   807
 is_lambda :: "[i=>o, i, [i,i]=>o, i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   808
    "is_lambda(M, A, is_b, z) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   809
       \<forall>p[M]. p \<in> z \<longleftrightarrow>
13418
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   810
        (\<exists>u[M]. \<exists>v[M]. u\<in>A & pair(M,u,v,p) & is_b(u,v))"
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   811
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
   812
lemma (in M_trivial) lam_closed:
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   813
     "[| strong_replacement(M, \<lambda>x y. y = <x,b(x)>); M(A); \<forall>x\<in>A. M(b(x)) |]
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   814
      ==> M(\<lambda>x\<in>A. b(x))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   815
by (simp add: lam_def, blast intro: RepFun_closed dest: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   816
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
   817
text\<open>Better than \<open>lam_closed\<close>: has the formula \<^term>\<open>x\<in>A\<close>\<close>
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
   818
lemma (in M_trivial) lam_closed2:
13418
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   819
  "[|strong_replacement(M, \<lambda>x y. x\<in>A & y = \<langle>x, b(x)\<rangle>);
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   820
     M(A); \<forall>m[M]. m\<in>A \<longrightarrow> M(b(m))|] ==> M(Lambda(A,b))"
13418
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   821
apply (simp add: lam_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   822
apply (blast intro: RepFun_closed2 dest: transM)
13418
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   823
done
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   824
13702
c7cf8fa66534 Polishing.
paulson
parents: 13687
diff changeset
   825
lemma (in M_trivial) lambda_abs2:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   826
     "[| Relation1(M,A,is_b,b); M(A); \<forall>m[M]. m\<in>A \<longrightarrow> M(b(m)); M(z) |]
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   827
      ==> is_lambda(M,A,is_b,z) \<longleftrightarrow> z = Lambda(A,b)"
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   828
apply (simp add: Relation1_def is_lambda_def)
13418
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   829
apply (rule iffI)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   830
 prefer 2 apply (simp add: lam_def)
13702
c7cf8fa66534 Polishing.
paulson
parents: 13687
diff changeset
   831
apply (rule equality_iffI)
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   832
apply (simp add: lam_def)
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   833
apply (rule iffI)
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   834
 apply (blast dest: transM)
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   835
apply (auto simp add: transM [of _ A])
13418
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   836
done
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   837
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
   838
lemma is_lambda_cong [cong]:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   839
     "[| A=A';  z=z';
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   840
         !!x y. [| x\<in>A; M(x); M(y) |] ==> is_b(x,y) \<longleftrightarrow> is_b'(x,y) |]
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   841
      ==> is_lambda(M, A, %x y. is_b(x,y), z) \<longleftrightarrow>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   842
          is_lambda(M, A', %x y. is_b'(x,y), z')"
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   843
by (simp add: is_lambda_def)
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
   844
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   845
lemma (in M_trans) image_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   846
     "[| M(r); M(A); M(z) |] ==> image(M,r,A,z) \<longleftrightarrow> z = r``A"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   847
apply (simp add: image_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   848
apply (rule iffI)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   849
 apply (blast intro!: equalityI dest: transM, blast)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   850
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   851
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   852
subsubsection\<open>Relativization of Powerset\<close>
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   853
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   854
text\<open>What about \<open>Pow_abs\<close>?  Powerset is NOT absolute!
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   855
      This result is one direction of absoluteness.\<close>
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   856
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   857
lemma (in M_trans) powerset_Pow:
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   858
     "powerset(M, x, Pow(x))"
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   859
by (simp add: powerset_def)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   860
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
   861
text\<open>But we can't prove that the powerset in \<open>M\<close> includes the
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   862
      real powerset.\<close>
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   863
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   864
lemma (in M_trans) powerset_imp_subset_Pow:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   865
     "[| powerset(M,x,y); M(y) |] ==> y \<subseteq> Pow(x)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   866
apply (simp add: powerset_def)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   867
apply (blast dest: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   868
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   869
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   870
lemma (in M_trans) powerset_abs:
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   871
  assumes
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   872
     "M(y)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   873
  shows
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   874
    "powerset(M,x,y) \<longleftrightarrow> y = {a\<in>Pow(x) . M(a)}"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   875
proof (intro iffI equalityI)
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   876
  (* First show the converse implication by double inclusion *)
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   877
  assume "powerset(M,x,y)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   878
  with \<open>M(y)\<close>  
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   879
  show "y \<subseteq> {a\<in>Pow(x) . M(a)}"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   880
    using powerset_imp_subset_Pow transM by blast
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   881
  from \<open>powerset(M,x,y)\<close>
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   882
  show "{a\<in>Pow(x) . M(a)} \<subseteq> y"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   883
    using transM unfolding powerset_def by auto
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   884
next (* we show the direct implication *)
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   885
  assume
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   886
    "y = {a \<in> Pow(x) . M(a)}"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   887
  then
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   888
  show "powerset(M, x, y)"
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   889
    unfolding powerset_def subset_def using transM by blast
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   890
qed
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   891
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   892
subsubsection\<open>Absoluteness for the Natural Numbers\<close>
13418
7c0ba9dba978 tweaks, aiming towards relativization of "satisfies"
paulson
parents: 13397
diff changeset
   893
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
   894
lemma (in M_trivial) nat_into_M [intro]:
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   895
     "n \<in> nat ==> M(n)"
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   896
by (induct n rule: nat_induct, simp_all)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   897
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   898
lemma (in M_trans) nat_case_closed [intro,simp]:
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   899
  "[|M(k); M(a); \<forall>m[M]. M(b(m))|] ==> M(nat_case(a,b,k))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   900
apply (case_tac "k=0", simp)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   901
apply (case_tac "\<exists>m. k = succ(m)", force)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   902
apply (simp add: nat_case_def)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   903
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   904
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   905
lemma (in M_trivial) quasinat_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   906
     "M(z) ==> is_quasinat(M,z) \<longleftrightarrow> quasinat(z)"
13350
paulson
parents: 13348
diff changeset
   907
by (auto simp add: is_quasinat_def quasinat_def)
paulson
parents: 13348
diff changeset
   908
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   909
lemma (in M_trivial) nat_case_abs [simp]:
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   910
     "[| relation1(M,is_b,b); M(k); M(z) |]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   911
      ==> is_nat_case(M,a,is_b,k,z) \<longleftrightarrow> z = nat_case(a,b,k)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   912
apply (case_tac "quasinat(k)")
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   913
 prefer 2
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   914
 apply (simp add: is_nat_case_def non_nat_case)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   915
 apply (force simp add: quasinat_def)
13350
paulson
parents: 13348
diff changeset
   916
apply (simp add: quasinat_def is_nat_case_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   917
apply (elim disjE exE)
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
   918
 apply (simp_all add: relation1_def)
13350
paulson
parents: 13348
diff changeset
   919
done
paulson
parents: 13348
diff changeset
   920
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   921
(*NOT for the simplifier.  The assumption M(z') is apparently necessary, but
13363
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13353
diff changeset
   922
  causes the error "Failed congruence proof!"  It may be better to replace
c26eeb000470 instantiation of locales M_trancl and M_wfrank;
paulson
parents: 13353
diff changeset
   923
  is_nat_case by nat_case before attempting congruence reasoning.*)
13434
78b93a667c01 better sats rules for higher-order operators
paulson
parents: 13428
diff changeset
   924
lemma is_nat_case_cong:
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   925
     "[| a = a'; k = k';  z = z';  M(z');
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   926
       !!x y. [| M(x); M(y) |] ==> is_b(x,y) \<longleftrightarrow> is_b'(x,y) |]
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   927
      ==> is_nat_case(M, a, is_b, k, z) \<longleftrightarrow> is_nat_case(M, a', is_b', k', z')"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   928
by (simp add: is_nat_case_def)
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
   929
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   930
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   931
subsection\<open>Absoluteness for Ordinals\<close>
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   932
text\<open>These results constitute Theorem IV 5.1 of Kunen (page 126).\<close>
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   933
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   934
lemma (in M_trans) lt_closed:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   935
     "[| j<i; M(i) |] ==> M(j)"
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   936
by (blast dest: ltD intro: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   937
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   938
lemma (in M_trans) transitive_set_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   939
     "M(a) ==> transitive_set(M,a) \<longleftrightarrow> Transset(a)"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   940
by (simp add: transitive_set_def Transset_def)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   941
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
   942
lemma (in M_trans) ordinal_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   943
     "M(a) ==> ordinal(M,a) \<longleftrightarrow> Ord(a)"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   944
by (simp add: ordinal_def Ord_def)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   945
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   946
lemma (in M_trivial) limit_ordinal_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   947
     "M(a) ==> limit_ordinal(M,a) \<longleftrightarrow> Limit(a)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   948
apply (unfold Limit_def limit_ordinal_def)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   949
apply (simp add: Ord_0_lt_iff)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   950
apply (simp add: lt_def, blast)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   951
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   952
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   953
lemma (in M_trivial) successor_ordinal_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   954
     "M(a) ==> successor_ordinal(M,a) \<longleftrightarrow> Ord(a) & (\<exists>b[M]. a = succ(b))"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   955
apply (simp add: successor_ordinal_def, safe)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   956
apply (drule Ord_cases_disj, auto)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   957
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   958
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   959
lemma finite_Ord_is_nat:
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   960
      "[| Ord(a); ~ Limit(a); \<forall>x\<in>a. ~ Limit(x) |] ==> a \<in> nat"
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   961
by (induct a rule: trans_induct3, simp_all)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   962
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   963
lemma (in M_trivial) finite_ordinal_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   964
     "M(a) ==> finite_ordinal(M,a) \<longleftrightarrow> a \<in> nat"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   965
apply (simp add: finite_ordinal_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   966
apply (blast intro: finite_Ord_is_nat intro: nat_into_Ord
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   967
             dest: Ord_trans naturals_not_limit)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   968
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   969
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   970
lemma Limit_non_Limit_implies_nat:
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   971
     "[| Limit(a); \<forall>x\<in>a. ~ Limit(x) |] ==> a = nat"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   972
apply (rule le_anti_sym)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   973
apply (rule all_lt_imp_le, blast, blast intro: Limit_is_Ord)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   974
 apply (simp add: lt_def)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   975
 apply (blast intro: Ord_in_Ord Ord_trans finite_Ord_is_nat)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   976
apply (erule nat_le_Limit)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   977
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   978
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   979
lemma (in M_trivial) omega_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   980
     "M(a) ==> omega(M,a) \<longleftrightarrow> a = nat"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   981
apply (simp add: omega_def)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   982
apply (blast intro: Limit_non_Limit_implies_nat dest: naturals_not_limit)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   983
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   984
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   985
lemma (in M_trivial) number1_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   986
     "M(a) ==> number1(M,a) \<longleftrightarrow> a = 1"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   987
by (simp add: number1_def)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   988
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   989
lemma (in M_trivial) number2_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   990
     "M(a) ==> number2(M,a) \<longleftrightarrow> a = succ(1)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   991
by (simp add: number2_def)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   992
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   993
lemma (in M_trivial) number3_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
   994
     "M(a) ==> number3(M,a) \<longleftrightarrow> a = succ(succ(1))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   995
by (simp add: number3_def)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   996
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
   997
text\<open>Kunen continued to 20...\<close>
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
   998
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
   999
(*Could not get this to work.  The \<lambda>x\<in>nat is essential because everything
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1000
  but the recursion variable must stay unchanged.  But then the recursion
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1001
  equations only hold for x\<in>nat (or in some other set) and not for the
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1002
  whole of the class M.
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1003
  consts
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1004
    natnumber_aux :: "[i=>o,i] => i"
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1005
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1006
  primrec
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1007
      "natnumber_aux(M,0) = (\<lambda>x\<in>nat. if empty(M,x) then 1 else 0)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1008
      "natnumber_aux(M,succ(n)) =
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22710
diff changeset
  1009
           (\<lambda>x\<in>nat. if (\<exists>y[M]. natnumber_aux(M,n)`y=1 & successor(M,y,x))
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22710
diff changeset
  1010
                     then 1 else 0)"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1011
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
  1012
  definition
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1013
    natnumber :: "[i=>o,i,i] => o"
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1014
      "natnumber(M,n,x) == natnumber_aux(M,n)`x = 1"
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1015
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1016
  lemma (in M_trivial) [simp]:
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1017
       "natnumber(M,0,x) == x=0"
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1018
*)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1019
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1020
subsection\<open>Some instances of separation and strong replacement\<close>
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1021
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1022
locale M_basic = M_trivial +
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1023
assumes Inter_separation:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1024
     "M(A) ==> separation(M, \<lambda>x. \<forall>y[M]. y\<in>A \<longrightarrow> x\<in>y)"
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1025
  and Diff_separation:
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1026
     "M(B) ==> separation(M, \<lambda>x. x \<notin> B)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1027
  and cartprod_separation:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1028
     "[| M(A); M(B) |]
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13290
diff changeset
  1029
      ==> separation(M, \<lambda>z. \<exists>x[M]. x\<in>A & (\<exists>y[M]. y\<in>B & pair(M,x,y,z)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1030
  and image_separation:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1031
     "[| M(A); M(r) |]
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1032
      ==> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r & (\<exists>x[M]. x\<in>A & pair(M,x,y,p)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1033
  and converse_separation:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1034
     "M(r) ==> separation(M,
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13290
diff changeset
  1035
         \<lambda>z. \<exists>p[M]. p\<in>r & (\<exists>x[M]. \<exists>y[M]. pair(M,x,y,p) & pair(M,y,x,z)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1036
  and restrict_separation:
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1037
     "M(A) ==> separation(M, \<lambda>z. \<exists>x[M]. x\<in>A & (\<exists>y[M]. pair(M,x,y,z)))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1038
  and comp_separation:
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1039
     "[| M(r); M(s) |]
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1040
      ==> separation(M, \<lambda>xz. \<exists>x[M]. \<exists>y[M]. \<exists>z[M]. \<exists>xy[M]. \<exists>yz[M].
32960
69916a850301 eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents: 22710
diff changeset
  1041
                  pair(M,x,z,xz) & pair(M,x,y,xy) & pair(M,y,z,yz) &
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1042
                  xy\<in>s & yz\<in>r)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1043
  and pred_separation:
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13290
diff changeset
  1044
     "[| M(r); M(x) |] ==> separation(M, \<lambda>y. \<exists>p[M]. p\<in>r & pair(M,y,x,p))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1045
  and Memrel_separation:
13298
b4f370679c65 Constructible: some separation axioms
paulson
parents: 13290
diff changeset
  1046
     "separation(M, \<lambda>z. \<exists>x[M]. \<exists>y[M]. pair(M,x,y,z) & x \<in> y)"
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1047
  and funspace_succ_replacement:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1048
     "M(n) ==>
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1049
      strong_replacement(M, \<lambda>p z. \<exists>f[M]. \<exists>b[M]. \<exists>nb[M]. \<exists>cnbf[M].
13306
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
  1050
                pair(M,f,b,p) & pair(M,n,b,nb) & is_cons(M,nb,f,cnbf) &
6eebcddee32b more internalized formulas and separation proofs
paulson
parents: 13299
diff changeset
  1051
                upair(M,cnbf,cnbf,z))"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1052
  and is_recfun_separation:
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 65449
diff changeset
  1053
     \<comment> \<open>for well-founded recursion: used to prove \<open>is_recfun_equal\<close>\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1054
     "[| M(r); M(f); M(g); M(a); M(b) |]
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1055
     ==> separation(M,
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1056
            \<lambda>x. \<exists>xa[M]. \<exists>xb[M].
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1057
                pair(M,x,a,xa) & xa \<in> r & pair(M,x,b,xb) & xb \<in> r &
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1058
                (\<exists>fx[M]. \<exists>gx[M]. fun_apply(M,f,x,fx) & fun_apply(M,g,x,gx) &
13319
23de7b3af453 More Separation proofs
paulson
parents: 13316
diff changeset
  1059
                                   fx \<noteq> gx))"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1060
     and power_ax:         "power_ax(M)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1061
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1062
lemma (in M_trivial) cartprod_iff_lemma:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1063
     "[| M(C);  \<forall>u[M]. u \<in> C \<longleftrightarrow> (\<exists>x\<in>A. \<exists>y\<in>B. u = {{x}, {x,y}});
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
  1064
         powerset(M, A \<union> B, p1); powerset(M, p1, p2);  M(p2) |]
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1065
       ==> C = {u \<in> p2 . \<exists>x\<in>A. \<exists>y\<in>B. u = {{x}, {x,y}}}"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1066
apply (simp add: powerset_def)
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
  1067
apply (rule equalityI, clarify, simp)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1068
 apply (frule transM, assumption)
13611
2edf034c902a Adapted to new simplifier.
berghofe
parents: 13564
diff changeset
  1069
 apply (frule transM, assumption, simp (no_asm_simp))
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1070
 apply blast
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1071
apply clarify
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1072
apply (frule transM, assumption, force)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1073
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1074
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1075
lemma (in M_basic) cartprod_iff:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1076
     "[| M(A); M(B); M(C) |]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1077
      ==> cartprod(M,A,B,C) \<longleftrightarrow>
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1078
          (\<exists>p1[M]. \<exists>p2[M]. powerset(M,A \<union> B,p1) & powerset(M,p1,p2) &
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1079
                   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
  1080
apply (simp add: Pair_def cartprod_def, safe)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1081
defer 1
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1082
  apply (simp add: powerset_def)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1083
 apply blast
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1084
txt\<open>Final, difficult case: the left-to-right direction of the theorem.\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1085
apply (insert power_ax, simp add: power_ax_def)
59788
6f7b6adac439 prefer local fixes;
wenzelm
parents: 58871
diff changeset
  1086
apply (frule_tac x="A \<union> B" and P="\<lambda>x. rex(M,Q(x))" for Q in rspec)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1087
apply (blast, clarify)
59788
6f7b6adac439 prefer local fixes;
wenzelm
parents: 58871
diff changeset
  1088
apply (drule_tac x=z and P="\<lambda>x. rex(M,Q(x))" for Q in rspec)
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
  1089
apply assumption
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1090
apply (blast intro: cartprod_iff_lemma)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1091
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1092
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1093
lemma (in M_basic) cartprod_closed_lemma:
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
  1094
     "[| M(A); M(B) |] ==> \<exists>C[M]. cartprod(M,A,B,C)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1095
apply (simp del: cartprod_abs add: cartprod_iff)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1096
apply (insert power_ax, simp add: power_ax_def)
59788
6f7b6adac439 prefer local fixes;
wenzelm
parents: 58871
diff changeset
  1097
apply (frule_tac x="A \<union> B" and P="\<lambda>x. rex(M,Q(x))" for Q in rspec)
13299
3a932abf97e8 More use of relativized quantifiers
paulson
parents: 13298
diff changeset
  1098
apply (blast, clarify)
59788
6f7b6adac439 prefer local fixes;
wenzelm
parents: 58871
diff changeset
  1099
apply (drule_tac x=z and P="\<lambda>x. rex(M,Q(x))" for Q in rspec, auto)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1100
apply (intro rexI conjI, simp+)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1101
apply (insert cartprod_separation [of A B], simp)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1102
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1103
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1104
text\<open>All the lemmas above are necessary because Powerset is not absolute.
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1105
      I should have used Replacement instead!\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1106
lemma (in M_basic) cartprod_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1107
     "[| M(A); M(B) |] ==> M(A*B)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1108
by (frule cartprod_closed_lemma, assumption, force)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1109
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1110
lemma (in M_basic) sum_closed [intro,simp]:
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1111
     "[| M(A); M(B) |] ==> M(A+B)"
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1112
by (simp add: sum_def)
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1113
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1114
lemma (in M_basic) sum_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1115
     "[| M(A); M(B); M(Z) |] ==> is_sum(M,A,B,Z) \<longleftrightarrow> (Z = A+B)"
13350
paulson
parents: 13348
diff changeset
  1116
by (simp add: is_sum_def sum_def singleton_0 nat_into_M)
paulson
parents: 13348
diff changeset
  1117
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1118
lemma (in M_trivial) Inl_in_M_iff [iff]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1119
     "M(Inl(a)) \<longleftrightarrow> M(a)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1120
by (simp add: Inl_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1121
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1122
lemma (in M_trivial) Inl_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1123
     "M(Z) ==> is_Inl(M,a,Z) \<longleftrightarrow> (Z = Inl(a))"
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1124
by (simp add: is_Inl_def Inl_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1125
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1126
lemma (in M_trivial) Inr_in_M_iff [iff]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1127
     "M(Inr(a)) \<longleftrightarrow> M(a)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1128
by (simp add: Inr_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1129
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1130
lemma (in M_trivial) Inr_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1131
     "M(Z) ==> is_Inr(M,a,Z) \<longleftrightarrow> (Z = Inr(a))"
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1132
by (simp add: is_Inr_def Inr_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1133
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1134
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1135
subsubsection \<open>converse of a relation\<close>
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1136
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1137
lemma (in M_basic) M_converse_iff:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1138
     "M(r) ==>
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1139
      converse(r) =
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1140
      {z \<in> \<Union>(\<Union>(r)) * \<Union>(\<Union>(r)).
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1141
       \<exists>p\<in>r. \<exists>x[M]. \<exists>y[M]. p = \<langle>x,y\<rangle> & z = \<langle>y,x\<rangle>}"
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1142
apply (rule equalityI)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1143
 prefer 2 apply (blast dest: transM, clarify, simp)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1144
apply (simp add: Pair_def)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1145
apply (blast dest: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1146
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1147
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1148
lemma (in M_basic) converse_closed [intro,simp]:
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1149
     "M(r) ==> M(converse(r))"
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1150
apply (simp add: M_converse_iff)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1151
apply (insert converse_separation [of r], simp)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1152
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1153
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1154
lemma (in M_basic) converse_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1155
     "[| M(r); M(z) |] ==> is_converse(M,r,z) \<longleftrightarrow> z = converse(r)"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1156
apply (simp add: is_converse_def)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1157
apply (rule iffI)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1158
 prefer 2 apply blast
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1159
apply (rule M_equalityI)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1160
  apply simp
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1161
  apply (blast dest: transM)+
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1162
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1163
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1164
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1165
subsubsection \<open>image, preimage, domain, range\<close>
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1166
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1167
lemma (in M_basic) image_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1168
     "[| M(A); M(r) |] ==> M(r``A)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1169
apply (simp add: image_iff_Collect)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1170
apply (insert image_separation [of A r], simp)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1171
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1172
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1173
lemma (in M_basic) vimage_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1174
     "[| M(r); M(A); M(z) |] ==> pre_image(M,r,A,z) \<longleftrightarrow> z = r-``A"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1175
apply (simp add: pre_image_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1176
apply (rule iffI)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1177
 apply (blast intro!: equalityI dest: transM, blast)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1178
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1179
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1180
lemma (in M_basic) vimage_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1181
     "[| M(A); M(r) |] ==> M(r-``A)"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1182
by (simp add: vimage_def)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1183
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1184
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1185
subsubsection\<open>Domain, range and field\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1186
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1187
lemma (in M_trans) domain_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1188
     "[| M(r); M(z) |] ==> is_domain(M,r,z) \<longleftrightarrow> z = domain(r)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1189
apply (simp add: is_domain_def)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1190
apply (blast intro!: equalityI dest: transM)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1191
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1192
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1193
lemma (in M_basic) domain_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1194
     "M(r) ==> M(domain(r))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1195
apply (simp add: domain_eq_vimage)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1196
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1197
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1198
lemma (in M_trans) range_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1199
     "[| M(r); M(z) |] ==> is_range(M,r,z) \<longleftrightarrow> z = range(r)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1200
apply (simp add: is_range_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1201
apply (blast intro!: equalityI dest: transM)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1202
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1203
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1204
lemma (in M_basic) range_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1205
     "M(r) ==> M(range(r))"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1206
apply (simp add: range_eq_image)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1207
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1208
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1209
lemma (in M_basic) field_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1210
     "[| M(r); M(z) |] ==> is_field(M,r,z) \<longleftrightarrow> z = field(r)"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1211
by (simp add: is_field_def field_def)
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
  1212
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1213
lemma (in M_basic) field_closed [intro,simp]:
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
  1214
     "M(r) ==> M(field(r))"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1215
by (simp add: field_def)
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
  1216
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
  1217
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1218
subsubsection\<open>Relations, functions and application\<close>
13254
5146ccaedf42 class quantifiers (some)
paulson
parents: 13251
diff changeset
  1219
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1220
lemma (in M_trans) relation_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1221
     "M(r) ==> is_relation(M,r) \<longleftrightarrow> relation(r)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1222
apply (simp add: is_relation_def relation_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1223
apply (blast dest!: bspec dest: pair_components_in_M)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1224
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1225
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1226
lemma (in M_trivial) function_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1227
     "M(r) ==> is_function(M,r) \<longleftrightarrow> function(r)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1228
apply (simp add: is_function_def function_def, safe)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1229
   apply (frule transM, assumption)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1230
  apply (blast dest: pair_components_in_M)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1231
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1232
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1233
lemma (in M_basic) apply_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1234
     "[|M(f); M(a)|] ==> M(f`a)"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1235
by (simp add: apply_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1236
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1237
lemma (in M_basic) apply_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1238
     "[| M(f); M(x); M(y) |] ==> fun_apply(M,f,x,y) \<longleftrightarrow> f`x = y"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1239
apply (simp add: fun_apply_def apply_def, blast)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1240
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1241
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1242
lemma (in M_trivial) typed_function_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1243
     "[| M(A); M(f) |] ==> typed_function(M,A,B,f) \<longleftrightarrow> f \<in> A -> B"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1244
apply (auto simp add: typed_function_def relation_def Pi_iff)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1245
apply (blast dest: pair_components_in_M)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1246
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1247
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1248
lemma (in M_basic) injection_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1249
     "[| M(A); M(f) |] ==> injection(M,A,B,f) \<longleftrightarrow> f \<in> inj(A,B)"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1250
apply (simp add: injection_def apply_iff inj_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1251
apply (blast dest: transM [of _ A])
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1252
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1253
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1254
lemma (in M_basic) surjection_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1255
     "[| M(A); M(B); M(f) |] ==> surjection(M,A,B,f) \<longleftrightarrow> f \<in> surj(A,B)"
13352
3cd767f8d78b new definitions of fun_apply and M_is_recfun
paulson
parents: 13350
diff changeset
  1256
by (simp add: surjection_def surj_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1257
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1258
lemma (in M_basic) bijection_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1259
     "[| M(A); M(B); M(f) |] ==> bijection(M,A,B,f) \<longleftrightarrow> f \<in> bij(A,B)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1260
by (simp add: bijection_def bij_def)
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1261
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1262
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1263
subsubsection\<open>Composition of relations\<close>
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1264
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1265
lemma (in M_basic) M_comp_iff:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1266
     "[| M(r); M(s) |]
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1267
      ==> r O s =
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1268
          {xz \<in> domain(s) * range(r).
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1269
            \<exists>x[M]. \<exists>y[M]. \<exists>z[M]. 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
  1270
apply (simp add: comp_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1271
apply (rule equalityI)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1272
 apply clarify
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1273
 apply simp
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1274
 apply  (blast dest:  transM)+
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1275
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1276
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1277
lemma (in M_basic) comp_closed [intro,simp]:
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1278
     "[| M(r); M(s) |] ==> M(r O s)"
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1279
apply (simp add: M_comp_iff)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1280
apply (insert comp_separation [of r s], simp)
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
  1281
done
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
  1282
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1283
lemma (in M_basic) composition_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1284
     "[| M(r); M(s); M(t) |] ==> composition(M,r,s,t) \<longleftrightarrow> t = r O s"
13247
e3c289f0724b towards absoluteness of wfrec-defined functions
paulson
parents: 13245
diff changeset
  1285
apply safe
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1286
 txt\<open>Proving \<^term>\<open>composition(M, r, s, r O s)\<close>\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1287
 prefer 2
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
  1288
 apply (simp add: composition_def comp_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1289
 apply (blast dest: transM)
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1290
txt\<open>Opposite implication\<close>
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
  1291
apply (rule M_equalityI)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
  1292
  apply (simp add: composition_def comp_def)
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
  1293
  apply (blast del: allE dest: transM)+
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1294
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1295
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1296
text\<open>no longer needed\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1297
lemma (in M_basic) restriction_is_function:
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1298
     "[| restriction(M,f,A,z); function(f); M(f); M(A); M(z) |]
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1299
      ==> function(z)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1300
apply (simp add: restriction_def ball_iff_equiv)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1301
apply (unfold function_def, blast)
13269
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
  1302
done
3ba9be497c33 Tidying and introduction of various new theorems
paulson
parents: 13268
diff changeset
  1303
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1304
lemma (in M_trans) restriction_abs [simp]:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1305
     "[| M(f); M(A); M(z) |]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1306
      ==> restriction(M,f,A,z) \<longleftrightarrow> z = restrict(f,A)"
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1307
apply (simp add: ball_iff_equiv restriction_def restrict_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1308
apply (blast intro!: equalityI dest: transM)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1309
done
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1310
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1311
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1312
lemma (in M_trans) M_restrict_iff:
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1313
     "M(r) ==> restrict(r,A) = {z \<in> r . \<exists>x\<in>A. \<exists>y[M]. z = \<langle>x, y\<rangle>}"
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1314
by (simp add: restrict_def, blast dest: transM)
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1315
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1316
lemma (in M_basic) restrict_closed [intro,simp]:
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1317
     "[| M(A); M(r) |] ==> M(restrict(r,A))"
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1318
apply (simp add: M_restrict_iff)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1319
apply (insert restrict_separation [of A], simp)
13290
28ce81eff3de separation of M_axioms into M_triv_axioms and M_axioms
paulson
parents: 13269
diff changeset
  1320
done
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1321
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1322
lemma (in M_trans) Inter_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1323
     "[| M(A); M(z) |] ==> big_inter(M,A,z) \<longleftrightarrow> z = \<Inter>(A)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1324
apply (simp add: big_inter_def Inter_def)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1325
apply (blast intro!: equalityI dest: transM)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1326
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1327
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1328
lemma (in M_basic) Inter_closed [intro,simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1329
     "M(A) ==> M(\<Inter>(A))"
13245
714f7a423a15 development and tweaks
paulson
parents: 13223
diff changeset
  1330
by (insert Inter_separation, simp add: Inter_def)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1331
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1332
lemma (in M_basic) Int_closed [intro,simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1333
     "[| M(A); M(B) |] ==> M(A \<inter> B)"
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1334
apply (subgoal_tac "M({A,B})")
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1335
apply (frule Inter_closed, force+)
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1336
done
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1337
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1338
lemma (in M_basic) Diff_closed [intro,simp]:
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1339
     "[|M(A); M(B)|] ==> M(A-B)"
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1340
by (insert Diff_separation, simp add: Diff_def)
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1341
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1342
subsubsection\<open>Some Facts About Separation Axioms\<close>
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1343
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1344
lemma (in M_basic) separation_conj:
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1345
     "[|separation(M,P); separation(M,Q)|] ==> separation(M, \<lambda>z. P(z) & Q(z))"
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1346
by (simp del: separation_closed
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1347
         add: separation_iff Collect_Int_Collect_eq [symmetric])
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1348
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1349
(*???equalities*)
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1350
lemma Collect_Un_Collect_eq:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1351
     "Collect(A,P) \<union> Collect(A,Q) = Collect(A, %x. P(x) | Q(x))"
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1352
by blast
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1353
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1354
lemma Diff_Collect_eq:
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1355
     "A - Collect(A,P) = Collect(A, %x. ~ P(x))"
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1356
by blast
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1357
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1358
lemma (in M_trans) Collect_rall_eq:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1359
     "M(Y) ==> Collect(A, %x. \<forall>y[M]. y\<in>Y \<longrightarrow> P(x,y)) =
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1360
               (if Y=0 then A else (\<Inter>y \<in> Y. {x \<in> A. P(x,y)}))"
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1361
  by (simp,blast dest: transM)
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1362
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1363
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1364
lemma (in M_basic) separation_disj:
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1365
     "[|separation(M,P); separation(M,Q)|] ==> separation(M, \<lambda>z. P(z) | Q(z))"
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1366
by (simp del: separation_closed
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1367
         add: separation_iff Collect_Un_Collect_eq [symmetric])
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1368
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1369
lemma (in M_basic) separation_neg:
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1370
     "separation(M,P) ==> separation(M, \<lambda>z. ~P(z))"
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1371
by (simp del: separation_closed
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1372
         add: separation_iff Diff_Collect_eq [symmetric])
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1373
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1374
lemma (in M_basic) separation_imp:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1375
     "[|separation(M,P); separation(M,Q)|]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1376
      ==> separation(M, \<lambda>z. P(z) \<longrightarrow> Q(z))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1377
by (simp add: separation_neg separation_disj not_disj_iff_imp [symmetric])
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1378
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1379
text\<open>This result is a hint of how little can be done without the Reflection
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1380
  Theorem.  The quantifier has to be bounded by a set.  We also need another
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1381
  instance of Separation!\<close>
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1382
lemma (in M_basic) separation_rall:
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1383
     "[|M(Y); \<forall>y[M]. separation(M, \<lambda>x. P(x,y));
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1384
        \<forall>z[M]. strong_replacement(M, \<lambda>x y. y = {u \<in> z . P(u,x)})|]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1385
      ==> separation(M, \<lambda>x. \<forall>y[M]. y\<in>Y \<longrightarrow> P(x,y))"
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1386
apply (simp del: separation_closed rall_abs
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1387
         add: separation_iff Collect_rall_eq)
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1388
apply (blast intro!:  RepFun_closed dest: transM)
13436
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1389
done
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1390
8fd1d803a3d3 tweaks involving Separation
paulson
parents: 13434
diff changeset
  1391
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1392
subsubsection\<open>Functions and function space\<close>
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1393
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1394
text\<open>The assumption \<^term>\<open>M(A->B)\<close> is unusual, but essential: in
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1395
all but trivial cases, A->B cannot be expected to belong to \<^term>\<open>M\<close>.\<close>
71417
89d05db6dd1f Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents: 69593
diff changeset
  1396
lemma (in M_trivial) is_funspace_abs [simp]:
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 52458
diff changeset
  1397
     "[|M(A); M(B); M(F); M(A->B)|] ==> is_funspace(M,A,B,F) \<longleftrightarrow> F = A->B"
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1398
apply (simp add: is_funspace_def)
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1399
apply (rule iffI)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1400
 prefer 2 apply blast
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1401
apply (rule M_equalityI)
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1402
  apply simp_all
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1403
done
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1404
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1405
lemma (in M_basic) succ_fun_eq2:
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1406
     "[|M(B); M(n->B)|] ==>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1407
      succ(n) -> B =
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1408
      \<Union>{z. p \<in> (n->B)*B, \<exists>f[M]. \<exists>b[M]. p = <f,b> & z = {cons(<n,b>, f)}}"
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1409
apply (simp add: succ_fun_eq)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1410
apply (blast dest: transM)
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1411
done
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1412
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1413
lemma (in M_basic) funspace_succ:
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1414
     "[|M(n); M(B); M(n->B) |] ==> M(succ(n) -> B)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1415
apply (insert funspace_succ_replacement [of n], simp)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1416
apply (force simp add: succ_fun_eq2 univalent_def)
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1417
done
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1418
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1419
text\<open>\<^term>\<open>M\<close> contains all finite function spaces.  Needed to prove the
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1420
absoluteness of transitive closure.  See the definition of
61798
27f3c10b0b50 isabelle update_cartouches -c -t;
wenzelm
parents: 60770
diff changeset
  1421
\<open>rtrancl_alt\<close> in in \<open>WF_absolute.thy\<close>.\<close>
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1422
lemma (in M_basic) finite_funspace_closed [intro,simp]:
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1423
     "[|n\<in>nat; M(B)|] ==> M(n->B)"
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1424
apply (induct_tac n, simp)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1425
apply (simp add: funspace_succ nat_into_M)
13268
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1426
done
240509babf00 more use of relativized quantifiers
paulson
parents: 13254
diff changeset
  1427
13350
paulson
parents: 13348
diff changeset
  1428
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1429
subsection\<open>Relativization and Absoluteness for Boolean Operators\<close>
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1430
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
  1431
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1432
  is_bool_of_o :: "[i=>o, o, i] => o" where
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1433
   "is_bool_of_o(M,P,z) == (P & number1(M,z)) | (~P & empty(M,z))"
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1434
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1435
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1436
  is_not :: "[i=>o, i, i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1437
   "is_not(M,a,z) == (number1(M,a)  & empty(M,z)) |
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1438
                     (~number1(M,a) & number1(M,z))"
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1439
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1440
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1441
  is_and :: "[i=>o, i, i, i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1442
   "is_and(M,a,b,z) == (number1(M,a)  & z=b) |
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1443
                       (~number1(M,a) & empty(M,z))"
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1444
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1445
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1446
  is_or :: "[i=>o, i, i, i] => o" where
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1447
   "is_or(M,a,b,z) == (number1(M,a)  & number1(M,z)) |
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1448
                      (~number1(M,a) & z=b)"
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1449
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1450
lemma (in M_trivial) bool_of_o_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1451
     "M(z) ==> is_bool_of_o(M,P,z) \<longleftrightarrow> z = bool_of_o(P)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1452
by (simp add: is_bool_of_o_def bool_of_o_def)
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1453
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1454
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1455
lemma (in M_trivial) not_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1456
     "[| M(a); M(z)|] ==> is_not(M,a,z) \<longleftrightarrow> z = not(a)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1457
by (simp add: Bool.not_def cond_def is_not_def)
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1458
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1459
lemma (in M_trivial) and_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1460
     "[| M(a); M(b); M(z)|] ==> is_and(M,a,b,z) \<longleftrightarrow> z = a and b"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1461
by (simp add: Bool.and_def cond_def is_and_def)
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1462
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1463
lemma (in M_trivial) or_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1464
     "[| M(a); M(b); M(z)|] ==> is_or(M,a,b,z) \<longleftrightarrow> z = a or b"
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1465
by (simp add: Bool.or_def cond_def is_or_def)
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1466
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1467
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1468
lemma (in M_trivial) bool_of_o_closed [intro,simp]:
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1469
     "M(bool_of_o(P))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1470
by (simp add: bool_of_o_def)
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1471
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1472
lemma (in M_trivial) and_closed [intro,simp]:
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1473
     "[| M(p); M(q) |] ==> M(p and q)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1474
by (simp add: and_def cond_def)
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1475
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1476
lemma (in M_trivial) or_closed [intro,simp]:
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1477
     "[| M(p); M(q) |] ==> M(p or q)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1478
by (simp add: or_def cond_def)
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1479
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1480
lemma (in M_trivial) not_closed [intro,simp]:
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1481
     "M(p) ==> M(not(p))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1482
by (simp add: Bool.not_def cond_def)
13423
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1483
7ec771711c09 More lemmas, working towards relativization of "satisfies"
paulson
parents: 13418
diff changeset
  1484
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1485
subsection\<open>Relativization and Absoluteness for List Operators\<close>
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1486
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
  1487
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1488
  is_Nil :: "[i=>o, i] => o" where
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1489
     \<comment> \<open>because \<^prop>\<open>[] \<equiv> Inl(0)\<close>\<close>
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1490
    "is_Nil(M,xs) == \<exists>zero[M]. empty(M,zero) & is_Inl(M,zero,xs)"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1491
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1492
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1493
  is_Cons :: "[i=>o,i,i,i] => o" where
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1494
     \<comment> \<open>because \<^prop>\<open>Cons(a, l) \<equiv> Inr(\<langle>a,l\<rangle>)\<close>\<close>
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1495
    "is_Cons(M,a,l,Z) == \<exists>p[M]. pair(M,a,l,p) & is_Inr(M,p,Z)"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1496
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1497
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1498
lemma (in M_trivial) Nil_in_M [intro,simp]: "M(Nil)"
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1499
by (simp add: Nil_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1500
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1501
lemma (in M_trivial) Nil_abs [simp]: "M(Z) ==> is_Nil(M,Z) \<longleftrightarrow> (Z = Nil)"
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1502
by (simp add: is_Nil_def Nil_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1503
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1504
lemma (in M_trivial) Cons_in_M_iff [iff]: "M(Cons(a,l)) \<longleftrightarrow> M(a) & M(l)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1505
by (simp add: Cons_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1506
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1507
lemma (in M_trivial) Cons_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1508
     "[|M(a); M(l); M(Z)|] ==> is_Cons(M,a,l,Z) \<longleftrightarrow> (Z = Cons(a,l))"
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1509
by (simp add: is_Cons_def Cons_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1510
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1511
21233
5a5c8ea5f66a tuned specifications;
wenzelm
parents: 16417
diff changeset
  1512
definition
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1513
  quasilist :: "i => o" where
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1514
    "quasilist(xs) == xs=Nil | (\<exists>x l. xs = Cons(x,l))"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1515
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1516
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1517
  is_quasilist :: "[i=>o,i] => o" where
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1518
    "is_quasilist(M,z) == is_Nil(M,z) | (\<exists>x[M]. \<exists>l[M]. is_Cons(M,x,l,z))"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1519
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1520
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1521
  list_case' :: "[i, [i,i]=>i, i] => i" where
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1522
    \<comment> \<open>A version of \<^term>\<open>list_case\<close> that's always defined.\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1523
    "list_case'(a,b,xs) ==
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1524
       if quasilist(xs) then list_case(a,b,xs) else 0"
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1525
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1526
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1527
  is_list_case :: "[i=>o, i, [i,i,i]=>o, i, i] => o" where
67443
3abf6a722518 standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents: 65449
diff changeset
  1528
    \<comment> \<open>Returns 0 for non-lists\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1529
    "is_list_case(M, a, is_b, xs, z) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1530
       (is_Nil(M,xs) \<longrightarrow> z=a) &
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1531
       (\<forall>x[M]. \<forall>l[M]. is_Cons(M,x,l,xs) \<longrightarrow> is_b(x,l,z)) &
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1532
       (is_quasilist(M,xs) | empty(M,z))"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1533
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1534
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1535
  hd' :: "i => i" where
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1536
    \<comment> \<open>A version of \<^term>\<open>hd\<close> that's always defined.\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1537
    "hd'(xs) == if quasilist(xs) then hd(xs) else 0"
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1538
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1539
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1540
  tl' :: "i => i" where
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1541
    \<comment> \<open>A version of \<^term>\<open>tl\<close> that's always defined.\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1542
    "tl'(xs) == if quasilist(xs) then tl(xs) else 0"
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1543
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1544
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1545
  is_hd :: "[i=>o,i,i] => o" where
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1546
     \<comment> \<open>\<^term>\<open>hd([]) = 0\<close> no constraints if not a list.
60770
240563fbf41d isabelle update_cartouches;
wenzelm
parents: 59788
diff changeset
  1547
          Avoiding implication prevents the simplifier's looping.\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1548
    "is_hd(M,xs,H) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1549
       (is_Nil(M,xs) \<longrightarrow> empty(M,H)) &
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1550
       (\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | H=x) &
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1551
       (is_quasilist(M,xs) | empty(M,H))"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1552
21404
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1553
definition
eb85850d3eb7 more robust syntax for definition/abbreviation/notation;
wenzelm
parents: 21233
diff changeset
  1554
  is_tl :: "[i=>o,i,i] => o" where
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1555
     \<comment> \<open>\<^term>\<open>tl([]) = []\<close>; see comments about \<^term>\<open>is_hd\<close>\<close>
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1556
    "is_tl(M,xs,T) ==
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1557
       (is_Nil(M,xs) \<longrightarrow> T=xs) &
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1558
       (\<forall>x[M]. \<forall>l[M]. ~ is_Cons(M,x,l,xs) | T=l) &
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1559
       (is_quasilist(M,xs) | empty(M,T))"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1560
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1561
subsubsection\<open>\<^term>\<open>quasilist\<close>: For Case-Splitting with \<^term>\<open>list_case'\<close>\<close>
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1562
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1563
lemma [iff]: "quasilist(Nil)"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1564
by (simp add: quasilist_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1565
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1566
lemma [iff]: "quasilist(Cons(x,l))"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1567
by (simp add: quasilist_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1568
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1569
lemma list_imp_quasilist: "l \<in> list(A) ==> quasilist(l)"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1570
by (erule list.cases, simp_all)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1571
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1572
subsubsection\<open>\<^term>\<open>list_case'\<close>, the Modified Version of \<^term>\<open>list_case\<close>\<close>
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1573
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1574
lemma list_case'_Nil [simp]: "list_case'(a,b,Nil) = a"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1575
by (simp add: list_case'_def quasilist_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1576
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1577
lemma list_case'_Cons [simp]: "list_case'(a,b,Cons(x,l)) = b(x,l)"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1578
by (simp add: list_case'_def quasilist_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1579
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1580
lemma non_list_case: "~ quasilist(x) ==> list_case'(a,b,x) = 0"
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1581
by (simp add: quasilist_def list_case'_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1582
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1583
lemma list_case'_eq_list_case [simp]:
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1584
     "xs \<in> list(A) ==>list_case'(a,b,xs) = list_case(a,b,xs)"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1585
by (erule list.cases, simp_all)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1586
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1587
lemma (in M_basic) list_case'_closed [intro,simp]:
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1588
  "[|M(k); M(a); \<forall>x[M]. \<forall>y[M]. M(b(x,y))|] ==> M(list_case'(a,b,k))"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1589
apply (case_tac "quasilist(k)")
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1590
 apply (simp add: quasilist_def, force)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1591
apply (simp add: non_list_case)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1592
done
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1593
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1594
lemma (in M_trivial) quasilist_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1595
     "M(z) ==> is_quasilist(M,z) \<longleftrightarrow> quasilist(z)"
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1596
by (auto simp add: is_quasilist_def quasilist_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1597
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1598
lemma (in M_trivial) list_case_abs [simp]:
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
  1599
     "[| relation2(M,is_b,b); M(k); M(z) |]
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1600
      ==> is_list_case(M,a,is_b,k,z) \<longleftrightarrow> z = list_case'(a,b,k)"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1601
apply (case_tac "quasilist(k)")
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1602
 prefer 2
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1603
 apply (simp add: is_list_case_def non_list_case)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1604
 apply (force simp add: quasilist_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1605
apply (simp add: quasilist_def is_list_case_def)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1606
apply (elim disjE exE)
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
  1607
 apply (simp_all add: relation2_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1608
done
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1609
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1610
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67443
diff changeset
  1611
subsubsection\<open>The Modified Operators \<^term>\<open>hd'\<close> and \<^term>\<open>tl'\<close>\<close>
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1612
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1613
lemma (in M_trivial) is_hd_Nil: "is_hd(M,[],Z) \<longleftrightarrow> empty(M,Z)"
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
  1614
by (simp add: is_hd_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1615
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1616
lemma (in M_trivial) is_hd_Cons:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1617
     "[|M(a); M(l)|] ==> is_hd(M,Cons(a,l),Z) \<longleftrightarrow> Z = a"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1618
by (force simp add: is_hd_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1619
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1620
lemma (in M_trivial) hd_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1621
     "[|M(x); M(y)|] ==> is_hd(M,x,y) \<longleftrightarrow> y = hd'(x)"
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1622
apply (simp add: hd'_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1623
apply (intro impI conjI)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1624
 prefer 2 apply (force simp add: is_hd_def)
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
  1625
apply (simp add: quasilist_def is_hd_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1626
apply (elim disjE exE, auto)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1627
done
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1628
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1629
lemma (in M_trivial) is_tl_Nil: "is_tl(M,[],Z) \<longleftrightarrow> Z = []"
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
  1630
by (simp add: is_tl_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1631
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1632
lemma (in M_trivial) is_tl_Cons:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1633
     "[|M(a); M(l)|] ==> is_tl(M,Cons(a,l),Z) \<longleftrightarrow> Z = l"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1634
by (force simp add: is_tl_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1635
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1636
lemma (in M_trivial) tl_abs [simp]:
46823
57bf0cecb366 More mathematical symbols for ZF examples
paulson
parents: 32960
diff changeset
  1637
     "[|M(x); M(y)|] ==> is_tl(M,x,y) \<longleftrightarrow> y = tl'(x)"
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1638
apply (simp add: tl'_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1639
apply (intro impI conjI)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1640
 prefer 2 apply (force simp add: is_tl_def)
13505
52a16cb7fefb Relativized right up to L satisfies V=L!
paulson
parents: 13436
diff changeset
  1641
apply (simp add: quasilist_def is_tl_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1642
apply (elim disjE exE, auto)
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1643
done
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1644
13634
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
  1645
lemma (in M_trivial) relation1_tl: "relation1(M, is_tl(M), tl')"
99a593b49b04 Re-organization of Constructible theories
paulson
parents: 13628
diff changeset
  1646
by (simp add: relation1_def)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1647
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1648
lemma hd'_Nil: "hd'([]) = 0"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1649
by (simp add: hd'_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1650
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1651
lemma hd'_Cons: "hd'(Cons(a,l)) = a"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1652
by (simp add: hd'_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1653
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1654
lemma tl'_Nil: "tl'([]) = []"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1655
by (simp add: tl'_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1656
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1657
lemma tl'_Cons: "tl'(Cons(a,l)) = l"
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1658
by (simp add: tl'_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1659
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1660
lemma iterates_tl_Nil: "n \<in> nat ==> tl'^n ([]) = []"
13628
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1661
apply (induct_tac n)
87482b5e3f2e Various simplifications of the Constructible theories
paulson
parents: 13615
diff changeset
  1662
apply (simp_all add: tl'_Nil)
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1663
done
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1664
13564
1500a2e48d44 renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents: 13563
diff changeset
  1665
lemma (in M_basic) tl'_closed: "M(x) ==> M(tl'(x))"
13397
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1666
apply (simp add: tl'_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1667
apply (force simp add: quasilist_def)
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1668
done
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1669
6e5f4d911435 Absoluteness of the function "nth"
paulson
parents: 13382
diff changeset
  1670
13223
45be08fbdcff new theory of inner models
paulson
parents:
diff changeset
  1671
end