# HG changeset patch # User huffman # Date 1274718772 25200 # Node ID 00f13d3ad47473d7769ba4a2c2ab856fc31c91d0 # Parent 1535aa1c943a2c9eb0acb389ba5c8953fdfa79d5 rename type 'a maybe to 'a match; rename Fixrec.return to Fixrec.succeed diff -r 1535aa1c943a -r 00f13d3ad474 src/HOLCF/Fixrec.thy --- a/src/HOLCF/Fixrec.thy Mon May 24 13:48:57 2010 +0200 +++ b/src/HOLCF/Fixrec.thy Mon May 24 09:32:52 2010 -0700 @@ -11,65 +11,65 @@ ("Tools/fixrec.ML") begin -subsection {* Maybe monad type *} +subsection {* Pattern-match monad *} default_sort cpo -pcpodef (open) 'a maybe = "UNIV::(one ++ 'a u) set" +pcpodef (open) 'a match = "UNIV::(one ++ 'a u) set" by simp_all definition - fail :: "'a maybe" where - "fail = Abs_maybe (sinl\ONE)" + fail :: "'a match" where + "fail = Abs_match (sinl\ONE)" definition - return :: "'a \ 'a maybe" where - "return = (\ x. Abs_maybe (sinr\(up\x)))" + succeed :: "'a \ 'a match" where + "succeed = (\ x. Abs_match (sinr\(up\x)))" definition - maybe_when :: "'b \ ('a \ 'b) \ 'a maybe \ 'b::pcpo" where - "maybe_when = (\ f r m. sscase\(\ x. f)\(fup\r)\(Rep_maybe m))" + match_case :: "'b \ ('a \ 'b) \ 'a match \ 'b::pcpo" where + "match_case = (\ f r m. sscase\(\ x. f)\(fup\r)\(Rep_match m))" -lemma maybeE: - "\p = \ \ Q; p = fail \ Q; \x. p = return\x \ Q\ \ Q" -apply (unfold fail_def return_def) +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_maybe_strict) +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_maybe) +apply (rule_tac p=y in upE, simp, simp add: cont_Abs_match) done -lemma return_defined [simp]: "return\x \ \" -by (simp add: return_def cont_Abs_maybe Abs_maybe_defined) +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_maybe_defined) +by (simp add: fail_def Abs_match_defined) -lemma return_eq [simp]: "(return\x = return\y) = (x = y)" -by (simp add: return_def cont_Abs_maybe Abs_maybe_inject) +lemma succeed_eq [simp]: "(succeed\x = succeed\y) = (x = y)" +by (simp add: succeed_def cont_Abs_match Abs_match_inject) -lemma return_neq_fail [simp]: - "return\x \ fail" "fail \ return\x" -by (simp_all add: return_def fail_def cont_Abs_maybe Abs_maybe_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 maybe_when_rews [simp]: - "maybe_when\f\r\\ = \" - "maybe_when\f\r\fail = f" - "maybe_when\f\r\(return\x) = r\x" -by (simp_all add: return_def fail_def maybe_when_def cont_Rep_maybe +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_maybe Abs_maybe_inverse Rep_maybe_strict) + cont_Abs_match Abs_match_inverse Rep_match_strict) translations - "case m of XCONST fail \ t1 | XCONST return\x \ t2" - == "CONST maybe_when\t1\(\ x. t2)\m" + "case m of XCONST fail \ t1 | XCONST succeed\x \ t2" + == "CONST match_case\t1\(\ x. t2)\m" subsubsection {* Run operator *} definition - run :: "'a maybe \ 'a::pcpo" where - "run = maybe_when\\\ID" + run :: "'a match \ 'a::pcpo" where + "run = match_case\\\ID" text {* rewrite rules for run *} @@ -79,17 +79,17 @@ lemma run_fail [simp]: "run\fail = \" by (simp add: run_def) -lemma run_return [simp]: "run\(return\x) = x" +lemma run_succeed [simp]: "run\(succeed\x) = x" by (simp add: run_def) subsubsection {* Monad plus operator *} definition - mplus :: "'a maybe \ 'a maybe \ 'a maybe" where - "mplus = (\ m1 m2. case m1 of fail \ m2 | return\x \ m1)" + mplus :: "'a match \ 'a match \ 'a match" where + "mplus = (\ m1 m2. case m1 of fail \ m2 | succeed\x \ m1)" abbreviation - mplus_syn :: "['a maybe, 'a maybe] \ 'a maybe" (infixr "+++" 65) where + mplus_syn :: "['a match, 'a match] \ 'a match" (infixr "+++" 65) where "m1 +++ m2 == mplus\m1\m2" text {* rewrite rules for mplus *} @@ -100,23 +100,23 @@ lemma mplus_fail [simp]: "fail +++ m = m" by (simp add: mplus_def) -lemma mplus_return [simp]: "return\x +++ m = return\x" +lemma mplus_succeed [simp]: "succeed\x +++ m = succeed\x" by (simp add: mplus_def) lemma mplus_fail2 [simp]: "m +++ fail = m" -by (rule_tac p=m in maybeE, simp_all) +by (cases m, simp_all) lemma mplus_assoc: "(x +++ y) +++ z = x +++ (y +++ z)" -by (rule_tac p=x in maybeE, simp_all) +by (cases x, simp_all) subsubsection {* Fatbar combinator *} definition - fatbar :: "('a \ 'b maybe) \ ('a \ 'b maybe) \ ('a \ 'b maybe)" where + fatbar :: "('a \ 'b match) \ ('a \ 'b match) \ ('a \ 'b match)" where "fatbar = (\ a b x. a\x +++ b\x)" abbreviation - fatbar_syn :: "['a \ 'b maybe, 'a \ 'b maybe] \ 'a \ 'b maybe" (infixr "\" 60) where + fatbar_syn :: "['a \ 'b match, 'a \ 'b match] \ 'a \ 'b match" (infixr "\" 60) where "m1 \ m2 == fatbar\m1\m2" lemma fatbar1: "m\x = \ \ (m \ ms)\x = \" @@ -125,7 +125,7 @@ lemma fatbar2: "m\x = fail \ (m \ ms)\x = ms\x" by (simp add: fatbar_def) -lemma fatbar3: "m\x = return\y \ (m \ ms)\x = return\y" +lemma fatbar3: "m\x = succeed\y \ (m \ ms)\x = succeed\y" by (simp add: fatbar_def) lemmas fatbar_simps = fatbar1 fatbar2 fatbar3 @@ -136,7 +136,7 @@ lemma run_fatbar2: "m\x = fail \ run\((m \ ms)\x) = run\(ms\x)" by (simp add: fatbar_def) -lemma run_fatbar3: "m\x = return\y \ run\((m \ ms)\x) = y" +lemma run_fatbar3: "m\x = succeed\y \ run\((m \ ms)\x) = y" by (simp add: fatbar_def) lemmas run_fatbar_simps [simp] = run_fatbar1 run_fatbar2 run_fatbar3 @@ -144,23 +144,23 @@ subsection {* Case branch combinator *} definition - branch :: "('a \ 'b maybe) \ ('b \ 'c) \ ('a \ 'c maybe)" where - "branch p \ \ r x. maybe_when\fail\(\ y. return\(r\y))\(p\x)" + branch :: "('a \ 'b match) \ ('b \ 'c) \ ('a \ 'c match)" where + "branch p \ \ r x. match_case\fail\(\ y. succeed\(r\y))\(p\x)" -lemma branch_rews: +lemma branch_simps: "p\x = \ \ branch p\r\x = \" "p\x = fail \ branch p\r\x = fail" - "p\x = return\y \ branch p\r\x = return\(r\y)" + "p\x = succeed\y \ branch p\r\x = succeed\(r\y)" by (simp_all add: branch_def) -lemma branch_return [simp]: "branch return\r\x = return\(r\x)" +lemma branch_succeed [simp]: "branch succeed\r\x = succeed\(r\x)" by (simp add: branch_def) subsubsection {* Cases operator *} definition - cases :: "'a maybe \ 'a::pcpo" where - "cases = maybe_when\\\ID" + cases :: "'a match \ 'a::pcpo" where + "cases = match_case\\\ID" text {* rewrite rules for cases *} @@ -170,7 +170,7 @@ lemma cases_fail [simp]: "cases\fail = \" by (simp add: cases_def) -lemma cases_return [simp]: "cases\(return\x) = x" +lemma cases_succeed [simp]: "cases\(succeed\x) = x" by (simp add: cases_def) subsection {* Case syntax *} @@ -204,9 +204,9 @@ "_variable _noargs r" => "CONST unit_when\r" parse_translation {* -(* rewrite (_pat x) => (return) *) +(* rewrite (_pat x) => (succeed) *) (* rewrite (_variable x t) => (Abs_CFun (%x. t)) *) - [(@{syntax_const "_pat"}, fn _ => Syntax.const @{const_syntax Fixrec.return}), + [(@{syntax_const "_pat"}, fn _ => Syntax.const @{const_syntax Fixrec.succeed}), mk_binder_tr (@{syntax_const "_variable"}, @{const_syntax Abs_CFun})]; *} @@ -243,17 +243,17 @@ *} translations - "x" <= "_match (CONST Fixrec.return) (_variable x)" + "x" <= "_match (CONST Fixrec.succeed) (_variable x)" subsection {* Pattern combinators for data constructors *} -types ('a, 'b) pat = "'a \ 'b maybe" +types ('a, 'b) pat = "'a \ 'b match" definition cpair_pat :: "('a, 'c) pat \ ('b, 'd) pat \ ('a \ 'b, 'c \ 'd) pat" where "cpair_pat p1 p2 = (\(x, y). - maybe_when\fail\(\ a. maybe_when\fail\(\ b. return\(a, b))\(p2\y))\(p1\x))" + match_case\fail\(\ a. match_case\fail\(\ b. succeed\(a, b))\(p2\y))\(p1\x))" definition spair_pat :: @@ -274,15 +274,15 @@ definition TT_pat :: "(tr, unit) pat" where - "TT_pat = (\ b. If b then return\() else fail fi)" + "TT_pat = (\ b. If b then succeed\() else fail fi)" definition FF_pat :: "(tr, unit) pat" where - "FF_pat = (\ b. If b then fail else return\() fi)" + "FF_pat = (\ b. If b then fail else succeed\() fi)" definition ONE_pat :: "(one, unit) pat" where - "ONE_pat = (\ ONE. return\())" + "ONE_pat = (\ ONE. succeed\())" text {* Parse translations (patterns) *} translations @@ -331,21 +331,21 @@ lemma cpair_pat1: "branch p\r\x = \ \ branch (cpair_pat p q)\(csplit\r)\(x, y) = \" apply (simp add: branch_def cpair_pat_def) -apply (rule_tac p="p\x" in maybeE, simp_all) +apply (cases "p\x", simp_all) done lemma cpair_pat2: "branch p\r\x = fail \ branch (cpair_pat p q)\(csplit\r)\(x, y) = fail" apply (simp add: branch_def cpair_pat_def) -apply (rule_tac p="p\x" in maybeE, simp_all) +apply (cases "p\x", simp_all) done lemma cpair_pat3: - "branch p\r\x = return\s \ + "branch p\r\x = succeed\s \ branch (cpair_pat p q)\(csplit\r)\(x, y) = branch q\s\y" apply (simp add: branch_def cpair_pat_def) -apply (rule_tac p="p\x" in maybeE, simp_all) -apply (rule_tac p="q\y" in maybeE, simp_all) +apply (cases "p\x", simp_all) +apply (cases "q\y", simp_all) done lemmas cpair_pat [simp] = @@ -377,35 +377,35 @@ lemma TT_pat [simp]: "branch TT_pat\(unit_when\r)\\ = \" - "branch TT_pat\(unit_when\r)\TT = return\r" + "branch TT_pat\(unit_when\r)\TT = succeed\r" "branch TT_pat\(unit_when\r)\FF = fail" by (simp_all add: branch_def TT_pat_def) lemma FF_pat [simp]: "branch FF_pat\(unit_when\r)\\ = \" "branch FF_pat\(unit_when\r)\TT = fail" - "branch FF_pat\(unit_when\r)\FF = return\r" + "branch FF_pat\(unit_when\r)\FF = succeed\r" by (simp_all add: branch_def FF_pat_def) lemma ONE_pat [simp]: "branch ONE_pat\(unit_when\r)\\ = \" - "branch ONE_pat\(unit_when\r)\ONE = return\r" + "branch ONE_pat\(unit_when\r)\ONE = succeed\r" by (simp_all add: branch_def ONE_pat_def) subsection {* Wildcards, as-patterns, and lazy patterns *} definition - wild_pat :: "'a \ unit maybe" where - "wild_pat = (\ x. return\())" + wild_pat :: "'a \ unit match" where + "wild_pat = (\ x. succeed\())" definition - as_pat :: "('a \ 'b maybe) \ 'a \ ('a \ 'b) maybe" where - "as_pat p = (\ x. maybe_when\fail\(\ a. return\(x, a))\(p\x))" + as_pat :: "('a \ 'b match) \ 'a \ ('a \ 'b) match" where + "as_pat p = (\ x. match_case\fail\(\ a. succeed\(x, a))\(p\x))" definition - lazy_pat :: "('a \ 'b::pcpo maybe) \ ('a \ 'b maybe)" where - "lazy_pat p = (\ x. return\(cases\(p\x)))" + lazy_pat :: "('a \ 'b::pcpo match) \ ('a \ 'b match)" where + "lazy_pat p = (\ x. succeed\(cases\(p\x)))" text {* Parse translations (patterns) *} translations @@ -419,21 +419,21 @@ translations "_" <= "_match (CONST wild_pat) _noargs" -lemma wild_pat [simp]: "branch wild_pat\(unit_when\r)\x = return\r" +lemma wild_pat [simp]: "branch wild_pat\(unit_when\r)\x = succeed\r" by (simp add: branch_def wild_pat_def) lemma as_pat [simp]: "branch (as_pat p)\(csplit\r)\x = branch p\(r\x)\x" apply (simp add: branch_def as_pat_def) -apply (rule_tac p="p\x" in maybeE, simp_all) +apply (cases "p\x", simp_all) done lemma lazy_pat [simp]: - "branch p\r\x = \ \ branch (lazy_pat p)\r\x = return\(r\\)" - "branch p\r\x = fail \ branch (lazy_pat p)\r\x = return\(r\\)" - "branch p\r\x = return\s \ branch (lazy_pat p)\r\x = return\s" + "branch p\r\x = \ \ branch (lazy_pat p)\r\x = succeed\(r\\)" + "branch p\r\x = fail \ branch (lazy_pat p)\r\x = succeed\(r\\)" + "branch p\r\x = succeed\s \ branch (lazy_pat p)\r\x = succeed\s" apply (simp_all add: branch_def lazy_pat_def) -apply (rule_tac [!] p="p\x" in maybeE, simp_all) +apply (cases "p\x", simp_all)+ done @@ -442,47 +442,47 @@ default_sort pcpo definition - match_UU :: "'a \ 'c maybe \ 'c maybe" + match_UU :: "'a \ 'c match \ 'c match" where "match_UU = strictify\(\ x k. fail)" definition - match_cpair :: "'a::cpo \ 'b::cpo \ ('a \ 'b \ 'c maybe) \ 'c maybe" + 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 maybe) \ 'c maybe" + 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 maybe) \ 'c maybe" + 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 maybe) \ 'c maybe" + 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 maybe) \ 'c maybe" + match_up :: "'a::cpo u \ ('a \ 'c match) \ 'c match" where "match_up = (\ x k. fup\k\x)" definition - match_ONE :: "one \ 'c maybe \ 'c maybe" + match_ONE :: "one \ 'c match \ 'c match" where "match_ONE = (\ ONE k. k)" definition - match_TT :: "tr \ 'c maybe \ 'c maybe" + match_TT :: "tr \ 'c match \ 'c match" where "match_TT = (\ x k. If x then k else fail fi)" definition - match_FF :: "tr \ 'c maybe \ 'c maybe" + match_FF :: "tr \ 'c match \ 'c match" where "match_FF = (\ x k. If x then fail else k fi)" @@ -584,6 +584,6 @@ (@{const_name UU}, @{const_name match_UU}) ] *} -hide_const (open) return fail run cases +hide_const (open) succeed fail run cases end diff -r 1535aa1c943a -r 00f13d3ad474 src/HOLCF/Tools/Domain/domain_constructors.ML --- a/src/HOLCF/Tools/Domain/domain_constructors.ML Mon May 24 13:48:57 2010 +0200 +++ b/src/HOLCF/Tools/Domain/domain_constructors.ML Mon May 24 09:32:52 2010 -0700 @@ -850,7 +850,7 @@ in pat_const $ p1 $ p2 end; - fun mk_tuple_pat [] = return_const HOLogic.unitT + fun mk_tuple_pat [] = succeed_const HOLogic.unitT | mk_tuple_pat ps = foldr1 mk_pair_pat ps; fun branch_const (T,U,V) = Const (@{const_name branch}, @@ -951,7 +951,7 @@ val (fun1, fun2, taken) = pat_lhs (pat, args); val defs = @{thm branch_def} :: pat_defs; val goal = mk_trp (mk_strict fun1); - val rules = @{thms maybe_when_rews} @ case_rews; + val rules = @{thms match_case_simps} @ case_rews; val tacs = [simp_tac (beta_ss addsimps rules) 1]; in prove thy defs goal (K tacs) end; fun pat_apps (i, (pat, (con, args))) = @@ -966,7 +966,7 @@ val concl = mk_trp (mk_eq (fun1 ` con_app, rhs)); val goal = Logic.list_implies (assms, concl); val defs = @{thm branch_def} :: pat_defs; - val rules = @{thms maybe_when_rews} @ case_rews; + val rules = @{thms match_case_simps} @ case_rews; val tacs = [asm_simp_tac (beta_ss addsimps rules) 1]; in prove thy defs goal (K tacs) end; in map_index pat_app spec end; diff -r 1535aa1c943a -r 00f13d3ad474 src/HOLCF/Tools/fixrec.ML --- a/src/HOLCF/Tools/fixrec.ML Mon May 24 13:48:57 2010 +0200 +++ b/src/HOLCF/Tools/fixrec.ML Mon May 24 09:32:52 2010 -0700 @@ -85,7 +85,7 @@ in case t of Const(@{const_name Rep_CFun}, _) $ - Const(@{const_name Fixrec.return}, _) $ u => u + Const(@{const_name Fixrec.succeed}, _) $ u => u | _ => run ` t end; @@ -267,7 +267,7 @@ let val (lhs,rhs) = dest_eqs (Logic.strip_imp_concl (strip_alls eq)); in - building match_name lhs (mk_return rhs) [] (taken_names eq) + building match_name lhs (mk_succeed rhs) [] (taken_names eq) end; (* returns the sum (using +++) of the terms in ms *) diff -r 1535aa1c943a -r 00f13d3ad474 src/HOLCF/Tools/holcf_library.ML --- a/src/HOLCF/Tools/holcf_library.ML Mon May 24 13:48:57 2010 +0200 +++ b/src/HOLCF/Tools/holcf_library.ML Mon May 24 09:32:52 2010 -0700 @@ -252,15 +252,15 @@ (*** pattern match monad type ***) -fun mk_matchT T = Type (@{type_name "maybe"}, [T]); +fun mk_matchT T = Type (@{type_name "match"}, [T]); -fun dest_matchT (Type(@{type_name "maybe"}, [T])) = T +fun dest_matchT (Type(@{type_name "match"}, [T])) = T | dest_matchT T = raise TYPE ("dest_matchT", [T], []); fun mk_fail T = Const (@{const_name "Fixrec.fail"}, mk_matchT T); -fun return_const T = Const (@{const_name "Fixrec.return"}, T ->> mk_matchT T); -fun mk_return t = return_const (fastype_of t) ` t; +fun succeed_const T = Const (@{const_name "Fixrec.succeed"}, T ->> mk_matchT T); +fun mk_succeed t = succeed_const (fastype_of t) ` t; (*** lifted boolean type ***)