1(*  Title:      HOL/HOLCF/One.thy
2    Author:     Oscar Slotosch
3*)
4
5section \<open>The unit domain\<close>
6
7theory One
8  imports Lift
9begin
10
11type_synonym one = "unit lift"
12
13translations
14  (type) "one" \<leftharpoondown> (type) "unit lift"
15
16definition ONE :: "one"
17  where "ONE \<equiv> Def ()"
18
19text \<open>Exhaustion and Elimination for type @{typ one}\<close>
20
21lemma Exh_one: "t = \<bottom> \<or> t = ONE"
22  by (induct t) (simp_all add: ONE_def)
23
24lemma oneE [case_names bottom ONE]: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = ONE \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
25  by (induct p) (simp_all add: ONE_def)
26
27lemma one_induct [case_names bottom ONE]: "P \<bottom> \<Longrightarrow> P ONE \<Longrightarrow> P x"
28  by (cases x rule: oneE) simp_all
29
30lemma dist_below_one [simp]: "ONE \<notsqsubseteq> \<bottom>"
31  by (simp add: ONE_def)
32
33lemma below_ONE [simp]: "x \<sqsubseteq> ONE"
34  by (induct x rule: one_induct) simp_all
35
36lemma ONE_below_iff [simp]: "ONE \<sqsubseteq> x \<longleftrightarrow> x = ONE"
37  by (induct x rule: one_induct) simp_all
38
39lemma ONE_defined [simp]: "ONE \<noteq> \<bottom>"
40  by (simp add: ONE_def)
41
42lemma 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"
47  by (induct x rule: one_induct) simp_all
48
49lemma compact_ONE: "compact ONE"
50  by (rule compact_chfin)
51
52text \<open>Case analysis function for type @{typ one}\<close>
53
54definition one_case :: "'a::pcpo \<rightarrow> one \<rightarrow> 'a"
55  where "one_case = (\<Lambda> a x. seq\<cdot>x\<cdot>a)"
56
57translations
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"
61
62lemma one_case1 [simp]: "(case \<bottom> of ONE \<Rightarrow> t) = \<bottom>"
63  by (simp add: one_case_def)
64
65lemma one_case2 [simp]: "(case ONE of ONE \<Rightarrow> t) = t"
66  by (simp add: one_case_def)
67
68lemma one_case3 [simp]: "(case x of ONE \<Rightarrow> ONE) = x"
69  by (induct x rule: one_induct) simp_all
70
71end
72