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