simplify fixrec proofs for mutually-recursive definitions; generate better fixpoint induction rules
1 (* Title: HOLCF/One.thy
5 header {* The unit domain *}
11 types one = "unit lift"
13 "one" <= (type) "unit lift"
20 text {* Exhaustion and Elimination for type @{typ one} *}
22 lemma Exh_one: "t = \<bottom> \<or> t = ONE"
23 unfolding ONE_def by (induct t) simp_all
25 lemma oneE: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = ONE \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
26 unfolding ONE_def by (induct p) simp_all
28 lemma one_induct: "\<lbrakk>P \<bottom>; P ONE\<rbrakk> \<Longrightarrow> P x"
29 by (cases x rule: oneE) simp_all
31 lemma dist_below_one [simp]: "\<not> ONE \<sqsubseteq> \<bottom>"
32 unfolding ONE_def by simp
34 lemma below_ONE [simp]: "x \<sqsubseteq> ONE"
35 by (induct x rule: one_induct) simp_all
37 lemma ONE_below_iff [simp]: "ONE \<sqsubseteq> x \<longleftrightarrow> x = ONE"
38 by (induct x rule: one_induct) simp_all
40 lemma ONE_defined [simp]: "ONE \<noteq> \<bottom>"
41 unfolding ONE_def by simp
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
50 lemma compact_ONE: "compact ONE"
51 by (rule compact_chfin)
53 text {* Case analysis function for type @{typ one} *}
56 one_when :: "'a::pcpo \<rightarrow> one \<rightarrow> 'a" where
57 "one_when = (\<Lambda> a. strictify\<cdot>(\<Lambda> _. a))"
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"
63 lemma one_when1 [simp]: "(case \<bottom> of ONE \<Rightarrow> t) = \<bottom>"
64 by (simp add: one_when_def)
66 lemma one_when2 [simp]: "(case ONE of ONE \<Rightarrow> t) = t"
67 by (simp add: one_when_def)
69 lemma one_when3 [simp]: "(case x of ONE \<Rightarrow> ONE) = x"
70 by (induct x rule: one_induct) simp_all