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