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