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