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