author | paulson |
Tue, 01 Oct 2002 13:26:10 +0200 | |
changeset 13615 | 449a70d88b38 |
parent 13564 | 1500a2e48d44 |
child 13634 | 99a593b49b04 |
permissions | -rw-r--r-- |
13505 | 1 |
(* Title: ZF/Constructible/WF_absolute.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 2002 University of Cambridge |
|
5 |
*) |
|
6 |
||
13306 | 7 |
header {*Absoluteness for Well-Founded Relations and Well-Founded Recursion*} |
8 |
||
13242 | 9 |
theory WF_absolute = WFrec: |
13223 | 10 |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
11 |
subsection{*Every well-founded relation is a subset of some inverse image of |
13247 | 12 |
an ordinal*} |
13 |
||
14 |
lemma wf_rvimage_Ord: "Ord(i) \<Longrightarrow> wf(rvimage(A, f, Memrel(i)))" |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
15 |
by (blast intro: wf_rvimage wf_Memrel) |
13247 | 16 |
|
17 |
||
18 |
constdefs |
|
19 |
wfrank :: "[i,i]=>i" |
|
20 |
"wfrank(r,a) == wfrec(r, a, %x f. \<Union>y \<in> r-``{x}. succ(f`y))" |
|
21 |
||
22 |
constdefs |
|
23 |
wftype :: "i=>i" |
|
24 |
"wftype(r) == \<Union>y \<in> range(r). succ(wfrank(r,y))" |
|
25 |
||
26 |
lemma wfrank: "wf(r) ==> wfrank(r,a) = (\<Union>y \<in> r-``{a}. succ(wfrank(r,y)))" |
|
27 |
by (subst wfrank_def [THEN def_wfrec], simp_all) |
|
28 |
||
29 |
lemma Ord_wfrank: "wf(r) ==> Ord(wfrank(r,a))" |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
30 |
apply (rule_tac a=a in wf_induct, assumption) |
13247 | 31 |
apply (subst wfrank, assumption) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
32 |
apply (rule Ord_succ [THEN Ord_UN], blast) |
13247 | 33 |
done |
34 |
||
35 |
lemma wfrank_lt: "[|wf(r); <a,b> \<in> r|] ==> wfrank(r,a) < wfrank(r,b)" |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
36 |
apply (rule_tac a1 = b in wfrank [THEN ssubst], assumption) |
13247 | 37 |
apply (rule UN_I [THEN ltI]) |
38 |
apply (simp add: Ord_wfrank vimage_iff)+ |
|
39 |
done |
|
40 |
||
41 |
lemma Ord_wftype: "wf(r) ==> Ord(wftype(r))" |
|
42 |
by (simp add: wftype_def Ord_wfrank) |
|
43 |
||
44 |
lemma wftypeI: "\<lbrakk>wf(r); x \<in> field(r)\<rbrakk> \<Longrightarrow> wfrank(r,x) \<in> wftype(r)" |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
45 |
apply (simp add: wftype_def) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
46 |
apply (blast intro: wfrank_lt [THEN ltD]) |
13247 | 47 |
done |
48 |
||
49 |
||
50 |
lemma wf_imp_subset_rvimage: |
|
51 |
"[|wf(r); r \<subseteq> A*A|] ==> \<exists>i f. Ord(i) & r <= rvimage(A, f, Memrel(i))" |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
52 |
apply (rule_tac x="wftype(r)" in exI) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
53 |
apply (rule_tac x="\<lambda>x\<in>A. wfrank(r,x)" in exI) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
54 |
apply (simp add: Ord_wftype, clarify) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
55 |
apply (frule subsetD, assumption, clarify) |
13247 | 56 |
apply (simp add: rvimage_iff wfrank_lt [THEN ltD]) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
57 |
apply (blast intro: wftypeI) |
13247 | 58 |
done |
59 |
||
60 |
theorem wf_iff_subset_rvimage: |
|
61 |
"relation(r) ==> wf(r) <-> (\<exists>i f A. Ord(i) & r <= rvimage(A, f, Memrel(i)))" |
|
62 |
by (blast dest!: relation_field_times_field wf_imp_subset_rvimage |
|
63 |
intro: wf_rvimage_Ord [THEN wf_subset]) |
|
64 |
||
65 |
||
13223 | 66 |
subsection{*Transitive closure without fixedpoints*} |
67 |
||
68 |
constdefs |
|
69 |
rtrancl_alt :: "[i,i]=>i" |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
70 |
"rtrancl_alt(A,r) == |
13223 | 71 |
{p \<in> A*A. \<exists>n\<in>nat. \<exists>f \<in> succ(n) -> A. |
13242 | 72 |
(\<exists>x y. p = <x,y> & f`0 = x & f`n = y) & |
13223 | 73 |
(\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)}" |
74 |
||
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
75 |
lemma alt_rtrancl_lemma1 [rule_format]: |
13223 | 76 |
"n \<in> nat |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
77 |
==> \<forall>f \<in> succ(n) -> field(r). |
13223 | 78 |
(\<forall>i\<in>n. \<langle>f`i, f ` succ(i)\<rangle> \<in> r) --> \<langle>f`0, f`n\<rangle> \<in> r^*" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
79 |
apply (induct_tac n) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
80 |
apply (simp_all add: apply_funtype rtrancl_refl, clarify) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
81 |
apply (rename_tac n f) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
82 |
apply (rule rtrancl_into_rtrancl) |
13223 | 83 |
prefer 2 apply assumption |
84 |
apply (drule_tac x="restrict(f,succ(n))" in bspec) |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
85 |
apply (blast intro: restrict_type2) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
86 |
apply (simp add: Ord_succ_mem_iff nat_0_le [THEN ltD] leI [THEN ltD] ltI) |
13223 | 87 |
done |
88 |
||
89 |
lemma rtrancl_alt_subset_rtrancl: "rtrancl_alt(field(r),r) <= r^*" |
|
90 |
apply (simp add: rtrancl_alt_def) |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
91 |
apply (blast intro: alt_rtrancl_lemma1) |
13223 | 92 |
done |
93 |
||
94 |
lemma rtrancl_subset_rtrancl_alt: "r^* <= rtrancl_alt(field(r),r)" |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
95 |
apply (simp add: rtrancl_alt_def, clarify) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
96 |
apply (frule rtrancl_type [THEN subsetD], clarify, simp) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
97 |
apply (erule rtrancl_induct) |
13223 | 98 |
txt{*Base case, trivial*} |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
99 |
apply (rule_tac x=0 in bexI) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
100 |
apply (rule_tac x="lam x:1. xa" in bexI) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
101 |
apply simp_all |
13223 | 102 |
txt{*Inductive step*} |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
103 |
apply clarify |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
104 |
apply (rename_tac n f) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
105 |
apply (rule_tac x="succ(n)" in bexI) |
13223 | 106 |
apply (rule_tac x="lam i:succ(succ(n)). if i=succ(n) then z else f`i" in bexI) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
107 |
apply (simp add: Ord_succ_mem_iff nat_0_le [THEN ltD] leI [THEN ltD] ltI) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
108 |
apply (blast intro: mem_asym) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
109 |
apply typecheck |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
110 |
apply auto |
13223 | 111 |
done |
112 |
||
113 |
lemma rtrancl_alt_eq_rtrancl: "rtrancl_alt(field(r),r) = r^*" |
|
114 |
by (blast del: subsetI |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
115 |
intro: rtrancl_alt_subset_rtrancl rtrancl_subset_rtrancl_alt) |
13223 | 116 |
|
117 |
||
13242 | 118 |
constdefs |
119 |
||
13324 | 120 |
rtran_closure_mem :: "[i=>o,i,i,i] => o" |
121 |
--{*The property of belonging to @{text "rtran_closure(r)"}*} |
|
122 |
"rtran_closure_mem(M,A,r,p) == |
|
123 |
\<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M]. |
|
124 |
omega(M,nnat) & n\<in>nnat & successor(M,n,n') & |
|
125 |
(\<exists>f[M]. typed_function(M,n',A,f) & |
|
126 |
(\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) & |
|
127 |
fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) & |
|
128 |
(\<forall>j[M]. j\<in>n --> |
|
129 |
(\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M]. |
|
130 |
fun_apply(M,f,j,fj) & successor(M,j,sj) & |
|
131 |
fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))" |
|
132 |
||
13242 | 133 |
rtran_closure :: "[i=>o,i,i] => o" |
13324 | 134 |
"rtran_closure(M,r,s) == |
135 |
\<forall>A[M]. is_field(M,r,A) --> |
|
136 |
(\<forall>p[M]. p \<in> s <-> rtran_closure_mem(M,A,r,p))" |
|
13242 | 137 |
|
138 |
tran_closure :: "[i=>o,i,i] => o" |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
139 |
"tran_closure(M,r,t) == |
13268 | 140 |
\<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" |
13242 | 141 |
|
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
142 |
lemma (in M_basic) rtran_closure_mem_iff: |
13324 | 143 |
"[|M(A); M(r); M(p)|] |
144 |
==> rtran_closure_mem(M,A,r,p) <-> |
|
145 |
(\<exists>n[M]. n\<in>nat & |
|
146 |
(\<exists>f[M]. f \<in> succ(n) -> A & |
|
147 |
(\<exists>x[M]. \<exists>y[M]. p = <x,y> & f`0 = x & f`n = y) & |
|
148 |
(\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)))" |
|
13352 | 149 |
by (simp add: rtran_closure_mem_def Ord_succ_mem_iff nat_0_le [THEN ltD]) |
150 |
||
13242 | 151 |
|
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
152 |
locale M_trancl = M_basic + |
13242 | 153 |
assumes rtrancl_separation: |
13324 | 154 |
"[| M(r); M(A) |] ==> separation (M, rtran_closure_mem(M,A,r))" |
13242 | 155 |
and wellfounded_trancl_separation: |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13306
diff
changeset
|
156 |
"[| M(r); M(Z) |] ==> |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13306
diff
changeset
|
157 |
separation (M, \<lambda>x. |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13306
diff
changeset
|
158 |
\<exists>w[M]. \<exists>wx[M]. \<exists>rp[M]. |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13306
diff
changeset
|
159 |
w \<in> Z & pair(M,w,x,wx) & tran_closure(M,r,rp) & wx \<in> rp)" |
13242 | 160 |
|
161 |
||
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
162 |
lemma (in M_trancl) rtran_closure_rtrancl: |
13242 | 163 |
"M(r) ==> rtran_closure(M,r,rtrancl(r))" |
13324 | 164 |
apply (simp add: rtran_closure_def rtran_closure_mem_iff |
165 |
rtrancl_alt_eq_rtrancl [symmetric] rtrancl_alt_def) |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
166 |
apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) |
13242 | 167 |
done |
168 |
||
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
169 |
lemma (in M_trancl) rtrancl_closed [intro,simp]: |
13242 | 170 |
"M(r) ==> M(rtrancl(r))" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
171 |
apply (insert rtrancl_separation [of r "field(r)"]) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
172 |
apply (simp add: rtrancl_alt_eq_rtrancl [symmetric] |
13324 | 173 |
rtrancl_alt_def rtran_closure_mem_iff) |
13242 | 174 |
done |
175 |
||
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
176 |
lemma (in M_trancl) rtrancl_abs [simp]: |
13242 | 177 |
"[| M(r); M(z) |] ==> rtran_closure(M,r,z) <-> z = rtrancl(r)" |
178 |
apply (rule iffI) |
|
179 |
txt{*Proving the right-to-left implication*} |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
180 |
prefer 2 apply (blast intro: rtran_closure_rtrancl) |
13242 | 181 |
apply (rule M_equalityI) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
182 |
apply (simp add: rtran_closure_def rtrancl_alt_eq_rtrancl [symmetric] |
13324 | 183 |
rtrancl_alt_def rtran_closure_mem_iff) |
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
184 |
apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) |
13242 | 185 |
done |
186 |
||
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
187 |
lemma (in M_trancl) trancl_closed [intro,simp]: |
13242 | 188 |
"M(r) ==> M(trancl(r))" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
189 |
by (simp add: trancl_def comp_closed rtrancl_closed) |
13242 | 190 |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
191 |
lemma (in M_trancl) trancl_abs [simp]: |
13242 | 192 |
"[| M(r); M(z) |] ==> tran_closure(M,r,z) <-> z = trancl(r)" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
193 |
by (simp add: tran_closure_def trancl_def) |
13242 | 194 |
|
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13306
diff
changeset
|
195 |
lemma (in M_trancl) wellfounded_trancl_separation': |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13306
diff
changeset
|
196 |
"[| M(r); M(Z) |] ==> separation (M, \<lambda>x. \<exists>w[M]. w \<in> Z & <w,x> \<in> r^+)" |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13306
diff
changeset
|
197 |
by (insert wellfounded_trancl_separation [of r Z], simp) |
13242 | 198 |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
199 |
text{*Alternative proof of @{text wf_on_trancl}; inspiration for the |
13242 | 200 |
relativized version. Original version is on theory WF.*} |
201 |
lemma "[| wf[A](r); r-``A <= A |] ==> wf[A](r^+)" |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
202 |
apply (simp add: wf_on_def wf_def) |
13242 | 203 |
apply (safe intro!: equalityI) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
204 |
apply (drule_tac x = "{x\<in>A. \<exists>w. \<langle>w,x\<rangle> \<in> r^+ & w \<in> Z}" in spec) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
205 |
apply (blast elim: tranclE) |
13242 | 206 |
done |
207 |
||
208 |
lemma (in M_trancl) wellfounded_on_trancl: |
|
209 |
"[| wellfounded_on(M,A,r); r-``A <= A; M(r); M(A) |] |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
210 |
==> wellfounded_on(M,A,r^+)" |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
211 |
apply (simp add: wellfounded_on_def) |
13242 | 212 |
apply (safe intro!: equalityI) |
213 |
apply (rename_tac Z x) |
|
13268 | 214 |
apply (subgoal_tac "M({x\<in>A. \<exists>w[M]. w \<in> Z & \<langle>w,x\<rangle> \<in> r^+})") |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
215 |
prefer 2 |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13306
diff
changeset
|
216 |
apply (blast intro: wellfounded_trancl_separation') |
13299 | 217 |
apply (drule_tac x = "{x\<in>A. \<exists>w[M]. w \<in> Z & \<langle>w,x\<rangle> \<in> r^+}" in rspec, safe) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
218 |
apply (blast dest: transM, simp) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
219 |
apply (rename_tac y w) |
13242 | 220 |
apply (drule_tac x=w in bspec, assumption, clarify) |
221 |
apply (erule tranclE) |
|
222 |
apply (blast dest: transM) (*transM is needed to prove M(xa)*) |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
223 |
apply blast |
13242 | 224 |
done |
225 |
||
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
226 |
lemma (in M_trancl) wellfounded_trancl: |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
227 |
"[|wellfounded(M,r); M(r)|] ==> wellfounded(M,r^+)" |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
228 |
apply (simp add: wellfounded_iff_wellfounded_on_field) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
229 |
apply (rule wellfounded_on_subset_A, erule wellfounded_on_trancl) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
230 |
apply blast |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
231 |
apply (simp_all add: trancl_type [THEN field_rel_subset]) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
232 |
done |
13242 | 233 |
|
13223 | 234 |
text{*Relativized to M: Every well-founded relation is a subset of some |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
235 |
inverse image of an ordinal. Key step is the construction (in M) of a |
13223 | 236 |
rank function.*} |
237 |
||
238 |
||
13428 | 239 |
locale M_wfrank = M_trancl + |
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
240 |
assumes wfrank_separation: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
241 |
"M(r) ==> |
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
242 |
separation (M, \<lambda>x. |
13348 | 243 |
\<forall>rplus[M]. tran_closure(M,r,rplus) --> |
13352 | 244 |
~ (\<exists>f[M]. M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f)))" |
13348 | 245 |
and wfrank_strong_replacement: |
13242 | 246 |
"M(r) ==> |
13348 | 247 |
strong_replacement(M, \<lambda>x z. |
248 |
\<forall>rplus[M]. tran_closure(M,r,rplus) --> |
|
249 |
(\<exists>y[M]. \<exists>f[M]. pair(M,x,y,z) & |
|
13352 | 250 |
M_is_recfun(M, %x f y. is_range(M,f,y), rplus, x, f) & |
13348 | 251 |
is_range(M,f,y)))" |
13242 | 252 |
and Ord_wfrank_separation: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
253 |
"M(r) ==> |
13348 | 254 |
separation (M, \<lambda>x. |
255 |
\<forall>rplus[M]. tran_closure(M,r,rplus) --> |
|
256 |
~ (\<forall>f[M]. \<forall>rangef[M]. |
|
257 |
is_range(M,f,rangef) --> |
|
13352 | 258 |
M_is_recfun(M, \<lambda>x f y. is_range(M,f,y), rplus, x, f) --> |
13348 | 259 |
ordinal(M,rangef)))" |
260 |
||
261 |
text{*Proving that the relativized instances of Separation or Replacement |
|
262 |
agree with the "real" ones.*} |
|
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
263 |
|
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
264 |
lemma (in M_wfrank) wfrank_separation': |
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
265 |
"M(r) ==> |
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
266 |
separation |
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
267 |
(M, \<lambda>x. ~ (\<exists>f[M]. is_recfun(r^+, x, %x f. range(f), f)))" |
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
268 |
apply (insert wfrank_separation [of r]) |
13353 | 269 |
apply (simp add: relativize2_def is_recfun_abs [of "%x. range"]) |
13339
0f89104dd377
Fixed quantified variable name preservation for ball and bex (bounded quants)
paulson
parents:
13324
diff
changeset
|
270 |
done |
13223 | 271 |
|
13348 | 272 |
lemma (in M_wfrank) wfrank_strong_replacement': |
273 |
"M(r) ==> |
|
274 |
strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>f[M]. |
|
275 |
pair(M,x,y,z) & is_recfun(r^+, x, %x f. range(f), f) & |
|
276 |
y = range(f))" |
|
277 |
apply (insert wfrank_strong_replacement [of r]) |
|
13353 | 278 |
apply (simp add: relativize2_def is_recfun_abs [of "%x. range"]) |
13348 | 279 |
done |
280 |
||
281 |
lemma (in M_wfrank) Ord_wfrank_separation': |
|
282 |
"M(r) ==> |
|
283 |
separation (M, \<lambda>x. |
|
284 |
~ (\<forall>f[M]. is_recfun(r^+, x, \<lambda>x. range, f) --> Ord(range(f))))" |
|
285 |
apply (insert Ord_wfrank_separation [of r]) |
|
13353 | 286 |
apply (simp add: relativize2_def is_recfun_abs [of "%x. range"]) |
13348 | 287 |
done |
288 |
||
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
289 |
text{*This function, defined using replacement, is a rank function for |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
290 |
well-founded relations within the class M.*} |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
291 |
constdefs |
13242 | 292 |
wellfoundedrank :: "[i=>o,i,i] => i" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
293 |
"wellfoundedrank(M,r,A) == |
13268 | 294 |
{p. x\<in>A, \<exists>y[M]. \<exists>f[M]. |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
295 |
p = <x,y> & is_recfun(r^+, x, %x f. range(f), f) & |
13242 | 296 |
y = range(f)}" |
13223 | 297 |
|
13268 | 298 |
lemma (in M_wfrank) exists_wfrank: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
299 |
"[| wellfounded(M,r); M(a); M(r) |] |
13268 | 300 |
==> \<exists>f[M]. is_recfun(r^+, a, %x f. range(f), f)" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
301 |
apply (rule wellfounded_exists_is_recfun) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
302 |
apply (blast intro: wellfounded_trancl) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
303 |
apply (rule trans_trancl) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
304 |
apply (erule wfrank_separation') |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
305 |
apply (erule wfrank_strong_replacement') |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
306 |
apply (simp_all add: trancl_subset_times) |
13223 | 307 |
done |
308 |
||
13268 | 309 |
lemma (in M_wfrank) M_wellfoundedrank: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
310 |
"[| wellfounded(M,r); M(r); M(A) |] ==> M(wellfoundedrank(M,r,A))" |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
311 |
apply (insert wfrank_strong_replacement' [of r]) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
312 |
apply (simp add: wellfoundedrank_def) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
313 |
apply (rule strong_replacement_closed) |
13242 | 314 |
apply assumption+ |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
315 |
apply (rule univalent_is_recfun) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
316 |
apply (blast intro: wellfounded_trancl) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
317 |
apply (rule trans_trancl) |
13505 | 318 |
apply (simp add: trancl_subset_times) |
319 |
apply (blast dest: transM) |
|
13223 | 320 |
done |
321 |
||
13268 | 322 |
lemma (in M_wfrank) Ord_wfrank_range [rule_format]: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
323 |
"[| wellfounded(M,r); a\<in>A; M(r); M(A) |] |
13348 | 324 |
==> \<forall>f[M]. is_recfun(r^+, a, %x f. range(f), f) --> Ord(range(f))" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
325 |
apply (drule wellfounded_trancl, assumption) |
13428 | 326 |
apply (rule wellfounded_induct, assumption, erule (1) transM) |
13254 | 327 |
apply simp |
13348 | 328 |
apply (blast intro: Ord_wfrank_separation', clarify) |
13242 | 329 |
txt{*The reasoning in both cases is that we get @{term y} such that |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
330 |
@{term "\<langle>y, x\<rangle> \<in> r^+"}. We find that |
13242 | 331 |
@{term "f`y = restrict(f, r^+ -`` {y})"}. *} |
332 |
apply (rule OrdI [OF _ Ord_is_Transset]) |
|
333 |
txt{*An ordinal is a transitive set...*} |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
334 |
apply (simp add: Transset_def) |
13242 | 335 |
apply clarify |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
336 |
apply (frule apply_recfun2, assumption) |
13242 | 337 |
apply (force simp add: restrict_iff) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
338 |
txt{*...of ordinals. This second case requires the induction hyp.*} |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
339 |
apply clarify |
13242 | 340 |
apply (rename_tac i y) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
341 |
apply (frule apply_recfun2, assumption) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
342 |
apply (frule is_recfun_imp_in_r, assumption) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
343 |
apply (frule is_recfun_restrict) |
13242 | 344 |
(*simp_all won't work*) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
345 |
apply (simp add: trans_trancl trancl_subset_times)+ |
13242 | 346 |
apply (drule spec [THEN mp], assumption) |
347 |
apply (subgoal_tac "M(restrict(f, r^+ -`` {y}))") |
|
13348 | 348 |
apply (drule_tac x="restrict(f, r^+ -`` {y})" in rspec) |
349 |
apply assumption |
|
13242 | 350 |
apply (simp add: function_apply_equality [OF _ is_recfun_imp_function]) |
351 |
apply (blast dest: pair_components_in_M) |
|
13223 | 352 |
done |
353 |
||
13268 | 354 |
lemma (in M_wfrank) Ord_range_wellfoundedrank: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
355 |
"[| wellfounded(M,r); r \<subseteq> A*A; M(r); M(A) |] |
13242 | 356 |
==> Ord (range(wellfoundedrank(M,r,A)))" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
357 |
apply (frule wellfounded_trancl, assumption) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
358 |
apply (frule trancl_subset_times) |
13242 | 359 |
apply (simp add: wellfoundedrank_def) |
360 |
apply (rule OrdI [OF _ Ord_is_Transset]) |
|
361 |
prefer 2 |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
362 |
txt{*by our previous result the range consists of ordinals.*} |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
363 |
apply (blast intro: Ord_wfrank_range) |
13242 | 364 |
txt{*We still must show that the range is a transitive set.*} |
13247 | 365 |
apply (simp add: Transset_def, clarify, simp) |
13293 | 366 |
apply (rename_tac x i f u) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
367 |
apply (frule is_recfun_imp_in_r, assumption) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
368 |
apply (subgoal_tac "M(u) & M(i) & M(x)") |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
369 |
prefer 2 apply (blast dest: transM, clarify) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
370 |
apply (rule_tac a=u in rangeI) |
13293 | 371 |
apply (rule_tac x=u in ReplaceI) |
372 |
apply simp |
|
373 |
apply (rule_tac x="restrict(f, r^+ -`` {u})" in rexI) |
|
374 |
apply (blast intro: is_recfun_restrict trans_trancl dest: apply_recfun2) |
|
375 |
apply simp |
|
376 |
apply blast |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
377 |
txt{*Unicity requirement of Replacement*} |
13242 | 378 |
apply clarify |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
379 |
apply (frule apply_recfun2, assumption) |
13293 | 380 |
apply (simp add: trans_trancl is_recfun_cut) |
13223 | 381 |
done |
382 |
||
13268 | 383 |
lemma (in M_wfrank) function_wellfoundedrank: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
384 |
"[| wellfounded(M,r); M(r); M(A)|] |
13242 | 385 |
==> function(wellfoundedrank(M,r,A))" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
386 |
apply (simp add: wellfoundedrank_def function_def, clarify) |
13242 | 387 |
txt{*Uniqueness: repeated below!*} |
388 |
apply (drule is_recfun_functional, assumption) |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
389 |
apply (blast intro: wellfounded_trancl) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
390 |
apply (simp_all add: trancl_subset_times trans_trancl) |
13223 | 391 |
done |
392 |
||
13268 | 393 |
lemma (in M_wfrank) domain_wellfoundedrank: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
394 |
"[| wellfounded(M,r); M(r); M(A)|] |
13242 | 395 |
==> domain(wellfoundedrank(M,r,A)) = A" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
396 |
apply (simp add: wellfoundedrank_def function_def) |
13242 | 397 |
apply (rule equalityI, auto) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
398 |
apply (frule transM, assumption) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
399 |
apply (frule_tac a=x in exists_wfrank, assumption+, clarify) |
13293 | 400 |
apply (rule_tac b="range(f)" in domainI) |
401 |
apply (rule_tac x=x in ReplaceI) |
|
402 |
apply simp |
|
13268 | 403 |
apply (rule_tac x=f in rexI, blast, simp_all) |
13242 | 404 |
txt{*Uniqueness (for Replacement): repeated above!*} |
405 |
apply clarify |
|
406 |
apply (drule is_recfun_functional, assumption) |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
407 |
apply (blast intro: wellfounded_trancl) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
408 |
apply (simp_all add: trancl_subset_times trans_trancl) |
13223 | 409 |
done |
410 |
||
13268 | 411 |
lemma (in M_wfrank) wellfoundedrank_type: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
412 |
"[| wellfounded(M,r); M(r); M(A)|] |
13242 | 413 |
==> wellfoundedrank(M,r,A) \<in> A -> range(wellfoundedrank(M,r,A))" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
414 |
apply (frule function_wellfoundedrank [of r A], assumption+) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
415 |
apply (frule function_imp_Pi) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
416 |
apply (simp add: wellfoundedrank_def relation_def) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
417 |
apply blast |
13242 | 418 |
apply (simp add: domain_wellfoundedrank) |
13223 | 419 |
done |
420 |
||
13268 | 421 |
lemma (in M_wfrank) Ord_wellfoundedrank: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
422 |
"[| wellfounded(M,r); a \<in> A; r \<subseteq> A*A; M(r); M(A) |] |
13242 | 423 |
==> Ord(wellfoundedrank(M,r,A) ` a)" |
424 |
by (blast intro: apply_funtype [OF wellfoundedrank_type] |
|
425 |
Ord_in_Ord [OF Ord_range_wellfoundedrank]) |
|
13223 | 426 |
|
13268 | 427 |
lemma (in M_wfrank) wellfoundedrank_eq: |
13242 | 428 |
"[| is_recfun(r^+, a, %x. range, f); |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
429 |
wellfounded(M,r); a \<in> A; M(f); M(r); M(A)|] |
13242 | 430 |
==> wellfoundedrank(M,r,A) ` a = range(f)" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
431 |
apply (rule apply_equality) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
432 |
prefer 2 apply (blast intro: wellfoundedrank_type) |
13242 | 433 |
apply (simp add: wellfoundedrank_def) |
434 |
apply (rule ReplaceI) |
|
13268 | 435 |
apply (rule_tac x="range(f)" in rexI) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
436 |
apply blast |
13268 | 437 |
apply simp_all |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
438 |
txt{*Unicity requirement of Replacement*} |
13242 | 439 |
apply clarify |
440 |
apply (drule is_recfun_functional, assumption) |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
441 |
apply (blast intro: wellfounded_trancl) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
442 |
apply (simp_all add: trancl_subset_times trans_trancl) |
13223 | 443 |
done |
444 |
||
13247 | 445 |
|
13268 | 446 |
lemma (in M_wfrank) wellfoundedrank_lt: |
13247 | 447 |
"[| <a,b> \<in> r; |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
448 |
wellfounded(M,r); r \<subseteq> A*A; M(r); M(A)|] |
13247 | 449 |
==> wellfoundedrank(M,r,A) ` a < wellfoundedrank(M,r,A) ` b" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
450 |
apply (frule wellfounded_trancl, assumption) |
13247 | 451 |
apply (subgoal_tac "a\<in>A & b\<in>A") |
452 |
prefer 2 apply blast |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
453 |
apply (simp add: lt_def Ord_wellfoundedrank, clarify) |
13428 | 454 |
apply (frule exists_wfrank [of concl: _ b], erule (1) transM, assumption) |
455 |
apply clarify |
|
13247 | 456 |
apply (rename_tac fb) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
457 |
apply (frule is_recfun_restrict [of concl: "r^+" a]) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
458 |
apply (rule trans_trancl, assumption) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
459 |
apply (simp_all add: r_into_trancl trancl_subset_times) |
13247 | 460 |
txt{*Still the same goal, but with new @{text is_recfun} assumptions.*} |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
461 |
apply (simp add: wellfoundedrank_eq) |
13247 | 462 |
apply (frule_tac a=a in wellfoundedrank_eq, assumption+) |
463 |
apply (simp_all add: transM [of a]) |
|
464 |
txt{*We have used equations for wellfoundedrank and now must use some |
|
465 |
for @{text is_recfun}. *} |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
466 |
apply (rule_tac a=a in rangeI) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
467 |
apply (simp add: is_recfun_type [THEN apply_iff] vimage_singleton_iff |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
468 |
r_into_trancl apply_recfun r_into_trancl) |
13247 | 469 |
done |
470 |
||
471 |
||
13268 | 472 |
lemma (in M_wfrank) wellfounded_imp_subset_rvimage: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
473 |
"[|wellfounded(M,r); r \<subseteq> A*A; M(r); M(A)|] |
13247 | 474 |
==> \<exists>i f. Ord(i) & r <= rvimage(A, f, Memrel(i))" |
475 |
apply (rule_tac x="range(wellfoundedrank(M,r,A))" in exI) |
|
476 |
apply (rule_tac x="wellfoundedrank(M,r,A)" in exI) |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
477 |
apply (simp add: Ord_range_wellfoundedrank, clarify) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
478 |
apply (frule subsetD, assumption, clarify) |
13247 | 479 |
apply (simp add: rvimage_iff wellfoundedrank_lt [THEN ltD]) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
480 |
apply (blast intro: apply_rangeI wellfoundedrank_type) |
13247 | 481 |
done |
482 |
||
13268 | 483 |
lemma (in M_wfrank) wellfounded_imp_wf: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
484 |
"[|wellfounded(M,r); relation(r); M(r)|] ==> wf(r)" |
13247 | 485 |
by (blast dest!: relation_field_times_field wellfounded_imp_subset_rvimage |
486 |
intro: wf_rvimage_Ord [THEN wf_subset]) |
|
487 |
||
13268 | 488 |
lemma (in M_wfrank) wellfounded_on_imp_wf_on: |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
489 |
"[|wellfounded_on(M,A,r); relation(r); M(r); M(A)|] ==> wf[A](r)" |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
490 |
apply (simp add: wellfounded_on_iff_wellfounded wf_on_def) |
13247 | 491 |
apply (rule wellfounded_imp_wf) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
492 |
apply (simp_all add: relation_def) |
13247 | 493 |
done |
494 |
||
495 |
||
13268 | 496 |
theorem (in M_wfrank) wf_abs [simp]: |
13247 | 497 |
"[|relation(r); M(r)|] ==> wellfounded(M,r) <-> wf(r)" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
498 |
by (blast intro: wellfounded_imp_wf wf_imp_relativized) |
13247 | 499 |
|
13268 | 500 |
theorem (in M_wfrank) wf_on_abs [simp]: |
13247 | 501 |
"[|relation(r); M(r); M(A)|] ==> wellfounded_on(M,A,r) <-> wf[A](r)" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
502 |
by (blast intro: wellfounded_on_imp_wf_on wf_on_imp_relativized) |
13247 | 503 |
|
13254 | 504 |
|
505 |
text{*absoluteness for wfrec-defined functions.*} |
|
506 |
||
507 |
(*first use is_recfun, then M_is_recfun*) |
|
508 |
||
509 |
lemma (in M_trancl) wfrec_relativize: |
|
510 |
"[|wf(r); M(a); M(r); |
|
13268 | 511 |
strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M]. |
13254 | 512 |
pair(M,x,y,z) & |
513 |
is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & |
|
514 |
y = H(x, restrict(g, r -`` {x}))); |
|
515 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] |
|
516 |
==> wfrec(r,a,H) = z <-> |
|
13268 | 517 |
(\<exists>f[M]. is_recfun(r^+, a, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & |
13254 | 518 |
z = H(a,restrict(f,r-``{a})))" |
519 |
apply (frule wf_trancl) |
|
520 |
apply (simp add: wftrec_def wfrec_def, safe) |
|
521 |
apply (frule wf_exists_is_recfun |
|
522 |
[of concl: "r^+" a "\<lambda>x f. H(x, restrict(f, r -`` {x}))"]) |
|
523 |
apply (simp_all add: trans_trancl function_restrictI trancl_subset_times) |
|
13268 | 524 |
apply (clarify, rule_tac x=x in rexI) |
13254 | 525 |
apply (simp_all add: the_recfun_eq trans_trancl trancl_subset_times) |
526 |
done |
|
527 |
||
528 |
||
529 |
text{*Assuming @{term r} is transitive simplifies the occurrences of @{text H}. |
|
530 |
The premise @{term "relation(r)"} is necessary |
|
531 |
before we can replace @{term "r^+"} by @{term r}. *} |
|
532 |
theorem (in M_trancl) trans_wfrec_relativize: |
|
533 |
"[|wf(r); trans(r); relation(r); M(r); M(a); |
|
13353 | 534 |
wfrec_replacement(M,MH,r); relativize2(M,MH,H); |
13254 | 535 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] |
13268 | 536 |
==> wfrec(r,a,H) = z <-> (\<exists>f[M]. is_recfun(r,a,H,f) & z = H(a,f))" |
13353 | 537 |
apply (frule wfrec_replacement', assumption+) |
538 |
apply (simp cong: is_recfun_cong |
|
539 |
add: wfrec_relativize trancl_eq_r |
|
540 |
is_recfun_restrict_idem domain_restrict_idem) |
|
541 |
done |
|
13254 | 542 |
|
13353 | 543 |
theorem (in M_trancl) trans_wfrec_abs: |
544 |
"[|wf(r); trans(r); relation(r); M(r); M(a); M(z); |
|
545 |
wfrec_replacement(M,MH,r); relativize2(M,MH,H); |
|
546 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] |
|
547 |
==> is_wfrec(M,MH,r,a,z) <-> z=wfrec(r,a,H)" |
|
548 |
apply (simp add: trans_wfrec_relativize [THEN iff_sym] is_wfrec_abs, blast) |
|
549 |
done |
|
13254 | 550 |
|
551 |
lemma (in M_trancl) trans_eq_pair_wfrec_iff: |
|
552 |
"[|wf(r); trans(r); relation(r); M(r); M(y); |
|
13353 | 553 |
wfrec_replacement(M,MH,r); relativize2(M,MH,H); |
13254 | 554 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] |
555 |
==> y = <x, wfrec(r, x, H)> <-> |
|
13268 | 556 |
(\<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)" |
13293 | 557 |
apply safe |
558 |
apply (simp add: trans_wfrec_relativize [THEN iff_sym, of concl: _ x]) |
|
13254 | 559 |
txt{*converse direction*} |
560 |
apply (rule sym) |
|
561 |
apply (simp add: trans_wfrec_relativize, blast) |
|
562 |
done |
|
563 |
||
564 |
||
565 |
subsection{*M is closed under well-founded recursion*} |
|
566 |
||
567 |
text{*Lemma with the awkward premise mentioning @{text wfrec}.*} |
|
13268 | 568 |
lemma (in M_wfrank) wfrec_closed_lemma [rule_format]: |
13254 | 569 |
"[|wf(r); M(r); |
570 |
strong_replacement(M, \<lambda>x y. y = \<langle>x, wfrec(r, x, H)\<rangle>); |
|
571 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] |
|
572 |
==> M(a) --> M(wfrec(r,a,H))" |
|
573 |
apply (rule_tac a=a in wf_induct, assumption+) |
|
574 |
apply (subst wfrec, assumption, clarify) |
|
575 |
apply (drule_tac x1=x and x="\<lambda>x\<in>r -`` {x}. wfrec(r, x, H)" |
|
576 |
in rspec [THEN rspec]) |
|
577 |
apply (simp_all add: function_lam) |
|
13505 | 578 |
apply (blast intro: lam_closed dest: pair_components_in_M) |
13254 | 579 |
done |
580 |
||
581 |
text{*Eliminates one instance of replacement.*} |
|
13268 | 582 |
lemma (in M_wfrank) wfrec_replacement_iff: |
13353 | 583 |
"strong_replacement(M, \<lambda>x z. |
584 |
\<exists>y[M]. pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g))) <-> |
|
13254 | 585 |
strong_replacement(M, |
13268 | 586 |
\<lambda>x y. \<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)" |
13254 | 587 |
apply simp |
588 |
apply (rule strong_replacement_cong, blast) |
|
589 |
done |
|
590 |
||
591 |
text{*Useful version for transitive relations*} |
|
13268 | 592 |
theorem (in M_wfrank) trans_wfrec_closed: |
13254 | 593 |
"[|wf(r); trans(r); relation(r); M(r); M(a); |
13353 | 594 |
wfrec_replacement(M,MH,r); relativize2(M,MH,H); |
13254 | 595 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] |
596 |
==> M(wfrec(r,a,H))" |
|
13353 | 597 |
apply (frule wfrec_replacement', assumption+) |
13254 | 598 |
apply (frule wfrec_replacement_iff [THEN iffD1]) |
599 |
apply (rule wfrec_closed_lemma, assumption+) |
|
600 |
apply (simp_all add: wfrec_replacement_iff trans_eq_pair_wfrec_iff) |
|
601 |
done |
|
602 |
||
13506 | 603 |
subsection{*Absoluteness without assuming transitivity*} |
13254 | 604 |
lemma (in M_trancl) eq_pair_wfrec_iff: |
605 |
"[|wf(r); M(r); M(y); |
|
13268 | 606 |
strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M]. |
13254 | 607 |
pair(M,x,y,z) & |
608 |
is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & |
|
609 |
y = H(x, restrict(g, r -`` {x}))); |
|
610 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g))|] |
|
611 |
==> y = <x, wfrec(r, x, H)> <-> |
|
13268 | 612 |
(\<exists>f[M]. is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & |
13254 | 613 |
y = <x, H(x,restrict(f,r-``{x}))>)" |
614 |
apply safe |
|
13293 | 615 |
apply (simp add: wfrec_relativize [THEN iff_sym, of concl: _ x]) |
13254 | 616 |
txt{*converse direction*} |
617 |
apply (rule sym) |
|
618 |
apply (simp add: wfrec_relativize, blast) |
|
619 |
done |
|
620 |
||
621 |
text{*Full version not assuming transitivity, but maybe not very useful.*} |
|
13268 | 622 |
theorem (in M_wfrank) wfrec_closed: |
13254 | 623 |
"[|wf(r); M(r); M(a); |
13353 | 624 |
wfrec_replacement(M,MH,r^+); |
625 |
relativize2(M,MH, \<lambda>x f. H(x, restrict(f, r -`` {x}))); |
|
13254 | 626 |
\<forall>x[M]. \<forall>g[M]. function(g) --> M(H(x,g)) |] |
627 |
==> M(wfrec(r,a,H))" |
|
13353 | 628 |
apply (frule wfrec_replacement' |
629 |
[of MH "r^+" "\<lambda>x f. H(x, restrict(f, r -`` {x}))"]) |
|
630 |
prefer 4 |
|
631 |
apply (frule wfrec_replacement_iff [THEN iffD1]) |
|
632 |
apply (rule wfrec_closed_lemma, assumption+) |
|
633 |
apply (simp_all add: eq_pair_wfrec_iff func.function_restrictI) |
|
13254 | 634 |
done |
635 |
||
13223 | 636 |
end |