Theory Nat_Discrete

(*  Title:      HOL/HOLCF/Library/Nat_Discrete.thy
    Author:     Brian Huffman
*)

section ‹Discrete cpo instance for naturals›

theory Nat_Discrete
imports HOLCF
begin

text ‹Discrete cpo instance for typnat.›

instantiation nat :: discrete_cpo
begin

definition below_nat_def:
  "(x::nat)  y  x = y"

instance proof
qed (rule below_nat_def)

end

text ‹
  TODO: implement a command to automate discrete predomain instances.
›

instantiation nat :: predomain
begin

definition
  "(liftemb :: nat u  udom u)  liftemb oo u_map(Λ x. Discr x)"

definition
  "(liftprj :: udom u  nat u)  u_map(Λ y. undiscr y) oo liftprj"

definition
  "liftdefl  (λ(t::nat itself). LIFTDEFL(nat discr))"

instance proof
  show "ep_pair liftemb (liftprj :: udom u  nat u)"
    unfolding liftemb_nat_def liftprj_nat_def
    apply (rule ep_pair_comp)
    apply (rule ep_pair_u_map)
    apply (simp add: ep_pair.intro)
    apply (rule predomain_ep)
    done
  show "castLIFTDEFL(nat) = liftemb oo (liftprj :: udom u  nat u)"
    unfolding liftemb_nat_def liftprj_nat_def liftdefl_nat_def
    apply (simp add: cast_liftdefl cfcomp1 u_map_map)
    apply (simp add: ID_def [symmetric] u_map_ID)
    done
qed

end

end