src/HOL/HOLCF/Up.thy
author wenzelm
Mon, 01 Jan 2018 23:07:24 +0100
changeset 67312 0d25e02759b7
parent 63040 eb4ddd18d635
child 67399 eab6ce8368fa
permissions -rw-r--r--
misc tuning and modernization;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
42151
4da4fc77664b tuned headers;
wenzelm
parents: 41430
diff changeset
     1
(*  Title:      HOL/HOLCF/Up.thy
40502
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents: 40327
diff changeset
     2
    Author:     Franz Regensburger
8e92772bc0e8 move map functions to new theory file Map_Functions; add theory file Plain_HOLCF
huffman
parents: 40327
diff changeset
     3
    Author:     Brian Huffman
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     4
*)
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     5
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
     6
section \<open>The type of lifted values\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
     7
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
     8
theory Up
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
     9
  imports Cfun
15577
e16da3068ad6 fix headers
huffman
parents: 15576
diff changeset
    10
begin
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    11
36452
d37c6eed8117 renamed command 'defaultsort' to 'default_sort';
wenzelm
parents: 35900
diff changeset
    12
default_sort cpo
15599
10cedbd5289e instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
huffman
parents: 15593
diff changeset
    13
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    14
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    15
subsection \<open>Definition of new type for lifting\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    16
61998
b66d2ca1f907 clarified print modes;
wenzelm
parents: 58880
diff changeset
    17
datatype 'a u  ("(_\<^sub>\<bottom>)" [1000] 999) = Ibottom | Iup 'a
18290
5fc309770840 add xsymbol syntax for u type constructor
huffman
parents: 18078
diff changeset
    18
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    19
primrec Ifup :: "('a \<rightarrow> 'b::pcpo) \<Rightarrow> 'a u \<Rightarrow> 'b"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    20
  where
34941
156925dd67af dropped some old primrecs and some constdefs
haftmann
parents: 33808
diff changeset
    21
    "Ifup f Ibottom = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    22
  | "Ifup f (Iup x) = f\<cdot>x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    23
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    24
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    25
subsection \<open>Ordering on lifted cpo\<close>
15593
24d770bbc44a reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    26
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29530
diff changeset
    27
instantiation u :: (cpo) below
25787
398dec10518e update sq_ord/po instance proofs
huffman
parents: 25785
diff changeset
    28
begin
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    29
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    30
definition below_up_def:
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    31
  "(op \<sqsubseteq>) \<equiv>
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    32
    (\<lambda>x y.
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    33
      (case x of
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    34
        Ibottom \<Rightarrow> True
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    35
      | Iup a \<Rightarrow> (case y of Ibottom \<Rightarrow> False | Iup b \<Rightarrow> a \<sqsubseteq> b)))"
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    36
25787
398dec10518e update sq_ord/po instance proofs
huffman
parents: 25785
diff changeset
    37
instance ..
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    38
25787
398dec10518e update sq_ord/po instance proofs
huffman
parents: 25785
diff changeset
    39
end
398dec10518e update sq_ord/po instance proofs
huffman
parents: 25785
diff changeset
    40
16753
fb6801c926d2 define 'a u with datatype package;
huffman
parents: 16553
diff changeset
    41
