Theory Char_Discrete
section ‹Discrete cpo instance for 8-bit char type›
theory Char_Discrete
imports HOLCF
begin
subsection ‹Discrete cpo instance for \<^typ>‹char›.›
instantiation char :: discrete_cpo
begin
definition below_char_def:
"(x::char) ⊑ y ⟷ x = y"
instance proof
qed (rule below_char_def)
end
text ‹
TODO: implement a command to automate discrete predomain instances.
›
instantiation char :: predomain
begin
definition
"(liftemb :: char u → udom u) ≡ liftemb oo u_map⋅(Λ x. Discr x)"
definition
"(liftprj :: udom u → char u) ≡ u_map⋅(Λ y. undiscr y) oo liftprj"
definition
"liftdefl ≡ (λ(t::char itself). LIFTDEFL(char discr))"
instance proof
show "ep_pair liftemb (liftprj :: udom u → char u)"
unfolding liftemb_char_def liftprj_char_def
apply (rule ep_pair_comp)
apply (rule ep_pair_u_map)
apply (simp add: ep_pair.intro)
apply (rule predomain_ep)
done
show "cast⋅LIFTDEFL(char) = liftemb oo (liftprj :: udom u → char u)"
unfolding liftemb_char_def liftprj_char_def liftdefl_char_def
apply (simp add: cast_liftdefl cfcomp1 u_map_map)
apply (simp add: ID_def [symmetric] u_map_ID)
done
qed
end
end