| author | wenzelm | 
| Sat, 26 Sep 2020 17:04:51 +0200 | |
| changeset 72312 | 0134a7d6ad56 | 
| parent 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 42151 | 1 | (* Title: HOL/HOLCF/One.thy | 
| 2640 | 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 | |
| 62175 | 5 | section \<open>The unit domain\<close> | 
| 15577 | 6 | |
| 7 | theory One | |
| 67312 | 8 | imports Lift | 
| 15577 | 9 | begin | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 10 | |
| 67312 | 11 | type_synonym one = "unit lift" | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 12 | |
| 41295 | 13 | translations | 
| 67312 | 14 | (type) "one" \<leftharpoondown> (type) "unit lift" | 
| 41295 | 15 | |
| 16 | definition ONE :: "one" | |
| 67312 | 17 | where "ONE \<equiv> Def ()" | 
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 18 | |
| 69597 | 19 | text \<open>Exhaustion and Elimination for type \<^typ>\<open>one\<close>\<close> | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
14981diff
changeset | 20 | |
| 16747 | 21 | lemma Exh_one: "t = \<bottom> \<or> t = ONE" | 
| 67312 | 22 | by (induct t) (simp_all add: ONE_def) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
14981diff
changeset | 23 | |
| 35783 | 24 | lemma oneE [case_names bottom ONE]: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = ONE \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q" | 
| 67312 | 25 | by (induct p) (simp_all add: ONE_def) | 
| 27293 | 26 | |
| 67312 | 27 | lemma one_induct [case_names bottom ONE]: "P \<bottom> \<Longrightarrow> P ONE \<Longrightarrow> P x" | 
| 28 | by (cases x rule: oneE) simp_all | |
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
14981diff
changeset | 29 | |
| 41182 | 30 | lemma dist_below_one [simp]: "ONE \<notsqsubseteq> \<bottom>" | 
| 67312 | 31 | by (simp add: ONE_def) | 
| 27293 | 32 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30911diff
changeset | 33 | lemma below_ONE [simp]: "x \<sqsubseteq> ONE" | 
| 67312 | 34 | by (induct x rule: one_induct) simp_all | 
| 27293 | 35 | |
| 31076 
99fe356cbbc2
rename constant sq_le to below; rename class sq_ord to below; less->below in many lemma names
 huffman parents: 
30911diff
changeset | 36 | lemma ONE_below_iff [simp]: "ONE \<sqsubseteq> x \<longleftrightarrow> x = ONE" | 
| 67312 | 37 | by (induct x rule: one_induct) simp_all | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
14981diff
changeset | 38 | |
| 30911 
7809cbaa1b61
domain package: simplify internal proofs of con_rews
 huffman parents: 
29141diff
changeset | 39 | lemma ONE_defined [simp]: "ONE \<noteq> \<bottom>" | 
| 67312 | 40 | by (simp add: ONE_def) | 
| 15576 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 huffman parents: 
14981diff
changeset | 41 | |
| 27293 | 42 | lemma one_neq_iffs [simp]: | 
| 43 | "x \<noteq> ONE \<longleftrightarrow> x = \<bottom>" | |
| 44 | "ONE \<noteq> x \<longleftrightarrow> x = \<bottom>" | |
| 45 | "x \<noteq> \<bottom> \<longleftrightarrow> x = ONE" | |
| 46 | "\<bottom> \<noteq> x \<longleftrightarrow> x = ONE" | |
| 67312 | 47 | by (induct x rule: one_induct) simp_all | 
| 27293 | 48 | |
| 49 | lemma compact_ONE: "compact ONE" | |
| 67312 | 50 | by (rule compact_chfin) | 
| 17838 | 51 | |
| 69597 | 52 | text \<open>Case analysis function for type \<^typ>\<open>one\<close>\<close> | 
| 18080 | 53 | |
| 67312 | 54 | definition one_case :: "'a::pcpo \<rightarrow> one \<rightarrow> 'a" | 
| 55 | where "one_case = (\<Lambda> a x. seq\<cdot>x\<cdot>a)" | |
| 18080 | 56 | |
| 57 | translations | |
| 67312 | 58 | "case x of XCONST ONE \<Rightarrow> t" \<rightleftharpoons> "CONST one_case\<cdot>t\<cdot>x" | 
| 59 | "case x of XCONST ONE :: 'a \<Rightarrow> t" \<rightharpoonup> "CONST one_case\<cdot>t\<cdot>x" | |
| 60 | "\<Lambda> (XCONST ONE). t" \<rightleftharpoons> "CONST one_case\<cdot>t" | |
| 18080 | 61 | |
| 40212 | 62 | lemma one_case1 [simp]: "(case \<bottom> of ONE \<Rightarrow> t) = \<bottom>" | 
| 67312 | 63 | by (simp add: one_case_def) | 
| 18080 | 64 | |
| 40212 | 65 | lemma one_case2 [simp]: "(case ONE of ONE \<Rightarrow> t) = t" | 
| 67312 | 66 | by (simp add: one_case_def) | 
| 18080 | 67 | |
| 40212 | 68 | lemma one_case3 [simp]: "(case x of ONE \<Rightarrow> ONE) = x" | 
| 67312 | 69 | by (induct x rule: one_induct) simp_all | 
| 18080 | 70 | |
| 243 
c22b85994e17
Franz Regensburger's Higher-Order Logic of Computable Functions embedding LCF
 nipkow parents: diff
changeset | 71 | end |