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