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