author | wenzelm |
Sun, 25 Oct 1998 12:33:27 +0100 | |
changeset 5769 | 6a422b22ba02 |
parent 5529 | 4a54acae6a15 |
child 6046 | 2c8a8be36c94 |
permissions | -rw-r--r-- |
1461 | 1 |
(* Title: ZF/ex/Limit |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2 |
ID: $Id$ |
1461 | 3 |
Author: Sten Agerholm |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
4 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
5 |
The inverse limit construction. |
5136 | 6 |
|
7 |
(Proofs tidied up considerably by lcp) |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
8 |
*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
9 |
|
2469 | 10 |
val nat_linear_le = [nat_into_Ord,nat_into_Ord] MRS Ord_linear_le; |
11 |
||
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
12 |
open Limit; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
13 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
14 |
(*----------------------------------------------------------------------*) |
1461 | 15 |
(* Useful goal commands. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
16 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
17 |
|
4152 | 18 |
val brr = fn thl => fn n => by (REPEAT(ares_tac thl n)); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
19 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
20 |
(*----------------------------------------------------------------------*) |
1461 | 21 |
(* Basic results. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
22 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
23 |
|
5136 | 24 |
Goalw [set_def] "x:fst(D) ==> x:set(D)"; |
25 |
by (assume_tac 1); |
|
3425 | 26 |
qed "set_I"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
27 |
|
5136 | 28 |
Goalw [rel_def] "<x,y>:snd(D) ==> rel(D,x,y)"; |
29 |
by (assume_tac 1); |
|
3425 | 30 |
qed "rel_I"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
31 |
|
5136 | 32 |
Goalw [rel_def] "rel(D,x,y) ==> <x,y>:snd(D)"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
33 |
by (assume_tac 1); |
3425 | 34 |
qed "rel_E"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
35 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
36 |
(*----------------------------------------------------------------------*) |
1461 | 37 |
(* I/E/D rules for po and cpo. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
38 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
39 |
|
5136 | 40 |
Goalw [po_def] "[|po(D); x:set(D)|] ==> rel(D,x,x)"; |
41 |
by (Blast_tac 1); |
|
3425 | 42 |
qed "po_refl"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
43 |
|
5136 | 44 |
Goalw [po_def] "[|po(D); rel(D,x,y); rel(D,y,z); x:set(D); \ |
45 |
\ y:set(D); z:set(D)|] ==> rel(D,x,z)"; |
|
46 |
by (Blast_tac 1); |
|
3425 | 47 |
qed "po_trans"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
48 |
|
5136 | 49 |
Goalw [po_def] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
50 |
"[|po(D); rel(D,x,y); rel(D,y,x); x:set(D); y:set(D)|] ==> x = y"; |
5136 | 51 |
by (Blast_tac 1); |
3425 | 52 |
qed "po_antisym"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
53 |
|
5136 | 54 |
val prems = Goalw [po_def] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
55 |
"[| !!x. x:set(D) ==> rel(D,x,x); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
56 |
\ !!x y z. [| rel(D,x,y); rel(D,y,z); x:set(D); y:set(D); z:set(D)|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
57 |
\ rel(D,x,z); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
58 |
\ !!x y. [| rel(D,x,y); rel(D,y,x); x:set(D); y:set(D)|] ==> x=y |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
59 |
\ po(D)"; |
4152 | 60 |
by Safe_tac; |
5136 | 61 |
by (REPEAT (ares_tac prems 1)); |
3425 | 62 |
qed "poI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
63 |
|
5136 | 64 |
val prems = Goalw [cpo_def] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
65 |
"[| po(D); !!X. chain(D,X) ==> islub(D,X,x(D,X))|] ==> cpo(D)"; |
4091 | 66 |
by (safe_tac (claset() addSIs [exI])); |
5136 | 67 |
by (REPEAT (ares_tac prems 1)); |
3425 | 68 |
qed "cpoI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
69 |
|
5136 | 70 |
Goalw [cpo_def] "cpo(D) ==> po(D)"; |
71 |
by (Blast_tac 1); |
|
3425 | 72 |
qed "cpo_po"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
73 |
|
5136 | 74 |
Goal "[|cpo(D); x:set(D)|] ==> rel(D,x,x)"; |
75 |
by (blast_tac (claset() addIs [po_refl, cpo_po]) 1); |
|
3425 | 76 |
qed "cpo_refl"; |
77 |
Addsimps [cpo_refl]; |
|
5136 | 78 |
AddSIs [cpo_refl]; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
79 |
|
5136 | 80 |
Goal "[|cpo(D); rel(D,x,y); rel(D,y,z); x:set(D); \ |
81 |
\ y:set(D); z:set(D)|] ==> rel(D,x,z)"; |
|
82 |
by (blast_tac (claset() addIs [cpo_po, po_trans]) 1); |
|
3425 | 83 |
qed "cpo_trans"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
84 |
|
5136 | 85 |
Goal "[|cpo(D); rel(D,x,y); rel(D,y,x); x:set(D); y:set(D)|] ==> x = y"; |
86 |
by (blast_tac (claset() addIs [cpo_po, po_antisym]) 1); |
|
3425 | 87 |
qed "cpo_antisym"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
88 |
|
5136 | 89 |
val [cpo,chain,ex] = Goalw [cpo_def] |
2469 | 90 |
"[|cpo(D); chain(D,X); !!x. islub(D,X,x) ==> R|] ==> R"; |
1623 | 91 |
by (rtac (chain RS (cpo RS conjunct2 RS spec RS mp) RS exE) 1); |
5136 | 92 |
by (etac ex 1); |
3425 | 93 |
qed "cpo_islub"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
94 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
95 |
(*----------------------------------------------------------------------*) |
1461 | 96 |
(* Theorems about isub and islub. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
97 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
98 |
|
5136 | 99 |
Goalw [islub_def] "islub(D,X,x) ==> isub(D,X,x)"; |
100 |
by (Asm_simp_tac 1); |
|
3425 | 101 |
qed "islub_isub"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
102 |
|
5136 | 103 |
Goalw [islub_def,isub_def] "islub(D,X,x) ==> x:set(D)"; |
104 |
by (Asm_simp_tac 1); |
|
3425 | 105 |
qed "islub_in"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
106 |
|
5136 | 107 |
Goalw [islub_def,isub_def] "[|islub(D,X,x); n:nat|] ==> rel(D,X`n,x)"; |
108 |
by (Asm_simp_tac 1); |
|
3425 | 109 |
qed "islub_ub"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
110 |
|
5136 | 111 |
Goalw [islub_def] "[|islub(D,X,x); isub(D,X,y)|] ==> rel(D,x,y)"; |
112 |
by (Blast_tac 1); |
|
3425 | 113 |
qed "islub_least"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
114 |
|
5136 | 115 |
val prems = Goalw [islub_def] (* islubI *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
116 |
"[|isub(D,X,x); !!y. isub(D,X,y) ==> rel(D,x,y)|] ==> islub(D,X,x)"; |
4152 | 117 |
by Safe_tac; |
1623 | 118 |
by (REPEAT(ares_tac prems 1)); |
3425 | 119 |
qed "islubI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
120 |
|
5136 | 121 |
val prems = Goalw [isub_def] (* isubI *) |
2469 | 122 |
"[|x:set(D); !!n. n:nat ==> rel(D,X`n,x)|] ==> isub(D,X,x)"; |
4152 | 123 |
by Safe_tac; |
1623 | 124 |
by (REPEAT(ares_tac prems 1)); |
3425 | 125 |
qed "isubI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
126 |
|
5136 | 127 |
val prems = Goalw [isub_def] (* isubE *) |
128 |
"[|isub(D,X,x); [|x:set(D); !!n. n:nat==>rel(D,X`n,x)|] ==> P \ |
|
129 |
\ |] ==> P"; |
|
130 |
by (asm_simp_tac (simpset() addsimps prems) 1); |
|
3425 | 131 |
qed "isubE"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
132 |
|
5136 | 133 |
Goalw [isub_def] "isub(D,X,x) ==> x:set(D)"; |
134 |
by (Asm_simp_tac 1); |
|
3425 | 135 |
qed "isubD1"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
136 |
|
5136 | 137 |
Goalw [isub_def] "[|isub(D,X,x); n:nat|]==>rel(D,X`n,x)"; |
138 |
by (Asm_simp_tac 1); |
|
3425 | 139 |
qed "isubD2"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
140 |
|
5136 | 141 |
Goal "[|islub(D,X,x); islub(D,X,y); cpo(D)|] ==> x = y"; |
142 |
by (blast_tac (claset() addIs [cpo_antisym,islub_least, |
|
143 |
islub_isub,islub_in]) 1); |
|
3425 | 144 |
qed "islub_unique"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
145 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
146 |
(*----------------------------------------------------------------------*) |
1461 | 147 |
(* lub gives the least upper bound of chains. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
148 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
149 |
|
5136 | 150 |
Goalw [lub_def] "[|chain(D,X); cpo(D)|] ==> islub(D,X,lub(D,X))"; |
151 |
by (best_tac (claset() addEs [cpo_islub] addIs [theI, islub_unique]) 1); |
|
3425 | 152 |
qed "cpo_lub"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
153 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
154 |
(*----------------------------------------------------------------------*) |
1461 | 155 |
(* Theorems about chains. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
156 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
157 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
158 |
val chainI = prove_goalw Limit.thy [chain_def] |
2469 | 159 |
"!!z.[|X:nat->set(D); !!n. n:nat ==> rel(D,X`n,X`succ(n))|] ==> chain(D,X)" |
160 |
(fn prems => [Asm_simp_tac 1]); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
161 |
|
5136 | 162 |
Goalw [chain_def] "chain(D,X) ==> X : nat -> set(D)"; |
163 |
by (Asm_simp_tac 1); |
|
3425 | 164 |
qed "chain_fun"; |
5136 | 165 |
|
166 |
Goalw [chain_def] "[|chain(D,X); n:nat|] ==> X`n : set(D)"; |
|
167 |
by (blast_tac (claset() addDs [apply_type]) 1); |
|
3425 | 168 |
qed "chain_in"; |
5136 | 169 |
|
170 |
Goalw [chain_def] "[|chain(D,X); n:nat|] ==> rel(D, X ` n, X ` succ(n))"; |
|
171 |
by (Blast_tac 1); |
|
3425 | 172 |
qed "chain_rel"; |
5136 | 173 |
|
174 |
Addsimps [chain_in, chain_rel]; |
|
175 |
||
176 |
Goal "[|chain(D,X); cpo(D); n:nat; m:nat|] ==> rel(D,X`n,(X`(m #+ n)))"; |
|
1623 | 177 |
by (res_inst_tac [("n","m")] nat_induct 1); |
2469 | 178 |
by (ALLGOALS Simp_tac); |
5136 | 179 |
by (rtac cpo_trans 2); (* loops if repeated *) |
180 |
by (REPEAT (ares_tac [cpo_refl,chain_in,chain_rel,nat_succI,add_type] 1)); |
|
3425 | 181 |
qed "chain_rel_gen_add"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
182 |
|
5136 | 183 |
Goal "[| n le succ(x); ~ n le x; x : nat; n:nat |] ==> n = succ(x)"; |
184 |
by (etac le_anti_sym 1); |
|
185 |
by (asm_simp_tac (simpset() addsimps [not_le_iff_lt RS iff_sym, |
|
186 |
nat_into_Ord]) 1); |
|
3425 | 187 |
qed "le_succ_eq"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
188 |
|
5136 | 189 |
Goal (* chain_rel_gen *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
190 |
"[|n le m; chain(D,X); cpo(D); n:nat; m:nat|] ==> rel(D,X`n,X`m)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
191 |
by (rtac impE 1); (* The first three steps prepare for the induction proof *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
192 |
by (assume_tac 3); |
5136 | 193 |
by (assume_tac 2); |
1623 | 194 |
by (res_inst_tac [("n","m")] nat_induct 1); |
4152 | 195 |
by Safe_tac; |
5136 | 196 |
by (Asm_full_simp_tac 1); |
197 |
by (rtac cpo_trans 2); |
|
198 |
by (rtac (le_succ_eq RS subst) 1); |
|
199 |
by (auto_tac (claset() addIs [chain_in,chain_rel], |
|
200 |
simpset())); |
|
3425 | 201 |
qed "chain_rel_gen"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
202 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
203 |
(*----------------------------------------------------------------------*) |
1461 | 204 |
(* Theorems about pcpos and bottom. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
205 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
206 |
|
5136 | 207 |
val prems = Goalw [pcpo_def] (* pcpoI *) |
3840 | 208 |
"[|!!y. y:set(D)==>rel(D,x,y); x:set(D); cpo(D)|]==>pcpo(D)"; |
5136 | 209 |
by (auto_tac (claset() addIs prems, simpset())); |
3425 | 210 |
qed "pcpoI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
211 |
|
5136 | 212 |
Goalw [pcpo_def] "pcpo(D) ==> cpo(D)"; |
213 |
by (etac conjunct1 1); |
|
214 |
qed "pcpo_cpo"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
215 |
|
5136 | 216 |
Goalw [pcpo_def] (* pcpo_bot_ex1 *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
217 |
"pcpo(D) ==> EX! x. x:set(D) & (ALL y:set(D). rel(D,x,y))"; |
5136 | 218 |
by (blast_tac (claset() addIs [cpo_antisym]) 1); |
3425 | 219 |
qed "pcpo_bot_ex1"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
220 |
|
5136 | 221 |
Goalw [bot_def] (* bot_least *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
222 |
"[| pcpo(D); y:set(D)|] ==> rel(D,bot(D),y)"; |
5136 | 223 |
by (best_tac (claset() addIs [pcpo_bot_ex1 RS theI2]) 1); |
3425 | 224 |
qed "bot_least"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
225 |
|
5136 | 226 |
Goalw [bot_def] (* bot_in *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
227 |
"pcpo(D) ==> bot(D):set(D)"; |
5136 | 228 |
by (best_tac (claset() addIs [pcpo_bot_ex1 RS theI2]) 1); |
3425 | 229 |
qed "bot_in"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
230 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
231 |
val prems = goal Limit.thy (* bot_unique *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
232 |
"[| pcpo(D); x:set(D); !!y. y:set(D) ==> rel(D,x,y)|] ==> x = bot(D)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
233 |
by (rtac cpo_antisym 1); |
1623 | 234 |
brr(pcpo_cpo::bot_in::bot_least::prems) 1; |
3425 | 235 |
qed "bot_unique"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
236 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
237 |
(*----------------------------------------------------------------------*) |
1461 | 238 |
(* Constant chains and lubs and cpos. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
239 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
240 |
|
5136 | 241 |
Goalw [chain_def] "[|x:set(D); cpo(D)|] ==> chain(D,(lam n:nat. x))"; |
242 |
by (asm_simp_tac (simpset() addsimps [lam_type, nat_succI]) 1); |
|
3425 | 243 |
qed "chain_const"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
244 |
|
5136 | 245 |
Goalw [islub_def,isub_def] |
246 |
"[|x:set(D); cpo(D)|] ==> islub(D,(lam n:nat. x),x)"; |
|
3425 | 247 |
by (Asm_simp_tac 1); |
248 |
by (Blast_tac 1); |
|
249 |
qed "islub_const"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
250 |
|
5136 | 251 |
Goal "[|x:set(D); cpo(D)|] ==> lub(D,lam n:nat. x) = x"; |
252 |
by (blast_tac (claset() addIs [islub_unique, cpo_lub, |
|
253 |
chain_const, islub_const]) 1); |
|
3425 | 254 |
qed "lub_const"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
255 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
256 |
(*----------------------------------------------------------------------*) |
1461 | 257 |
(* Taking the suffix of chains has no effect on ub's. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
258 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
259 |
|
5136 | 260 |
Goalw [isub_def,suffix_def] (* isub_suffix *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
261 |
"[|chain(D,X); cpo(D); n:nat|] ==> isub(D,suffix(X,n),x) <-> isub(D,X,x)"; |
5136 | 262 |
by (Asm_simp_tac 1); |
4152 | 263 |
by Safe_tac; |
5136 | 264 |
by (blast_tac (claset() addIs [chain_in, add_type]) 2); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
265 |
by (rtac cpo_trans 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
266 |
by (rtac chain_rel_gen_add 2); |
5136 | 267 |
by Auto_tac; |
3425 | 268 |
qed "isub_suffix"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
269 |
|
5136 | 270 |
Goalw [islub_def] (* islub_suffix *) |
271 |
"[|chain(D,X); cpo(D); n:nat|] ==> islub(D,suffix(X,n),x) <-> islub(D,X,x)"; |
|
272 |
by (asm_simp_tac (simpset() addsimps [isub_suffix]) 1); |
|
3425 | 273 |
qed "islub_suffix"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
274 |
|
5136 | 275 |
Goalw [lub_def] (* lub_suffix *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
276 |
"[|chain(D,X); cpo(D); n:nat|] ==> lub(D,suffix(X,n)) = lub(D,X)"; |
5136 | 277 |
by (asm_simp_tac (simpset() addsimps [islub_suffix]) 1); |
3425 | 278 |
qed "lub_suffix"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
279 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
280 |
(*----------------------------------------------------------------------*) |
1461 | 281 |
(* Dominate and subchain. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
282 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
283 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
284 |
val dominateI = prove_goalw Limit.thy [dominate_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
285 |
"[| !!m. m:nat ==> n(m):nat; !!m. m:nat ==> rel(D,X`m,Y`n(m))|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
286 |
\ dominate(D,X,Y)" |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
287 |
(fn prems => [rtac ballI 1,rtac bexI 1,REPEAT(ares_tac prems 1)]); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
288 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
289 |
val [dom,isub,cpo,X,Y] = goal Limit.thy |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
290 |
"[|dominate(D,X,Y); isub(D,Y,x); cpo(D); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
291 |
\ X:nat->set(D); Y:nat->set(D)|] ==> isub(D,X,x)"; |
1623 | 292 |
by (rewtac isub_def); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
293 |
by (rtac conjI 1); |
1623 | 294 |
by (rtac (rewrite_rule[isub_def]isub RS conjunct1) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
295 |
by (rtac ballI 1); |
1623 | 296 |
by (rtac (rewrite_rule[dominate_def]dom RS bspec RS bexE) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
297 |
by (assume_tac 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
298 |
by (rtac cpo_trans 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
299 |
by (rtac cpo 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
300 |
by (assume_tac 1); |
1623 | 301 |
by (rtac (rewrite_rule[isub_def]isub RS conjunct2 RS bspec) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
302 |
by (assume_tac 1); |
1623 | 303 |
by (etac (X RS apply_type) 1); |
304 |
by (etac (Y RS apply_type) 1); |
|
305 |
by (rtac (rewrite_rule[isub_def]isub RS conjunct1) 1); |
|
3425 | 306 |
qed "dominate_isub"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
307 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
308 |
val [dom,Xlub,Ylub,cpo,X,Y] = goal Limit.thy |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
309 |
"[|dominate(D,X,Y); islub(D,X,x); islub(D,Y,y); cpo(D); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
310 |
\ X:nat->set(D); Y:nat->set(D)|] ==> rel(D,x,y)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
311 |
val Xub = rewrite_rule[islub_def]Xlub RS conjunct1; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
312 |
val Yub = rewrite_rule[islub_def]Ylub RS conjunct1; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
313 |
val Xub_y = Yub RS (dom RS dominate_isub); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
314 |
val lem = Xub_y RS (rewrite_rule[islub_def]Xlub RS conjunct2 RS spec RS mp); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
315 |
val thm = Y RS (X RS (cpo RS lem)); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
316 |
by (rtac thm 1); |
3425 | 317 |
qed "dominate_islub"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
318 |
|
5136 | 319 |
val prems = Goalw [subchain_def] (* subchainE *) |
2469 | 320 |
"[|subchain(X,Y); n:nat; !!m. [|m:nat; X`n = Y`(n #+ m)|] ==> Q|] ==> Q"; |
1623 | 321 |
by (rtac (hd prems RS bspec RS bexE) 1); |
322 |
by (resolve_tac prems 2); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
323 |
by (assume_tac 3); |
1623 | 324 |
by (REPEAT(ares_tac prems 1)); |
3425 | 325 |
qed "subchainE"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
326 |
|
5136 | 327 |
Goal "[|subchain(Y,X); isub(D,X,x)|] ==> isub(D,Y,x)"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
328 |
by (rtac isubI 1); |
5136 | 329 |
by (safe_tac (claset() addSEs [isubE, subchainE])); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
330 |
by (assume_tac 1); |
2469 | 331 |
by (Asm_simp_tac 1); |
3425 | 332 |
qed "subchain_isub"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
333 |
|
5268 | 334 |
Goal "[|dominate(D,X,Y); subchain(Y,X); islub(D,X,x); islub(D,Y,y); cpo(D); \ |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
335 |
\ X:nat->set(D); Y:nat->set(D)|] ==> x = y"; |
5136 | 336 |
by (blast_tac (claset() addIs [cpo_antisym, dominate_islub, islub_least, |
337 |
subchain_isub, islub_isub, islub_in]) 1); |
|
3425 | 338 |
qed "dominate_islub_eq"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
339 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
340 |
(*----------------------------------------------------------------------*) |
1461 | 341 |
(* Matrix. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
342 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
343 |
|
5136 | 344 |
Goalw [matrix_def] (* matrix_fun *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
345 |
"matrix(D,M) ==> M : nat -> (nat -> set(D))"; |
5136 | 346 |
by (Asm_simp_tac 1); |
3425 | 347 |
qed "matrix_fun"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
348 |
|
5136 | 349 |
Goal "[|matrix(D,M); n:nat|] ==> M`n : nat -> set(D)"; |
350 |
by (blast_tac (claset() addIs [apply_funtype, matrix_fun]) 1); |
|
3425 | 351 |
qed "matrix_in_fun"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
352 |
|
5136 | 353 |
Goal "[|matrix(D,M); n:nat; m:nat|] ==> M`n`m : set(D)"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
354 |
by (rtac apply_type 1); |
5136 | 355 |
by (REPEAT(ares_tac[matrix_in_fun] 1)); |
3425 | 356 |
qed "matrix_in"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
357 |
|
5136 | 358 |
Goalw [matrix_def] (* matrix_rel_1_0 *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
359 |
"[|matrix(D,M); n:nat; m:nat|] ==> rel(D,M`n`m,M`succ(n)`m)"; |
5136 | 360 |
by (Asm_simp_tac 1); |
3425 | 361 |
qed "matrix_rel_1_0"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
362 |
|
5136 | 363 |
Goalw [matrix_def] (* matrix_rel_0_1 *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
364 |
"[|matrix(D,M); n:nat; m:nat|] ==> rel(D,M`n`m,M`n`succ(m))"; |
5136 | 365 |
by (Asm_simp_tac 1); |
3425 | 366 |
qed "matrix_rel_0_1"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
367 |
|
5136 | 368 |
Goalw [matrix_def] (* matrix_rel_1_1 *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
369 |
"[|matrix(D,M); n:nat; m:nat|] ==> rel(D,M`n`m,M`succ(n)`succ(m))"; |
5136 | 370 |
by (Asm_simp_tac 1); |
3425 | 371 |
qed "matrix_rel_1_1"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
372 |
|
5136 | 373 |
Goal "f:X->Y->Z ==> (lam y:Y. lam x:X. f`x`y):Y->X->Z"; |
374 |
by (blast_tac (claset() addIs [lam_type, apply_funtype]) 1); |
|
3425 | 375 |
qed "fun_swap"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
376 |
|
5136 | 377 |
Goalw [matrix_def] (* matrix_sym_axis *) |
378 |
"matrix(D,M) ==> matrix(D,lam m:nat. lam n:nat. M`n`m)"; |
|
379 |
by (asm_simp_tac (simpset() addsimps [fun_swap]) 1); |
|
3425 | 380 |
qed "matrix_sym_axis"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
381 |
|
5136 | 382 |
Goalw [chain_def] (* matrix_chain_diag *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
383 |
"matrix(D,M) ==> chain(D,lam n:nat. M`n`n)"; |
4152 | 384 |
by Safe_tac; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
385 |
by (rtac lam_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
386 |
by (rtac matrix_in 1); |
1623 | 387 |
by (REPEAT(ares_tac prems 1)); |
2469 | 388 |
by (Asm_simp_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
389 |
by (rtac matrix_rel_1_1 1); |
1623 | 390 |
by (REPEAT(ares_tac prems 1)); |
3425 | 391 |
qed "matrix_chain_diag"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
392 |
|
5136 | 393 |
Goalw [chain_def] (* matrix_chain_left *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
394 |
"[|matrix(D,M); n:nat|] ==> chain(D,M`n)"; |
4152 | 395 |
by Safe_tac; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
396 |
by (rtac apply_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
397 |
by (rtac matrix_fun 1); |
1623 | 398 |
by (REPEAT(ares_tac prems 1)); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
399 |
by (rtac matrix_rel_0_1 1); |
1623 | 400 |
by (REPEAT(ares_tac prems 1)); |
3425 | 401 |
qed "matrix_chain_left"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
402 |
|
5136 | 403 |
Goalw [chain_def] (* matrix_chain_right *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
404 |
"[|matrix(D,M); m:nat|] ==> chain(D,lam n:nat. M`n`m)"; |
5136 | 405 |
by (auto_tac (claset() addIs [lam_type,matrix_in,matrix_rel_1_0], |
406 |
simpset())); |
|
3425 | 407 |
qed "matrix_chain_right"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
408 |
|
5136 | 409 |
val prems = Goalw [matrix_def] (* matrix_chainI *) |
3840 | 410 |
"[|!!x. x:nat==>chain(D,M`x); !!y. y:nat==>chain(D,lam x:nat. M`x`y); \ |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
411 |
\ M:nat->nat->set(D); cpo(D)|] ==> matrix(D,M)"; |
4091 | 412 |
by (safe_tac (claset() addSIs [ballI])); |
1623 | 413 |
by (cut_inst_tac[("y1","m"),("n","n")](hd(tl prems) RS chain_rel) 2); |
2469 | 414 |
by (Asm_full_simp_tac 4); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
415 |
by (rtac cpo_trans 5); |
1623 | 416 |
by (cut_inst_tac[("y1","m"),("n","n")](hd(tl prems) RS chain_rel) 6); |
2469 | 417 |
by (Asm_full_simp_tac 8); |
1623 | 418 |
by (TRYALL(rtac (chain_fun RS apply_type))); |
419 |
brr(chain_rel::nat_succI::prems) 1; |
|
3425 | 420 |
qed "matrix_chainI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
421 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
422 |
val lemma = prove_goal Limit.thy |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
423 |
"!!z.[|m : nat; rel(D, (lam n:nat. M`n`n)`m, y)|] ==> rel(D,M`m`m, y)" |
2469 | 424 |
(fn prems => [Asm_full_simp_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
425 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
426 |
val lemma2 = prove_goal Limit.thy |
3840 | 427 |
"!!z.[|x:nat; m:nat; rel(D,(lam n:nat. M`n`m1)`x,(lam n:nat. M`n`m1)`m)|] ==> \ |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
428 |
\ rel(D,M`x`m1,M`m`m1)" |
2469 | 429 |
(fn prems => [Asm_full_simp_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
430 |
|
5136 | 431 |
Goalw [isub_def] (* isub_lemma *) |
432 |
"[|isub(D, lam n:nat. M`n`n, y); matrix(D,M); cpo(D)|] ==> \ |
|
433 |
\ isub(D, lam n:nat. lub(D,lam m:nat. M`n`m), y)"; |
|
4152 | 434 |
by Safe_tac; |
2469 | 435 |
by (Asm_simp_tac 1); |
5136 | 436 |
by (forward_tac [matrix_fun RS apply_type] 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
437 |
by (assume_tac 1); |
2469 | 438 |
by (Asm_simp_tac 1); |
5136 | 439 |
by (rtac (matrix_chain_left RS cpo_lub RS islub_least) 1); |
440 |
by (REPEAT (assume_tac 1)); |
|
1623 | 441 |
by (rewtac isub_def); |
4152 | 442 |
by Safe_tac; |
5136 | 443 |
by (excluded_middle_tac "n le na" 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
444 |
by (rtac cpo_trans 1); |
5136 | 445 |
by (assume_tac 1); |
1623 | 446 |
by (rtac (not_le_iff_lt RS iffD1 RS leI RS chain_rel_gen) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
447 |
by (assume_tac 3); |
5136 | 448 |
by (REPEAT(ares_tac [nat_into_Ord,matrix_chain_left] 1)); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
449 |
by (rtac lemma 1); |
5136 | 450 |
by (assume_tac 1); |
451 |
by (Blast_tac 1); |
|
452 |
by (REPEAT(ares_tac [matrix_in] 1)); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
453 |
by (rtac cpo_trans 1); |
5136 | 454 |
by (assume_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
455 |
by (rtac lemma2 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
456 |
by (rtac lemma 4); |
5136 | 457 |
by (Blast_tac 5); |
458 |
by (REPEAT(ares_tac [chain_rel_gen,matrix_chain_right,matrix_in,isubD1] 1)); |
|
3425 | 459 |
qed "isub_lemma"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
460 |
|
5136 | 461 |
Goalw [chain_def] (* matrix_chain_lub *) |
3840 | 462 |
"[|matrix(D,M); cpo(D)|] ==> chain(D,lam n:nat. lub(D,lam m:nat. M`n`m))"; |
4152 | 463 |
by Safe_tac; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
464 |
by (rtac lam_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
465 |
by (rtac islub_in 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
466 |
by (rtac cpo_lub 1); |
5136 | 467 |
by (assume_tac 2); |
2469 | 468 |
by (Asm_simp_tac 2); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
469 |
by (rtac chainI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
470 |
by (rtac lam_type 1); |
5136 | 471 |
by (REPEAT(ares_tac [matrix_in] 1)); |
2469 | 472 |
by (Asm_simp_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
473 |
by (rtac matrix_rel_0_1 1); |
5136 | 474 |
by (REPEAT(assume_tac 1)); |
4091 | 475 |
by (asm_simp_tac (simpset() addsimps |
5136 | 476 |
[matrix_chain_left RS chain_fun RS eta]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
477 |
by (rtac dominate_islub 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
478 |
by (rtac cpo_lub 3); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
479 |
by (rtac cpo_lub 2); |
1623 | 480 |
by (rewtac dominate_def); |
5136 | 481 |
by (REPEAT(ares_tac [matrix_chain_left,nat_succI,chain_fun] 2)); |
482 |
by (blast_tac (claset() addIs [matrix_rel_1_0]) 1); |
|
3425 | 483 |
qed "matrix_chain_lub"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
484 |
|
5136 | 485 |
Goal (* isub_eq *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
486 |
"[|matrix(D,M); cpo(D)|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
487 |
\ isub(D,(lam n:nat. lub(D,lam m:nat. M`n`m)),y) <-> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
488 |
\ isub(D,(lam n:nat. M`n`n),y)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
489 |
by (rtac iffI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
490 |
by (rtac dominate_isub 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
491 |
by (assume_tac 2); |
1623 | 492 |
by (rewtac dominate_def); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
493 |
by (rtac ballI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
494 |
by (rtac bexI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
495 |
by (assume_tac 2); |
2469 | 496 |
by (Asm_simp_tac 1); |
4091 | 497 |
by (asm_simp_tac (simpset() addsimps |
5136 | 498 |
[matrix_chain_left RS chain_fun RS eta]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
499 |
by (rtac islub_ub 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
500 |
by (rtac cpo_lub 1); |
1623 | 501 |
by (REPEAT(ares_tac |
5136 | 502 |
[matrix_chain_left,matrix_chain_diag,chain_fun,matrix_chain_lub] 1)); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
503 |
by (rtac isub_lemma 1); |
5136 | 504 |
by (REPEAT(assume_tac 1)); |
3425 | 505 |
qed "isub_eq"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
506 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
507 |
val lemma1 = prove_goalw Limit.thy [lub_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
508 |
"lub(D,(lam n:nat. lub(D,lam m:nat. M`n`m))) = \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
509 |
\ (THE x. islub(D, (lam n:nat. lub(D,lam m:nat. M`n`m)), x))" |
5136 | 510 |
(fn _ => [Fast_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
511 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
512 |
val lemma2 = prove_goalw Limit.thy [lub_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
513 |
"lub(D,(lam n:nat. M`n`n)) = \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
514 |
\ (THE x. islub(D, (lam n:nat. M`n`n), x))" |
5136 | 515 |
(fn _ => [Fast_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
516 |
|
5136 | 517 |
Goal (* lub_matrix_diag *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
518 |
"[|matrix(D,M); cpo(D)|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
519 |
\ lub(D,(lam n:nat. lub(D,lam m:nat. M`n`m))) = \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
520 |
\ lub(D,(lam n:nat. M`n`n))"; |
4091 | 521 |
by (simp_tac (simpset() addsimps [lemma1,lemma2]) 1); |
1623 | 522 |
by (rewtac islub_def); |
5136 | 523 |
by (asm_simp_tac (simpset() addsimps [isub_eq]) 1); |
3425 | 524 |
qed "lub_matrix_diag"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
525 |
|
5136 | 526 |
Goal (* lub_matrix_diag_sym *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
527 |
"[|matrix(D,M); cpo(D)|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
528 |
\ lub(D,(lam m:nat. lub(D,lam n:nat. M`n`m))) = \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
529 |
\ lub(D,(lam n:nat. M`n`n))"; |
5136 | 530 |
by (dtac (matrix_sym_axis RS lub_matrix_diag) 1); |
531 |
by Auto_tac; |
|
3425 | 532 |
qed "lub_matrix_diag_sym"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
533 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
534 |
(*----------------------------------------------------------------------*) |
1461 | 535 |
(* I/E/D rules for mono and cont. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
536 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
537 |
|
5136 | 538 |
val prems = Goalw [mono_def] (* monoI *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
539 |
"[|f:set(D)->set(E); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
540 |
\ !!x y. [|rel(D,x,y); x:set(D); y:set(D)|] ==> rel(E,f`x,f`y)|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
541 |
\ f:mono(D,E)"; |
5136 | 542 |
by (blast_tac(claset() addSIs prems) 1); |
3425 | 543 |
qed "monoI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
544 |
|
5136 | 545 |
Goalw [mono_def] "f:mono(D,E) ==> f:set(D)->set(E)"; |
546 |
by (Fast_tac 1); |
|
3425 | 547 |
qed "mono_fun"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
548 |
|
5136 | 549 |
Goal "[|f:mono(D,E); x:set(D)|] ==> f`x:set(E)"; |
550 |
by (blast_tac(claset() addSIs [mono_fun RS apply_type]) 1); |
|
3425 | 551 |
qed "mono_map"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
552 |
|
5136 | 553 |
Goalw [mono_def] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
554 |
"[|f:mono(D,E); rel(D,x,y); x:set(D); y:set(D)|] ==> rel(E,f`x,f`y)"; |
5136 | 555 |
by (Blast_tac 1); |
3425 | 556 |
qed "mono_mono"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
557 |
|
5136 | 558 |
val prems = Goalw [cont_def,mono_def] (* contI *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
559 |
"[|f:set(D)->set(E); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
560 |
\ !!x y. [|rel(D,x,y); x:set(D); y:set(D)|] ==> rel(E,f`x,f`y); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
561 |
\ !!X. chain(D,X) ==> f`lub(D,X) = lub(E,lam n:nat. f`(X`n))|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
562 |
\ f:cont(D,E)"; |
4091 | 563 |
by (fast_tac(claset() addSIs prems) 1); |
3425 | 564 |
qed "contI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
565 |
|
5136 | 566 |
Goalw [cont_def] "f:cont(D,E) ==> f:mono(D,E)"; |
567 |
by (Blast_tac 1); |
|
3425 | 568 |
qed "cont2mono"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
569 |
|
5136 | 570 |
Goalw [cont_def] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
571 |
"f:cont(D,E) ==> f:set(D)->set(E)"; |
5136 | 572 |
by (rtac mono_fun 1); |
573 |
by (Blast_tac 1); |
|
3425 | 574 |
qed "cont_fun"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
575 |
|
5136 | 576 |
Goal "[|f:cont(D,E); x:set(D)|] ==> f`x:set(E)"; |
577 |
by (blast_tac(claset() addSIs [cont_fun RS apply_type]) 1); |
|
3425 | 578 |
qed "cont_map"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
579 |
|
5136 | 580 |
Goalw [cont_def] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
581 |
"[|f:cont(D,E); rel(D,x,y); x:set(D); y:set(D)|] ==> rel(E,f`x,f`y)"; |
5136 | 582 |
by (blast_tac(claset() addSIs [mono_mono]) 1); |
3425 | 583 |
qed "cont_mono"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
584 |
|
5136 | 585 |
Goalw [cont_def] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
586 |
"[|f:cont(D,E); chain(D,X)|] ==> f`(lub(D,X)) = lub(E,lam n:nat. f`(X`n))"; |
5136 | 587 |
by (Blast_tac 1); |
3425 | 588 |
qed "cont_lub"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
589 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
590 |
(*----------------------------------------------------------------------*) |
1461 | 591 |
(* Continuity and chains. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
592 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
593 |
|
5136 | 594 |
Goal "[|f:mono(D,E); chain(D,X)|] ==> chain(E,lam n:nat. f`(X`n))"; |
595 |
by (simp_tac (simpset() addsimps [chain_def]) 1); |
|
596 |
by (blast_tac(claset() addIs [lam_type, mono_map, chain_in, |
|
597 |
mono_mono, chain_rel]) 1); |
|
3425 | 598 |
qed "mono_chain"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
599 |
|
5136 | 600 |
Goal "[|f:cont(D,E); chain(D,X)|] ==> chain(E,lam n:nat. f`(X`n))"; |
601 |
by (blast_tac(claset() addIs [mono_chain, cont2mono]) 1); |
|
3425 | 602 |
qed "cont_chain"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
603 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
604 |
(*----------------------------------------------------------------------*) |
1461 | 605 |
(* I/E/D rules about (set+rel) cf, the continuous function space. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
606 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
607 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
608 |
(* The following development more difficult with cpo-as-relation approach. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
609 |
|
5136 | 610 |
Goalw [set_def,cf_def] "f:set(cf(D,E)) ==> f:cont(D,E)"; |
2469 | 611 |
by (Asm_full_simp_tac 1); |
3425 | 612 |
qed "cf_cont"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
613 |
|
5136 | 614 |
Goalw [set_def,cf_def] (* Non-trivial with relation *) |
615 |
"f:cont(D,E) ==> f:set(cf(D,E))"; |
|
2469 | 616 |
by (Asm_full_simp_tac 1); |
3425 | 617 |
qed "cont_cf"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
618 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
619 |
(* rel_cf originally an equality. Now stated as two rules. Seemed easiest. |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
620 |
Besides, now complicated by typing assumptions. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
621 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
622 |
val prems = goal Limit.thy |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
623 |
"[|!!x. x:set(D) ==> rel(E,f`x,g`x); f:cont(D,E); g:cont(D,E)|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
624 |
\ rel(cf(D,E),f,g)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
625 |
by (rtac rel_I 1); |
4091 | 626 |
by (simp_tac (simpset() addsimps [cf_def]) 1); |
4152 | 627 |
by Safe_tac; |
5136 | 628 |
by (REPEAT (ares_tac prems 1)); |
3425 | 629 |
qed "rel_cfI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
630 |
|
5136 | 631 |
Goalw [rel_def,cf_def] "[|rel(cf(D,E),f,g); x:set(D)|] ==> rel(E,f`x,g`x)"; |
2469 | 632 |
by (Asm_full_simp_tac 1); |
3425 | 633 |
qed "rel_cf"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
634 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
635 |
(*----------------------------------------------------------------------*) |
1461 | 636 |
(* Theorems about the continuous function space. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
637 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
638 |
|
5136 | 639 |
Goal (* chain_cf *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
640 |
"[| chain(cf(D,E),X); x:set(D)|] ==> chain(E,lam n:nat. X`n`x)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
641 |
by (rtac chainI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
642 |
by (rtac lam_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
643 |
by (rtac apply_type 1); |
5136 | 644 |
by (assume_tac 2); |
645 |
by (REPEAT(ares_tac[cont_fun,cf_cont,chain_in] 1)); |
|
2469 | 646 |
by (Asm_simp_tac 1); |
5136 | 647 |
by (REPEAT(ares_tac[rel_cf,chain_rel] 1)); |
3425 | 648 |
qed "chain_cf"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
649 |
|
5136 | 650 |
Goal (* matrix_lemma *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
651 |
"[|chain(cf(D,E),X); chain(D,Xa); cpo(D); cpo(E) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
652 |
\ matrix(E,lam x:nat. lam xa:nat. X`x`(Xa`xa))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
653 |
by (rtac matrix_chainI 1); |
2469 | 654 |
by (Asm_simp_tac 1); |
655 |
by (Asm_simp_tac 2); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
656 |
by (rtac chainI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
657 |
by (rtac lam_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
658 |
by (rtac apply_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
659 |
by (rtac (chain_in RS cf_cont RS cont_fun) 1); |
5136 | 660 |
by (REPEAT(assume_tac 1)); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
661 |
by (rtac chain_in 1); |
5136 | 662 |
by (REPEAT(assume_tac 1)); |
2469 | 663 |
by (Asm_simp_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
664 |
by (rtac cont_mono 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
665 |
by (rtac (chain_in RS cf_cont) 1); |
5136 | 666 |
by (REPEAT (assume_tac 1)); |
667 |
brr [chain_rel,chain_in,nat_succI] 1; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
668 |
by (rtac chainI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
669 |
by (rtac lam_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
670 |
by (rtac apply_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
671 |
by (rtac (chain_in RS cf_cont RS cont_fun) 1); |
5136 | 672 |
by (REPEAT(assume_tac 1)); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
673 |
by (rtac chain_in 1); |
5136 | 674 |
by (REPEAT(assume_tac 1)); |
2469 | 675 |
by (Asm_simp_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
676 |
by (rtac rel_cf 1); |
5136 | 677 |
brr [chain_in,chain_rel] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
678 |
by (rtac lam_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
679 |
by (rtac lam_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
680 |
by (rtac apply_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
681 |
by (rtac (chain_in RS cf_cont RS cont_fun) 1); |
5136 | 682 |
by (REPEAT (assume_tac 1)); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
683 |
by (rtac chain_in 1); |
5136 | 684 |
by (REPEAT (assume_tac 1)); |
3425 | 685 |
qed "matrix_lemma"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
686 |
|
5136 | 687 |
Goal (* chain_cf_lub_cont *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
688 |
"[|chain(cf(D,E),X); cpo(D); cpo(E) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
689 |
\ (lam x:set(D). lub(E, lam n:nat. X ` n ` x)) : cont(D, E)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
690 |
by (rtac contI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
691 |
by (rtac lam_type 1); |
5136 | 692 |
by (REPEAT(ares_tac[chain_cf RS cpo_lub RS islub_in] 1)); |
2469 | 693 |
by (Asm_simp_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
694 |
by (rtac dominate_islub 1); |
5136 | 695 |
by (REPEAT(ares_tac[chain_cf RS cpo_lub] 2)); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
696 |
by (rtac dominateI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
697 |
by (assume_tac 1); |
2469 | 698 |
by (Asm_simp_tac 1); |
5136 | 699 |
by (REPEAT(ares_tac [chain_in RS cf_cont RS cont_mono] 1)); |
700 |
by (REPEAT(ares_tac [chain_cf RS chain_fun] 1)); |
|
2034 | 701 |
by (stac beta 1); |
5136 | 702 |
by (REPEAT(ares_tac [cpo_lub RS islub_in] 1)); |
703 |
by (asm_simp_tac(simpset() addsimps[chain_in RS cf_cont RS cont_lub]) 1); |
|
704 |
by (forward_tac[matrix_lemma RS lub_matrix_diag]1); |
|
705 |
by (REPEAT (assume_tac 1)); |
|
2469 | 706 |
by (Asm_full_simp_tac 1); |
4091 | 707 |
by (asm_simp_tac(simpset() addsimps[chain_in RS beta]) 1); |
5136 | 708 |
by (dtac (matrix_lemma RS lub_matrix_diag_sym) 1); |
709 |
by (REPEAT (assume_tac 1)); |
|
2469 | 710 |
by (Asm_full_simp_tac 1); |
3425 | 711 |
qed "chain_cf_lub_cont"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
712 |
|
5136 | 713 |
Goal (* islub_cf *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
714 |
"[| chain(cf(D,E),X); cpo(D); cpo(E)|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
715 |
\ islub(cf(D,E), X, lam x:set(D). lub(E,lam n:nat. X`n`x))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
716 |
by (rtac islubI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
717 |
by (rtac isubI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
718 |
by (rtac (chain_cf_lub_cont RS cont_cf) 1); |
5136 | 719 |
by (REPEAT (assume_tac 1)); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
720 |
by (rtac rel_cfI 1); |
5136 | 721 |
by (fast_tac (claset() addSDs [chain_cf RS cpo_lub RS islub_ub] |
722 |
addss simpset()) 1); |
|
723 |
by (blast_tac (claset() addIs [cf_cont,chain_in]) 1); |
|
724 |
by (blast_tac (claset() addIs [cont_cf,chain_cf_lub_cont]) 1); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
725 |
by (rtac rel_cfI 1); |
2469 | 726 |
by (Asm_simp_tac 1); |
5136 | 727 |
by (REPEAT (blast_tac (claset() addIs [chain_cf_lub_cont,isubD1,cf_cont]) 2)); |
728 |
by (best_tac (claset() addIs [chain_cf RS cpo_lub RS islub_least, |
|
729 |
cf_cont RS cont_fun RS apply_type, isubI] |
|
730 |
addEs [isubD2 RS rel_cf, isubD1] |
|
731 |
addss simpset()) 1); |
|
3425 | 732 |
qed "islub_cf"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
733 |
|
5136 | 734 |
Goal (* cpo_cf *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
735 |
"[| cpo(D); cpo(E)|] ==> cpo(cf(D,E))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
736 |
by (rtac (poI RS cpoI) 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
737 |
by (rtac rel_cfI 1); |
5136 | 738 |
brr[cpo_refl, cf_cont RS cont_fun RS apply_type, cf_cont] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
739 |
by (rtac rel_cfI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
740 |
by (rtac cpo_trans 1); |
5136 | 741 |
by (assume_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
742 |
by (etac rel_cf 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
743 |
by (assume_tac 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
744 |
by (rtac rel_cf 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
745 |
by (assume_tac 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
746 |
brr[cf_cont RS cont_fun RS apply_type,cf_cont]1; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
747 |
by (rtac fun_extension 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
748 |
brr[cf_cont RS cont_fun]1; |
5136 | 749 |
by (fast_tac (claset() addIs [islub_cf]) 2); |
750 |
by (blast_tac (claset() addIs [cpo_antisym,rel_cf, |
|
751 |
cf_cont RS cont_fun RS apply_type]) 1); |
|
752 |
||
3425 | 753 |
qed "cpo_cf"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
754 |
|
5136 | 755 |
Goal "[| chain(cf(D,E),X); cpo(D); cpo(E)|] ==> \ |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
756 |
\ lub(cf(D,E), X) = (lam x:set(D). lub(E,lam n:nat. X`n`x))"; |
5136 | 757 |
by (blast_tac (claset() addIs [islub_unique,cpo_lub,islub_cf,cpo_cf]) 1); |
3425 | 758 |
qed "lub_cf"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
759 |
|
5136 | 760 |
Goal "[|y:set(E); cpo(D); cpo(E)|] ==> (lam x:set(D).y) : cont(D,E)"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
761 |
by (rtac contI 1); |
2469 | 762 |
by (Asm_simp_tac 2); |
5136 | 763 |
by (blast_tac (claset() addIs [lam_type]) 1); |
764 |
by (asm_simp_tac(simpset() addsimps [chain_in, cpo_lub RS islub_in, |
|
765 |
lub_const]) 1); |
|
3425 | 766 |
qed "const_cont"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
767 |
|
5136 | 768 |
Goal "[|cpo(D); pcpo(E); y:cont(D,E)|]==>rel(cf(D,E),(lam x:set(D).bot(E)),y)"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
769 |
by (rtac rel_cfI 1); |
2469 | 770 |
by (Asm_simp_tac 1); |
5136 | 771 |
brr[bot_least, bot_in, apply_type, cont_fun, const_cont, cpo_cf, pcpo_cpo] 1; |
3425 | 772 |
qed "cf_least"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
773 |
|
5136 | 774 |
Goal (* pcpo_cf *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
775 |
"[|cpo(D); pcpo(E)|] ==> pcpo(cf(D,E))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
776 |
by (rtac pcpoI 1); |
5136 | 777 |
brr[cf_least, bot_in, const_cont RS cont_cf, cf_cont, cpo_cf, pcpo_cpo] 1; |
3425 | 778 |
qed "pcpo_cf"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
779 |
|
5136 | 780 |
Goal (* bot_cf *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
781 |
"[|cpo(D); pcpo(E)|] ==> bot(cf(D,E)) = (lam x:set(D).bot(E))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
782 |
by (rtac (bot_unique RS sym) 1); |
5136 | 783 |
brr[pcpo_cf, cf_least, bot_in RS const_cont RS cont_cf, cf_cont, pcpo_cpo] 1; |
3425 | 784 |
qed "bot_cf"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
785 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
786 |
(*----------------------------------------------------------------------*) |
1461 | 787 |
(* Identity and composition. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
788 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
789 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
790 |
val id_thm = prove_goalw Perm.thy [id_def] "x:X ==> (id(X)`x) = x" |
4091 | 791 |
(fn prems => [simp_tac(simpset() addsimps prems) 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
792 |
|
5136 | 793 |
Goal (* id_cont *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
794 |
"cpo(D) ==> id(set(D)):cont(D,D)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
795 |
by (rtac contI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
796 |
by (rtac id_type 1); |
4091 | 797 |
by (asm_simp_tac (simpset() addsimps[id_thm]) 1); |
5136 | 798 |
by (asm_simp_tac(simpset() addsimps[id_thm, cpo_lub RS islub_in, chain_in, chain_fun RS eta]) 1); |
3425 | 799 |
qed "id_cont"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
800 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
801 |
val comp_cont_apply = cont_fun RSN(2,cont_fun RS comp_fun_apply); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
802 |
|
5136 | 803 |
Goal (* comp_pres_cont *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
804 |
"[| f:cont(D',E); g:cont(D,D'); cpo(D)|] ==> f O g : cont(D,E)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
805 |
by (rtac contI 1); |
2034 | 806 |
by (stac comp_cont_apply 2); |
807 |
by (stac comp_cont_apply 5); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
808 |
by (rtac cont_mono 8); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
809 |
by (rtac cont_mono 9); (* 15 subgoals *) |
5136 | 810 |
brr[comp_fun,cont_fun,cont_map] 1; (* proves all but the lub case *) |
2034 | 811 |
by (stac comp_cont_apply 1); |
812 |
by (stac cont_lub 4); |
|
813 |
by (stac cont_lub 6); |
|
5136 | 814 |
by (asm_full_simp_tac(simpset() addsimps [comp_cont_apply,chain_in]) 8); |
815 |
by (auto_tac (claset() addIs [cpo_lub RS islub_in, cont_chain], simpset())); |
|
3425 | 816 |
qed "comp_pres_cont"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
817 |
|
5136 | 818 |
Goal (* comp_mono *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
819 |
"[| f:cont(D',E); g:cont(D,D'); f':cont(D',E); g':cont(D,D'); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
820 |
\ rel(cf(D',E),f,f'); rel(cf(D,D'),g,g'); cpo(D); cpo(E) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
821 |
\ rel(cf(D,E),f O g,f' O g')"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
822 |
by (rtac rel_cfI 1); (* extra proof obl: f O g and f' O g' cont. Extra asm cpo(D). *) |
2034 | 823 |
by (stac comp_cont_apply 1); |
824 |
by (stac comp_cont_apply 4); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
825 |
by (rtac cpo_trans 7); |
5136 | 826 |
by (REPEAT (ares_tac [rel_cf,cont_mono,cont_map,comp_pres_cont] 1)); |
3425 | 827 |
qed "comp_mono"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
828 |
|
5136 | 829 |
Goal (* chain_cf_comp *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
830 |
"[| chain(cf(D',E),X); chain(cf(D,D'),Y); cpo(D); cpo(E)|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
831 |
\ chain(cf(D,E),lam n:nat. X`n O Y`n)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
832 |
by (rtac chainI 1); |
2469 | 833 |
by (Asm_simp_tac 2); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
834 |
by (rtac rel_cfI 2); |
2034 | 835 |
by (stac comp_cont_apply 2); |
836 |
by (stac comp_cont_apply 5); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
837 |
by (rtac cpo_trans 8); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
838 |
by (rtac rel_cf 9); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
839 |
by (rtac cont_mono 11); |
5136 | 840 |
brr[lam_type, comp_pres_cont, cont_cf, chain_in RS cf_cont, cont_map, chain_rel,rel_cf,nat_succI] 1; |
3425 | 841 |
qed "chain_cf_comp"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
842 |
|
5136 | 843 |
Goal (* comp_lubs *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
844 |
"[| chain(cf(D',E),X); chain(cf(D,D'),Y); cpo(D); cpo(D'); cpo(E)|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
845 |
\ lub(cf(D',E),X) O lub(cf(D,D'),Y) = lub(cf(D,E),lam n:nat. X`n O Y`n)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
846 |
by (rtac fun_extension 1); |
2034 | 847 |
by (stac lub_cf 3); |
5136 | 848 |
brr[comp_fun, cf_cont RS cont_fun, cpo_lub RS islub_in, cpo_cf, chain_cf_comp] 1; |
849 |
by (asm_simp_tac(simpset() |
|
850 |
addsimps[chain_in RS |
|
851 |
cf_cont RSN(3,chain_in RS |
|
852 |
cf_cont RS comp_cont_apply)]) 1); |
|
2034 | 853 |
by (stac comp_cont_apply 1); |
5136 | 854 |
brr[cpo_lub RS islub_in RS cf_cont, cpo_cf] 1; |
855 |
by (asm_simp_tac(simpset() addsimps |
|
856 |
[lub_cf,chain_cf, chain_in RS cf_cont RS cont_lub, |
|
857 |
chain_cf RS cpo_lub RS islub_in]) 1); |
|
1623 | 858 |
by (cut_inst_tac[("M","lam xa:nat. lam xb:nat. X`xa`(Y`xb`x)")] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
859 |
lub_matrix_diag 1); |
2469 | 860 |
by (Asm_full_simp_tac 3); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
861 |
by (rtac matrix_chainI 1); |
2469 | 862 |
by (Asm_simp_tac 1); |
5525 | 863 |
by (Asm_simp_tac 2); |
864 |
by (dtac (chain_in RS cf_cont) 1 THEN atac 1); |
|
865 |
by (fast_tac (claset() addDs [chain_cf RSN(2,cont_chain)] |
|
5136 | 866 |
addss simpset()) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
867 |
by (rtac chain_cf 1); |
5136 | 868 |
by (REPEAT (ares_tac [cont_fun RS apply_type, chain_in RS cf_cont, |
869 |
lam_type] 1)); |
|
3425 | 870 |
qed "comp_lubs"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
871 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
872 |
(*----------------------------------------------------------------------*) |
1461 | 873 |
(* Theorems about projpair. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
874 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
875 |
|
5136 | 876 |
Goalw [projpair_def] (* projpairI *) |
877 |
"[| e:cont(D,E); p:cont(E,D); p O e = id(set(D)); \ |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
878 |
\ rel(cf(E,E))(e O p)(id(set(E)))|] ==> projpair(D,E,e,p)"; |
2469 | 879 |
by (Fast_tac 1); |
3425 | 880 |
qed "projpairI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
881 |
|
5136 | 882 |
val prems = Goalw [projpair_def] (* projpairE *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
883 |
"[| projpair(D,E,e,p); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
884 |
\ [| e:cont(D,E); p:cont(E,D); p O e = id(set(D)); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
885 |
\ rel(cf(E,E))(e O p)(id(set(E)))|] ==> Q |] ==> Q"; |
1623 | 886 |
by (rtac (hd(tl prems)) 1); |
4091 | 887 |
by (REPEAT(asm_simp_tac(simpset() addsimps[hd prems]) 1)); |
3425 | 888 |
qed "projpairE"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
889 |
|
5136 | 890 |
Goal "projpair(D,E,e,p) ==> e:cont(D,E)"; |
891 |
by (etac projpairE 1); |
|
892 |
by (assume_tac 1); |
|
3425 | 893 |
qed "projpair_e_cont"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
894 |
|
5136 | 895 |
Goal (* projpair_p_cont *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
896 |
"projpair(D,E,e,p) ==> p:cont(E,D)"; |
5136 | 897 |
by (etac projpairE 1); |
898 |
by (assume_tac 1); |
|
3425 | 899 |
qed "projpair_p_cont"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
900 |
|
5136 | 901 |
Goal "projpair(D,E,e,p) ==> p O e = id(set(D))"; |
902 |
by (etac projpairE 1); |
|
903 |
by (assume_tac 1); |
|
3425 | 904 |
qed "projpair_eq"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
905 |
|
5136 | 906 |
Goal (* projpair_rel *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
907 |
"projpair(D,E,e,p) ==> rel(cf(E,E))(e O p)(id(set(E)))"; |
5136 | 908 |
by (etac projpairE 1); |
909 |
by (assume_tac 1); |
|
3425 | 910 |
qed "projpair_rel"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
911 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
912 |
val projpairDs = [projpair_e_cont,projpair_p_cont,projpair_eq,projpair_rel]; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
913 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
914 |
(*----------------------------------------------------------------------*) |
1461 | 915 |
(* NB! projpair_e_cont and projpair_p_cont cannot be used repeatedly *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
916 |
(* at the same time since both match a goal of the form f:cont(X,Y).*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
917 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
918 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
919 |
(*----------------------------------------------------------------------*) |
1461 | 920 |
(* Uniqueness of embedding projection pairs. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
921 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
922 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
923 |
val id_comp = fun_is_rel RS left_comp_id; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
924 |
val comp_id = fun_is_rel RS right_comp_id; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
925 |
|
5136 | 926 |
val prems = goal thy (* lemma1 *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
927 |
"[|cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p'); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
928 |
\ rel(cf(D,E),e,e')|] ==> rel(cf(E,D),p',p)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
929 |
val [_,_,p1,p2,_] = prems; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
930 |
(* The two theorems proj_e_cont and proj_p_cont are useless unless they |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
931 |
are used manually, one at a time. Therefore the following contl. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
932 |
val contl = [p1 RS projpair_e_cont,p1 RS projpair_p_cont, |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
933 |
p2 RS projpair_e_cont,p2 RS projpair_p_cont]; |
1623 | 934 |
by (rtac (p2 RS projpair_p_cont RS cont_fun RS id_comp RS subst) 1); |
935 |
by (rtac (p1 RS projpair_eq RS subst) 1); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
936 |
by (rtac cpo_trans 1); |
1623 | 937 |
brr(cpo_cf::prems) 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
938 |
(* The following corresponds to EXISTS_TAC, non-trivial instantiation. *) |
1623 | 939 |
by (res_inst_tac[("f","p O (e' O p')")]cont_cf 4); |
2034 | 940 |
by (stac comp_assoc 1); |
1623 | 941 |
brr(cpo_refl::cpo_cf::cont_cf::comp_mono::comp_pres_cont::(contl@prems)) 1; |
942 |
by (res_inst_tac[("P","%x. rel(cf(E,D),p O e' O p',x)")] |
|
943 |
(p1 RS projpair_p_cont RS cont_fun RS comp_id RS subst) 1); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
944 |
by (rtac comp_mono 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
945 |
brr(cpo_refl::cpo_cf::cont_cf::comp_mono::comp_pres_cont::id_cont:: |
1623 | 946 |
projpair_rel::(contl@prems)) 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
947 |
val lemma1 = result(); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
948 |
|
5136 | 949 |
val prems = goal thy (* lemma2 *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
950 |
"[|cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p'); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
951 |
\ rel(cf(E,D),p',p)|] ==> rel(cf(D,E),e,e')"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
952 |
val [_,_,p1,p2,_] = prems; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
953 |
val contl = [p1 RS projpair_e_cont,p1 RS projpair_p_cont, |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
954 |
p2 RS projpair_e_cont,p2 RS projpair_p_cont]; |
1623 | 955 |
by (rtac (p1 RS projpair_e_cont RS cont_fun RS comp_id RS subst) 1); |
956 |
by (rtac (p2 RS projpair_eq RS subst) 1); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
957 |
by (rtac cpo_trans 1); |
1623 | 958 |
brr(cpo_cf::prems) 1; |
959 |
by (res_inst_tac[("f","(e O p) O e'")]cont_cf 4); |
|
2034 | 960 |
by (stac comp_assoc 1); |
1623 | 961 |
brr((cpo_cf RS cpo_refl)::cont_cf::comp_mono::comp_pres_cont::(contl@prems)) 1; |
962 |
by (res_inst_tac[("P","%x. rel(cf(D,E),(e O p) O e',x)")] |
|
963 |
(p2 RS projpair_e_cont RS cont_fun RS id_comp RS subst) 1); |
|
5136 | 964 |
brr((cpo_cf RS cpo_refl)::cont_cf::comp_mono::id_cont::comp_pres_cont::projpair_rel::(contl@prems)) 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
965 |
val lemma2 = result(); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
966 |
|
5136 | 967 |
val prems = goal thy (* projpair_unique *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
968 |
"[|cpo(D); cpo(E); projpair(D,E,e,p); projpair(D,E,e',p')|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
969 |
\ (e=e')<->(p=p')"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
970 |
val [_,_,p1,p2] = prems; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
971 |
val contl = [p1 RS projpair_e_cont,p1 RS projpair_p_cont, |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
972 |
p2 RS projpair_e_cont,p2 RS projpair_p_cont]; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
973 |
by (rtac iffI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
974 |
by (rtac cpo_antisym 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
975 |
by (rtac lemma1 2); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
976 |
(* First some existentials are instantiated. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
977 |
by (resolve_tac prems 4); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
978 |
by (resolve_tac prems 4); |
2469 | 979 |
by (Asm_simp_tac 4); |
5136 | 980 |
brr([cpo_cf,cpo_refl,cont_cf,projpair_e_cont]@prems) 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
981 |
by (rtac lemma1 1); |
5136 | 982 |
by (REPEAT (ares_tac prems 1)); |
2469 | 983 |
by (Asm_simp_tac 1); |
1623 | 984 |
brr(cpo_cf::cpo_refl::cont_cf::(contl @ prems)) 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
985 |
by (rtac cpo_antisym 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
986 |
by (rtac lemma2 2); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
987 |
(* First some existentials are instantiated. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
988 |
by (resolve_tac prems 4); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
989 |
by (resolve_tac prems 4); |
2469 | 990 |
by (Asm_simp_tac 4); |
5136 | 991 |
brr([cpo_cf,cpo_refl,cont_cf,projpair_p_cont]@prems) 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
992 |
by (rtac lemma2 1); |
5136 | 993 |
by (REPEAT (ares_tac prems 1)); |
2469 | 994 |
by (Asm_simp_tac 1); |
1623 | 995 |
brr(cpo_cf::cpo_refl::cont_cf::(contl @ prems)) 1; |
3425 | 996 |
qed "projpair_unique"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
997 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
998 |
(* Slightly different, more asms, since THE chooses the unique element. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
999 |
|
5136 | 1000 |
Goalw [emb_def,Rp_def] (* embRp *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1001 |
"[|emb(D,E,e); cpo(D); cpo(E)|] ==> projpair(D,E,e,Rp(D,E,e))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1002 |
by (rtac theI2 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1003 |
by (assume_tac 2); |
5136 | 1004 |
by (blast_tac (claset() addIs [projpair_unique RS iffD1]) 1); |
3425 | 1005 |
qed "embRp"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1006 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1007 |
val embI = prove_goalw Limit.thy [emb_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1008 |
"!!x. projpair(D,E,e,p) ==> emb(D,E,e)" |
2469 | 1009 |
(fn prems => [Fast_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1010 |
|
5136 | 1011 |
Goal "[|projpair(D,E,e,p); cpo(D); cpo(E)|] ==> Rp(D,E,e) = p"; |
1012 |
by (blast_tac (claset() addIs [embRp, embI, projpair_unique RS iffD1]) 1); |
|
3425 | 1013 |
qed "Rp_unique"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1014 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1015 |
val emb_cont = prove_goalw Limit.thy [emb_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1016 |
"emb(D,E,e) ==> e:cont(D,E)" |
1623 | 1017 |
(fn prems => [rtac(hd prems RS exE) 1,rtac projpair_e_cont 1,atac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1018 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1019 |
(* The following three theorems have cpo asms due to THE (uniqueness). *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1020 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1021 |
val Rp_cont = embRp RS projpair_p_cont; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1022 |
val embRp_eq = embRp RS projpair_eq; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1023 |
val embRp_rel = embRp RS projpair_rel; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1024 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1025 |
val id_apply = prove_goalw Perm.thy [id_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1026 |
"!!z. x:A ==> id(A)`x = x" |
2469 | 1027 |
(fn prems => [Asm_simp_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1028 |
|
5136 | 1029 |
Goal (* embRp_eq_thm *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1030 |
"[|emb(D,E,e); x:set(D); cpo(D); cpo(E)|] ==> Rp(D,E,e)`(e`x) = x"; |
1623 | 1031 |
by (rtac (comp_fun_apply RS subst) 1); |
5136 | 1032 |
brr[Rp_cont,emb_cont,cont_fun] 1; |
2034 | 1033 |
by (stac embRp_eq 1); |
5136 | 1034 |
by (auto_tac (claset() addIs [id_apply], simpset())); |
3425 | 1035 |
qed "embRp_eq_thm"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1036 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1037 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1038 |
(*----------------------------------------------------------------------*) |
1461 | 1039 |
(* The identity embedding. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1040 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1041 |
|
5136 | 1042 |
Goalw [projpair_def] (* projpair_id *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1043 |
"cpo(D) ==> projpair(D,D,id(set(D)),id(set(D)))"; |
4152 | 1044 |
by Safe_tac; |
5136 | 1045 |
brr[id_cont,id_comp,id_type] 1; |
2034 | 1046 |
by (stac id_comp 1); (* Matches almost anything *) |
5136 | 1047 |
brr[id_cont,id_type,cpo_refl,cpo_cf,cont_cf] 1; |
3425 | 1048 |
qed "projpair_id"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1049 |
|
5136 | 1050 |
Goal (* emb_id *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1051 |
"cpo(D) ==> emb(D,D,id(set(D)))"; |
5136 | 1052 |
by (auto_tac (claset() addIs [embI,projpair_id], simpset())); |
3425 | 1053 |
qed "emb_id"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1054 |
|
5136 | 1055 |
Goal (* Rp_id *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1056 |
"cpo(D) ==> Rp(D,D,id(set(D))) = id(set(D))"; |
5136 | 1057 |
by (auto_tac (claset() addIs [Rp_unique,projpair_id], simpset())); |
3425 | 1058 |
qed "Rp_id"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1059 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1060 |
(*----------------------------------------------------------------------*) |
1461 | 1061 |
(* Composition preserves embeddings. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1062 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1063 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1064 |
(* Considerably shorter, only partly due to a simpler comp_assoc. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1065 |
(* Proof in HOL-ST: 70 lines (minus 14 due to comp_assoc complication). *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1066 |
(* Proof in Isa/ZF: 23 lines (compared to 56: 60% reduction). *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1067 |
|
5136 | 1068 |
Goalw [projpair_def] (* lemma *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1069 |
"[|emb(D,D',e); emb(D',E,e'); cpo(D); cpo(D'); cpo(E)|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1070 |
\ projpair(D,E,e' O e,(Rp(D,D',e)) O (Rp(D',E,e')))"; |
4152 | 1071 |
by Safe_tac; |
5136 | 1072 |
brr[comp_pres_cont,Rp_cont,emb_cont] 1; |
1623 | 1073 |
by (rtac (comp_assoc RS subst) 1); |
1074 |
by (res_inst_tac[("t1","e'")](comp_assoc RS ssubst) 1); |
|
2034 | 1075 |
by (stac embRp_eq 1); (* Matches everything due to subst/ssubst. *) |
5136 | 1076 |
by (REPEAT (assume_tac 1)); |
2034 | 1077 |
by (stac comp_id 1); |
5136 | 1078 |
brr[cont_fun,Rp_cont,embRp_eq] 1; |
1623 | 1079 |
by (rtac (comp_assoc RS subst) 1); |
1080 |
by (res_inst_tac[("t1","Rp(D,D',e)")](comp_assoc RS ssubst) 1); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1081 |
by (rtac cpo_trans 1); |
5136 | 1082 |
brr[cpo_cf] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1083 |
by (rtac comp_mono 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1084 |
by (rtac cpo_refl 6); |
5136 | 1085 |
brr[cont_cf,Rp_cont] 7; |
1086 |
brr[cpo_cf] 6; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1087 |
by (rtac comp_mono 5); |
5136 | 1088 |
brr[embRp_rel] 10; |
1089 |
brr[cpo_cf RS cpo_refl, cont_cf,Rp_cont] 9; |
|
2034 | 1090 |
by (stac comp_id 10); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1091 |
by (rtac embRp_rel 11); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1092 |
(* There are 16 subgoals at this point. All are proved immediately by: *) |
5136 | 1093 |
by (REPEAT (ares_tac [comp_pres_cont,Rp_cont,id_cont, |
1094 |
emb_cont,cont_fun,cont_cf] 1)); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1095 |
val lemma = result(); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1096 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1097 |
(* The use of RS is great in places like the following, both ugly in HOL. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1098 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1099 |
val emb_comp = lemma RS embI; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1100 |
val Rp_comp = lemma RS Rp_unique; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1101 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1102 |
(*----------------------------------------------------------------------*) |
1461 | 1103 |
(* Infinite cartesian product. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1104 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1105 |
|
5136 | 1106 |
Goalw [set_def,iprod_def] (* iprodI *) |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5136
diff
changeset
|
1107 |
"x:(PROD n:nat. set(DD`n)) ==> x:set(iprod(DD))"; |
2469 | 1108 |
by (Asm_full_simp_tac 1); |
3425 | 1109 |
qed "iprodI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1110 |
|
5136 | 1111 |
Goalw [set_def,iprod_def] (* iprodE *) |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5136
diff
changeset
|
1112 |
"x:set(iprod(DD)) ==> x:(PROD n:nat. set(DD`n))"; |
2469 | 1113 |
by (Asm_full_simp_tac 1); |
3425 | 1114 |
qed "iprodE"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1115 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1116 |
(* Contains typing conditions in contrast to HOL-ST *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1117 |
|
5136 | 1118 |
val prems = Goalw [iprod_def] (* rel_iprodI *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1119 |
"[|!!n. n:nat ==> rel(DD`n,f`n,g`n); f:(PROD n:nat. set(DD`n)); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1120 |
\ g:(PROD n:nat. set(DD`n))|] ==> rel(iprod(DD),f,g)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1121 |
by (rtac rel_I 1); |
2469 | 1122 |
by (Simp_tac 1); |
4152 | 1123 |
by Safe_tac; |
5136 | 1124 |
by (REPEAT (ares_tac prems 1)); |
3425 | 1125 |
qed "rel_iprodI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1126 |
|
5136 | 1127 |
Goalw [iprod_def] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1128 |
"[|rel(iprod(DD),f,g); n:nat|] ==> rel(DD`n,f`n,g`n)"; |
5136 | 1129 |
by (fast_tac (claset() addDs [rel_E] addss simpset()) 1); |
3425 | 1130 |
qed "rel_iprodE"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1131 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1132 |
(* Some special theorems like dProdApIn_cpo and other `_cpo' |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1133 |
probably not needed in Isabelle, wait and see. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1134 |
|
5136 | 1135 |
val prems = Goalw [chain_def] (* chain_iprod *) |
2469 | 1136 |
"[|chain(iprod(DD),X); !!n. n:nat ==> cpo(DD`n); n:nat|] ==> \ |
3840 | 1137 |
\ chain(DD`n,lam m:nat. X`m`n)"; |
4152 | 1138 |
by Safe_tac; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1139 |
by (rtac lam_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1140 |
by (rtac apply_type 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1141 |
by (rtac iprodE 1); |
1623 | 1142 |
by (etac (hd prems RS conjunct1 RS apply_type) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1143 |
by (resolve_tac prems 1); |
4091 | 1144 |
by (asm_simp_tac(simpset() addsimps prems) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1145 |
by (rtac rel_iprodE 1); |
4091 | 1146 |
by (asm_simp_tac (simpset() addsimps prems) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1147 |
by (resolve_tac prems 1); |
3425 | 1148 |
qed "chain_iprod"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1149 |
|
5136 | 1150 |
val prems = Goalw [islub_def,isub_def] (* islub_iprod *) |
2469 | 1151 |
"[|chain(iprod(DD),X); !!n. n:nat ==> cpo(DD`n)|] ==> \ |
3840 | 1152 |
\ islub(iprod(DD),X,lam n:nat. lub(DD`n,lam m:nat. X`m`n))"; |
4152 | 1153 |
by Safe_tac; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1154 |
by (rtac iprodI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1155 |
by (rtac lam_type 1); |
1623 | 1156 |
brr((chain_iprod RS cpo_lub RS islub_in)::prems) 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1157 |
by (rtac rel_iprodI 1); |
2469 | 1158 |
by (Asm_simp_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1159 |
(* Here, HOL resolution is handy, Isabelle resolution bad. *) |
1623 | 1160 |
by (res_inst_tac[("P","%t. rel(DD`na,t,lub(DD`na,lam x:nat. X`x`na))"), |
1161 |
("b1","%n. X`n`na")](beta RS subst) 1); |
|
1162 |
brr((chain_iprod RS cpo_lub RS islub_ub)::iprodE::chain_in::prems) 1; |
|
1163 |
brr(iprodI::lam_type::(chain_iprod RS cpo_lub RS islub_in)::prems) 1; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1164 |
by (rtac rel_iprodI 1); |
2469 | 1165 |
by (Asm_simp_tac 1); |
1623 | 1166 |
brr(islub_least::(chain_iprod RS cpo_lub)::prems) 1; |
1167 |
by (rewtac isub_def); |
|
4152 | 1168 |
by Safe_tac; |
1623 | 1169 |
by (etac (iprodE RS apply_type) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1170 |
by (assume_tac 1); |
2469 | 1171 |
by (Asm_simp_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1172 |
by (dtac bspec 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1173 |
by (etac rel_iprodE 2); |
1623 | 1174 |
brr(lam_type::(chain_iprod RS cpo_lub RS islub_in)::iprodE::prems) 1; |
3425 | 1175 |
qed "islub_iprod"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1176 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1177 |
val prems = goal Limit.thy (* cpo_iprod *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1178 |
"(!!n. n:nat ==> cpo(DD`n)) ==> cpo(iprod(DD))"; |
5136 | 1179 |
brr[cpoI,poI] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1180 |
by (rtac rel_iprodI 1); (* not repeated: want to solve 1 and leave 2 unchanged *) |
1623 | 1181 |
brr(cpo_refl::(iprodE RS apply_type)::iprodE::prems) 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1182 |
by (rtac rel_iprodI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1183 |
by (dtac rel_iprodE 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1184 |
by (dtac rel_iprodE 2); |
1623 | 1185 |
brr(cpo_trans::(iprodE RS apply_type)::iprodE::prems) 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1186 |
by (rtac fun_extension 1); |
1623 | 1187 |
brr(cpo_antisym::rel_iprodE::(iprodE RS apply_type)::iprodE::prems) 1; |
1188 |
brr(islub_iprod::prems) 1; |
|
3425 | 1189 |
qed "cpo_iprod"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1190 |
|
5136 | 1191 |
val prems = Goalw [islub_def,isub_def] (* lub_iprod *) |
2469 | 1192 |
"[|chain(iprod(DD),X); !!n. n:nat ==> cpo(DD`n)|] ==> \ |
3840 | 1193 |
\ lub(iprod(DD),X) = (lam n:nat. lub(DD`n,lam m:nat. X`m`n))"; |
1623 | 1194 |
brr((cpo_lub RS islub_unique)::islub_iprod::cpo_iprod::prems) 1; |
3425 | 1195 |
qed "lub_iprod"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1196 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1197 |
(*----------------------------------------------------------------------*) |
1461 | 1198 |
(* The notion of subcpo. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1199 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1200 |
|
5136 | 1201 |
val prems = Goalw [subcpo_def] (* subcpoI *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1202 |
"[|set(D)<=set(E); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1203 |
\ !!x y. [|x:set(D); y:set(D)|] ==> rel(D,x,y)<->rel(E,x,y); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1204 |
\ !!X. chain(D,X) ==> lub(E,X) : set(D)|] ==> subcpo(D,E)"; |
4152 | 1205 |
by Safe_tac; |
4091 | 1206 |
by (asm_full_simp_tac(simpset() addsimps prems) 2); |
1207 |
by (asm_simp_tac(simpset() addsimps prems) 2); |
|
1623 | 1208 |
brr(prems@[subsetD]) 1; |
3425 | 1209 |
qed "subcpoI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1210 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1211 |
val subcpo_subset = prove_goalw Limit.thy [subcpo_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1212 |
"!!x. subcpo(D,E) ==> set(D)<=set(E)" |
2469 | 1213 |
(fn prems => [Fast_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1214 |
|
5136 | 1215 |
Goalw [subcpo_def] |
1216 |
"[|subcpo(D,E); x:set(D); y:set(D)|] ==> rel(D,x,y)<->rel(E,x,y)"; |
|
1217 |
by (Blast_tac 1); |
|
1218 |
qed "subcpo_rel_eq"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1219 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1220 |
val subcpo_relD1 = subcpo_rel_eq RS iffD1; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1221 |
val subcpo_relD2 = subcpo_rel_eq RS iffD2; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1222 |
|
5136 | 1223 |
Goalw [subcpo_def] "[|subcpo(D,E); chain(D,X)|] ==> lub(E,X) : set(D)"; |
1224 |
by (Blast_tac 1); |
|
1225 |
qed "subcpo_lub"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1226 |
|
5136 | 1227 |
Goal "[|subcpo(D,E); chain(D,X)|] ==> chain(E,X)"; |
1228 |
by (rtac (Pi_type RS chainI) 1); |
|
1229 |
by (REPEAT |
|
1230 |
(blast_tac (claset() addIs [chain_fun, subcpo_relD1, |
|
1231 |
subcpo_subset RS subsetD, |
|
1232 |
chain_in, chain_rel]) 1)); |
|
3425 | 1233 |
qed "chain_subcpo"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1234 |
|
5136 | 1235 |
Goal "[|subcpo(D,E); chain(D,X); isub(D,X,x)|] ==> isub(E,X,x)"; |
1236 |
by (blast_tac (claset() addIs [isubI, subcpo_relD1,subcpo_relD1, |
|
1237 |
chain_in, isubD1, isubD2, |
|
1238 |
subcpo_subset RS subsetD, |
|
1239 |
chain_in, chain_rel]) 1); |
|
3425 | 1240 |
qed "ub_subcpo"; |
1461 | 1241 |
|
5136 | 1242 |
Goal "[|subcpo(D,E); cpo(E); chain(D,X)|] ==> islub(D,X,lub(E,X))"; |
1243 |
by (blast_tac (claset() addIs [islubI, isubI, subcpo_lub, |
|
1244 |
subcpo_relD2, chain_in, |
|
1245 |
islub_ub, islub_least, cpo_lub, |
|
1246 |
chain_subcpo, isubD1, ub_subcpo]) 1); |
|
3425 | 1247 |
qed "islub_subcpo"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1248 |
|
5136 | 1249 |
Goal "[|subcpo(D,E); cpo(E)|] ==> cpo(D)"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1250 |
brr[cpoI,poI]1; |
5136 | 1251 |
by (asm_full_simp_tac(simpset() addsimps[subcpo_rel_eq]) 1); |
1252 |
brr[cpo_refl, subcpo_subset RS subsetD] 1; |
|
1253 |
by (rotate_tac ~3 1); |
|
1254 |
by (asm_full_simp_tac(simpset() addsimps[subcpo_rel_eq]) 1); |
|
1255 |
by (blast_tac (claset() addIs [subcpo_subset RS subsetD, cpo_trans]) 1); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1256 |
(* Changing the order of the assumptions, otherwise full_simp doesn't work. *) |
5136 | 1257 |
by (rotate_tac ~2 1); |
1258 |
by (asm_full_simp_tac(simpset() addsimps[subcpo_rel_eq]) 1); |
|
1259 |
by (blast_tac (claset() addIs [cpo_antisym, subcpo_subset RS subsetD]) 1); |
|
1260 |
by (fast_tac (claset() addIs [islub_subcpo]) 1); |
|
3425 | 1261 |
qed "subcpo_cpo"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1262 |
|
5136 | 1263 |
Goal "[|subcpo(D,E); cpo(E); chain(D,X)|] ==> lub(D,X) = lub(E,X)"; |
1264 |
by (blast_tac (claset() addIs [cpo_lub RS islub_unique, |
|
1265 |
islub_subcpo, subcpo_cpo]) 1); |
|
3425 | 1266 |
qed "lub_subcpo"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1267 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1268 |
(*----------------------------------------------------------------------*) |
1461 | 1269 |
(* Making subcpos using mkcpo. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1270 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1271 |
|
5136 | 1272 |
Goalw [set_def,mkcpo_def] "[|x:set(D); P(x)|] ==> x:set(mkcpo(D,P))"; |
1273 |
by Auto_tac; |
|
3425 | 1274 |
qed "mkcpoI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1275 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1276 |
(* Old proof where cpos are non-reflexive relations. |
1623 | 1277 |
by (rewtac set_def); (* Annoying, cannot just rewrite once. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1278 |
by (rtac CollectI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1279 |
by (rtac domainI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1280 |
by (rtac CollectI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1281 |
(* Now, work on subgoal 2 (and 3) to instantiate unknown. *) |
2469 | 1282 |
by (Simp_tac 2); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1283 |
by (rtac conjI 2); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1284 |
by (rtac conjI 3); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1285 |
by (resolve_tac prems 3); |
4091 | 1286 |
by (simp_tac(simpset() addsimps [rewrite_rule[set_def](hd prems)]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1287 |
by (resolve_tac prems 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1288 |
by (rtac cpo_refl 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1289 |
by (resolve_tac prems 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1290 |
by (rtac rel_I 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1291 |
by (rtac CollectI 1); |
4091 | 1292 |
by (fast_tac(claset() addSIs [rewrite_rule[set_def](hd prems)]) 1); |
2469 | 1293 |
by (Simp_tac 1); |
5136 | 1294 |
brr[conjI,cpo_refl] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1295 |
*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1296 |
|
5136 | 1297 |
Goalw [set_def,mkcpo_def] (* mkcpoD1 *) |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5136
diff
changeset
|
1298 |
"x:set(mkcpo(D,P))==> x:set(D)"; |
2469 | 1299 |
by (Asm_full_simp_tac 1); |
3425 | 1300 |
qed "mkcpoD1"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1301 |
|
5136 | 1302 |
Goalw [set_def,mkcpo_def] (* mkcpoD2 *) |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5136
diff
changeset
|
1303 |
"x:set(mkcpo(D,P))==> P(x)"; |
2469 | 1304 |
by (Asm_full_simp_tac 1); |
3425 | 1305 |
qed "mkcpoD2"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1306 |
|
5136 | 1307 |
Goalw [rel_def,mkcpo_def] (* rel_mkcpoE *) |
5147
825877190618
More tidying and removal of "\!\!... from Goal commands
paulson
parents:
5136
diff
changeset
|
1308 |
"rel(mkcpo(D,P),x,y) ==> rel(D,x,y)"; |
2469 | 1309 |
by (Asm_full_simp_tac 1); |
3425 | 1310 |
qed "rel_mkcpoE"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1311 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1312 |
val rel_mkcpo = prove_goalw Limit.thy [mkcpo_def,rel_def,set_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1313 |
"!!z. [|x:set(D); y:set(D)|] ==> rel(mkcpo(D,P),x,y) <-> rel(D,x,y)" |
2469 | 1314 |
(fn prems => [Asm_simp_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1315 |
|
2469 | 1316 |
(* The HOL proof is simpler, problems due to cpos as purely in upair. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1317 |
(* And chains as set functions. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1318 |
|
5136 | 1319 |
Goal (* chain_mkcpo *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1320 |
"chain(mkcpo(D,P),X) ==> chain(D,X)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1321 |
by (rtac chainI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1322 |
(*---begin additional---*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1323 |
by (rtac Pi_type 1); |
5136 | 1324 |
brr[chain_fun] 1; |
1325 |
brr[chain_in RS mkcpoD1] 1; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1326 |
(*---end additional---*) |
1623 | 1327 |
by (rtac (rel_mkcpo RS iffD1) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1328 |
(*---begin additional---*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1329 |
by (rtac mkcpoD1 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1330 |
by (rtac mkcpoD1 2); |
5136 | 1331 |
brr[chain_in,nat_succI] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1332 |
(*---end additional---*) |
5136 | 1333 |
by (auto_tac (claset() addIs [chain_rel], simpset())); |
3425 | 1334 |
qed "chain_mkcpo"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1335 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1336 |
val prems = goal Limit.thy (* subcpo_mkcpo *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1337 |
"[|!!X. chain(mkcpo(D,P),X) ==> P(lub(D,X)); cpo(D)|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1338 |
\ subcpo(mkcpo(D,P),D)"; |
1623 | 1339 |
brr(subcpoI::subsetI::prems) 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1340 |
by (rtac rel_mkcpo 2); |
1623 | 1341 |
by (REPEAT(etac mkcpoD1 1)); |
1342 |
brr(mkcpoI::(cpo_lub RS islub_in)::chain_mkcpo::prems) 1; |
|
3425 | 1343 |
qed "subcpo_mkcpo"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1344 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1345 |
(*----------------------------------------------------------------------*) |
1461 | 1346 |
(* Embedding projection chains of cpos. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1347 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1348 |
|
5136 | 1349 |
val prems = Goalw [emb_chain_def] (* emb_chainI *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1350 |
"[|!!n. n:nat ==> cpo(DD`n); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1351 |
\ !!n. n:nat ==> emb(DD`n,DD`succ(n),ee`n)|] ==> emb_chain(DD,ee)"; |
4152 | 1352 |
by Safe_tac; |
5136 | 1353 |
by (REPEAT (ares_tac prems 1)); |
3425 | 1354 |
qed "emb_chainI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1355 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1356 |
val emb_chain_cpo = prove_goalw Limit.thy [emb_chain_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1357 |
"!!x. [|emb_chain(DD,ee); n:nat|] ==> cpo(DD`n)" |
2469 | 1358 |
(fn prems => [Fast_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1359 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1360 |
val emb_chain_emb = prove_goalw Limit.thy [emb_chain_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1361 |
"!!x. [|emb_chain(DD,ee); n:nat|] ==> emb(DD`n,DD`succ(n),ee`n)" |
2469 | 1362 |
(fn prems => [Fast_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1363 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1364 |
(*----------------------------------------------------------------------*) |
1461 | 1365 |
(* Dinf, the inverse Limit. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1366 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1367 |
|
5136 | 1368 |
val prems = Goalw [Dinf_def] (* DinfI *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1369 |
"[|x:(PROD n:nat. set(DD`n)); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1370 |
\ !!n. n:nat ==> Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1371 |
\ x:set(Dinf(DD,ee))"; |
1623 | 1372 |
brr(mkcpoI::iprodI::ballI::prems) 1; |
3425 | 1373 |
qed "DinfI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1374 |
|
5136 | 1375 |
Goalw [Dinf_def] "x:set(Dinf(DD,ee)) ==> x:(PROD n:nat. set(DD`n))"; |
1376 |
by (etac (mkcpoD1 RS iprodE) 1); |
|
1377 |
qed "Dinf_prod"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1378 |
|
5136 | 1379 |
Goalw [Dinf_def] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1380 |
"[|x:set(Dinf(DD,ee)); n:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1381 |
\ Rp(DD`n,DD`succ(n),ee`n)`(x`succ(n)) = x`n"; |
5136 | 1382 |
by (blast_tac (claset() addDs [mkcpoD2]) 1); |
1383 |
qed "Dinf_eq"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1384 |
|
5136 | 1385 |
val prems = Goalw [Dinf_def] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1386 |
"[|!!n. n:nat ==> rel(DD`n,x`n,y`n); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1387 |
\ x:(PROD n:nat. set(DD`n)); y:(PROD n:nat. set(DD`n))|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1388 |
\ rel(Dinf(DD,ee),x,y)"; |
1623 | 1389 |
by (rtac (rel_mkcpo RS iffD2) 1); |
1390 |
brr(rel_iprodI::iprodI::prems) 1; |
|
3425 | 1391 |
qed "rel_DinfI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1392 |
|
5136 | 1393 |
Goalw [Dinf_def] "[|rel(Dinf(DD,ee),x,y); n:nat|] ==> rel(DD`n,x`n,y`n)"; |
1394 |
by (etac (rel_mkcpoE RS rel_iprodE) 1); |
|
1395 |
by (assume_tac 1); |
|
3425 | 1396 |
qed "rel_Dinf"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1397 |
|
5136 | 1398 |
Goalw [Dinf_def] "chain(Dinf(DD,ee),X) ==> chain(iprod(DD),X)"; |
1399 |
by (etac chain_mkcpo 1); |
|
1400 |
qed "chain_Dinf"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1401 |
|
5136 | 1402 |
Goalw [Dinf_def] (* subcpo_Dinf *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1403 |
"emb_chain(DD,ee) ==> subcpo(Dinf(DD,ee),iprod(DD))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1404 |
by (rtac subcpo_mkcpo 1); |
1623 | 1405 |
by (fold_tac [Dinf_def]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1406 |
by (rtac ballI 1); |
2034 | 1407 |
by (stac lub_iprod 1); |
5136 | 1408 |
brr[chain_Dinf, emb_chain_cpo] 1; |
2469 | 1409 |
by (Asm_simp_tac 1); |
2034 | 1410 |
by (stac (Rp_cont RS cont_lub) 1); |
5136 | 1411 |
brr[emb_chain_cpo,emb_chain_emb,nat_succI,chain_iprod,chain_Dinf] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1412 |
(* Useful simplification, ugly in HOL. *) |
5136 | 1413 |
by (asm_simp_tac(simpset() addsimps[Dinf_eq,chain_in]) 1); |
1414 |
by (auto_tac (claset() addIs [cpo_iprod,emb_chain_cpo], simpset())); |
|
3425 | 1415 |
qed "subcpo_Dinf"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1416 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1417 |
(* Simple example of existential reasoning in Isabelle versus HOL. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1418 |
|
5136 | 1419 |
Goal "emb_chain(DD,ee) ==> cpo(Dinf(DD,ee))"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1420 |
by (rtac subcpo_cpo 1); |
5136 | 1421 |
by (auto_tac (claset() addIs [subcpo_Dinf,cpo_iprod,emb_chain_cpo], simpset())); |
3425 | 1422 |
qed "cpo_Dinf"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1423 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1424 |
(* Again and again the proofs are much easier to WRITE in Isabelle, but |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1425 |
the proof steps are essentially the same (I think). *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1426 |
|
5136 | 1427 |
Goal (* lub_Dinf *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1428 |
"[|chain(Dinf(DD,ee),X); emb_chain(DD,ee)|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1429 |
\ lub(Dinf(DD,ee),X) = (lam n:nat. lub(DD`n,lam m:nat. X`m`n))"; |
2034 | 1430 |
by (stac (subcpo_Dinf RS lub_subcpo) 1); |
5136 | 1431 |
by (auto_tac (claset() addIs [cpo_iprod,emb_chain_cpo,lub_iprod,chain_Dinf], simpset())); |
3425 | 1432 |
qed "lub_Dinf"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1433 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1434 |
(*----------------------------------------------------------------------*) |
1461 | 1435 |
(* Generalising embedddings D_m -> D_{m+1} to embeddings D_m -> D_n, *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1436 |
(* defined as eps(DD,ee,m,n), via e_less and e_gr. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1437 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1438 |
|
5136 | 1439 |
Goalw [e_less_def] (* e_less_eq *) |
1440 |
"m:nat ==> e_less(DD,ee,m,m) = id(set(DD`m))"; |
|
4091 | 1441 |
by (asm_simp_tac (simpset() addsimps[diff_self_eq_0]) 1); |
3425 | 1442 |
qed "e_less_eq"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1443 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1444 |
(* ARITH_CONV proves the following in HOL. Would like something similar |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1445 |
in Isabelle/ZF. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1446 |
|
5136 | 1447 |
Goal "[|n:nat; m:nat|] ==> succ(m#+n)#-m = succ(n)"; |
1614 | 1448 |
(*Uses add_succ_right the wrong way round!*) |
2469 | 1449 |
by (asm_simp_tac |
4091 | 1450 |
(simpset_of Nat.thy addsimps [add_succ_right RS sym, diff_add_inverse]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1451 |
val lemma_succ_sub = result(); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1452 |
|
5136 | 1453 |
Goalw [e_less_def] (* e_less_add *) |
1454 |
"[|m:nat; k:nat|] ==> \ |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1455 |
\ e_less(DD,ee,m,succ(m#+k)) = (ee`(m#+k))O(e_less(DD,ee,m,m#+k))"; |
4091 | 1456 |
by (asm_simp_tac (simpset() addsimps [lemma_succ_sub,diff_add_inverse]) 1); |
3425 | 1457 |
qed "e_less_add"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1458 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1459 |
(* Again, would like more theorems about arithmetic. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1460 |
(* Well, HOL has much better support and automation of natural numbers. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1461 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1462 |
val add1 = prove_goal Limit.thy |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1463 |
"!!x. n:nat ==> succ(n) = n #+ 1" |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1464 |
(fn prems => |
4091 | 1465 |
[asm_simp_tac (simpset() addsimps[add_succ_right,add_0_right]) 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1466 |
|
5136 | 1467 |
Goal (* succ_sub1 *) |
5514
324e1560a5c9
inserted space in #-1 to prevent confusion with an integer constant
paulson
parents:
5268
diff
changeset
|
1468 |
"x:nat ==> 0 < x --> succ(x #- 1)=x"; |
1623 | 1469 |
by (res_inst_tac[("n","x")]nat_induct 1); |
5136 | 1470 |
by (assume_tac 1); |
2469 | 1471 |
by (Fast_tac 1); |
4152 | 1472 |
by Safe_tac; |
2469 | 1473 |
by (Asm_simp_tac 1); |
1474 |
by (Asm_simp_tac 1); |
|
3425 | 1475 |
qed "succ_sub1"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1476 |
|
5136 | 1477 |
Goal (* succ_le_pos *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1478 |
"[|m:nat; k:nat|] ==> succ(m) le m #+ k --> 0 < k"; |
1623 | 1479 |
by (res_inst_tac[("n","m")]nat_induct 1); |
5136 | 1480 |
by (assume_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1481 |
by (rtac impI 1); |
5136 | 1482 |
by (Asm_full_simp_tac 1); |
4152 | 1483 |
by Safe_tac; |
5136 | 1484 |
by (Asm_full_simp_tac 1); (* Surprise, surprise. *) |
3425 | 1485 |
qed "succ_le_pos"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1486 |
|
5136 | 1487 |
Goal "[|n:nat; m:nat|] ==> m le n --> (EX k:nat. n = m #+ k)"; |
1623 | 1488 |
by (res_inst_tac[("n","m")]nat_induct 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1489 |
by (assume_tac 1); |
4152 | 1490 |
by Safe_tac; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1491 |
by (rtac bexI 1); |
1623 | 1492 |
by (rtac (add_0 RS sym) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1493 |
by (assume_tac 1); |
2469 | 1494 |
by (Asm_full_simp_tac 1); |
1495 |
(* Great, by luck I found le_cs. Such cs's and ss's should be documented. *) |
|
1496 |
by (fast_tac le_cs 1); |
|
1497 |
by (asm_simp_tac |
|
4091 | 1498 |
(simpset_of Nat.thy addsimps[add_succ, add_succ_right RS sym]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1499 |
by (rtac bexI 1); |
2034 | 1500 |
by (stac (succ_sub1 RS mp) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1501 |
(* Instantiation. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1502 |
by (rtac refl 3); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1503 |
by (assume_tac 1); |
1623 | 1504 |
by (rtac (succ_le_pos RS mp) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1505 |
by (assume_tac 3); (* Instantiation *) |
5136 | 1506 |
by (ALLGOALS Asm_simp_tac); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1507 |
val lemma_le_exists = result(); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1508 |
|
5136 | 1509 |
val prems = goal thy (* le_exists *) |
2469 | 1510 |
"[|m le n; !!x. [|n=m#+x; x:nat|] ==> Q; m:nat; n:nat|] ==> Q"; |
1623 | 1511 |
by (rtac (lemma_le_exists RS mp RS bexE) 1); |
5136 | 1512 |
by (DEPTH_SOLVE (ares_tac prems 1)); |
3425 | 1513 |
qed "le_exists"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1514 |
|
5136 | 1515 |
Goal (* e_less_le *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1516 |
"[|m le n; m:nat; n:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1517 |
\ e_less(DD,ee,m,succ(n)) = ee`n O e_less(DD,ee,m,n)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1518 |
by (rtac le_exists 1); |
5136 | 1519 |
by (assume_tac 1); |
1520 |
by (asm_simp_tac(simpset() addsimps[e_less_add]) 1); |
|
1521 |
by (REPEAT (assume_tac 1)); |
|
3425 | 1522 |
qed "e_less_le"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1523 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1524 |
(* All theorems assume variables m and n are natural numbers. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1525 |
|
5136 | 1526 |
Goal "m:nat ==> e_less(DD,ee,m,succ(m)) = ee`m O id(set(DD`m))"; |
1527 |
by (asm_simp_tac(simpset() addsimps[e_less_le,e_less_eq]) 1); |
|
3425 | 1528 |
qed "e_less_succ"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1529 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1530 |
val prems = goal Limit.thy (* e_less_succ_emb *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1531 |
"[|!!n. n:nat ==> emb(DD`n,DD`succ(n),ee`n); m:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1532 |
\ e_less(DD,ee,m,succ(m)) = ee`m"; |
5529 | 1533 |
by (asm_simp_tac(simpset() addsimps e_less_succ::prems) 1); |
2034 | 1534 |
by (stac comp_id 1); |
1623 | 1535 |
brr(emb_cont::cont_fun::refl::prems) 1; |
3425 | 1536 |
qed "e_less_succ_emb"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1537 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1538 |
(* Compare this proof with the HOL one, here we do type checking. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1539 |
(* In any case the one below was very easy to write. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1540 |
|
5136 | 1541 |
Goal "[|emb_chain(DD,ee); m:nat; k:nat|] ==> \ |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1542 |
\ emb(DD`m,DD`(m#+k),e_less(DD,ee,m,m#+k))"; |
1623 | 1543 |
by (res_inst_tac[("n","k")]nat_induct 1); |
5136 | 1544 |
by (assume_tac 1); |
1545 |
by (asm_simp_tac(simpset() addsimps[add_0_right,e_less_eq]) 1); |
|
1546 |
brr[emb_id,emb_chain_cpo] 1; |
|
1547 |
by (asm_simp_tac(simpset() addsimps[add_succ_right,e_less_add]) 1); |
|
1548 |
by (auto_tac (claset() addIs [emb_comp,emb_chain_emb,emb_chain_cpo,add_type], |
|
1549 |
simpset())); |
|
3425 | 1550 |
qed "emb_e_less_add"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1551 |
|
5136 | 1552 |
Goal "[|m le n; emb_chain(DD,ee); m:nat; n:nat|] ==> \ |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1553 |
\ emb(DD`m,DD`n,e_less(DD,ee,m,n))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1554 |
(* same proof as e_less_le *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1555 |
by (rtac le_exists 1); |
5136 | 1556 |
by (assume_tac 1); |
1557 |
by (asm_simp_tac(simpset() addsimps[emb_e_less_add]) 1); |
|
1558 |
by (REPEAT (assume_tac 1)); |
|
3425 | 1559 |
qed "emb_e_less"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1560 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1561 |
val comp_mono_eq = prove_goal Limit.thy |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1562 |
"!!z.[|f=f'; g=g'|] ==> f O g = f' O g'" |
2469 | 1563 |
(fn prems => [Asm_simp_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1564 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1565 |
(* Typing, typing, typing, three irritating assumptions. Extra theorems |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1566 |
needed in proof, but no real difficulty. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1567 |
(* Note also the object-level implication for induction on k. This |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1568 |
must be removed later to allow the theorems to be used for simp. |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1569 |
Therefore this theorem is only a lemma. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1570 |
|
5136 | 1571 |
Goal (* e_less_split_add_lemma *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1572 |
"[| emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1573 |
\ n le k --> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1574 |
\ e_less(DD,ee,m,m#+k) = e_less(DD,ee,m#+n,m#+k) O e_less(DD,ee,m,m#+n)"; |
5136 | 1575 |
by (eres_inst_tac[("n","k")]nat_induct 1); |
1576 |
by (asm_full_simp_tac(simpset() addsimps [e_less_eq, id_type RS id_comp]) 1); |
|
1623 | 1577 |
by (asm_simp_tac(ZF_ss addsimps[le_succ_iff]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1578 |
by (rtac impI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1579 |
by (etac disjE 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1580 |
by (etac impE 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1581 |
by (assume_tac 1); |
5136 | 1582 |
by (asm_simp_tac(ZF_ss addsimps[add_succ_right, e_less_add, add_type,nat_succI]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1583 |
(* Again and again, simplification is a pain. When does it work, when not? *) |
2034 | 1584 |
by (stac e_less_le 1); |
5136 | 1585 |
brr[add_le_mono,nat_le_refl,add_type,nat_succI] 1; |
2034 | 1586 |
by (stac comp_assoc 1); |
5136 | 1587 |
brr[comp_mono_eq,refl] 1; |
1588 |
by (asm_simp_tac(ZF_ss addsimps[e_less_eq,add_type,nat_succI]) 1); |
|
2034 | 1589 |
by (stac id_comp 1); (* simp cannot unify/inst right, use brr below(?). *) |
5136 | 1590 |
by (REPEAT (ares_tac [emb_e_less_add RS emb_cont RS cont_fun, refl, |
1591 |
nat_succI] 1)); |
|
3425 | 1592 |
qed "e_less_split_add_lemma"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1593 |
|
5136 | 1594 |
Goal "[| n le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
1595 |
\ e_less(DD,ee,m,m#+k) = e_less(DD,ee,m#+n,m#+k) O e_less(DD,ee,m,m#+n)"; |
|
1596 |
by (blast_tac (claset() addIs [e_less_split_add_lemma RS mp]) 1); |
|
1597 |
qed "e_less_split_add"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1598 |
|
5136 | 1599 |
Goalw [e_gr_def] (* e_gr_eq *) |
1600 |
"m:nat ==> e_gr(DD,ee,m,m) = id(set(DD`m))"; |
|
4091 | 1601 |
by (asm_simp_tac (simpset() addsimps[diff_self_eq_0]) 1); |
3425 | 1602 |
qed "e_gr_eq"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1603 |
|
5136 | 1604 |
Goalw [e_gr_def] (* e_gr_add *) |
1605 |
"[|n:nat; k:nat|] ==> \ |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1606 |
\ e_gr(DD,ee,succ(n#+k),n) = \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1607 |
\ e_gr(DD,ee,n#+k,n) O Rp(DD`(n#+k),DD`succ(n#+k),ee`(n#+k))"; |
4091 | 1608 |
by (asm_simp_tac (simpset() addsimps [lemma_succ_sub,diff_add_inverse]) 1); |
3425 | 1609 |
qed "e_gr_add"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1610 |
|
5136 | 1611 |
Goal "[|n le m; m:nat; n:nat|] ==> \ |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1612 |
\ e_gr(DD,ee,succ(m),n) = e_gr(DD,ee,m,n) O Rp(DD`m,DD`succ(m),ee`m)"; |
5136 | 1613 |
by (etac le_exists 1); |
1614 |
by (asm_simp_tac(simpset() addsimps[e_gr_add]) 1); |
|
1615 |
by (REPEAT (assume_tac 1)); |
|
3425 | 1616 |
qed "e_gr_le"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1617 |
|
5136 | 1618 |
Goal "m:nat ==> \ |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1619 |
\ e_gr(DD,ee,succ(m),m) = id(set(DD`m)) O Rp(DD`m,DD`succ(m),ee`m)"; |
5136 | 1620 |
by (asm_simp_tac(simpset() addsimps[e_gr_le,e_gr_eq]) 1); |
3425 | 1621 |
qed "e_gr_succ"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1622 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1623 |
(* Cpo asm's due to THE uniqueness. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1624 |
|
5136 | 1625 |
Goal "[|emb_chain(DD,ee); m:nat|] ==> \ |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1626 |
\ e_gr(DD,ee,succ(m),m) = Rp(DD`m,DD`succ(m),ee`m)"; |
5136 | 1627 |
by (asm_simp_tac(simpset() addsimps[e_gr_succ]) 1); |
1628 |
by (blast_tac (claset() addIs [id_comp, Rp_cont,cont_fun, |
|
1629 |
emb_chain_cpo,emb_chain_emb]) 1); |
|
3425 | 1630 |
qed "e_gr_succ_emb"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1631 |
|
5136 | 1632 |
Goal (* e_gr_fun_add *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1633 |
"[|emb_chain(DD,ee); n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1634 |
\ e_gr(DD,ee,n#+k,n): set(DD`(n#+k))->set(DD`n)"; |
1623 | 1635 |
by (res_inst_tac[("n","k")]nat_induct 1); |
5136 | 1636 |
by (assume_tac 1); |
1637 |
by (asm_simp_tac(simpset() addsimps[add_0_right,e_gr_eq,id_type]) 1); |
|
1638 |
by (asm_simp_tac(simpset() addsimps[add_succ_right,e_gr_add]) 1); |
|
1639 |
brr[comp_fun, Rp_cont, cont_fun, emb_chain_emb, emb_chain_cpo, add_type, nat_succI] 1; |
|
3425 | 1640 |
qed "e_gr_fun_add"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1641 |
|
5136 | 1642 |
Goal (* e_gr_fun *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1643 |
"[|n le m; emb_chain(DD,ee); m:nat; n:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1644 |
\ e_gr(DD,ee,m,n): set(DD`m)->set(DD`n)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1645 |
by (rtac le_exists 1); |
5136 | 1646 |
by (assume_tac 1); |
1647 |
by (asm_simp_tac(simpset() addsimps[e_gr_fun_add]) 1); |
|
1648 |
by (REPEAT (assume_tac 1)); |
|
3425 | 1649 |
qed "e_gr_fun"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1650 |
|
5136 | 1651 |
Goal (* e_gr_split_add_lemma *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1652 |
"[| emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1653 |
\ m le k --> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1654 |
\ e_gr(DD,ee,n#+k,n) = e_gr(DD,ee,n#+m,n) O e_gr(DD,ee,n#+k,n#+m)"; |
1623 | 1655 |
by (res_inst_tac[("n","k")]nat_induct 1); |
5136 | 1656 |
by (assume_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1657 |
by (rtac impI 1); |
1623 | 1658 |
by (asm_full_simp_tac(ZF_ss addsimps |
5136 | 1659 |
[le0_iff, add_0_right, e_gr_eq, id_type RS comp_id]) 1); |
1623 | 1660 |
by (asm_simp_tac(ZF_ss addsimps[le_succ_iff]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1661 |
by (rtac impI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1662 |
by (etac disjE 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1663 |
by (etac impE 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1664 |
by (assume_tac 1); |
5136 | 1665 |
by (asm_simp_tac(ZF_ss addsimps[add_succ_right, e_gr_add, add_type,nat_succI]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1666 |
(* Again and again, simplification is a pain. When does it work, when not? *) |
2034 | 1667 |
by (stac e_gr_le 1); |
5136 | 1668 |
brr[add_le_mono,nat_le_refl,add_type,nat_succI] 1; |
2034 | 1669 |
by (stac comp_assoc 1); |
5136 | 1670 |
brr[comp_mono_eq,refl] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1671 |
(* New direct subgoal *) |
5136 | 1672 |
by (asm_simp_tac(ZF_ss addsimps[e_gr_eq,add_type,nat_succI]) 1); |
2034 | 1673 |
by (stac comp_id 1); (* simp cannot unify/inst right, use brr below(?). *) |
5136 | 1674 |
by (REPEAT (ares_tac [e_gr_fun,add_type,refl,add_le_self,nat_succI] 1)); |
3425 | 1675 |
qed "e_gr_split_add_lemma"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1676 |
|
5136 | 1677 |
Goal "[| m le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
1678 |
\ e_gr(DD,ee,n#+k,n) = e_gr(DD,ee,n#+m,n) O e_gr(DD,ee,n#+k,n#+m)"; |
|
1679 |
by (blast_tac (claset() addIs [e_gr_split_add_lemma RS mp]) 1); |
|
1680 |
qed "e_gr_split_add"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1681 |
|
5136 | 1682 |
Goal "[|m le n; emb_chain(DD,ee); m:nat; n:nat|] ==> \ |
1683 |
\ e_less(DD,ee,m,n):cont(DD`m,DD`n)"; |
|
1684 |
by (blast_tac (claset() addIs [emb_cont, emb_e_less]) 1); |
|
1685 |
qed "e_less_cont"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1686 |
|
5136 | 1687 |
Goal (* e_gr_cont_lemma *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1688 |
"[|emb_chain(DD,ee); m:nat; n:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1689 |
\ n le m --> e_gr(DD,ee,m,n):cont(DD`m,DD`n)"; |
1623 | 1690 |
by (res_inst_tac[("n","m")]nat_induct 1); |
5136 | 1691 |
by (assume_tac 1); |
4091 | 1692 |
by (asm_full_simp_tac(simpset() addsimps |
5136 | 1693 |
[le0_iff,e_gr_eq,nat_0I]) 1); |
1694 |
brr[impI,id_cont,emb_chain_cpo,nat_0I] 1; |
|
4091 | 1695 |
by (asm_full_simp_tac(simpset() addsimps[le_succ_iff]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1696 |
by (etac disjE 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1697 |
by (etac impE 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1698 |
by (assume_tac 1); |
5136 | 1699 |
by (asm_simp_tac(simpset() addsimps[e_gr_le]) 1); |
1700 |
brr[comp_pres_cont,Rp_cont,emb_chain_cpo,emb_chain_emb,nat_succI] 1; |
|
1701 |
by (asm_simp_tac(simpset() addsimps[e_gr_eq,nat_succI]) 1); |
|
1702 |
by (auto_tac (claset() addIs [id_cont,emb_chain_cpo], simpset())); |
|
3425 | 1703 |
qed "e_gr_cont_lemma"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1704 |
|
5136 | 1705 |
Goal (* e_gr_cont *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1706 |
"[|n le m; emb_chain(DD,ee); m:nat; n:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1707 |
\ e_gr(DD,ee,m,n):cont(DD`m,DD`n)"; |
5136 | 1708 |
brr[e_gr_cont_lemma RS mp] 1; |
3425 | 1709 |
qed "e_gr_cont"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1710 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1711 |
(* Considerably shorter.... 57 against 26 *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1712 |
|
5136 | 1713 |
Goal (* e_less_e_gr_split_add *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1714 |
"[|n le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1715 |
\ e_less(DD,ee,m,m#+n) = e_gr(DD,ee,m#+k,m#+n) O e_less(DD,ee,m,m#+k)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1716 |
(* Use mp to prepare for induction. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1717 |
by (rtac mp 1); |
5136 | 1718 |
by (assume_tac 2); |
1623 | 1719 |
by (res_inst_tac[("n","k")]nat_induct 1); |
5136 | 1720 |
by (assume_tac 1); |
1623 | 1721 |
by (asm_full_simp_tac(ZF_ss addsimps |
5136 | 1722 |
[le0_iff, add_0_right, e_gr_eq, e_less_eq, id_type RS id_comp]) 1);by (simp_tac(ZF_ss addsimps[le_succ_iff]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1723 |
by (rtac impI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1724 |
by (etac disjE 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1725 |
by (etac impE 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1726 |
by (assume_tac 1); |
5136 | 1727 |
by (asm_simp_tac(ZF_ss addsimps[add_succ_right, e_gr_le, e_less_le, add_le_self,nat_le_refl,add_le_mono,add_type]) 1); |
2034 | 1728 |
by (stac comp_assoc 1); |
1623 | 1729 |
by (res_inst_tac[("s1","ee`(m#+x)")](comp_assoc RS subst) 1); |
2034 | 1730 |
by (stac embRp_eq 1); |
5136 | 1731 |
brr[emb_chain_emb,add_type,emb_chain_cpo,nat_succI] 1; |
2034 | 1732 |
by (stac id_comp 1); |
5136 | 1733 |
brr[e_less_cont RS cont_fun, add_type,add_le_self,refl] 1; |
1734 |
by (asm_full_simp_tac(ZF_ss addsimps[e_gr_eq,nat_succI,add_type]) 1); |
|
2034 | 1735 |
by (stac id_comp 1); |
5136 | 1736 |
by (REPEAT (ares_tac [e_less_cont RS cont_fun, add_type, |
1737 |
nat_succI,add_le_self,refl] 1)); |
|
3425 | 1738 |
qed "e_less_e_gr_split_add"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1739 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1740 |
(* Again considerably shorter, and easy to obtain from the previous thm. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1741 |
|
5136 | 1742 |
Goal (* e_gr_e_less_split_add *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1743 |
"[|m le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1744 |
\ e_gr(DD,ee,n#+m,n) = e_gr(DD,ee,n#+k,n) O e_less(DD,ee,n#+m,n#+k)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1745 |
(* Use mp to prepare for induction. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1746 |
by (rtac mp 1); |
5136 | 1747 |
by (assume_tac 2); |
1623 | 1748 |
by (res_inst_tac[("n","k")]nat_induct 1); |
5136 | 1749 |
by (assume_tac 1); |
4091 | 1750 |
by (asm_full_simp_tac(simpset() addsimps |
5136 | 1751 |
[add_0_right, e_gr_eq, e_less_eq, id_type RS id_comp]) 1); |
1623 | 1752 |
by (simp_tac(ZF_ss addsimps[le_succ_iff]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1753 |
by (rtac impI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1754 |
by (etac disjE 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1755 |
by (etac impE 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1756 |
by (assume_tac 1); |
5136 | 1757 |
by (asm_simp_tac(ZF_ss addsimps[add_succ_right, e_gr_le, e_less_le, add_le_self,nat_le_refl,add_le_mono,add_type]) 1); |
2034 | 1758 |
by (stac comp_assoc 1); |
1623 | 1759 |
by (res_inst_tac[("s1","ee`(n#+x)")](comp_assoc RS subst) 1); |
2034 | 1760 |
by (stac embRp_eq 1); |
5136 | 1761 |
brr[emb_chain_emb,add_type,emb_chain_cpo,nat_succI] 1; |
2034 | 1762 |
by (stac id_comp 1); |
5136 | 1763 |
brr[e_less_cont RS cont_fun, add_type, add_le_mono, nat_le_refl, refl] 1; |
1764 |
by (asm_full_simp_tac(ZF_ss addsimps[e_less_eq,nat_succI,add_type]) 1); |
|
2034 | 1765 |
by (stac comp_id 1); |
5136 | 1766 |
by (REPEAT (ares_tac [e_gr_cont RS cont_fun, add_type,nat_succI,add_le_self, |
1767 |
refl] 1)); |
|
3425 | 1768 |
qed "e_gr_e_less_split_add"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1769 |
|
2469 | 1770 |
|
5136 | 1771 |
Goalw [eps_def] (* emb_eps *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1772 |
"[|m le n; emb_chain(DD,ee); m:nat; n:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1773 |
\ emb(DD`m,DD`n,eps(DD,ee,m,n))"; |
5136 | 1774 |
by (asm_simp_tac(simpset()) 1); |
1775 |
brr[emb_e_less] 1; |
|
3425 | 1776 |
qed "emb_eps"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1777 |
|
5136 | 1778 |
Goalw [eps_def] (* eps_fun *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1779 |
"[|emb_chain(DD,ee); m:nat; n:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1780 |
\ eps(DD,ee,m,n): set(DD`m)->set(DD`n)"; |
5116
8eb343ab5748
Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents:
5068
diff
changeset
|
1781 |
by (rtac (split_if RS iffD2) 1); |
4152 | 1782 |
by Safe_tac; |
5136 | 1783 |
brr[e_less_cont RS cont_fun] 1; |
1784 |
by (auto_tac (claset() addIs [not_le_iff_lt RS iffD1 RS leI, e_gr_fun,nat_into_Ord], simpset())); |
|
3425 | 1785 |
qed "eps_fun"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1786 |
|
5136 | 1787 |
Goalw [eps_def] "n:nat ==> eps(DD,ee,n,n) = id(set(DD`n))"; |
1788 |
by (asm_simp_tac(simpset() addsimps [e_less_eq]) 1); |
|
1789 |
qed "eps_id"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1790 |
|
5136 | 1791 |
Goalw [eps_def] |
1792 |
"[|m:nat; n:nat|] ==> eps(DD,ee,m,m#+n) = e_less(DD,ee,m,m#+n)"; |
|
1793 |
by (asm_simp_tac(simpset() addsimps [add_le_self]) 1); |
|
1794 |
qed "eps_e_less_add"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1795 |
|
5136 | 1796 |
Goalw [eps_def] |
1797 |
"[|m le n; m:nat; n:nat|] ==> eps(DD,ee,m,n) = e_less(DD,ee,m,n)"; |
|
1798 |
by (Asm_simp_tac 1); |
|
1799 |
qed "eps_e_less"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1800 |
|
5136 | 1801 |
Goalw [eps_def] (* eps_e_gr_add *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1802 |
"[|n:nat; k:nat|] ==> eps(DD,ee,n#+k,n) = e_gr(DD,ee,n#+k,n)"; |
5116
8eb343ab5748
Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents:
5068
diff
changeset
|
1803 |
by (rtac (split_if RS iffD2) 1); |
4152 | 1804 |
by Safe_tac; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1805 |
by (etac leE 1); |
5136 | 1806 |
by (asm_simp_tac(simpset() addsimps[e_less_eq,e_gr_eq]) 2); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1807 |
(* Must control rewriting by instantiating a variable. *) |
4091 | 1808 |
by (asm_full_simp_tac(simpset() addsimps |
5136 | 1809 |
[read_instantiate [("i1","n")] (nat_into_Ord RS not_le_iff_lt RS iff_sym), |
1810 |
nat_into_Ord, |
|
1811 |
add_le_self]) 1); |
|
3425 | 1812 |
qed "eps_e_gr_add"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1813 |
|
5136 | 1814 |
Goal (* eps_e_gr *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1815 |
"[|n le m; m:nat; n:nat|] ==> eps(DD,ee,m,n) = e_gr(DD,ee,m,n)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1816 |
by (rtac le_exists 1); |
5136 | 1817 |
by (assume_tac 1); |
1818 |
by (asm_simp_tac(simpset() addsimps[eps_e_gr_add]) 1); |
|
1819 |
by (REPEAT (assume_tac 1)); |
|
3425 | 1820 |
qed "eps_e_gr"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1821 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1822 |
val prems = goal Limit.thy (* eps_succ_ee *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1823 |
"[|!!n. n:nat ==> emb(DD`n,DD`succ(n),ee`n); m:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1824 |
\ eps(DD,ee,m,succ(m)) = ee`m"; |
5529 | 1825 |
by (asm_simp_tac(simpset() addsimps eps_e_less::le_succ_iff::e_less_succ_emb:: |
1826 |
prems) 1); |
|
3425 | 1827 |
qed "eps_succ_ee"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1828 |
|
5136 | 1829 |
Goal (* eps_succ_Rp *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1830 |
"[|emb_chain(DD,ee); m:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1831 |
\ eps(DD,ee,succ(m),m) = Rp(DD`m,DD`succ(m),ee`m)"; |
5529 | 1832 |
by (asm_simp_tac(simpset() addsimps eps_e_gr::le_succ_iff::e_gr_succ_emb:: |
1833 |
prems) 1); |
|
3425 | 1834 |
qed "eps_succ_Rp"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1835 |
|
5136 | 1836 |
Goal (* eps_cont *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1837 |
"[|emb_chain(DD,ee); m:nat; n:nat|] ==> eps(DD,ee,m,n): cont(DD`m,DD`n)"; |
5136 | 1838 |
by (res_inst_tac [("i","m"),("j","n")] nat_linear_le 1); |
1839 |
by (ALLGOALS (asm_simp_tac(simpset() addsimps [eps_e_less,e_less_cont, |
|
1840 |
eps_e_gr,e_gr_cont]))); |
|
3425 | 1841 |
qed "eps_cont"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1842 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1843 |
(* Theorems about splitting. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1844 |
|
5136 | 1845 |
Goal (* eps_split_add_left *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1846 |
"[|n le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1847 |
\ eps(DD,ee,m,m#+k) = eps(DD,ee,m#+n,m#+k) O eps(DD,ee,m,m#+n)"; |
4091 | 1848 |
by (asm_simp_tac(simpset() addsimps |
5136 | 1849 |
[eps_e_less,add_le_self,add_le_mono]) 1); |
1850 |
by (auto_tac (claset() addIs [e_less_split_add], simpset())); |
|
3425 | 1851 |
qed "eps_split_add_left"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1852 |
|
5136 | 1853 |
Goal (* eps_split_add_left_rev *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1854 |
"[|n le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1855 |
\ eps(DD,ee,m,m#+n) = eps(DD,ee,m#+k,m#+n) O eps(DD,ee,m,m#+k)"; |
4091 | 1856 |
by (asm_simp_tac(simpset() addsimps |
5136 | 1857 |
[eps_e_less_add,eps_e_gr,add_le_self,add_le_mono]) 1); |
1858 |
by (auto_tac (claset() addIs [e_less_e_gr_split_add], simpset())); |
|
3425 | 1859 |
qed "eps_split_add_left_rev"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1860 |
|
5136 | 1861 |
Goal (* eps_split_add_right *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1862 |
"[|m le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1863 |
\ eps(DD,ee,n#+k,n) = eps(DD,ee,n#+m,n) O eps(DD,ee,n#+k,n#+m)"; |
4091 | 1864 |
by (asm_simp_tac(simpset() addsimps |
5136 | 1865 |
[eps_e_gr,add_le_self,add_le_mono]) 1); |
1866 |
by (auto_tac (claset() addIs [e_gr_split_add], simpset())); |
|
3425 | 1867 |
qed "eps_split_add_right"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1868 |
|
5136 | 1869 |
Goal (* eps_split_add_right_rev *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1870 |
"[|m le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1871 |
\ eps(DD,ee,n#+m,n) = eps(DD,ee,n#+k,n) O eps(DD,ee,n#+m,n#+k)"; |
4091 | 1872 |
by (asm_simp_tac(simpset() addsimps |
5136 | 1873 |
[eps_e_gr_add,eps_e_less,add_le_self,add_le_mono]) 1); |
1874 |
by (auto_tac (claset() addIs [e_gr_e_less_split_add], simpset())); |
|
3425 | 1875 |
qed "eps_split_add_right_rev"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1876 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1877 |
(* Arithmetic, little support in Isabelle/ZF. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1878 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1879 |
val prems = goal Limit.thy (* le_exists_lemma *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1880 |
"[|n le k; k le m; \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1881 |
\ !!p q. [|p le q; k=n#+p; m=n#+q; p:nat; q:nat|] ==> R; \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1882 |
\ m:nat; n:nat; k:nat|]==>R"; |
1623 | 1883 |
by (rtac (hd prems RS le_exists) 1); |
1884 |
by (rtac (le_exists) 1); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1885 |
by (rtac le_trans 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1886 |
(* Careful *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1887 |
by (resolve_tac prems 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1888 |
by (resolve_tac prems 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1889 |
by (resolve_tac prems 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1890 |
by (assume_tac 2); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1891 |
by (assume_tac 2); |
1623 | 1892 |
by (cut_facts_tac[hd prems,hd(tl prems)]1); |
2469 | 1893 |
by (Asm_full_simp_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1894 |
by (etac add_le_elim1 1); |
5136 | 1895 |
by (REPEAT (ares_tac prems 1)); |
3425 | 1896 |
qed "le_exists_lemma"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1897 |
|
5136 | 1898 |
Goal (* eps_split_left_le *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1899 |
"[|m le k; k le n; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1900 |
\ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1901 |
by (rtac le_exists_lemma 1); |
5136 | 1902 |
by (REPEAT (assume_tac 1)); |
2469 | 1903 |
by (Asm_simp_tac 1); |
5136 | 1904 |
by (auto_tac (claset() addIs [eps_split_add_left], simpset())); |
3425 | 1905 |
qed "eps_split_left_le"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1906 |
|
5136 | 1907 |
Goal (* eps_split_left_le_rev *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1908 |
"[|m le n; n le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1909 |
\ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1910 |
by (rtac le_exists_lemma 1); |
5136 | 1911 |
by (REPEAT (assume_tac 1)); |
2469 | 1912 |
by (Asm_simp_tac 1); |
5136 | 1913 |
by (auto_tac (claset() addIs [eps_split_add_left_rev], simpset())); |
3425 | 1914 |
qed "eps_split_left_le_rev"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1915 |
|
5136 | 1916 |
Goal (* eps_split_right_le *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1917 |
"[|n le k; k le m; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1918 |
\ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1919 |
by (rtac le_exists_lemma 1); |
5136 | 1920 |
by (REPEAT (assume_tac 1)); |
2469 | 1921 |
by (Asm_simp_tac 1); |
5136 | 1922 |
by (auto_tac (claset() addIs [eps_split_add_right], simpset())); |
3425 | 1923 |
qed "eps_split_right_le"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1924 |
|
5136 | 1925 |
Goal (* eps_split_right_le_rev *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1926 |
"[|n le m; m le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1927 |
\ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1928 |
by (rtac le_exists_lemma 1); |
5136 | 1929 |
by (REPEAT (assume_tac 1)); |
2469 | 1930 |
by (Asm_simp_tac 1); |
5136 | 1931 |
by (auto_tac (claset() addIs [eps_split_add_right_rev], simpset())); |
3425 | 1932 |
qed "eps_split_right_le_rev"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1933 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1934 |
(* The desired two theorems about `splitting'. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1935 |
|
5136 | 1936 |
Goal (* eps_split_left *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1937 |
"[|m le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1938 |
\ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"; |
2469 | 1939 |
by (rtac nat_linear_le 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1940 |
by (rtac eps_split_right_le_rev 4); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1941 |
by (assume_tac 4); |
2469 | 1942 |
by (rtac nat_linear_le 3); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1943 |
by (rtac eps_split_left_le 5); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1944 |
by (assume_tac 6); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1945 |
by (rtac eps_split_left_le_rev 10); |
5136 | 1946 |
by (REPEAT (assume_tac 1)); (* 20 trivial subgoals *) |
3425 | 1947 |
qed "eps_split_left"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1948 |
|
5136 | 1949 |
Goal (* eps_split_right *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1950 |
"[|n le k; emb_chain(DD,ee); m:nat; n:nat; k:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1951 |
\ eps(DD,ee,m,n) = eps(DD,ee,k,n) O eps(DD,ee,m,k)"; |
2469 | 1952 |
by (rtac nat_linear_le 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1953 |
by (rtac eps_split_left_le_rev 3); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1954 |
by (assume_tac 3); |
2469 | 1955 |
by (rtac nat_linear_le 8); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1956 |
by (rtac eps_split_right_le 10); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1957 |
by (assume_tac 11); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1958 |
by (rtac eps_split_right_le_rev 15); |
5136 | 1959 |
by (REPEAT (assume_tac 1)); (* 20 trivial subgoals *) |
3425 | 1960 |
qed "eps_split_right"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1961 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1962 |
(*----------------------------------------------------------------------*) |
1461 | 1963 |
(* That was eps: D_m -> D_n, NEXT rho_emb: D_n -> Dinf. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1964 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1965 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1966 |
(* Considerably shorter. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1967 |
|
5136 | 1968 |
Goalw [rho_emb_def] (* rho_emb_fun *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1969 |
"[|emb_chain(DD,ee); n:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1970 |
\ rho_emb(DD,ee,n): set(DD`n) -> set(Dinf(DD,ee))"; |
5136 | 1971 |
brr[lam_type, DinfI, eps_cont RS cont_fun RS apply_type] 1; |
2469 | 1972 |
by (Asm_simp_tac 1); |
5136 | 1973 |
by (res_inst_tac [("i","succ(na)"),("j","n")] nat_linear_le 1); |
1974 |
by (Blast_tac 1); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1975 |
by (assume_tac 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1976 |
(* The easiest would be to apply add1 everywhere also in the assumptions, |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1977 |
but since x le y is x<succ(y) simplification does too much with this thm. *) |
2034 | 1978 |
by (stac eps_split_right_le 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1979 |
by (assume_tac 2); |
2469 | 1980 |
by (asm_simp_tac(ZF_ss addsimps [add1]) 1); |
5136 | 1981 |
brr[add_le_self,nat_0I,nat_succI] 1; |
1982 |
by (asm_simp_tac(simpset() addsimps[eps_succ_Rp]) 1); |
|
2034 | 1983 |
by (stac comp_fun_apply 1); |
5136 | 1984 |
brr[eps_fun, nat_succI, Rp_cont RS cont_fun, emb_chain_emb, emb_chain_cpo,refl] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1985 |
(* Now the second part of the proof. Slightly different than HOL. *) |
5136 | 1986 |
by (asm_simp_tac(simpset() addsimps[eps_e_less,nat_succI]) 1); |
1623 | 1987 |
by (etac (le_iff RS iffD1 RS disjE) 1); |
5136 | 1988 |
by (asm_simp_tac(simpset() addsimps[e_less_le]) 1); |
2034 | 1989 |
by (stac comp_fun_apply 1); |
5136 | 1990 |
brr[e_less_cont,cont_fun,emb_chain_emb,emb_cont] 1; |
2034 | 1991 |
by (stac embRp_eq_thm 1); |
5136 | 1992 |
brr[emb_chain_emb, e_less_cont RS cont_fun RS apply_type, emb_chain_cpo, nat_succI] 1; |
1993 |
by (asm_simp_tac(simpset() addsimps[eps_e_less]) 1); |
|
1994 |
by (dtac asm_rl 1); |
|
1995 |
by (asm_full_simp_tac(simpset() addsimps[eps_succ_Rp, e_less_eq, id_apply, nat_succI]) 1); |
|
3425 | 1996 |
qed "rho_emb_fun"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1997 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1998 |
val rho_emb_apply1 = prove_goalw Limit.thy [rho_emb_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
1999 |
"!!z. x:set(DD`n) ==> rho_emb(DD,ee,n)`x = (lam m:nat. eps(DD,ee,n,m)`x)" |
2469 | 2000 |
(fn prems => [Asm_simp_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2001 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2002 |
val rho_emb_apply2 = prove_goalw Limit.thy [rho_emb_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2003 |
"!!z. [|x:set(DD`n); m:nat|] ==> rho_emb(DD,ee,n)`x`m = eps(DD,ee,n,m)`x" |
2469 | 2004 |
(fn prems => [Asm_simp_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2005 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2006 |
val rho_emb_id = prove_goal Limit.thy |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2007 |
"!!z. [| x:set(DD`n); n:nat|] ==> rho_emb(DD,ee,n)`x`n = x" |
4091 | 2008 |
(fn prems => [asm_simp_tac(simpset() addsimps[rho_emb_apply2,eps_id,id_thm]) 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2009 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2010 |
(* Shorter proof, 23 against 62. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2011 |
|
5136 | 2012 |
Goal (* rho_emb_cont *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2013 |
"[|emb_chain(DD,ee); n:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2014 |
\ rho_emb(DD,ee,n): cont(DD`n,Dinf(DD,ee))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2015 |
by (rtac contI 1); |
5136 | 2016 |
brr[rho_emb_fun] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2017 |
by (rtac rel_DinfI 1); |
1623 | 2018 |
by (SELECT_GOAL(rewtac rho_emb_def) 1); |
2469 | 2019 |
by (Asm_simp_tac 1); |
5136 | 2020 |
brr[eps_cont RS cont_mono, Dinf_prod,apply_type,rho_emb_fun] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2021 |
(* Continuity, different order, slightly different proofs. *) |
2034 | 2022 |
by (stac lub_Dinf 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2023 |
by (rtac chainI 1); |
5136 | 2024 |
brr[lam_type, rho_emb_fun RS apply_type, chain_in] 1; |
2469 | 2025 |
by (Asm_simp_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2026 |
by (rtac rel_DinfI 1); |
5136 | 2027 |
by (asm_simp_tac(simpset() addsimps [rho_emb_apply2,chain_in]) 1); |
2028 |
brr[eps_cont RS cont_mono, chain_rel, Dinf_prod, rho_emb_fun RS apply_type, chain_in,nat_succI] 1; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2029 |
(* Now, back to the result of applying lub_Dinf *) |
5136 | 2030 |
by (asm_simp_tac(simpset() addsimps [rho_emb_apply2,chain_in]) 1); |
2034 | 2031 |
by (stac rho_emb_apply1 1); |
5136 | 2032 |
brr[cpo_lub RS islub_in, emb_chain_cpo] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2033 |
by (rtac fun_extension 1); |
5136 | 2034 |
brr[lam_type, eps_cont RS cont_fun RS apply_type, cpo_lub RS islub_in, emb_chain_cpo] 1; |
2035 |
brr[cont_chain,eps_cont,emb_chain_cpo] 1; |
|
2469 | 2036 |
by (Asm_simp_tac 1); |
5136 | 2037 |
by (asm_simp_tac(simpset() addsimps[eps_cont RS cont_lub]) 1); |
3425 | 2038 |
qed "rho_emb_cont"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2039 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2040 |
(* 32 vs 61, using safe_tac with imp in asm would be unfortunate (5steps) *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2041 |
|
5136 | 2042 |
Goal (* lemma1 *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2043 |
"[|m le n; emb_chain(DD,ee); x:set(Dinf(DD,ee)); m:nat; n:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2044 |
\ rel(DD`n,e_less(DD,ee,m,n)`(x`m),x`n)"; |
5136 | 2045 |
by (etac rev_mp 1); (* For induction proof *) |
1623 | 2046 |
by (res_inst_tac[("n","n")]nat_induct 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2047 |
by (rtac impI 2); |
5136 | 2048 |
by (asm_full_simp_tac (simpset() addsimps [e_less_eq]) 2); |
2034 | 2049 |
by (stac id_thm 2); |
5136 | 2050 |
brr[apply_type,Dinf_prod,cpo_refl,emb_chain_cpo,nat_0I] 1; |
4091 | 2051 |
by (asm_full_simp_tac (simpset() addsimps [le_succ_iff]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2052 |
by (rtac impI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2053 |
by (etac disjE 1); |
1623 | 2054 |
by (dtac mp 1 THEN atac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2055 |
by (rtac cpo_trans 1); |
2034 | 2056 |
by (stac e_less_le 2); |
5136 | 2057 |
brr[emb_chain_cpo,nat_succI] 1; |
2034 | 2058 |
by (stac comp_fun_apply 1); |
5136 | 2059 |
brr[emb_chain_emb RS emb_cont, e_less_cont, cont_fun, apply_type, Dinf_prod] 1; |
1623 | 2060 |
by (res_inst_tac[("y","x`xa")](emb_chain_emb RS emb_cont RS cont_mono) 1); |
5136 | 2061 |
brr[e_less_cont RS cont_fun, apply_type,Dinf_prod] 1; |
1623 | 2062 |
by (res_inst_tac[("x1","x"),("n1","xa")](Dinf_eq RS subst) 1); |
2063 |
by (rtac (comp_fun_apply RS subst) 3); |
|
2064 |
by (res_inst_tac |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2065 |
[("P", |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2066 |
"%z. rel(DD ` succ(xa), \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2067 |
\ (ee ` xa O Rp(?DD46(xa) ` xa,?DD46(xa) ` succ(xa),?ee46(xa) ` xa)) ` \ |
1623 | 2068 |
\ (x ` succ(xa)),z)")](id_thm RS subst) 6); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2069 |
by (rtac rel_cf 7); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2070 |
(* Dinf and cont_fun doesn't go well together, both Pi(_,%x._). *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2071 |
(* brr solves 11 of 12 subgoals *) |
5136 | 2072 |
brr[Dinf_prod RS apply_type, cont_fun, Rp_cont, e_less_cont, emb_cont, emb_chain_emb, emb_chain_cpo, apply_type, embRp_rel, disjI1 RS (le_succ_iff RS iffD2), nat_succI] 1; |
2073 |
by (asm_full_simp_tac (simpset() addsimps [e_less_eq]) 1); |
|
2034 | 2074 |
by (stac id_thm 1); |
5136 | 2075 |
by (auto_tac (claset() addIs [apply_type,Dinf_prod,emb_chain_cpo], simpset())); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2076 |
val lemma1 = result(); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2077 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2078 |
(* 18 vs 40 *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2079 |
|
5136 | 2080 |
Goal (* lemma2 *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2081 |
"[|n le m; emb_chain(DD,ee); x:set(Dinf(DD,ee)); m:nat; n:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2082 |
\ rel(DD`n,e_gr(DD,ee,m,n)`(x`m),x`n)"; |
5136 | 2083 |
by (etac rev_mp 1); (* For induction proof *) |
1623 | 2084 |
by (res_inst_tac[("n","m")]nat_induct 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2085 |
by (rtac impI 2); |
5136 | 2086 |
by (asm_full_simp_tac (simpset() addsimps [e_gr_eq]) 2); |
2034 | 2087 |
by (stac id_thm 2); |
5136 | 2088 |
brr[apply_type,Dinf_prod,cpo_refl,emb_chain_cpo,nat_0I] 1; |
4091 | 2089 |
by (asm_full_simp_tac (simpset() addsimps [le_succ_iff]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2090 |
by (rtac impI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2091 |
by (etac disjE 1); |
1623 | 2092 |
by (dtac mp 1 THEN atac 1); |
2034 | 2093 |
by (stac e_gr_le 1); |
2094 |
by (stac comp_fun_apply 4); |
|
2095 |
by (stac Dinf_eq 7); |
|
5136 | 2096 |
brr[emb_chain_emb, emb_chain_cpo, Rp_cont, e_gr_cont, cont_fun, emb_cont, apply_type,Dinf_prod,nat_succI] 1; |
2097 |
by (asm_full_simp_tac (simpset() addsimps [e_gr_eq]) 1); |
|
2034 | 2098 |
by (stac id_thm 1); |
5136 | 2099 |
by (auto_tac (claset() addIs [apply_type,Dinf_prod,emb_chain_cpo], simpset())); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2100 |
val lemma2 = result(); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2101 |
|
5136 | 2102 |
Goalw [eps_def] (* eps1 *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2103 |
"[|emb_chain(DD,ee); x:set(Dinf(DD,ee)); m:nat; n:nat|] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2104 |
\ rel(DD`n,eps(DD,ee,m,n)`(x`m),x`n)"; |
5116
8eb343ab5748
Renamed expand_if to split_if and setloop split_tac to addsplits,
paulson
parents:
5068
diff
changeset
|
2105 |
by (split_tac [split_if] 1); |
5136 | 2106 |
brr[conjI, impI, lemma1, not_le_iff_lt RS iffD1 RS leI RS lemma2, nat_into_Ord] 1; |
3425 | 2107 |
qed "eps1"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2108 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2109 |
(* The following theorem is needed/useful due to type check for rel_cfI, |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2110 |
but also elsewhere. |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2111 |
Look for occurences of rel_cfI, rel_DinfI, etc to evaluate the problem. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2112 |
|
5136 | 2113 |
Goal (* lam_Dinf_cont *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2114 |
"[| emb_chain(DD,ee); n:nat |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2115 |
\ (lam x:set(Dinf(DD,ee)). x`n) : cont(Dinf(DD,ee),DD`n)"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2116 |
by (rtac contI 1); |
5136 | 2117 |
brr[lam_type,apply_type,Dinf_prod] 1; |
2469 | 2118 |
by (Asm_simp_tac 1); |
5136 | 2119 |
brr[rel_Dinf] 1; |
2034 | 2120 |
by (stac beta 1); |
5136 | 2121 |
by (auto_tac (claset() addIs [cpo_Dinf,islub_in,cpo_lub], simpset())); |
2122 |
by (asm_simp_tac(simpset() addsimps[chain_in,lub_Dinf]) 1); |
|
3425 | 2123 |
qed "lam_Dinf_cont"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2124 |
|
5136 | 2125 |
Goalw [rho_proj_def] (* rho_projpair *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2126 |
"[| emb_chain(DD,ee); n:nat |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2127 |
\ projpair(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n),rho_proj(DD,ee,n))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2128 |
by (rtac projpairI 1); |
5136 | 2129 |
brr[rho_emb_cont] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2130 |
(* lemma used, introduced because same fact needed below due to rel_cfI. *) |
5136 | 2131 |
brr[lam_Dinf_cont] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2132 |
(*-----------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2133 |
(* This part is 7 lines, but 30 in HOL (75% reduction!) *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2134 |
by (rtac fun_extension 1); |
2034 | 2135 |
by (stac id_thm 3); |
2136 |
by (stac comp_fun_apply 4); |
|
2137 |
by (stac beta 7); |
|
2138 |
by (stac rho_emb_id 8); |
|
5136 | 2139 |
brr[comp_fun, id_type, lam_type, rho_emb_fun, Dinf_prod RS apply_type, apply_type,refl] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2140 |
(*^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^^*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2141 |
by (rtac rel_cfI 1); (* ------------------>>>Yields type cond, not in HOL *) |
2034 | 2142 |
by (stac id_thm 1); |
2143 |
by (stac comp_fun_apply 2); |
|
2144 |
by (stac beta 5); |
|
2145 |
by (stac rho_emb_apply1 6); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2146 |
by (rtac rel_DinfI 7); (* ------------------>>>Yields type cond, not in HOL *) |
2034 | 2147 |
by (stac beta 7); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2148 |
brr(eps1::lam_type::rho_emb_fun::eps_fun:: (* Dinf_prod bad with lam_type *) |
5136 | 2149 |
[Dinf_prod RS apply_type, refl]) 1; |
2150 |
brr[apply_type, eps_fun, Dinf_prod, comp_pres_cont, rho_emb_cont, lam_Dinf_cont,id_cont,cpo_Dinf,emb_chain_cpo] 1; |
|
3425 | 2151 |
qed "rho_projpair"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2152 |
|
5136 | 2153 |
Goalw [emb_def] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2154 |
"[| emb_chain(DD,ee); n:nat |] ==> emb(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n))"; |
5136 | 2155 |
by (auto_tac (claset() addIs [exI,rho_projpair], simpset())); |
3425 | 2156 |
qed "emb_rho_emb"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2157 |
|
5268 | 2158 |
Goal "[| emb_chain(DD,ee); n:nat |] ==> \ |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2159 |
\ rho_proj(DD,ee,n) : cont(Dinf(DD,ee),DD`n)"; |
5136 | 2160 |
by (auto_tac (claset() addIs [rho_projpair,projpair_p_cont], simpset())); |
3425 | 2161 |
qed "rho_proj_cont"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2162 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2163 |
(*----------------------------------------------------------------------*) |
1461 | 2164 |
(* Commutivity and universality. *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2165 |
(*----------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2166 |
|
5136 | 2167 |
val prems = Goalw [commute_def] (* commuteI *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2168 |
"[| !!n. n:nat ==> emb(DD`n,E,r(n)); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2169 |
\ !!m n. [|m le n; m:nat; n:nat|] ==> r(n) O eps(DD,ee,m,n) = r(m) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2170 |
\ commute(DD,ee,E,r)"; |
4152 | 2171 |
by Safe_tac; |
5136 | 2172 |
by (REPEAT (ares_tac prems 1)); |
3425 | 2173 |
qed "commuteI"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2174 |
|
5136 | 2175 |
Goalw [commute_def] (* commute_emb *) |
2176 |
"[| commute(DD,ee,E,r); n:nat |] ==> emb(DD`n,E,r(n))"; |
|
2469 | 2177 |
by (Fast_tac 1); |
3425 | 2178 |
qed "commute_emb"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2179 |
|
5136 | 2180 |
Goalw [commute_def] (* commute_eq *) |
2181 |
"[| commute(DD,ee,E,r); m le n; m:nat; n:nat |] ==> \ |
|
2182 |
\ r(n) O eps(DD,ee,m,n) = r(m) "; |
|
2183 |
by (Blast_tac 1); |
|
3425 | 2184 |
qed "commute_eq"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2185 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2186 |
(* Shorter proof: 11 vs 46 lines. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2187 |
|
5136 | 2188 |
Goal (* rho_emb_commute *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2189 |
"emb_chain(DD,ee) ==> commute(DD,ee,Dinf(DD,ee),rho_emb(DD,ee))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2190 |
by (rtac commuteI 1); |
5136 | 2191 |
brr[emb_rho_emb] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2192 |
by (rtac fun_extension 1); (* Manual instantiation in HOL. *) |
2034 | 2193 |
by (stac comp_fun_apply 3); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2194 |
by (rtac fun_extension 6); (* Next, clean up and instantiate unknowns *) |
5136 | 2195 |
brr[comp_fun,rho_emb_fun,eps_fun,Dinf_prod,apply_type] 1; |
1623 | 2196 |
by (asm_simp_tac |
5136 | 2197 |
(simpset() addsimps[rho_emb_apply2, eps_fun RS apply_type]) 1); |
1623 | 2198 |
by (rtac (comp_fun_apply RS subst) 1); |
2199 |
by (rtac (eps_split_left RS subst) 4); |
|
5136 | 2200 |
by (auto_tac (claset() addIs [eps_fun], simpset())); |
3425 | 2201 |
qed "rho_emb_commute"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2202 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2203 |
val le_succ = prove_goal Arith.thy "n:nat ==> n le succ(n)" |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2204 |
(fn prems => |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2205 |
[REPEAT (ares_tac |
1623 | 2206 |
((disjI1 RS(le_succ_iff RS iffD2))::le_refl::nat_into_Ord::prems) 1)]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2207 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2208 |
(* Shorter proof: 21 vs 83 (106 - 23, due to OAssoc complication) *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2209 |
|
5136 | 2210 |
Goal (* commute_chain *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2211 |
"[| commute(DD,ee,E,r); emb_chain(DD,ee); cpo(E) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2212 |
\ chain(cf(E,E),lam n:nat. r(n) O Rp(DD`n,E,r(n)))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2213 |
by (rtac chainI 1); |
5136 | 2214 |
by (blast_tac (claset() addIs [lam_type, cont_cf, comp_pres_cont, commute_emb, Rp_cont, emb_cont, emb_chain_cpo]) 1); |
2469 | 2215 |
by (Asm_simp_tac 1); |
1623 | 2216 |
by (res_inst_tac[("r1","r"),("m1","n")](commute_eq RS subst) 1); |
5136 | 2217 |
brr[le_succ,nat_succI] 1; |
2034 | 2218 |
by (stac Rp_comp 1); |
5136 | 2219 |
brr[emb_eps,commute_emb,emb_chain_cpo,le_succ,nat_succI] 1; |
1623 | 2220 |
by (rtac (comp_assoc RS subst) 1); (* Remember that comp_assoc is simpler in Isa *) |
2221 |
by (res_inst_tac[("r1","r(succ(n))")](comp_assoc RS ssubst) 1); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2222 |
by (rtac comp_mono 1); |
5136 | 2223 |
by (REPEAT |
2224 |
(blast_tac (claset() addIs [comp_pres_cont, eps_cont, emb_eps, |
|
2225 |
commute_emb, Rp_cont, emb_cont, |
|
2226 |
emb_chain_cpo,le_succ]) 1)); |
|
1623 | 2227 |
by (res_inst_tac[("b","r(succ(n))")](comp_id RS subst) 1); (* 1 subst too much *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2228 |
by (rtac comp_mono 2); |
5136 | 2229 |
by (REPEAT |
2230 |
(blast_tac (claset() addIs [comp_pres_cont, eps_cont, emb_eps, emb_id, |
|
2231 |
commute_emb, Rp_cont, emb_cont, cont_fun, |
|
2232 |
emb_chain_cpo,le_succ]) 1)); |
|
2233 |
by (stac comp_id 1); (* Undoes "1 subst too much", typing next anyway *) |
|
2234 |
by (REPEAT |
|
2235 |
(blast_tac (claset() addIs [cont_fun, Rp_cont, emb_cont, commute_emb, |
|
2236 |
cont_cf, cpo_cf, emb_chain_cpo, |
|
2237 |
embRp_rel,emb_eps,le_succ]) 1)); |
|
3425 | 2238 |
qed "commute_chain"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2239 |
|
5136 | 2240 |
Goal (* rho_emb_chain *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2241 |
"emb_chain(DD,ee) ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2242 |
\ chain(cf(Dinf(DD,ee),Dinf(DD,ee)), \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2243 |
\ lam n:nat. rho_emb(DD,ee,n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n)))"; |
5136 | 2244 |
by (auto_tac (claset() addIs [commute_chain,rho_emb_commute,cpo_Dinf], simpset())); |
3425 | 2245 |
qed "rho_emb_chain"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2246 |
|
5136 | 2247 |
Goal "[| emb_chain(DD,ee); x:set(Dinf(DD,ee)) |] ==> \ |
2248 |
\ chain(Dinf(DD,ee), \ |
|
2249 |
\ lam n:nat. \ |
|
2250 |
\ (rho_emb(DD,ee,n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n)))`x)"; |
|
2251 |
by (dtac (rho_emb_chain RS chain_cf) 1); |
|
2252 |
by (assume_tac 1); |
|
2469 | 2253 |
by (Asm_full_simp_tac 1); |
3425 | 2254 |
qed "rho_emb_chain_apply1"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2255 |
|
5136 | 2256 |
Goal "[| chain(iprod(DD),X); emb_chain(DD,ee); n:nat |] ==> \ |
2257 |
\ chain(DD`n,lam m:nat. X `m `n)"; |
|
2258 |
by (auto_tac (claset() addIs [chain_iprod,emb_chain_cpo], simpset())); |
|
3425 | 2259 |
qed "chain_iprod_emb_chain"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2260 |
|
5136 | 2261 |
Goal (* rho_emb_chain_apply2 *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2262 |
"[| emb_chain(DD,ee); x:set(Dinf(DD,ee)); n:nat |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2263 |
\ chain \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2264 |
\ (DD`n, \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2265 |
\ lam xa:nat. \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2266 |
\ (rho_emb(DD, ee, xa) O Rp(DD ` xa, Dinf(DD, ee),rho_emb(DD, ee, xa))) ` \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2267 |
\ x ` n)"; |
5136 | 2268 |
by (forward_tac [rho_emb_chain_apply1 RS chain_Dinf RS chain_iprod_emb_chain] 1); |
2269 |
by Auto_tac; |
|
3425 | 2270 |
qed "rho_emb_chain_apply2"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2271 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2272 |
(* Shorter proof: 32 vs 72 (roughly), Isabelle proof has lemmas. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2273 |
|
5136 | 2274 |
Goal (* rho_emb_lub *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2275 |
"emb_chain(DD,ee) ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2276 |
\ lub(cf(Dinf(DD,ee),Dinf(DD,ee)), \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2277 |
\ lam n:nat. rho_emb(DD,ee,n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n))) = \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2278 |
\ id(set(Dinf(DD,ee)))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2279 |
by (rtac cpo_antisym 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2280 |
by (rtac cpo_cf 1); (* Instantiate variable, continued below (would loop otherwise) *) |
5136 | 2281 |
brr[cpo_Dinf] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2282 |
by (rtac islub_least 1); |
5136 | 2283 |
brr[cpo_lub,rho_emb_chain,cpo_cf,cpo_Dinf,isubI,cont_cf,id_cont] 1; |
2469 | 2284 |
by (Asm_simp_tac 1); |
5136 | 2285 |
brr[embRp_rel,emb_rho_emb,emb_chain_cpo,cpo_Dinf] 1; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2286 |
by (rtac rel_cfI 1); |
1623 | 2287 |
by (asm_simp_tac |
5136 | 2288 |
(simpset() addsimps[id_thm,lub_cf,rho_emb_chain,cpo_Dinf]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2289 |
by (rtac rel_DinfI 1); (* Addtional assumptions *) |
2034 | 2290 |
by (stac lub_Dinf 1); |
5136 | 2291 |
brr[rho_emb_chain_apply1] 1; |
2292 |
brr[Dinf_prod, cpo_lub RS islub_in, id_cont, cpo_Dinf, cpo_cf, cf_cont, rho_emb_chain, rho_emb_chain_apply1, id_cont RS cont_cf] 2; |
|
2469 | 2293 |
by (Asm_simp_tac 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2294 |
by (rtac dominate_islub 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2295 |
by (rtac cpo_lub 3); |
5136 | 2296 |
brr[rho_emb_chain_apply2,emb_chain_cpo] 3; |
1623 | 2297 |
by (res_inst_tac[("x1","x`n")](chain_const RS chain_fun) 3); |
5136 | 2298 |
brr[islub_const, apply_type, Dinf_prod, emb_chain_cpo, chain_fun, rho_emb_chain_apply2] 2; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2299 |
by (rtac dominateI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2300 |
by (assume_tac 1); |
2469 | 2301 |
by (Asm_simp_tac 1); |
2034 | 2302 |
by (stac comp_fun_apply 1); |
5136 | 2303 |
brr[cont_fun,Rp_cont,emb_cont,emb_rho_emb,cpo_Dinf,emb_chain_cpo] 1; |
2034 | 2304 |
by (stac ((rho_projpair RS Rp_unique)) 1); |
1623 | 2305 |
by (SELECT_GOAL(rewtac rho_proj_def) 5); |
2469 | 2306 |
by (Asm_simp_tac 5); |
2034 | 2307 |
by (stac rho_emb_id 5); |
5136 | 2308 |
by (auto_tac (claset() addIs [cpo_Dinf,apply_type,Dinf_prod,emb_chain_cpo], |
2309 |
simpset())); |
|
3425 | 2310 |
qed "rho_emb_lub"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2311 |
|
5136 | 2312 |
Goal (* theta_chain, almost same prf as commute_chain *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2313 |
"[| commute(DD,ee,E,r); commute(DD,ee,G,f); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2314 |
\ emb_chain(DD,ee); cpo(E); cpo(G) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2315 |
\ chain(cf(E,G),lam n:nat. f(n) O Rp(DD`n,E,r(n)))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2316 |
by (rtac chainI 1); |
5136 | 2317 |
by (blast_tac (claset() addIs [lam_type, cont_cf, comp_pres_cont, commute_emb, Rp_cont,emb_cont,emb_chain_cpo]) 1); |
2469 | 2318 |
by (Asm_simp_tac 1); |
1623 | 2319 |
by (res_inst_tac[("r1","r"),("m1","n")](commute_eq RS subst) 1); |
2320 |
by (res_inst_tac[("r1","f"),("m1","n")](commute_eq RS subst) 5); |
|
5136 | 2321 |
brr[le_succ,nat_succI] 1; |
2034 | 2322 |
by (stac Rp_comp 1); |
5136 | 2323 |
brr[emb_eps,commute_emb,emb_chain_cpo,le_succ,nat_succI] 1; |
1623 | 2324 |
by (rtac (comp_assoc RS subst) 1); (* Remember that comp_assoc is simpler in Isa *) |
2325 |
by (res_inst_tac[("r1","f(succ(n))")](comp_assoc RS ssubst) 1); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2326 |
by (rtac comp_mono 1); |
5136 | 2327 |
by (REPEAT (blast_tac (claset() addIs [comp_pres_cont, eps_cont, emb_eps, commute_emb, Rp_cont, emb_cont,emb_chain_cpo,le_succ]) 1)); |
1623 | 2328 |
by (res_inst_tac[("b","f(succ(n))")](comp_id RS subst) 1); (* 1 subst too much *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2329 |
by (rtac comp_mono 2); |
5136 | 2330 |
by (REPEAT (blast_tac (claset() addIs[comp_pres_cont, eps_cont, emb_eps, emb_id, commute_emb, Rp_cont, emb_cont,cont_fun,emb_chain_cpo,le_succ]) 1)); |
2331 |
by (stac comp_id 1); (* Undoes "1 subst too much", typing next anyway *) |
|
2332 |
by (REPEAT |
|
2333 |
(blast_tac (claset() addIs[cont_fun, Rp_cont, emb_cont, commute_emb, |
|
2334 |
cont_cf, cpo_cf,emb_chain_cpo, |
|
2335 |
embRp_rel,emb_eps,le_succ]) 1)); |
|
3425 | 2336 |
qed "theta_chain"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2337 |
|
5136 | 2338 |
Goal (* theta_proj_chain, same prf as theta_chain *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2339 |
"[| commute(DD,ee,E,r); commute(DD,ee,G,f); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2340 |
\ emb_chain(DD,ee); cpo(E); cpo(G) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2341 |
\ chain(cf(G,E),lam n:nat. r(n) O Rp(DD`n,G,f(n)))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2342 |
by (rtac chainI 1); |
5136 | 2343 |
by (blast_tac (claset() addIs [lam_type, cont_cf, comp_pres_cont, commute_emb, Rp_cont,emb_cont,emb_chain_cpo]) 1); |
2469 | 2344 |
by (Asm_simp_tac 1); |
1623 | 2345 |
by (res_inst_tac[("r1","r"),("m1","n")](commute_eq RS subst) 1); |
2346 |
by (res_inst_tac[("r1","f"),("m1","n")](commute_eq RS subst) 5); |
|
5136 | 2347 |
brr[le_succ,nat_succI] 1; |
2034 | 2348 |
by (stac Rp_comp 1); |
5136 | 2349 |
brr[emb_eps,commute_emb,emb_chain_cpo,le_succ,nat_succI] 1; |
1623 | 2350 |
by (rtac (comp_assoc RS subst) 1); (* Remember that comp_assoc is simpler in Isa *) |
2351 |
by (res_inst_tac[("r1","r(succ(n))")](comp_assoc RS ssubst) 1); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2352 |
by (rtac comp_mono 1); |
5136 | 2353 |
by (REPEAT (blast_tac (claset() addIs [comp_pres_cont, eps_cont, emb_eps, commute_emb, Rp_cont, emb_cont,emb_chain_cpo,le_succ]) 1)); |
1623 | 2354 |
by (res_inst_tac[("b","r(succ(n))")](comp_id RS subst) 1); (* 1 subst too much *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2355 |
by (rtac comp_mono 2); |
5136 | 2356 |
by (REPEAT (blast_tac (claset() addIs[comp_pres_cont, eps_cont, emb_eps, emb_id, commute_emb, Rp_cont, emb_cont,cont_fun,emb_chain_cpo,le_succ]) 1)); |
2357 |
by (stac comp_id 1); (* Undoes "1 subst too much", typing next anyway *) |
|
2358 |
by (REPEAT |
|
2359 |
(blast_tac (claset() addIs[cont_fun, Rp_cont, emb_cont, commute_emb, |
|
2360 |
cont_cf, cpo_cf,emb_chain_cpo,embRp_rel, |
|
2361 |
emb_eps,le_succ]) 1)); |
|
3425 | 2362 |
qed "theta_proj_chain"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2363 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2364 |
(* Simplification with comp_assoc is possible inside a lam-abstraction, |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2365 |
because it does not have assumptions. If it had, as the HOL-ST theorem |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2366 |
too strongly has, we would be in deep trouble due to the lack of proper |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2367 |
conditional rewriting (a HOL contrib provides something that works). *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2368 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2369 |
(* Controlled simplification inside lambda: introduce lemmas *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2370 |
|
5136 | 2371 |
Goal "[| commute(DD,ee,E,r); commute(DD,ee,G,f); \ |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2372 |
\ emb_chain(DD,ee); cpo(E); cpo(G); x:nat |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2373 |
\ r(x) O Rp(DD ` x, G, f(x)) O f(x) O Rp(DD ` x, E, r(x)) = \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2374 |
\ r(x) O Rp(DD ` x, E, r(x))"; |
1623 | 2375 |
by (res_inst_tac[("s1","f(x)")](comp_assoc RS subst) 1); |
2034 | 2376 |
by (stac embRp_eq 1); |
2377 |
by (stac id_comp 4); |
|
5136 | 2378 |
by (auto_tac (claset() addIs [cont_fun,Rp_cont,commute_emb,emb_chain_cpo], |
2379 |
simpset())); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2380 |
val lemma = result(); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2381 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2382 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2383 |
(* Shorter proof (but lemmas): 19 vs 79 (103 - 24, due to OAssoc) *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2384 |
|
5136 | 2385 |
Goalw [projpair_def,rho_proj_def] (* theta_projpair *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2386 |
"[| lub(cf(E,E), lam n:nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2387 |
\ commute(DD,ee,E,r); commute(DD,ee,G,f); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2388 |
\ emb_chain(DD,ee); cpo(E); cpo(G) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2389 |
\ projpair \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2390 |
\ (E,G, \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2391 |
\ lub(cf(E,G), lam n:nat. f(n) O Rp(DD`n,E,r(n))), \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2392 |
\ lub(cf(G,E), lam n:nat. r(n) O Rp(DD`n,G,f(n))))"; |
4152 | 2393 |
by Safe_tac; |
2034 | 2394 |
by (stac comp_lubs 3); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2395 |
(* The following one line is 15 lines in HOL, and includes existentials. *) |
5136 | 2396 |
brr[cf_cont,islub_in,cpo_lub,cpo_cf,theta_chain,theta_proj_chain] 1; |
4091 | 2397 |
by (simp_tac (simpset() addsimps[comp_assoc]) 1); |
5136 | 2398 |
by (asm_simp_tac (simpset() addsimps[lemma]) 1); |
2399 |
by (stac comp_lubs 1); |
|
2400 |
brr[cf_cont,islub_in,cpo_lub,cpo_cf,theta_chain,theta_proj_chain] 1; |
|
4091 | 2401 |
by (simp_tac (simpset() addsimps[comp_assoc]) 1); |
5136 | 2402 |
by (asm_simp_tac (simpset() addsimps[lemma]) 1); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2403 |
by (rtac dominate_islub 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2404 |
by (rtac cpo_lub 2); |
5136 | 2405 |
brr[commute_chain, commute_emb, islub_const, cont_cf, id_cont, cpo_cf, chain_fun,chain_const] 2; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2406 |
by (rtac dominateI 1); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2407 |
by (assume_tac 1); |
2469 | 2408 |
by (Asm_simp_tac 1); |
5136 | 2409 |
by (blast_tac (claset() addIs [embRp_rel,commute_emb,emb_chain_cpo]) 1); |
3425 | 2410 |
qed "theta_projpair"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2411 |
|
5136 | 2412 |
Goalw [emb_def] |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2413 |
"[| lub(cf(E,E), lam n:nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2414 |
\ commute(DD,ee,E,r); commute(DD,ee,G,f); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2415 |
\ emb_chain(DD,ee); cpo(E); cpo(G) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2416 |
\ emb(E,G,lub(cf(E,G), lam n:nat. f(n) O Rp(DD`n,E,r(n))))"; |
5136 | 2417 |
by (blast_tac (claset() addIs [theta_projpair]) 1); |
3425 | 2418 |
qed "emb_theta"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2419 |
|
5136 | 2420 |
Goal (* mono_lemma *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2421 |
"[| g:cont(D,D'); cpo(D); cpo(D'); cpo(E) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2422 |
\ (lam f : cont(D',E). f O g) : mono(cf(D',E),cf(D,E))"; |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2423 |
by (rtac monoI 1); |
1623 | 2424 |
by (REPEAT(dtac cf_cont 2)); |
2469 | 2425 |
by (Asm_simp_tac 2); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2426 |
by (rtac comp_mono 2); |
1623 | 2427 |
by (SELECT_GOAL(rewrite_goals_tac[set_def,cf_def]) 1); |
2469 | 2428 |
by (Asm_simp_tac 1); |
5136 | 2429 |
by (auto_tac (claset() addIs [lam_type,comp_pres_cont,cpo_cf,cont_cf], |
2430 |
simpset())); |
|
3425 | 2431 |
qed "mono_lemma"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2432 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2433 |
(* PAINFUL: wish condrew with difficult conds on term bound in lam-abs. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2434 |
(* Introduces need for lemmas. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2435 |
|
5136 | 2436 |
Goal "[| commute(DD,ee,E,r); commute(DD,ee,G,f); \ |
2437 |
\ emb_chain(DD,ee); cpo(E); cpo(G); n:nat |] ==> \ |
|
2438 |
\ (lam na:nat. (lam f:cont(E, G). f O r(n)) ` \ |
|
2439 |
\ ((lam n:nat. f(n) O Rp(DD ` n, E, r(n))) ` na)) = \ |
|
2440 |
\ (lam na:nat. (f(na) O Rp(DD ` na, E, r(na))) O r(n))"; |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2441 |
by (rtac fun_extension 1); |
2034 | 2442 |
by (stac beta 3); |
2443 |
by (stac beta 4); |
|
2444 |
by (stac beta 5); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2445 |
by (rtac lam_type 1); |
2034 | 2446 |
by (stac beta 1); |
5136 | 2447 |
by (ALLGOALS Asm_simp_tac); |
2448 |
by (ALLGOALS (fast_tac (claset() addIs [lam_type, comp_pres_cont, Rp_cont, emb_cont, commute_emb, emb_chain_cpo]))); |
|
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2449 |
val lemma = result(); |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2450 |
|
5136 | 2451 |
Goal "[| commute(DD,ee,E,r); commute(DD,ee,G,f); \ |
2452 |
\ emb_chain(DD,ee); cpo(E); cpo(G); n:nat |] ==> \ |
|
2453 |
\ chain(cf(DD`n,G),lam x:nat. (f(x) O Rp(DD ` x, E, r(x))) O r(n))"; |
|
2454 |
by (rtac (lemma RS subst) 1); |
|
2455 |
by (REPEAT |
|
2456 |
(blast_tac (claset() addIs[theta_chain,emb_chain_cpo, |
|
2457 |
commute_emb RS emb_cont RS mono_lemma RS mono_chain]) 1)); |
|
3425 | 2458 |
qed "chain_lemma"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2459 |
|
5136 | 2460 |
Goalw [suffix_def] (* suffix_lemma *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2461 |
"[| commute(DD,ee,E,r); commute(DD,ee,G,f); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2462 |
\ emb_chain(DD,ee); cpo(E); cpo(G); cpo(DD`x); x:nat |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2463 |
\ suffix(lam n:nat. (f(n) O Rp(DD`n,E,r(n))) O r(x),x) = (lam n:nat. f(x))"; |
5136 | 2464 |
by (Asm_simp_tac 1); |
2465 |
by (rtac (lam_type RS fun_extension) 1); |
|
2466 |
by (REPEAT (blast_tac (claset() addIs [lam_type, comp_fun, cont_fun, Rp_cont, emb_cont, commute_emb, add_type,emb_chain_cpo]) 1)); |
|
2469 | 2467 |
by (Asm_simp_tac 1); |
5136 | 2468 |
by (subgoal_tac "f(x #+ xa) O \ |
2469 |
\ (Rp(DD ` (x #+ xa), E, r(x #+ xa)) O r(x #+ xa)) O \ |
|
2470 |
\ eps(DD, ee, x, x #+ xa) = f(x)" 1); |
|
2471 |
by (asm_simp_tac (simpset() addsimps [embRp_eq,eps_fun RS id_comp,commute_emb, |
|
2472 |
emb_chain_cpo]) 2); |
|
2473 |
by (blast_tac (claset() addIs [commute_eq,add_type,add_le_self]) 2); |
|
2474 |
by (asm_full_simp_tac |
|
2475 |
(simpset() addsimps [comp_assoc,commute_eq,add_le_self]) 1); |
|
3425 | 2476 |
qed "suffix_lemma"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2477 |
|
5136 | 2478 |
|
2479 |
||
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2480 |
val mediatingI = prove_goalw Limit.thy [mediating_def] |
3840 | 2481 |
"[|emb(E,G,t); !!n. n:nat ==> f(n) = t O r(n) |]==>mediating(E,G,r,f,t)" |
5136 | 2482 |
(fn prems => [Safe_tac,REPEAT (ares_tac prems 1)]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2483 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2484 |
val mediating_emb = prove_goalw Limit.thy [mediating_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2485 |
"!!z. mediating(E,G,r,f,t) ==> emb(E,G,t)" |
2469 | 2486 |
(fn prems => [Fast_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2487 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2488 |
val mediating_eq = prove_goalw Limit.thy [mediating_def] |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2489 |
"!!z. [| mediating(E,G,r,f,t); n:nat |] ==> f(n) = t O r(n)" |
5136 | 2490 |
(fn prems => [Blast_tac 1]); |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2491 |
|
5136 | 2492 |
Goal (* lub_universal_mediating *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2493 |
"[| lub(cf(E,E), lam n:nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2494 |
\ commute(DD,ee,E,r); commute(DD,ee,G,f); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2495 |
\ emb_chain(DD,ee); cpo(E); cpo(G) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2496 |
\ mediating(E,G,r,f,lub(cf(E,G), lam n:nat. f(n) O Rp(DD`n,E,r(n))))"; |
5136 | 2497 |
brr[mediatingI,emb_theta] 1; |
1623 | 2498 |
by (res_inst_tac[("b","r(n)")](lub_const RS subst) 1); |
2034 | 2499 |
by (stac comp_lubs 3); |
5136 | 2500 |
by (REPEAT (blast_tac (claset() addIs [cont_cf, emb_cont, commute_emb, cpo_cf, theta_chain, chain_const, emb_chain_cpo]) 1)); |
2469 | 2501 |
by (Simp_tac 1); |
5136 | 2502 |
by (stac (lub_suffix RS sym) 1); |
2503 |
brr[chain_lemma,cpo_cf,emb_chain_cpo] 1; |
|
2504 |
by (asm_simp_tac |
|
2505 |
(simpset() addsimps [suffix_lemma, lub_const, cont_cf, emb_cont, |
|
2506 |
commute_emb, cpo_cf, emb_chain_cpo]) 1); |
|
3425 | 2507 |
qed "lub_universal_mediating"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2508 |
|
5136 | 2509 |
Goal (* lub_universal_unique *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2510 |
"[| mediating(E,G,r,f,t); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2511 |
\ lub(cf(E,E), lam n:nat. r(n) O Rp(DD`n,E,r(n))) = id(set(E)); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2512 |
\ commute(DD,ee,E,r); commute(DD,ee,G,f); \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2513 |
\ emb_chain(DD,ee); cpo(E); cpo(G) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2514 |
\ t = lub(cf(E,G), lam n:nat. f(n) O Rp(DD`n,E,r(n)))"; |
1623 | 2515 |
by (res_inst_tac[("b","t")](comp_id RS subst) 1); |
5136 | 2516 |
by (etac subst 2); |
1623 | 2517 |
by (res_inst_tac[("b","t")](lub_const RS subst) 2); |
2034 | 2518 |
by (stac comp_lubs 4); |
5136 | 2519 |
by (asm_simp_tac (simpset() addsimps [comp_assoc, |
2520 |
read_instantiate [("f","f")] mediating_eq]) 9); |
|
2521 |
brr[cont_fun, emb_cont, mediating_emb, cont_cf, cpo_cf, chain_const, commute_chain,emb_chain_cpo] 1; |
|
3425 | 2522 |
qed "lub_universal_unique"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2523 |
|
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2524 |
(*---------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2525 |
(* Dinf yields the inverse_limit, stated as rho_emb_commute and *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2526 |
(* Dinf_universal. *) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2527 |
(*---------------------------------------------------------------------*) |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2528 |
|
5136 | 2529 |
Goal (* Dinf_universal *) |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2530 |
"[| commute(DD,ee,G,f); emb_chain(DD,ee); cpo(G) |] ==> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2531 |
\ mediating \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2532 |
\ (Dinf(DD,ee),G,rho_emb(DD,ee),f, \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2533 |
\ lub(cf(Dinf(DD,ee),G), \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2534 |
\ lam n:nat. f(n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n)))) & \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2535 |
\ (ALL t. mediating(Dinf(DD,ee),G,rho_emb(DD,ee),f,t) --> \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2536 |
\ t = lub(cf(Dinf(DD,ee),G), \ |
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2537 |
\ lam n:nat. f(n) O Rp(DD`n,Dinf(DD,ee),rho_emb(DD,ee,n))))"; |
4152 | 2538 |
by Safe_tac; |
5136 | 2539 |
brr[lub_universal_mediating,rho_emb_commute,rho_emb_lub,cpo_Dinf] 1; |
2540 |
by (auto_tac (claset() addIs [lub_universal_unique,rho_emb_commute,rho_emb_lub,cpo_Dinf], simpset())); |
|
3425 | 2541 |
qed "Dinf_universal"; |
1281
68f6be60ab1c
The inverse limit construction -- thanks to Sten Agerholm
paulson
parents:
diff
changeset
|
2542 |