Theory Fixrec

Up to index of Isabelle/HOLCF

theory Fixrec
imports Cprod Sprod Ssum One Fix
uses (Tools/holcf_library.ML) (Tools/fixrec.ML)

(*  Title:      HOLCF/Fixrec.thy
Author: Amber Telfer and Brian Huffman
*)


header "Package for defining recursive functions in HOLCF"

theory Fixrec
imports Cprod Sprod Ssum Up One Tr Fix
uses
("Tools/holcf_library.ML")
("Tools/fixrec.ML")
begin


subsection {* Pattern-match monad *}

default_sort cpo

pcpodef (open) 'a match = "UNIV::(one ++ 'a u) set"
by simp_all

definition
fail :: "'a match" where
"fail = Abs_match (sinl·ONE)"


definition
succeed :: "'a -> 'a match" where
"succeed = (Λ x. Abs_match (sinr·(up·x)))"


definition
match_case :: "'b -> ('a -> 'b) -> 'a match -> 'b::pcpo" where
"match_case = (Λ f r m. sscase·(Λ x. f)·(fup·r)·(Rep_match m))"


lemma matchE [case_names bottom fail succeed, cases type: match]:
"[|p = ⊥ ==> Q; p = fail ==> Q; !!x. p = succeed·x ==> Q|] ==> Q"

unfolding fail_def succeed_def
apply (cases p, rename_tac r)
apply (rule_tac p=r in ssumE, simp add: Abs_match_strict)
apply (rule_tac p=x in oneE, simp, simp)
apply (rule_tac p=y in upE, simp, simp add: cont_Abs_match)
done

lemma succeed_defined [simp]: "succeed·x ≠ ⊥"
by (simp add: succeed_def cont_Abs_match Abs_match_defined)

lemma fail_defined [simp]: "fail ≠ ⊥"
by (simp add: fail_def Abs_match_defined)

lemma succeed_eq [simp]: "(succeed·x = succeed·y) = (x = y)"
by (simp add: succeed_def cont_Abs_match Abs_match_inject)

lemma succeed_neq_fail [simp]:
"succeed·x ≠ fail" "fail ≠ succeed·x"

by (simp_all add: succeed_def fail_def cont_Abs_match Abs_match_inject)

lemma match_case_simps [simp]:
"match_case·f·r·⊥ = ⊥"
"match_case·f·r·fail = f"
"match_case·f·r·(succeed·x) = r·x"

by (simp_all add: succeed_def fail_def match_case_def cont_Rep_match
cont2cont_LAM
cont_Abs_match Abs_match_inverse Rep_match_strict)


translations
"case m of XCONST fail => t1 | XCONST succeed·x => t2"
== "CONST match_case·t1·(Λ x. t2)·m"


subsubsection {* Run operator *}

definition
run :: "'a match -> 'a::pcpo" where
"run = match_case·⊥·ID"


text {* rewrite rules for run *}

lemma run_strict [simp]: "run·⊥ = ⊥"
by (simp add: run_def)

lemma run_fail [simp]: "run·fail = ⊥"
by (simp add: run_def)

lemma run_succeed [simp]: "run·(succeed·x) = x"
by (simp add: run_def)

subsubsection {* Monad plus operator *}

definition
mplus :: "'a match -> 'a match -> 'a match" where
"mplus = (Λ m1 m2. case m1 of fail => m2 | succeed·x => m1)"


abbreviation
mplus_syn :: "['a match, 'a match] => 'a match" (infixr "+++" 65) where
"m1 +++ m2 == mplus·m1·m2"


text {* rewrite rules for mplus *}

lemma mplus_strict [simp]: "⊥ +++ m = ⊥"
by (simp add: mplus_def)

lemma mplus_fail [simp]: "fail +++ m = m"
by (simp add: mplus_def)

lemma mplus_succeed [simp]: "succeed·x +++ m = succeed·x"
by (simp add: mplus_def)

lemma mplus_fail2 [simp]: "m +++ fail = m"
by (cases m, simp_all)

lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)"
by (cases x, simp_all)

subsection {* Match functions for built-in types *}

default_sort pcpo

definition
match_UU :: "'a -> 'c match -> 'c match"
where
"match_UU = strictify·(Λ x k. fail)"


definition
match_cpair :: "'a::cpo × 'b::cpo -> ('a -> 'b -> 'c match) -> 'c match"
where
"match_cpair = (Λ x k. csplit·k·x)"


definition
match_spair :: "'a ⊗ 'b -> ('a -> 'b -> 'c match) -> 'c match"
where
"match_spair = (Λ x k. ssplit·k·x)"


