src/HOL/HOLCF/One.thy
changeset 67312 0d25e02759b7
parent 62175 8ffc4d0e652d
child 69597 ff784d5a5bfb
equal deleted inserted replaced
67311:3869b2400e22 67312:0d25e02759b7
     3 *)
     3 *)
     4 
     4 
     5 section \<open>The unit domain\<close>
     5 section \<open>The unit domain\<close>
     6 
     6 
     7 theory One
     7 theory One
     8 imports Lift
     8   imports Lift
     9 begin
     9 begin
    10 
    10 
    11 type_synonym
    11 type_synonym one = "unit lift"
    12   one = "unit lift"
       
    13 
    12 
    14 translations
    13 translations
    15   (type) "one" <= (type) "unit lift"
    14   (type) "one" \<leftharpoondown> (type) "unit lift"
    16 
    15 
    17 definition ONE :: "one"
    16 definition ONE :: "one"
    18   where "ONE == Def ()"
    17   where "ONE \<equiv> Def ()"
    19 
    18 
    20 text \<open>Exhaustion and Elimination for type @{typ one}\<close>
    19 text \<open>Exhaustion and Elimination for type @{typ one}\<close>
    21 
    20 
    22 lemma Exh_one: "t = \<bottom> \<or> t = ONE"
    21 lemma Exh_one: "t = \<bottom> \<or> t = ONE"
    23 unfolding ONE_def by (induct t) simp_all
    22   by (induct t) (simp_all add: ONE_def)
    24 
    23 
    25 lemma oneE [case_names bottom ONE]: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = ONE \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
    24 lemma oneE [case_names bottom ONE]: "\<lbrakk>p = \<bottom> \<Longrightarrow> Q; p = ONE \<Longrightarrow> Q\<rbrakk> \<Longrightarrow> Q"
    26 unfolding ONE_def by (induct p) simp_all
    25   by (induct p) (simp_all add: ONE_def)
    27 
    26 
    28 lemma one_induct [case_names bottom ONE]: "\<lbrakk>P \<bottom>; P ONE\<rbrakk> \<Longrightarrow> P x"
    27 lemma one_induct [case_names bottom ONE]: "P \<bottom> \<Longrightarrow> P ONE \<Longrightarrow> P x"
    29 by (cases x rule: oneE) simp_all
    28   by (cases x rule: oneE) simp_all
    30 
    29 
    31 lemma dist_below_one [simp]: "ONE \<notsqsubseteq> \<bottom>"
    30 lemma dist_below_one [simp]: "ONE \<notsqsubseteq> \<bottom>"
    32 unfolding ONE_def by simp
    31   by (simp add: ONE_def)
    33 
    32 
    34 lemma below_ONE [simp]: "x \<sqsubseteq> ONE"
    33 lemma below_ONE [simp]: "x \<sqsubseteq> ONE"
    35 by (induct x rule: one_induct) simp_all
    34   by (induct x rule: one_induct) simp_all
    36 
    35 
    37 lemma ONE_below_iff [simp]: "ONE \<sqsubseteq> x \<longleftrightarrow> x = ONE"
    36 lemma ONE_below_iff [simp]: "ONE \<sqsubseteq> x \<longleftrightarrow> x = ONE"
    38 by (induct x rule: one_induct) simp_all
    37   by (induct x rule: one_induct) simp_all
    39 
    38 
    40 lemma ONE_defined [simp]: "ONE \<noteq> \<bottom>"
    39 lemma ONE_defined [simp]: "ONE \<noteq> \<bottom>"
    41 unfolding ONE_def by simp
    40   by (simp add: ONE_def)
    42 
    41 
    43 lemma one_neq_iffs [simp]:
    42 lemma one_neq_iffs [simp]:
    44   "x \<noteq> ONE \<longleftrightarrow> x = \<bottom>"
    43   "x \<noteq> ONE \<longleftrightarrow> x = \<bottom>"
    45   "ONE \<noteq> x \<longleftrightarrow> x = \<bottom>"
    44   "ONE \<noteq> x \<longleftrightarrow> x = \<bottom>"
    46   "x \<noteq> \<bottom> \<longleftrightarrow> x = ONE"
    45   "x \<noteq> \<bottom> \<longleftrightarrow> x = ONE"
    47   "\<bottom> \<noteq> x \<longleftrightarrow> x = ONE"
    46   "\<bottom> \<noteq> x \<longleftrightarrow> x = ONE"
    48 by (induct x rule: one_induct) simp_all
    47   by (induct x rule: one_induct) simp_all
    49 
    48 
    50 lemma compact_ONE: "compact ONE"
    49 lemma compact_ONE: "compact ONE"
    51 by (rule compact_chfin)
    50   by (rule compact_chfin)
    52 
    51 
    53 text \<open>Case analysis function for type @{typ one}\<close>
    52 text \<open>Case analysis function for type @{typ one}\<close>
    54 
    53 
    55 definition
    54 definition one_case :: "'a::pcpo \<rightarrow> one \<rightarrow> 'a"
    56   one_case :: "'a::pcpo \<rightarrow> one \<rightarrow> 'a" where
    55   where "one_case = (\<Lambda> a x. seq\<cdot>x\<cdot>a)"
    57   "one_case = (\<Lambda> a x. seq\<cdot>x\<cdot>a)"
       
    58 
    56 
    59 translations
    57 translations
    60   "case x of XCONST ONE \<Rightarrow> t" == "CONST one_case\<cdot>t\<cdot>x"
    58   "case x of XCONST ONE \<Rightarrow> t" \<rightleftharpoons> "CONST one_case\<cdot>t\<cdot>x"
    61   "case x of XCONST ONE :: 'a \<Rightarrow> t" => "CONST one_case\<cdot>t\<cdot>x"
    59   "case x of XCONST ONE :: 'a \<Rightarrow> t" \<rightharpoonup> "CONST one_case\<cdot>t\<cdot>x"
    62   "\<Lambda> (XCONST ONE). t" == "CONST one_case\<cdot>t"
    60   "\<Lambda> (XCONST ONE). t" \<rightleftharpoons> "CONST one_case\<cdot>t"
    63 
    61 
    64 lemma one_case1 [simp]: "(case \<bottom> of ONE \<Rightarrow> t) = \<bottom>"
    62 lemma one_case1 [simp]: "(case \<bottom> of ONE \<Rightarrow> t) = \<bottom>"
    65 by (simp add: one_case_def)
    63   by (simp add: one_case_def)
    66 
    64 
    67 lemma one_case2 [simp]: "(case ONE of ONE \<Rightarrow> t) = t"
    65 lemma one_case2 [simp]: "(case ONE of ONE \<Rightarrow> t) = t"
    68 by (simp add: one_case_def)
    66   by (simp add: one_case_def)
    69 
    67 
    70 lemma one_case3 [simp]: "(case x of ONE \<Rightarrow> ONE) = x"
    68 lemma one_case3 [simp]: "(case x of ONE \<Rightarrow> ONE) = x"
    71 by (induct x rule: one_induct) simp_all
    69   by (induct x rule: one_induct) simp_all
    72 
    70 
    73 end
    71 end