| author | wenzelm | 
| Wed, 01 Apr 2009 12:19:15 +0200 | |
| changeset 30835 | 46e16145d4bd | 
| parent 29141 | d5582ab1311f | 
| child 30911 | 7809cbaa1b61 | 
| permissions | -rw-r--r-- | 
| 2640 | 1 | (* Title: HOLCF/One.thy | 
| 2 | Author: Oscar Slotosch | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 3 | *) | 
| 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 4 | |
| 15577 | 5 | header {* The unit domain *}
 | 
| 6 | ||
| 7 | theory One | |
| 8 | imports Lift | |
| 9 | begin | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 10 | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
14981diff
changeset | 11 | types one = "unit lift" | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18111diff
changeset | 12 | translations | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18111diff
changeset | 13 | "one" <= (type) "unit lift" | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 14 | |
| 29141 | 15 | definition | 
| 3717 | 16 | ONE :: "one" | 
| 29141 | 17 | where | 
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18111diff
changeset | 18 | "ONE == Def ()" | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 19 | |
| 16747 | 20 | text {* Exhaustion and Elimination for type @{typ one} *}
 | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
14981diff
changeset | 21 | |
| 16747 | 22 | lemma Exh_one: "t = \<bottom> \<or> t = ONE" | 
| 27293 | 23 | unfolding ONE_def by (induct t) simp_all | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
14981diff
changeset | 24 | |
| 16747 | 25 | lemma oneE: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = ONE \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 27293 | 26 | unfolding ONE_def by (induct p) simp_all | 
| 27 | ||
| 28 | lemma one_induct: "\<lbrakk>P \<bottom>; P ONE\<rbrakk> \<Longrightarrow> P x" | |
| 29 | by (cases x rule: oneE) simp_all | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
14981diff
changeset | 30 | |
| 16747 | 31 | lemma dist_less_one [simp]: "\<not> ONE \<sqsubseteq> \<bottom>" | 
| 27293 | 32 | unfolding ONE_def by simp | 
| 33 | ||
| 34 | lemma less_ONE [simp]: "x \<sqsubseteq> ONE" | |
| 35 | by (induct x rule: one_induct) simp_all | |
| 36 | ||
| 37 | lemma ONE_less_iff [simp]: "ONE \<sqsubseteq> x \<longleftrightarrow> x = ONE" | |
| 38 | by (induct x rule: one_induct) simp_all | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
14981diff
changeset | 39 | |
| 16747 | 40 | lemma dist_eq_one [simp]: "ONE \<noteq> \<bottom>" "\<bottom> \<noteq> ONE" | 
| 27293 | 41 | unfolding ONE_def by simp_all | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
14981diff
changeset | 42 | |
| 27293 | 43 | lemma one_neq_iffs [simp]: | 
| 44 | "x \<noteq> ONE \<longleftrightarrow> x = \<bottom>" | |
| 45 | "ONE \<noteq> x \<longleftrightarrow> x = \<bottom>" | |
| 46 | "x \<noteq> \<bottom> \<longleftrightarrow> x = ONE" | |
| 47 | "\<bottom> \<noteq> x \<longleftrightarrow> x = ONE" | |
| 48 | by (induct x rule: one_induct) simp_all | |
| 49 | ||
| 50 | lemma compact_ONE: "compact ONE" | |
| 17838 | 51 | by (rule compact_chfin) | 
| 52 | ||
| 18080 | 53 | text {* Case analysis function for type @{typ one} *}
 | 
| 54 | ||
| 25131 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18111diff
changeset | 55 | definition | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18111diff
changeset | 56 | one_when :: "'a::pcpo \<rightarrow> one \<rightarrow> 'a" where | 
| 
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
 wenzelm parents: 
18111diff
changeset | 57 | "one_when = (\<Lambda> a. strictify\<cdot>(\<Lambda> _. a))" | 
| 18080 | 58 | |
| 59 | translations | |
| 27293 | 60 | "case x of XCONST ONE \<Rightarrow> t" == "CONST one_when\<cdot>t\<cdot>x" | 
| 61 | "\<Lambda> (XCONST ONE). t" == "CONST one_when\<cdot>t" | |
| 18080 | 62 | |
| 18111 | 63 | lemma one_when1 [simp]: "(case \<bottom> of ONE \<Rightarrow> t) = \<bottom>" | 
| 18080 | 64 | by (simp add: one_when_def) | 
| 65 | ||
| 18111 | 66 | lemma one_when2 [simp]: "(case ONE of ONE \<Rightarrow> t) = t" | 
| 18080 | 67 | by (simp add: one_when_def) | 
| 68 | ||
| 18111 | 69 | lemma one_when3 [simp]: "(case x of ONE \<Rightarrow> ONE) = x" | 
| 27293 | 70 | by (induct x rule: one_induct) simp_all | 
| 18080 | 71 | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 72 | end |