lemma minimal_up [iff]: "Ibottom \<sqsubseteq> z"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    42
  by (simp add: below_up_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    43
41182
717404c7d59a add notsqsubseteq syntax
huffman
parents: 40774
diff changeset
    44
lemma not_Iup_below [iff]: "Iup x \<notsqsubseteq> Ibottom"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    45
  by (simp add: below_up_def)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    46
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29530
diff changeset
    47
lemma Iup_below [iff]: "(Iup x \<sqsubseteq> Iup y) = (x \<sqsubseteq> y)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    48
  by (simp add: below_up_def)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    49
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    50
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    51
subsection \<open>Lifted cpo is a partial order\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    52
15599
10cedbd5289e instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
huffman
parents: 15593
diff changeset
    53
instance u :: (cpo) po
25787
398dec10518e update sq_ord/po instance proofs
huffman
parents: 25785
diff changeset
    54
proof
398dec10518e update sq_ord/po instance proofs
huffman
parents: 25785
diff changeset
    55
  fix x :: "'a u"
398dec10518e update sq_ord/po instance proofs
huffman
parents: 25785
diff changeset
    56
  show "x \<sqsubseteq> x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    57
    by (simp add: below_up_def split: u.split)
25787
398dec10518e update sq_ord/po instance proofs
huffman
parents: 25785
diff changeset
    58
next
398dec10518e update sq_ord/po instance proofs
huffman
parents: 25785
diff changeset
    59
  fix x y :: "'a u"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    60
  assume "x \<sqsubseteq> y" "y \<sqsubseteq> x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    61
  then show "x = y"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    62
    by (auto simp: below_up_def split: u.split_asm intro: below_antisym)
25787
398dec10518e update sq_ord/po instance proofs
huffman
parents: 25785
diff changeset
    63
next
398dec10518e update sq_ord/po instance proofs
huffman
parents: 25785
diff changeset
    64
  fix x y z :: "'a u"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    65
  assume "x \<sqsubseteq> y" "y \<sqsubseteq> z"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    66
  then show "x \<sqsubseteq> z"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    67
    by (auto simp: below_up_def split: u.split_asm intro: below_trans)
25787
398dec10518e update sq_ord/po instance proofs
huffman
parents: 25785
diff changeset
    68
qed
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
    69
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    70
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
    71
subsection \<open>Lifted cpo is a cpo\<close>
15593
24d770bbc44a reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
    72
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    73
lemma is_lub_Iup: "range S <<| x \<Longrightarrow> range (\<lambda>i. Iup (S i)) <<| Iup x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    74
  by (auto simp: is_lub_def is_ub_def ball_simps below_up_def split: u.split)
15599
10cedbd5289e instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
huffman
parents: 15593
diff changeset
    75
17838
3032e90c4975 added compactness theorems
huffman
parents: 17585
diff changeset
    76
lemma up_chain_lemma:
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    77
  assumes Y: "chain Y"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    78
  obtains "\<forall>i. Y i = Ibottom"
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    79
  | A k where "\<forall>i. Iup (A i) = Y (i + k)" and "chain A" and "range Y <<| Iup (\<Squnion>i. A i)"
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    80
proof (cases "\<exists>k. Y k \<noteq> Ibottom")
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    81
  case True
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    82
  then obtain k where k: "Y k \<noteq> Ibottom" ..
63040
eb4ddd18d635 eliminated old 'def';
wenzelm
parents: 62175
diff changeset
    83
  define A where "A i = (THE a. Iup a = Y (i + k))" for i
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    84
  have Iup_A: "\<forall>i. Iup (A i) = Y (i + k)"
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    85
  proof
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    86
    fix i :: nat
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    87
    from Y le_add2 have "Y k \<sqsubseteq> Y (i + k)" by (rule chain_mono)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    88
    with k have "Y (i + k) \<noteq> Ibottom" by (cases "Y k") auto
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    89
    then show "Iup (A i) = Y (i + k)"
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    90
      by (cases "Y (i + k)", simp_all add: A_def)
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    91
  qed
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    92
  from Y have chain_A: "chain A"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    93
    by (simp add: chain_def Iup_below [symmetric] Iup_A)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    94
  then have "range A <<| (\<Squnion>i. A i)"
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    95
    by (rule cpo_lubI)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    96
  then have "range (\<lambda>i. Iup (A i)) <<| Iup (\<Squnion>i. A i)"
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    97
    by (rule is_lub_Iup)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
    98
  then have "range (\<lambda>i. Y (i + k)) <<| Iup (\<Squnion>i. A i)"
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
    99
    by (simp only: Iup_A)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   100
  then have "range (\<lambda>i. Y i) <<| Iup (\<Squnion>i. A i)"
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   101
    by (simp only: is_lub_range_shift [OF Y])
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   102
  with Iup_A chain_A show ?thesis ..
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   103
next
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   104
  case False
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   105
  then have "\<forall>i. Y i = Ibottom" by simp
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   106
  then show ?thesis ..
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   107
qed
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   108
15599
10cedbd5289e instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
huffman
parents: 15593
diff changeset
   109
instance u :: (cpo) cpo
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   110
proof
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   111
  fix S :: "nat \<Rightarrow> 'a u"
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   112
  assume S: "chain S"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   113
  then show "\<exists>x. range (\<lambda>i. S i) <<| x"
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   114
  proof (rule up_chain_lemma)
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   115
    assume "\<forall>i. S i = Ibottom"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   116
    then have "range (\<lambda>i. S i) <<| Ibottom"
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 40502
diff changeset
   117
      by (simp add: is_lub_const)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   118
    then show ?thesis ..
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   119
  next
40085
c157ff4d59a6 add type annotation to avoid warning
huffman
parents: 40084
diff changeset
   120
    fix A :: "nat \<Rightarrow> 'a"
c157ff4d59a6 add type annotation to avoid warning
huffman
parents: 40084
diff changeset
   121
    assume "range S <<| Iup (\<Squnion>i. A i)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   122
    then show ?thesis ..
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   123
  qed
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   124
qed
15593
24d770bbc44a reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   125
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   126
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   127
subsection \<open>Lifted cpo is pointed\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   128
15599
10cedbd5289e instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
huffman
parents: 15593
diff changeset
   129
instance u :: (cpo) pcpo
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   130
  by intro_classes fast
15593
24d770bbc44a reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   131
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   132
text \<open>for compatibility with old HOLCF-Version\<close>
16753
fb6801c926d2 define 'a u with datatype package;
huffman
parents: 16553
diff changeset
   133
lemma inst_up_pcpo: "\<bottom> = Ibottom"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   134
  by (rule minimal_up [THEN bottomI, symmetric])
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   135
15593
24d770bbc44a reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   136
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   137
subsection \<open>Continuity of \emph{Iup} and \emph{Ifup}\<close>
15593
24d770bbc44a reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   138
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   139
text \<open>continuity for @{term Iup}\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   140
16319
1ff2965cc2e7 major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents: 16215
diff changeset
   141
lemma cont_Iup: "cont Iup"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   142
  apply (rule contI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   143
  apply (rule is_lub_Iup)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   144
  apply (erule cpo_lubI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   145
  done
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   146
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   147
text \<open>continuity for @{term Ifup}\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   148
16319
1ff2965cc2e7 major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents: 16215
diff changeset
   149
lemma cont_Ifup1: "cont (\<lambda>f. Ifup f x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   150
  by (induct x) simp_all
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   151
16319
1ff2965cc2e7 major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents: 16215
diff changeset
   152
lemma monofun_Ifup2: "monofun (\<lambda>x. Ifup f x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   153
  apply (rule monofunI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   154
  apply (case_tac x, simp)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   155
  apply (case_tac y, simp)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   156
  apply (simp add: monofun_cfun_arg)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   157
  done
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   158
16319
1ff2965cc2e7 major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents: 16215
diff changeset
   159
lemma cont_Ifup2: "cont (\<lambda>x. Ifup f x)"
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   160
proof (rule contI2)
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   161
  fix Y
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   162
  assume Y: "chain Y" and Y': "chain (\<lambda>i. Ifup f (Y i))"
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   163
  from Y show "Ifup f (\<Squnion>i. Y i) \<sqsubseteq> (\<Squnion>i. Ifup f (Y i))"
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   164
  proof (rule up_chain_lemma)
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   165
    fix A and k
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   166
    assume A: "\<forall>i. Iup (A i) = Y (i + k)"
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   167
    assume "chain A" and "range Y <<| Iup (\<Squnion>i. A i)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   168
    then have "Ifup f (\<Squnion>i. Y i) = (\<Squnion>i. Ifup f (Iup (A i)))"
40771
1c6f7d4b110e renamed several HOLCF theorems (listed in NEWS)
huffman
parents: 40502
diff changeset
   169
      by (simp add: lub_eqI contlub_cfun_arg)
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   170
    also have "\<dots> = (\<Squnion>i. Ifup f (Y (i + k)))"
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   171
      by (simp add: A)
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   172
    also have "\<dots> = (\<Squnion>i. Ifup f (Y i))"
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   173
      using Y' by (rule lub_range_shift)
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   174
    finally show ?thesis by simp
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   175
  qed simp
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   176
qed (rule monofun_Ifup2)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   177
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   178
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   179
subsection \<open>Continuous versions of constants\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   180
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   181
definition up  :: "'a \<rightarrow> 'a u"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   182
  where "up = (\<Lambda> x. Iup x)"
16319
1ff2965cc2e7 major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents: 16215
diff changeset
   183
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   184
definition fup :: "('a \<rightarrow> 'b::pcpo) \<rightarrow> 'a u \<rightarrow> 'b"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   185
  where "fup = (\<Lambda> f p. Ifup f p)"
15593
24d770bbc44a reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   186
24d770bbc44a reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   187
translations
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   188
  "case l of XCONST up\<cdot>x \<Rightarrow> t" \<rightleftharpoons> "CONST fup\<cdot>(\<Lambda> x. t)\<cdot>l"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   189
  "case l of (XCONST up :: 'a)\<cdot>x \<Rightarrow> t" \<rightharpoonup> "CONST fup\<cdot>(\<Lambda> x. t)\<cdot>l"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   190
  "\<Lambda>(XCONST up\<cdot>x). t" \<rightleftharpoons> "CONST fup\<cdot>(\<Lambda> x. t)"
15593
24d770bbc44a reordered and arranged for document generation, cleaned up some proofs
huffman
parents: 15577
diff changeset
   191
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   192
text \<open>continuous versions of lemmas for @{typ "('a)u"}\<close>
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   193
16753
fb6801c926d2 define 'a u with datatype package;
huffman
parents: 16553
diff changeset
   194
lemma Exh_Up: "z = \<bottom> \<or> (\<exists>x. z = up\<cdot>x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   195
  by (induct z) (simp add: inst_up_pcpo, simp add: up_def cont_Iup)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   196
16753
fb6801c926d2 define 'a u with datatype package;
huffman
parents: 16553
diff changeset
   197
lemma up_eq [simp]: "(up\<cdot>x = up\<cdot>y) = (x = y)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   198
  by (simp add: up_def cont_Iup)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   199
16753
fb6801c926d2 define 'a u with datatype package;
huffman
parents: 16553
diff changeset
   200
lemma up_inject: "up\<cdot>x = up\<cdot>y \<Longrightarrow> x = y"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   201
  by simp
16319
1ff2965cc2e7 major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents: 16215
diff changeset
   202
17838
3032e90c4975 added compactness theorems
huffman
parents: 17585
diff changeset
   203
lemma up_defined [simp]: "up\<cdot>x \<noteq> \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   204
  by (simp add: up_def cont_Iup inst_up_pcpo)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   205
41182
717404c7d59a add notsqsubseteq syntax
huffman
parents: 40774
diff changeset
   206
lemma not_up_less_UU: "up\<cdot>x \<notsqsubseteq> \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   207
  by simp (* FIXME: remove? *)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   208
31076
99fe356cbbc2 rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
huffman
parents: 29530
diff changeset
   209
lemma up_below [simp]: "up\<cdot>x \<sqsubseteq> up\<cdot>y \<longleftrightarrow> x \<sqsubseteq> y"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   210
  by (simp add: up_def cont_Iup)
16319
1ff2965cc2e7 major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents: 16215
diff changeset
   211
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   212
lemma upE [case_names bottom up, cases type: u]: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; \<And>x. p = up\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   213
  by (cases p) (simp add: inst_up_pcpo, simp add: up_def cont_Iup)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   214
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   215
lemma up_induct [case_names bottom up, induct type: u]: "P \<bottom> \<Longrightarrow> (\<And>x. P (up\<cdot>x)) \<Longrightarrow> P x"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   216
  by (cases x) simp_all
25788
30cbe0764599 declare upE as cases rule; add new rule up_induct
huffman
parents: 25787
diff changeset
   217
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   218
text \<open>lifting preserves chain-finiteness\<close>
25827
c2adeb1bae5c new instance proofs for classes finite_po, chfin, flat
huffman
parents: 25789
diff changeset
   219
17838
3032e90c4975 added compactness theorems
huffman
parents: 17585
diff changeset
   220
lemma up_chain_cases:
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   221
  assumes Y: "chain Y"
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   222
  obtains "\<forall>i. Y i = \<bottom>"
40084
23a1cfdb5acb simplify some proofs, convert to Isar style
huffman
parents: 40002
diff changeset
   223
  | A k where "\<forall>i. up\<cdot>(A i) = Y (i + k)" and "chain A" and "(\<Squnion>i. Y i) = up\<cdot>(\<Squnion>i. A i)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   224
  by (rule up_chain_lemma [OF Y]) (simp_all add: inst_up_pcpo up_def cont_Iup lub_eqI)
17838
3032e90c4975 added compactness theorems
huffman
parents: 17585
diff changeset
   225
25879
98b93782c3b1 new compactness lemmas
huffman
parents: 25827
diff changeset
   226
lemma compact_up: "compact x \<Longrightarrow> compact (up\<cdot>x)"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   227
  apply (rule compactI2)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   228
  apply (erule up_chain_cases)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   229
   apply simp
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   230
  apply (drule (1) compactD2, simp)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   231
  apply (erule exE)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   232
  apply (drule_tac f="up" and x="x" in monofun_cfun_arg)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   233
  apply (simp, erule exI)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   234
  done
25879
98b93782c3b1 new compactness lemmas
huffman
parents: 25827
diff changeset
   235
98b93782c3b1 new compactness lemmas
huffman
parents: 25827
diff changeset
   236
lemma compact_upD: "compact (up\<cdot>x) \<Longrightarrow> compact x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   237
  unfolding compact_def
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   238
  by (drule adm_subst [OF cont_Rep_cfun2 [where f=up]], simp)
25879
98b93782c3b1 new compactness lemmas
huffman
parents: 25827
diff changeset
   239
98b93782c3b1 new compactness lemmas
huffman
parents: 25827
diff changeset
   240
lemma compact_up_iff [simp]: "compact (up\<cdot>x) = compact x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   241
  by (safe elim!: compact_up compact_upD)
25879
98b93782c3b1 new compactness lemmas
huffman
parents: 25827
diff changeset
   242
25827
c2adeb1bae5c new instance proofs for classes finite_po, chfin, flat
huffman
parents: 25789
diff changeset
   243
instance u :: (chfin) chfin
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   244
  apply intro_classes
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   245
  apply (erule compact_imp_max_in_chain)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   246
  apply (rule_tac p="\<Squnion>i. Y i" in upE, simp_all)
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   247
  done
17838
3032e90c4975 added compactness theorems
huffman
parents: 17585
diff changeset
   248
62175
8ffc4d0e652d isabelle update_cartouches -c -t;
wenzelm
parents: 61998
diff changeset
   249
text \<open>properties of fup\<close>
17838
3032e90c4975 added compactness theorems
huffman
parents: 17585
diff changeset
   250
16319
1ff2965cc2e7 major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents: 16215
diff changeset
   251
lemma fup1 [simp]: "fup\<cdot>f\<cdot>\<bottom> = \<bottom>"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   252
  by (simp add: fup_def cont_Ifup1 cont_Ifup2 inst_up_pcpo cont2cont_LAM)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   253
16319
1ff2965cc2e7 major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
huffman
parents: 16215
diff changeset
   254
lemma fup2 [simp]: "fup\<cdot>f\<cdot>(up\<cdot>x) = f\<cdot>x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   255
  by (simp add: up_def fup_def cont_Iup cont_Ifup1 cont_Ifup2 cont2cont_LAM)
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   256
16553
aa36d41e4263 add csplit3, ssplit3, fup3 as simp rules
huffman
parents: 16326
diff changeset
   257
lemma fup3 [simp]: "fup\<cdot>up\<cdot>x = x"
67312
0d25e02759b7 misc tuning and modernization;
wenzelm
parents: 63040
diff changeset
   258
  by (cases x) simp_all
15576
efb95d0d01f7 converted to new-style theories, and combined numbered files
huffman
parents:
diff changeset
   259
26962
c8b20f615d6c use new class package for classes profinite, bifinite; remove approx class
huffman
parents: 26407
diff changeset
   260
end