author | wenzelm |
Fri, 09 Oct 2020 12:01:35 +0200 | |
changeset 72408 | 2daa5f549687 |
parent 71417 | 89d05db6dd1f |
child 76213 | e44d86131648 |
permissions | -rw-r--r-- |
13505 | 1 |
(* Title: ZF/Constructible/WF_absolute.thy |
2 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
3 |
*) |
|
4 |
||
60770 | 5 |
section \<open>Absoluteness of Well-Founded Recursion\<close> |
13306 | 6 |
|
16417 | 7 |
theory WF_absolute imports WFrec begin |
13223 | 8 |
|
60770 | 9 |
subsection\<open>Transitive closure without fixedpoints\<close> |
13223 | 10 |
|
21233 | 11 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
12 |
rtrancl_alt :: "[i,i]=>i" where |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
13 |
"rtrancl_alt(A,r) == |
13223 | 14 |
{p \<in> A*A. \<exists>n\<in>nat. \<exists>f \<in> succ(n) -> A. |
13242 | 15 |
(\<exists>x y. p = <x,y> & f`0 = x & f`n = y) & |
13223 | 16 |
(\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)}" |
17 |
||
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
18 |
lemma alt_rtrancl_lemma1 [rule_format]: |
13223 | 19 |
"n \<in> nat |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
20 |
==> \<forall>f \<in> succ(n) -> field(r). |
46823 | 21 |
(\<forall>i\<in>n. \<langle>f`i, f ` succ(i)\<rangle> \<in> r) \<longrightarrow> \<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
|
22 |
apply (induct_tac n) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
23 |
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
|
24 |
apply (rename_tac n f) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
25 |
apply (rule rtrancl_into_rtrancl) |
13223 | 26 |
prefer 2 apply assumption |
27 |
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
|
28 |
apply (blast intro: restrict_type2) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
29 |
apply (simp add: Ord_succ_mem_iff nat_0_le [THEN ltD] leI [THEN ltD] ltI) |
13223 | 30 |
done |
31 |
||
46823 | 32 |
lemma rtrancl_alt_subset_rtrancl: "rtrancl_alt(field(r),r) \<subseteq> r^*" |
13223 | 33 |
apply (simp add: rtrancl_alt_def) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
34 |
apply (blast intro: alt_rtrancl_lemma1) |
13223 | 35 |
done |
36 |
||
46823 | 37 |
lemma rtrancl_subset_rtrancl_alt: "r^* \<subseteq> rtrancl_alt(field(r),r)" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
38 |
apply (simp add: rtrancl_alt_def, clarify) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
39 |
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
|
40 |
apply (erule rtrancl_induct) |
60770 | 41 |
txt\<open>Base case, trivial\<close> |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
42 |
apply (rule_tac x=0 in bexI) |
46823 | 43 |
apply (rule_tac x="\<lambda>x\<in>1. xa" in bexI) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
44 |
apply simp_all |
60770 | 45 |
txt\<open>Inductive step\<close> |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
46 |
apply clarify |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
47 |
apply (rename_tac n f) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
48 |
apply (rule_tac x="succ(n)" in bexI) |
46823 | 49 |
apply (rule_tac x="\<lambda>i\<in>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
|
50 |
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
|
51 |
apply (blast intro: mem_asym) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
52 |
apply typecheck |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
53 |
apply auto |
13223 | 54 |
done |
55 |
||
56 |
lemma rtrancl_alt_eq_rtrancl: "rtrancl_alt(field(r),r) = r^*" |
|
57 |
by (blast del: subsetI |
|
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
58 |
intro: rtrancl_alt_subset_rtrancl rtrancl_subset_rtrancl_alt) |
13223 | 59 |
|
60 |
||
21233 | 61 |
definition |
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
62 |
rtran_closure_mem :: "[i=>o,i,i,i] => o" where |
67443
3abf6a722518
standardized towards new-style formal comments: isabelle update_comments;
wenzelm
parents:
61798
diff
changeset
|
63 |
\<comment> \<open>The property of belonging to \<open>rtran_closure(r)\<close>\<close> |
13324 | 64 |
"rtran_closure_mem(M,A,r,p) == |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
65 |
\<exists>nnat[M]. \<exists>n[M]. \<exists>n'[M]. |
13324 | 66 |
omega(M,nnat) & n\<in>nnat & successor(M,n,n') & |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
67 |
(\<exists>f[M]. typed_function(M,n',A,f) & |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
68 |
(\<exists>x[M]. \<exists>y[M]. \<exists>zero[M]. pair(M,x,y,p) & empty(M,zero) & |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
69 |
fun_apply(M,f,zero,x) & fun_apply(M,f,n,y)) & |
46823 | 70 |
(\<forall>j[M]. j\<in>n \<longrightarrow> |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
71 |
(\<exists>fj[M]. \<exists>sj[M]. \<exists>fsj[M]. \<exists>ffp[M]. |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
72 |
fun_apply(M,f,j,fj) & successor(M,j,sj) & |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
73 |
fun_apply(M,f,sj,fsj) & pair(M,fj,fsj,ffp) & ffp \<in> r)))" |
13324 | 74 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
75 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
76 |
rtran_closure :: "[i=>o,i,i] => o" where |
13324 | 77 |
"rtran_closure(M,r,s) == |
46823 | 78 |
\<forall>A[M]. is_field(M,r,A) \<longrightarrow> |
79 |
(\<forall>p[M]. p \<in> s \<longleftrightarrow> rtran_closure_mem(M,A,r,p))" |
|
13242 | 80 |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
81 |
definition |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21233
diff
changeset
|
82 |
tran_closure :: "[i=>o,i,i] => o" where |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
83 |
"tran_closure(M,r,t) == |
13268 | 84 |
\<exists>s[M]. rtran_closure(M,r,s) & composition(M,r,s,t)" |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
85 |
|
13564
1500a2e48d44
renamed M_triv_axioms to M_trivial and M_axioms to M_basic
paulson
parents:
13506
diff
changeset
|
86 |
locale M_trancl = M_basic + |
13242 | 87 |
assumes rtrancl_separation: |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
88 |
"[| M(r); M(A) |] ==> separation (M, rtran_closure_mem(M,A,r))" |
13242 | 89 |
and wellfounded_trancl_separation: |
32960
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
90 |
"[| M(r); M(Z) |] ==> |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
91 |
separation (M, \<lambda>x. |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
92 |
\<exists>w[M]. \<exists>wx[M]. \<exists>rp[M]. |
69916a850301
eliminated hard tabulators, guessing at each author's individual tab-width;
wenzelm
parents:
21404
diff
changeset
|
93 |
w \<in> Z & pair(M,w,x,wx) & tran_closure(M,r,rp) & wx \<in> rp)" |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
94 |
and M_nat [iff] : "M(nat)" |
13242 | 95 |
|
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
96 |
lemma (in M_trancl) rtran_closure_mem_iff: |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
97 |
"[|M(A); M(r); M(p)|] |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
98 |
==> rtran_closure_mem(M,A,r,p) \<longleftrightarrow> |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
99 |
(\<exists>n[M]. n\<in>nat & |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
100 |
(\<exists>f[M]. f \<in> succ(n) -> A & |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
101 |
(\<exists>x[M]. \<exists>y[M]. p = <x,y> & f`0 = x & f`n = y) & |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
102 |
(\<forall>i\<in>n. <f`i, f`succ(i)> \<in> r)))" |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
103 |
apply (simp add: rtran_closure_mem_def Ord_succ_mem_iff nat_0_le [THEN ltD]) |
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
104 |
done |
13242 | 105 |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
106 |
lemma (in M_trancl) rtran_closure_rtrancl: |
13242 | 107 |
"M(r) ==> rtran_closure(M,r,rtrancl(r))" |
13324 | 108 |
apply (simp add: rtran_closure_def rtran_closure_mem_iff |
109 |
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
|
110 |
apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) |
13242 | 111 |
done |
112 |
||
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
113 |
lemma (in M_trancl) rtrancl_closed [intro,simp]: |
13242 | 114 |
"M(r) ==> M(rtrancl(r))" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
115 |
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
|
116 |
apply (simp add: rtrancl_alt_eq_rtrancl [symmetric] |
13324 | 117 |
rtrancl_alt_def rtran_closure_mem_iff) |
13242 | 118 |
done |
119 |
||
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
120 |
lemma (in M_trancl) rtrancl_abs [simp]: |
46823 | 121 |
"[| M(r); M(z) |] ==> rtran_closure(M,r,z) \<longleftrightarrow> z = rtrancl(r)" |
13242 | 122 |
apply (rule iffI) |
60770 | 123 |
txt\<open>Proving the right-to-left implication\<close> |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
124 |
prefer 2 apply (blast intro: rtran_closure_rtrancl) |
13242 | 125 |
apply (rule M_equalityI) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
126 |
apply (simp add: rtran_closure_def rtrancl_alt_eq_rtrancl [symmetric] |
13324 | 127 |
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
|
128 |
apply (auto simp add: nat_0_le [THEN ltD] apply_funtype) |
13242 | 129 |
done |
130 |
||
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
131 |
lemma (in M_trancl) trancl_closed [intro,simp]: |
13242 | 132 |
"M(r) ==> M(trancl(r))" |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
133 |
by (simp add: trancl_def) |
13242 | 134 |
|
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
135 |
lemma (in M_trancl) trancl_abs [simp]: |
46823 | 136 |
"[| M(r); M(z) |] ==> tran_closure(M,r,z) \<longleftrightarrow> z = trancl(r)" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
137 |
by (simp add: tran_closure_def trancl_def) |
13242 | 138 |
|
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13306
diff
changeset
|
139 |
lemma (in M_trancl) wellfounded_trancl_separation': |
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13306
diff
changeset
|
140 |
"[| 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
|
141 |
by (insert wellfounded_trancl_separation [of r Z], simp) |
13242 | 142 |
|
61798 | 143 |
text\<open>Alternative proof of \<open>wf_on_trancl\<close>; inspiration for the |
60770 | 144 |
relativized version. Original version is on theory WF.\<close> |
46823 | 145 |
lemma "[| wf[A](r); r-``A \<subseteq> A |] ==> wf[A](r^+)" |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
146 |
apply (simp add: wf_on_def wf_def) |
71417
89d05db6dd1f
Simplified, generalised version of Constructible due to E. Gunther, M. Pagano and P. Sánchez Terraf
paulson <lp15@cam.ac.uk>
parents:
69593
diff
changeset
|
147 |
apply (safe) |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
148 |
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
|
149 |
apply (blast elim: tranclE) |
13242 | 150 |
done |
151 |
||
152 |
lemma (in M_trancl) wellfounded_on_trancl: |
|
46823 | 153 |
"[| wellfounded_on(M,A,r); r-``A \<subseteq> A; M(r); M(A) |] |
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
154 |
==> wellfounded_on(M,A,r^+)" |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
155 |
apply (simp add: wellfounded_on_def) |
13242 | 156 |
apply (safe intro!: equalityI) |
157 |
apply (rename_tac Z x) |
|
13268 | 158 |
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
|
159 |
prefer 2 |
13323
2c287f50c9f3
More relativization, reflection and proofs of separation
paulson
parents:
13306
diff
changeset
|
160 |
apply (blast intro: wellfounded_trancl_separation') |
13299 | 161 |
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
|
162 |
apply (blast dest: transM, simp) |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
163 |
apply (rename_tac y w) |
13242 | 164 |
apply (drule_tac x=w in bspec, assumption, clarify) |
165 |
apply (erule tranclE) |
|
166 |
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
|
167 |
apply blast |
13242 | 168 |
done |
169 |
||
13251
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
170 |
lemma (in M_trancl) wellfounded_trancl: |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
171 |
"[|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
|
172 |
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
|
173 |
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
|
174 |
apply blast |
74cb2af8811e
new treatment of wfrec, replacing wf[A](r) by wf(r)
paulson
parents:
13247
diff
changeset
|
175 |
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
|
176 |
done |
13242 | 177 |
|
13223 | 178 |
|
60770 | 179 |
text\<open>Absoluteness for wfrec-defined functions.\<close> |
13254 | 180 |
|
181 |
(*first use is_recfun, then M_is_recfun*) |
|
182 |
||
183 |
lemma (in M_trancl) wfrec_relativize: |
|
184 |
"[|wf(r); M(a); M(r); |
|
13268 | 185 |
strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M]. |
13254 | 186 |
pair(M,x,y,z) & |
187 |
is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & |
|
188 |
y = H(x, restrict(g, r -`` {x}))); |
|
46823 | 189 |
\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] |
190 |
==> wfrec(r,a,H) = z \<longleftrightarrow> |
|
13268 | 191 |
(\<exists>f[M]. is_recfun(r^+, a, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & |
13254 | 192 |
z = H(a,restrict(f,r-``{a})))" |
193 |
apply (frule wf_trancl) |
|
194 |
apply (simp add: wftrec_def wfrec_def, safe) |
|
195 |
apply (frule wf_exists_is_recfun |
|
196 |
[of concl: "r^+" a "\<lambda>x f. H(x, restrict(f, r -`` {x}))"]) |
|
197 |
apply (simp_all add: trans_trancl function_restrictI trancl_subset_times) |
|
13268 | 198 |
apply (clarify, rule_tac x=x in rexI) |
13254 | 199 |
apply (simp_all add: the_recfun_eq trans_trancl trancl_subset_times) |
200 |
done |
|
201 |
||
202 |
||
69593 | 203 |
text\<open>Assuming \<^term>\<open>r\<close> is transitive simplifies the occurrences of \<open>H\<close>. |
204 |
The premise \<^term>\<open>relation(r)\<close> is necessary |
|
205 |
before we can replace \<^term>\<open>r^+\<close> by \<^term>\<open>r\<close>.\<close> |
|
13254 | 206 |
theorem (in M_trancl) trans_wfrec_relativize: |
207 |
"[|wf(r); trans(r); relation(r); M(r); M(a); |
|
13634 | 208 |
wfrec_replacement(M,MH,r); relation2(M,MH,H); |
46823 | 209 |
\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] |
210 |
==> wfrec(r,a,H) = z \<longleftrightarrow> (\<exists>f[M]. is_recfun(r,a,H,f) & z = H(a,f))" |
|
13353 | 211 |
apply (frule wfrec_replacement', assumption+) |
212 |
apply (simp cong: is_recfun_cong |
|
213 |
add: wfrec_relativize trancl_eq_r |
|
214 |
is_recfun_restrict_idem domain_restrict_idem) |
|
215 |
done |
|
13254 | 216 |
|
13353 | 217 |
theorem (in M_trancl) trans_wfrec_abs: |
218 |
"[|wf(r); trans(r); relation(r); M(r); M(a); M(z); |
|
13634 | 219 |
wfrec_replacement(M,MH,r); relation2(M,MH,H); |
46823 | 220 |
\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] |
221 |
==> is_wfrec(M,MH,r,a,z) \<longleftrightarrow> z=wfrec(r,a,H)" |
|
13634 | 222 |
by (simp add: trans_wfrec_relativize [THEN iff_sym] is_wfrec_abs, blast) |
223 |
||
13254 | 224 |
|
225 |
lemma (in M_trancl) trans_eq_pair_wfrec_iff: |
|
226 |
"[|wf(r); trans(r); relation(r); M(r); M(y); |
|
13634 | 227 |
wfrec_replacement(M,MH,r); relation2(M,MH,H); |
46823 | 228 |
\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] |
229 |
==> y = <x, wfrec(r, x, H)> \<longleftrightarrow> |
|
13268 | 230 |
(\<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)" |
13293 | 231 |
apply safe |
232 |
apply (simp add: trans_wfrec_relativize [THEN iff_sym, of concl: _ x]) |
|
60770 | 233 |
txt\<open>converse direction\<close> |
13254 | 234 |
apply (rule sym) |
235 |
apply (simp add: trans_wfrec_relativize, blast) |
|
236 |
done |
|
237 |
||
238 |
||
60770 | 239 |
subsection\<open>M is closed under well-founded recursion\<close> |
13254 | 240 |
|
61798 | 241 |
text\<open>Lemma with the awkward premise mentioning \<open>wfrec\<close>.\<close> |
13634 | 242 |
lemma (in M_trancl) wfrec_closed_lemma [rule_format]: |
13254 | 243 |
"[|wf(r); M(r); |
244 |
strong_replacement(M, \<lambda>x y. y = \<langle>x, wfrec(r, x, H)\<rangle>); |
|
46823 | 245 |
\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] |
246 |
==> M(a) \<longrightarrow> M(wfrec(r,a,H))" |
|
13254 | 247 |
apply (rule_tac a=a in wf_induct, assumption+) |
248 |
apply (subst wfrec, assumption, clarify) |
|
249 |
apply (drule_tac x1=x and x="\<lambda>x\<in>r -`` {x}. wfrec(r, x, H)" |
|
250 |
in rspec [THEN rspec]) |
|
251 |
apply (simp_all add: function_lam) |
|
13505 | 252 |
apply (blast intro: lam_closed dest: pair_components_in_M) |
13254 | 253 |
done |
254 |
||
60770 | 255 |
text\<open>Eliminates one instance of replacement.\<close> |
13634 | 256 |
lemma (in M_trancl) wfrec_replacement_iff: |
13353 | 257 |
"strong_replacement(M, \<lambda>x z. |
46823 | 258 |
\<exists>y[M]. pair(M,x,y,z) & (\<exists>g[M]. is_recfun(r,x,H,g) & y = H(x,g))) \<longleftrightarrow> |
13254 | 259 |
strong_replacement(M, |
13268 | 260 |
\<lambda>x y. \<exists>f[M]. is_recfun(r,x,H,f) & y = <x, H(x,f)>)" |
13254 | 261 |
apply simp |
262 |
apply (rule strong_replacement_cong, blast) |
|
263 |
done |
|
264 |
||
60770 | 265 |
text\<open>Useful version for transitive relations\<close> |
13634 | 266 |
theorem (in M_trancl) trans_wfrec_closed: |
13254 | 267 |
"[|wf(r); trans(r); relation(r); M(r); M(a); |
13634 | 268 |
wfrec_replacement(M,MH,r); relation2(M,MH,H); |
46823 | 269 |
\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] |
13254 | 270 |
==> M(wfrec(r,a,H))" |
13353 | 271 |
apply (frule wfrec_replacement', assumption+) |
13254 | 272 |
apply (frule wfrec_replacement_iff [THEN iffD1]) |
273 |
apply (rule wfrec_closed_lemma, assumption+) |
|
274 |
apply (simp_all add: wfrec_replacement_iff trans_eq_pair_wfrec_iff) |
|
275 |
done |
|
276 |
||
60770 | 277 |
subsection\<open>Absoluteness without assuming transitivity\<close> |
13254 | 278 |
lemma (in M_trancl) eq_pair_wfrec_iff: |
279 |
"[|wf(r); M(r); M(y); |
|
13268 | 280 |
strong_replacement(M, \<lambda>x z. \<exists>y[M]. \<exists>g[M]. |
13254 | 281 |
pair(M,x,y,z) & |
282 |
is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), g) & |
|
283 |
y = H(x, restrict(g, r -`` {x}))); |
|
46823 | 284 |
\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g))|] |
285 |
==> y = <x, wfrec(r, x, H)> \<longleftrightarrow> |
|
13268 | 286 |
(\<exists>f[M]. is_recfun(r^+, x, \<lambda>x f. H(x, restrict(f, r -`` {x})), f) & |
13254 | 287 |
y = <x, H(x,restrict(f,r-``{x}))>)" |
288 |
apply safe |
|
13293 | 289 |
apply (simp add: wfrec_relativize [THEN iff_sym, of concl: _ x]) |
60770 | 290 |
txt\<open>converse direction\<close> |
13254 | 291 |
apply (rule sym) |
292 |
apply (simp add: wfrec_relativize, blast) |
|
293 |
done |
|
294 |
||
60770 | 295 |
text\<open>Full version not assuming transitivity, but maybe not very useful.\<close> |
13634 | 296 |
theorem (in M_trancl) wfrec_closed: |
13254 | 297 |
"[|wf(r); M(r); M(a); |
13353 | 298 |
wfrec_replacement(M,MH,r^+); |
13634 | 299 |
relation2(M,MH, \<lambda>x f. H(x, restrict(f, r -`` {x}))); |
46823 | 300 |
\<forall>x[M]. \<forall>g[M]. function(g) \<longrightarrow> M(H(x,g)) |] |
13254 | 301 |
==> M(wfrec(r,a,H))" |
13353 | 302 |
apply (frule wfrec_replacement' |
303 |
[of MH "r^+" "\<lambda>x f. H(x, restrict(f, r -`` {x}))"]) |
|
304 |
prefer 4 |
|
305 |
apply (frule wfrec_replacement_iff [THEN iffD1]) |
|
306 |
apply (rule wfrec_closed_lemma, assumption+) |
|
307 |
apply (simp_all add: eq_pair_wfrec_iff func.function_restrictI) |
|
13254 | 308 |
done |
309 |
||
13223 | 310 |
end |