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