definition
match_sinl :: "'a ⊕ 'b -> ('a -> 'c match) -> 'c match"
where
"match_sinl = (Λ x k. sscase·k·(Λ b. fail)·x)"


definition
match_sinr :: "'a ⊕ 'b -> ('b -> 'c match) -> 'c match"
where
"match_sinr = (Λ x k. sscase·(Λ a. fail)·k·x)"


definition
match_up :: "'a::cpo u -> ('a -> 'c match) -> 'c match"
where
"match_up = (Λ x k. fup·k·x)"


definition
match_ONE :: "one -> 'c match -> 'c match"
where
"match_ONE = (Λ ONE k. k)"


definition
match_TT :: "tr -> 'c match -> 'c match"
where
"match_TT = (Λ x k. If x then k else fail fi)"


definition
match_FF :: "tr -> 'c match -> 'c match"
where
"match_FF = (Λ x k. If x then fail else k fi)"


lemma match_UU_simps [simp]:
"match_UU·⊥·k = ⊥"
"x ≠ ⊥ ==> match_UU·x·k = fail"

by (simp_all add: match_UU_def)

lemma match_cpair_simps [simp]:
"match_cpair·(x, y)·k = k·x·y"

by (simp_all add: match_cpair_def)

lemma match_spair_simps [simp]:
"[|x ≠ ⊥; y ≠ ⊥|] ==> match_spair·(:x, y:)·k = k·x·y"
"match_spair·⊥·k = ⊥"

by (simp_all add: match_spair_def)

lemma match_sinl_simps [simp]:
"x ≠ ⊥ ==> match_sinl·(sinl·x)·k = k·x"
"y ≠ ⊥ ==> match_sinl·(sinr·y)·k = fail"
"match_sinl·⊥·k = ⊥"

by (simp_all add: match_sinl_def)

lemma match_sinr_simps [simp]:
"x ≠ ⊥ ==> match_sinr·(sinl·x)·k = fail"
"y ≠ ⊥ ==> match_sinr·(sinr·y)·k = k·y"
"match_sinr·⊥·k = ⊥"

by (simp_all add: match_sinr_def)

lemma match_up_simps [simp]:
"match_up·(up·x)·k = k·x"
"match_up·⊥·k = ⊥"

by (simp_all add: match_up_def)

lemma match_ONE_simps [simp]:
"match_ONE·ONE·k = k"
"match_ONE·⊥·k = ⊥"

by (simp_all add: match_ONE_def)

lemma match_TT_simps [simp]:
"match_TT·TT·k = k"
"match_TT·FF·k = fail"
"match_TT·⊥·k = ⊥"

by (simp_all add: match_TT_def)

lemma match_FF_simps [simp]:
"match_FF·FF·k = k"
"match_FF·TT·k = fail"
"match_FF·⊥·k = ⊥"

by (simp_all add: match_FF_def)

subsection {* Mutual recursion *}

text {*
The following rules are used to prove unfolding theorems from
fixed-point definitions of mutually recursive functions.
*}


lemma Pair_equalI: "[|x ≡ fst p; y ≡ snd p|] ==> (x, y) ≡ p"
by simp

lemma Pair_eqD1: "(x, y) = (x', y') ==> x = x'"
by simp

lemma Pair_eqD2: "(x, y) = (x', y') ==> y = y'"
by simp

lemma def_cont_fix_eq:
"[|f ≡ fix·(Abs_CFun F); cont F|] ==> f = F f"

by (simp, subst fix_eq, simp)

lemma def_cont_fix_ind:
"[|f ≡ fix·(Abs_CFun F); cont F; adm P; P ⊥; !!x. P x ==> P (F x)|] ==> P f"

by (simp add: fix_ind)

text {* lemma for proving rewrite rules *}

lemma ssubst_lhs: "[|t = s; P s = Q|] ==> P t = Q"
by simp


subsection {* Initializing the fixrec package *}

use "Tools/holcf_library.ML"
use "Tools/fixrec.ML"

setup {* Fixrec.setup *}

setup {*
Fixrec.add_matchers
[ (@{const_name up}, @{const_name match_up}),
(@{const_name sinl}, @{const_name match_sinl}),
(@{const_name sinr}, @{const_name match_sinr}),
(@{const_name spair}, @{const_name match_spair}),
(@{const_name Pair}, @{const_name match_cpair}),
(@{const_name ONE}, @{const_name match_ONE}),
(@{const_name TT}, @{const_name match_TT}),
(@{const_name FF}, @{const_name match_FF}),
(@{const_name UU}, @{const_name match_UU}) ]
*}


hide_const (open) succeed fail run

end