| author | wenzelm | 
| Tue, 27 Oct 2009 10:54:25 +0100 | |
| changeset 33219 | a69147d95957 | 
| parent 31076 | 99fe356cbbc2 | 
| child 33504 | b4210cc3ac97 | 
| permissions | -rw-r--r-- | 
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 1 | (* Title: HOLCF/Up.thy | 
| 16070 
4a83dd540b88
removed LICENCE note -- everything is subject to Isabelle licence as
 wenzelm parents: 
15599diff
changeset | 2 | Author: Franz Regensburger and Brian Huffman | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 3 | *) | 
| 
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 | header {* The type of lifted values *}
 | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 6 | |
| 15577 | 7 | theory Up | 
| 25911 | 8 | imports Bifinite | 
| 15577 | 9 | begin | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 10 | |
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 11 | defaultsort cpo | 
| 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 12 | |
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 13 | subsection {* Definition of new type for lifting *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 14 | |
| 16753 | 15 | datatype 'a u = Ibottom | Iup 'a | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 16 | |
| 18290 | 17 | syntax (xsymbols) | 
| 18 |   "u" :: "type \<Rightarrow> type" ("(_\<^sub>\<bottom>)" [1000] 999)
 | |
| 19 | ||
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 20 | consts | 
| 16753 | 21 |   Ifup :: "('a \<rightarrow> 'b::pcpo) \<Rightarrow> 'a u \<Rightarrow> 'b"
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 22 | |
| 16753 | 23 | primrec | 
| 24 | "Ifup f Ibottom = \<bottom>" | |
| 25 | "Ifup f (Iup x) = f\<cdot>x" | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 26 | |
| 18290 | 27 | subsection {* Ordering on lifted cpo *}
 | 
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 28 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 29 | instantiation u :: (cpo) below | 
| 25787 | 30 | begin | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 31 | |
| 25787 | 32 | definition | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 33 | below_up_def: | 
| 16753 | 34 | "(op \<sqsubseteq>) \<equiv> (\<lambda>x y. case x of Ibottom \<Rightarrow> True | Iup a \<Rightarrow> | 
| 35 | (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 | 37 | instance .. | 
| 38 | end | |
| 39 | ||
| 16753 | 40 | lemma minimal_up [iff]: "Ibottom \<sqsubseteq> z" | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 41 | by (simp add: below_up_def) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 42 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 43 | lemma not_Iup_below [iff]: "\<not> Iup x \<sqsubseteq> Ibottom" | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 44 | by (simp add: below_up_def) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 45 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 46 | lemma Iup_below [iff]: "(Iup x \<sqsubseteq> Iup y) = (x \<sqsubseteq> y)" | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 47 | by (simp add: below_up_def) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 48 | |
| 18290 | 49 | subsection {* Lifted cpo is a partial order *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 50 | |
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 51 | instance u :: (cpo) po | 
| 25787 | 52 | proof | 
| 53 | fix x :: "'a u" | |
| 54 | show "x \<sqsubseteq> x" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 55 | unfolding below_up_def by (simp split: u.split) | 
| 25787 | 56 | next | 
| 57 | fix x y :: "'a u" | |
| 58 | assume "x \<sqsubseteq> y" "y \<sqsubseteq> x" thus "x = y" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 59 | unfolding below_up_def | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 60 | by (auto split: u.split_asm intro: below_antisym) | 
| 25787 | 61 | next | 
| 62 | fix x y z :: "'a u" | |
| 63 | assume "x \<sqsubseteq> y" "y \<sqsubseteq> z" thus "x \<sqsubseteq> z" | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 64 | unfolding below_up_def | 
| 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 65 | by (auto split: u.split_asm intro: below_trans) | 
| 25787 | 66 | qed | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 67 | |
| 25827 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25789diff
changeset | 68 | lemma u_UNIV: "UNIV = insert Ibottom (range Iup)" | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25789diff
changeset | 69 | by (auto, case_tac x, auto) | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25789diff
changeset | 70 | |
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25789diff
changeset | 71 | instance u :: (finite_po) finite_po | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25789diff
changeset | 72 | by (intro_classes, simp add: u_UNIV) | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25789diff
changeset | 73 | |
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25789diff
changeset | 74 | |
| 18290 | 75 | subsection {* Lifted cpo is a cpo *}
 | 
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 76 | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 77 | lemma is_lub_Iup: | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 78 | "range S <<| x \<Longrightarrow> range (\<lambda>i. Iup (S i)) <<| Iup x" | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 79 | apply (rule is_lubI) | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 80 | apply (rule ub_rangeI) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 81 | apply (subst Iup_below) | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 82 | apply (erule is_ub_lub) | 
| 16753 | 83 | apply (case_tac u) | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 84 | apply (drule ub_rangeD) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 85 | apply simp | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 86 | apply simp | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 87 | apply (erule is_lub_lub) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 88 | apply (rule ub_rangeI) | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 89 | apply (drule_tac i=i in ub_rangeD) | 
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 90 | apply simp | 
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 91 | done | 
| 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 92 | |
| 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 93 | text {* Now some lemmas about chains of @{typ "'a u"} elements *}
 | 
| 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 94 | |
| 16753 | 95 | lemma up_lemma1: "z \<noteq> Ibottom \<Longrightarrow> Iup (THE a. Iup a = z) = z" | 
| 96 | by (case_tac z, simp_all) | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 97 | |
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 98 | lemma up_lemma2: | 
| 16753 | 99 | "\<lbrakk>chain Y; Y j \<noteq> Ibottom\<rbrakk> \<Longrightarrow> Y (i + j) \<noteq> Ibottom" | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 100 | apply (erule contrapos_nn) | 
| 25922 
cb04d05e95fb
rename lemma chain_mono3 -> chain_mono, chain_mono -> chain_mono_less
 huffman parents: 
25921diff
changeset | 101 | apply (drule_tac i="j" and j="i + j" in chain_mono) | 
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 102 | apply (rule le_add2) | 
| 16753 | 103 | apply (case_tac "Y j") | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 104 | apply assumption | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 105 | apply simp | 
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 106 | done | 
| 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 107 | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 108 | lemma up_lemma3: | 
| 16753 | 109 | "\<lbrakk>chain Y; Y j \<noteq> Ibottom\<rbrakk> \<Longrightarrow> Iup (THE a. Iup a = Y (i + j)) = Y (i + j)" | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 110 | by (rule up_lemma1 [OF up_lemma2]) | 
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 111 | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 112 | lemma up_lemma4: | 
| 16753 | 113 | "\<lbrakk>chain Y; Y j \<noteq> Ibottom\<rbrakk> \<Longrightarrow> chain (\<lambda>i. THE a. Iup a = Y (i + j))" | 
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 114 | apply (rule chainI) | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 115 | apply (rule Iup_below [THEN iffD1]) | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 116 | apply (subst up_lemma3, assumption+)+ | 
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 117 | apply (simp add: chainE) | 
| 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 118 | done | 
| 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 119 | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 120 | lemma up_lemma5: | 
| 16753 | 121 | "\<lbrakk>chain Y; Y j \<noteq> Ibottom\<rbrakk> \<Longrightarrow> | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 122 | (\<lambda>i. Y (i + j)) = (\<lambda>i. Iup (THE a. Iup a = Y (i + j)))" | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 123 | by (rule ext, rule up_lemma3 [symmetric]) | 
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 124 | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 125 | lemma up_lemma6: | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19105diff
changeset | 126 | "\<lbrakk>chain Y; Y j \<noteq> Ibottom\<rbrakk> | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 127 | \<Longrightarrow> range Y <<| Iup (\<Squnion>i. THE a. Iup a = Y(i + j))" | 
| 16933 | 128 | apply (rule_tac j1 = j in is_lub_range_shift [THEN iffD1]) | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 129 | apply assumption | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 130 | apply (subst up_lemma5, assumption+) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 131 | apply (rule is_lub_Iup) | 
| 26027 | 132 | apply (rule cpo_lubI) | 
| 16753 | 133 | apply (erule (1) up_lemma4) | 
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 134 | done | 
| 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 135 | |
| 17838 | 136 | lemma up_chain_lemma: | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 137 | "chain Y \<Longrightarrow> | 
| 27413 | 138 | (\<exists>A. chain A \<and> (\<Squnion>i. Y i) = Iup (\<Squnion>i. A i) \<and> | 
| 16753 | 139 | (\<exists>j. \<forall>i. Y (i + j) = Iup (A i))) \<or> (Y = (\<lambda>i. Ibottom))" | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 140 | apply (rule disjCI) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 141 | apply (simp add: expand_fun_eq) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 142 | apply (erule exE, rename_tac j) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 143 | apply (rule_tac x="\<lambda>i. THE a. Iup a = Y (i + j)" in exI) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 144 | apply (simp add: up_lemma4) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 145 | apply (simp add: up_lemma6 [THEN thelubI]) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 146 | apply (rule_tac x=j in exI) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 147 | apply (simp add: up_lemma3) | 
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 148 | done | 
| 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 149 | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 150 | lemma cpo_up: "chain (Y::nat \<Rightarrow> 'a u) \<Longrightarrow> \<exists>x. range Y <<| x" | 
| 17838 | 151 | apply (frule up_chain_lemma, safe) | 
| 27413 | 152 | apply (rule_tac x="Iup (\<Squnion>i. A i)" in exI) | 
| 17838 | 153 | apply (erule_tac j="j" in is_lub_range_shift [THEN iffD1, standard]) | 
| 26027 | 154 | apply (simp add: is_lub_Iup cpo_lubI) | 
| 17585 | 155 | apply (rule exI, rule lub_const) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 156 | done | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 157 | |
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 158 | instance u :: (cpo) cpo | 
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 159 | by intro_classes (rule cpo_up) | 
| 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 160 | |
| 18290 | 161 | subsection {* Lifted cpo is pointed *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 162 | |
| 17585 | 163 | lemma least_up: "\<exists>x::'a u. \<forall>y. x \<sqsubseteq> y" | 
| 16753 | 164 | apply (rule_tac x = "Ibottom" in exI) | 
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 165 | apply (rule minimal_up [THEN allI]) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 166 | done | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 167 | |
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 168 | instance u :: (cpo) pcpo | 
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 169 | by intro_classes (rule least_up) | 
| 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 170 | |
| 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 171 | text {* for compatibility with old HOLCF-Version *}
 | 
| 16753 | 172 | lemma inst_up_pcpo: "\<bottom> = Ibottom" | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 173 | by (rule minimal_up [THEN UU_I, symmetric]) | 
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 174 | |
| 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 175 | subsection {* Continuity of @{term Iup} and @{term Ifup} *}
 | 
| 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 176 | |
| 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 177 | text {* continuity for @{term Iup} *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 178 | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 179 | lemma cont_Iup: "cont Iup" | 
| 16215 | 180 | apply (rule contI) | 
| 15599 
10cedbd5289e
instance u :: (cpo) pcpo -- argument type no longer needs to be pointed
 huffman parents: 
15593diff
changeset | 181 | apply (rule is_lub_Iup) | 
| 26027 | 182 | apply (erule cpo_lubI) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 183 | done | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 184 | |
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 185 | text {* continuity for @{term Ifup} *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 186 | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 187 | lemma cont_Ifup1: "cont (\<lambda>f. Ifup f x)" | 
| 16753 | 188 | by (induct x, simp_all) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 189 | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 190 | lemma monofun_Ifup2: "monofun (\<lambda>x. Ifup f x)" | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 191 | apply (rule monofunI) | 
| 16753 | 192 | apply (case_tac x, simp) | 
| 193 | apply (case_tac y, simp) | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 194 | apply (simp add: monofun_cfun_arg) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 195 | done | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 196 | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 197 | lemma cont_Ifup2: "cont (\<lambda>x. Ifup f x)" | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 198 | apply (rule contI) | 
| 17838 | 199 | apply (frule up_chain_lemma, safe) | 
| 200 | apply (rule_tac j="j" in is_lub_range_shift [THEN iffD1, standard]) | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 201 | apply (erule monofun_Ifup2 [THEN ch2ch_monofun]) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 202 | apply (simp add: cont_cfun_arg) | 
| 18078 
20e5a6440790
change syntax for LAM to use expressions as patterns; define LAM pattern syntax for cpair, spair, sinl, sinr, up
 huffman parents: 
17838diff
changeset | 203 | apply (simp add: lub_const) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 204 | done | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 205 | |
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 206 | subsection {* Continuous versions of constants *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 207 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19105diff
changeset | 208 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19105diff
changeset | 209 | up :: "'a \<rightarrow> 'a u" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19105diff
changeset | 210 | "up = (\<Lambda> x. Iup x)" | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 211 | |
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19105diff
changeset | 212 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19105diff
changeset | 213 |   fup :: "('a \<rightarrow> 'b::pcpo) \<rightarrow> 'a u \<rightarrow> 'b" where
 | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
19105diff
changeset | 214 | "fup = (\<Lambda> f p. Ifup f p)" | 
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 215 | |
| 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 216 | translations | 
| 26046 | 217 | "case l of XCONST up\<cdot>x \<Rightarrow> t" == "CONST fup\<cdot>(\<Lambda> x. t)\<cdot>l" | 
| 218 | "\<Lambda>(XCONST up\<cdot>x). t" == "CONST fup\<cdot>(\<Lambda> x. t)" | |
| 15593 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 219 | |
| 
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
 huffman parents: 
15577diff
changeset | 220 | text {* continuous versions of lemmas for @{typ "('a)u"} *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 221 | |
| 16753 | 222 | lemma Exh_Up: "z = \<bottom> \<or> (\<exists>x. z = up\<cdot>x)" | 
| 223 | apply (induct z) | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 224 | apply (simp add: inst_up_pcpo) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 225 | apply (simp add: up_def cont_Iup) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 226 | done | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 227 | |
| 16753 | 228 | lemma up_eq [simp]: "(up\<cdot>x = up\<cdot>y) = (x = y)" | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 229 | by (simp add: up_def cont_Iup) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 230 | |
| 16753 | 231 | lemma up_inject: "up\<cdot>x = up\<cdot>y \<Longrightarrow> x = y" | 
| 232 | by simp | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 233 | |
| 17838 | 234 | lemma up_defined [simp]: "up\<cdot>x \<noteq> \<bottom>" | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 235 | 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 | 236 | |
| 25785 | 237 | lemma not_up_less_UU: "\<not> up\<cdot>x \<sqsubseteq> \<bottom>" | 
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 238 | by simp (* FIXME: remove? *) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 239 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
29530diff
changeset | 240 | lemma up_below [simp]: "up\<cdot>x \<sqsubseteq> up\<cdot>y \<longleftrightarrow> x \<sqsubseteq> y" | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 241 | by (simp add: up_def cont_Iup) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 242 | |
| 25788 
30cbe0764599
declare upE as cases rule; add new rule up_induct
 huffman parents: 
25787diff
changeset | 243 | lemma upE [cases type: u]: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; \<And>x. p = up\<cdot>x \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 
30cbe0764599
declare upE as cases rule; add new rule up_induct
 huffman parents: 
25787diff
changeset | 244 | apply (cases p) | 
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 245 | apply (simp add: inst_up_pcpo) | 
| 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 246 | apply (simp add: up_def cont_Iup) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 247 | done | 
| 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 248 | |
| 25788 
30cbe0764599
declare upE as cases rule; add new rule up_induct
 huffman parents: 
25787diff
changeset | 249 | lemma up_induct [induct type: u]: "\<lbrakk>P \<bottom>; \<And>x. P (up\<cdot>x)\<rbrakk> \<Longrightarrow> P x" | 
| 
30cbe0764599
declare upE as cases rule; add new rule up_induct
 huffman parents: 
25787diff
changeset | 250 | by (cases x, simp_all) | 
| 
30cbe0764599
declare upE as cases rule; add new rule up_induct
 huffman parents: 
25787diff
changeset | 251 | |
| 25827 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25789diff
changeset | 252 | text {* lifting preserves chain-finiteness *}
 | 
| 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25789diff
changeset | 253 | |
| 17838 | 254 | lemma up_chain_cases: | 
| 255 | "chain Y \<Longrightarrow> | |
| 256 | (\<exists>A. chain A \<and> (\<Squnion>i. Y i) = up\<cdot>(\<Squnion>i. A i) \<and> | |
| 257 | (\<exists>j. \<forall>i. Y (i + j) = up\<cdot>(A i))) \<or> Y = (\<lambda>i. \<bottom>)" | |
| 258 | by (simp add: inst_up_pcpo up_def cont_Iup up_chain_lemma) | |
| 259 | ||
| 25879 | 260 | lemma compact_up: "compact x \<Longrightarrow> compact (up\<cdot>x)" | 
| 261 | apply (rule compactI2) | |
| 262 | apply (drule up_chain_cases, safe) | |
| 263 | apply (drule (1) compactD2, simp) | |
| 264 | apply (erule exE, rule_tac x="i + j" in exI) | |
| 265 | apply simp | |
| 266 | apply simp | |
| 267 | done | |
| 268 | ||
| 269 | lemma compact_upD: "compact (up\<cdot>x) \<Longrightarrow> compact x" | |
| 270 | unfolding compact_def | |
| 271 | by (drule adm_subst [OF cont_Rep_CFun2 [where f=up]], simp) | |
| 272 | ||
| 273 | lemma compact_up_iff [simp]: "compact (up\<cdot>x) = compact x" | |
| 274 | by (safe elim!: compact_up compact_upD) | |
| 275 | ||
| 25827 
c2adeb1bae5c
new instance proofs for classes finite_po, chfin, flat
 huffman parents: 
25789diff
changeset | 276 | instance u :: (chfin) chfin | 
| 25921 | 277 | apply intro_classes | 
| 25879 | 278 | apply (erule compact_imp_max_in_chain) | 
| 25898 | 279 | apply (rule_tac p="\<Squnion>i. Y i" in upE, simp_all) | 
| 17838 | 280 | done | 
| 281 | ||
| 282 | text {* properties of fup *}
 | |
| 283 | ||
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 284 | lemma fup1 [simp]: "fup\<cdot>f\<cdot>\<bottom> = \<bottom>" | 
| 29530 
9905b660612b
change to simpler, more extensible continuity simproc
 huffman parents: 
29138diff
changeset | 285 | 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 | 286 | |
| 16319 
1ff2965cc2e7
major cleanup: rewrote cpo proofs, removed obsolete lemmas, renamed some lemmas
 huffman parents: 
16215diff
changeset | 287 | lemma fup2 [simp]: "fup\<cdot>f\<cdot>(up\<cdot>x) = f\<cdot>x" | 
| 29530 
9905b660612b
change to simpler, more extensible continuity simproc
 huffman parents: 
29138diff
changeset | 288 | 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 | 289 | |
| 16553 | 290 | lemma fup3 [simp]: "fup\<cdot>up\<cdot>x = x" | 
| 25788 
30cbe0764599
declare upE as cases rule; add new rule up_induct
 huffman parents: 
25787diff
changeset | 291 | by (cases x, simp_all) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 292 | |
| 25911 | 293 | subsection {* Lifted cpo is a bifinite domain *}
 | 
| 294 | ||
| 26962 
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
 huffman parents: 
26407diff
changeset | 295 | instantiation u :: (profinite) bifinite | 
| 
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
 huffman parents: 
26407diff
changeset | 296 | begin | 
| 25911 | 297 | |
| 26962 
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
 huffman parents: 
26407diff
changeset | 298 | definition | 
| 25911 | 299 | approx_up_def: | 
| 26962 
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
 huffman parents: 
26407diff
changeset | 300 | "approx = (\<lambda>n. fup\<cdot>(\<Lambda> x. up\<cdot>(approx n\<cdot>x)))" | 
| 25911 | 301 | |
| 26962 
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
 huffman parents: 
26407diff
changeset | 302 | instance proof | 
| 25911 | 303 | fix i :: nat and x :: "'a u" | 
| 27310 | 304 | show "chain (approx :: nat \<Rightarrow> 'a u \<rightarrow> 'a u)" | 
| 25911 | 305 | unfolding approx_up_def by simp | 
| 306 | show "(\<Squnion>i. approx i\<cdot>x) = x" | |
| 307 | unfolding approx_up_def | |
| 308 | by (simp add: lub_distribs eta_cfun) | |
| 309 | show "approx i\<cdot>(approx i\<cdot>x) = approx i\<cdot>x" | |
| 310 | unfolding approx_up_def | |
| 311 | by (induct x, simp, simp) | |
| 312 |   have "{x::'a u. approx i\<cdot>x = x} \<subseteq>
 | |
| 313 |         insert \<bottom> ((\<lambda>x. up\<cdot>x) ` {x::'a. approx i\<cdot>x = x})"
 | |
| 314 | unfolding approx_up_def | |
| 27310 | 315 | by (rule subsetI, case_tac x, simp_all) | 
| 25911 | 316 |   thus "finite {x::'a u. approx i\<cdot>x = x}"
 | 
| 317 | by (rule finite_subset, simp add: finite_fixes_approx) | |
| 318 | qed | |
| 319 | ||
| 26962 
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
 huffman parents: 
26407diff
changeset | 320 | end | 
| 
c8b20f615d6c
use new class package for classes profinite, bifinite; remove approx class
 huffman parents: 
26407diff
changeset | 321 | |
| 25911 | 322 | lemma approx_up [simp]: "approx i\<cdot>(up\<cdot>x) = up\<cdot>(approx i\<cdot>x)" | 
| 323 | unfolding approx_up_def by simp | |
| 324 | ||
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: diff
changeset | 325 | end |