Theory Sfun

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

section ‹The Strict Function Type›

theory Sfun
  imports Cfun
begin

pcpodef ('a, 'b) sfun (infixr "→!" 0) = "{f :: 'a  'b. f = }"
  by simp_all

type_notation (ASCII)
  sfun  (infixr "->!" 0)

text ‹TODO: Define nice syntax for abstraction, application.›

definition sfun_abs :: "('a  'b)  ('a →! 'b)"
  where "sfun_abs = (Λ f. Abs_sfun (strictifyf))"

definition sfun_rep :: "('a →! 'b)  'a  'b"
  where "sfun_rep = (Λ f. Rep_sfun f)"

lemma sfun_rep_beta: "sfun_repf = Rep_sfun f"
  by (simp add: sfun_rep_def cont_Rep_sfun)

lemma sfun_rep_strict1 [simp]: "sfun_rep = "
  unfolding sfun_rep_beta by (rule Rep_sfun_strict)

lemma sfun_rep_strict2 [simp]: "sfun_repf = "
  unfolding sfun_rep_beta by (rule Rep_sfun [simplified])

lemma strictify_cancel: "f =   strictifyf = f"
  by (simp add: cfun_eq_iff strictify_conv_if)

lemma sfun_abs_sfun_rep [simp]: "sfun_abs(sfun_repf) = f"
  unfolding sfun_abs_def sfun_rep_def
  apply (simp add: cont_Abs_sfun cont_Rep_sfun)
  apply (simp add: Rep_sfun_inject [symmetric] Abs_sfun_inverse)
  apply (simp add: cfun_eq_iff strictify_conv_if)
  apply (simp add: Rep_sfun [simplified])
  done

lemma sfun_rep_sfun_abs [simp]: "sfun_rep(sfun_absf) = strictifyf"
  unfolding sfun_abs_def sfun_rep_def
  apply (simp add: cont_Abs_sfun cont_Rep_sfun)
  apply (simp add: Abs_sfun_inverse)
  done

lemma sfun_eq_iff: "f = g  sfun_repf = sfun_repg"
  by (simp add: sfun_rep_def cont_Rep_sfun Rep_sfun_inject)

lemma sfun_below_iff: "f  g  sfun_repf  sfun_repg"
  by (simp add: sfun_rep_def cont_Rep_sfun below_sfun_def)

end