# HG changeset patch # User huffman # Date 1140137198 -3600 # Node ID e32cf29f01fc92857b3f710cda12351647810c34 # Parent a6a15d3446ad486744247c5c17a852aed14b1dc4 make maybe into a real type constructor; remove monad syntax diff -r a6a15d3446ad -r e32cf29f01fc src/HOLCF/Fixrec.thy --- a/src/HOLCF/Fixrec.thy Thu Feb 16 23:40:32 2006 +0100 +++ b/src/HOLCF/Fixrec.thy Fri Feb 17 01:46:38 2006 +0100 @@ -14,84 +14,81 @@ defaultsort cpo -types 'a maybe = "one ++ 'a u" +pcpodef (open) 'a maybe = "UNIV::(one ++ 'a u) set" +by simp constdefs fail :: "'a maybe" - "fail \ sinl\ONE" + "fail \ Abs_maybe (sinl\ONE)" return :: "'a \ 'a maybe" - "return \ sinr oo up" + "return \ \ x. Abs_maybe (sinr\(up\x))" + + maybe_when :: "'b \ ('a \ 'b) \ 'a maybe \ 'b::pcpo" + "maybe_when \ \ f r m. sscase\(\ x. f)\(fup\r)\(Rep_maybe m)" lemma maybeE: "\p = \ \ Q; p = fail \ Q; \x. p = return\x \ Q\ \ Q" apply (unfold fail_def return_def) -apply (rule_tac p=p in ssumE, simp) +apply (cases p, rename_tac r) +apply (rule_tac p=r in ssumE, simp add: Abs_maybe_strict) apply (rule_tac p=x in oneE, simp, simp) -apply (rule_tac p=y in upE, simp, simp) +apply (rule_tac p=y in upE, simp, simp add: cont_Abs_maybe) done lemma return_defined [simp]: "return\x \ \" -by (simp add: return_def) +by (simp add: return_def cont_Abs_maybe Abs_maybe_defined) lemma fail_defined [simp]: "fail \ \" -by (simp add: fail_def) +by (simp add: fail_def Abs_maybe_defined) lemma return_eq [simp]: "(return\x = return\y) = (x = y)" -by (simp add: return_def) +by (simp add: return_def cont_Abs_maybe Abs_maybe_inject) lemma return_neq_fail [simp]: "return\x \ fail" "fail \ return\x" -by (simp_all add: return_def fail_def) +by (simp_all add: return_def fail_def cont_Abs_maybe Abs_maybe_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 + cont_Abs_maybe Abs_maybe_inverse Rep_maybe_strict) + +translations + "case m of fail \ t1 | return\x \ t2" == "maybe_when\t1\(\ x. t2)\m" subsubsection {* Monadic bind operator *} constdefs bind :: "'a maybe \ ('a \ 'b maybe) \ 'b maybe" - "bind \ \ m f. sscase\sinl\(fup\f)\m" - -syntax ">>=" :: "['a maybe, 'a \ 'b maybe] \ 'b maybe" (infixl ">>=" 50) -translations "m >>= f" == "bind\m\f" - -nonterminals - maybebind maybebinds - -syntax - "_MBIND" :: "pttrn \ 'a maybe \ maybebind" ("(2_ <-/ _)" 10) - "" :: "maybebind \ maybebinds" ("_") - - "_MBINDS" :: "[maybebind, maybebinds] \ maybebinds" ("_;/ _") - "_MDO" :: "[maybebinds, 'a maybe] \ 'a maybe" ("(do _;/ (_))" 10) - -translations - "_MDO (_MBINDS b bs) e" == "_MDO b (_MDO bs e)" - "do (x,y) <- m; e" == "m >>= (LAM . e)" - "do x <- m; e" == "m >>= (LAM x. e)" + "bind \ \ m f. case m of fail \ fail | return\x \ f\x" text {* monad laws *} -lemma bind_strict [simp]: "UU >>= f = UU" +lemma bind_strict [simp]: "bind\\\f = \" +by (simp add: bind_def) + +lemma bind_fail [simp]: "bind\fail\f = fail" by (simp add: bind_def) -lemma bind_fail [simp]: "fail >>= f = fail" -by (simp add: bind_def fail_def) +lemma left_unit [simp]: "bind\(return\a)\k = k\a" +by (simp add: bind_def) -lemma left_unit [simp]: "(return\a) >>= k = k\a" -by (simp add: bind_def return_def) - -lemma right_unit [simp]: "m >>= return = m" +lemma right_unit [simp]: "bind\m\return = m" by (rule_tac p=m in maybeE, simp_all) -lemma bind_assoc [simp]: - "(do b <- (do a <- m; k\a); h\b) = (do a <- m; b <- k\a; h\b)" +lemma bind_assoc: + "bind\(bind\m\k)\h = bind\m\(\ a. bind\(k\a)\h)" by (rule_tac p=m in maybeE, simp_all) subsubsection {* Run operator *} constdefs - run:: "'a::pcpo maybe \ 'a" - "run \ sscase\\\(fup\ID)" + run:: "'a maybe \ 'a::pcpo" + "run \ maybe_when\\\ID" text {* rewrite rules for run *} @@ -99,16 +96,16 @@ by (simp add: run_def) lemma run_fail [simp]: "run\fail = \" -by (simp add: run_def fail_def) +by (simp add: run_def) lemma run_return [simp]: "run\(return\x) = x" -by (simp add: run_def return_def) +by (simp add: run_def) subsubsection {* Monad plus operator *} constdefs mplus :: "'a maybe \ 'a maybe \ 'a maybe" - "mplus \ \ m1 m2. sscase\(\ x. m2)\(fup\return)\m1" + "mplus \ \ m1 m2. case m1 of fail \ m2 | return\x \ m1" syntax "+++" :: "['a maybe, 'a maybe] \ 'a maybe" (infixr "+++" 65) translations "m1 +++ m2" == "mplus\m1\m2" @@ -119,10 +116,10 @@ by (simp add: mplus_def) lemma mplus_fail [simp]: "fail +++ m = m" -by (simp add: mplus_def fail_def) +by (simp add: mplus_def) lemma mplus_return [simp]: "return\x +++ m = return\x" -by (simp add: mplus_def return_def) +by (simp add: mplus_def) lemma mplus_fail2 [simp]: "m +++ fail = m" by (rule_tac p=m in maybeE, simp_all) @@ -254,7 +251,8 @@ constdefs cpair_pat :: "('a, 'c) pat \ ('b, 'd) pat \ ('a \ 'b, 'c \ 'd) pat" - "cpair_pat p1 p2 \ \\x, y\. do a <- p1\x; b <- p2\y; return\\a, b\" + "cpair_pat p1 p2 \ \\x, y\. + bind\(p1\x)\(\ a. bind\(p2\y)\(\ b. return\\a, b\))" spair_pat :: "('a, 'c) pat \ ('b, 'd) pat \ ('a::pcpo \ 'b::pcpo, 'c \ 'd) pat" @@ -389,7 +387,7 @@ "wild_pat \ \ x. return\()" as_pat :: "('a \ 'b maybe) \ 'a \ ('a \ 'b) maybe" - "as_pat p \ \ x. do a <- p\x; return\\x, a\" + "as_pat p \ \ x. bind\(p\x)\(\ a. return\\x, a\)" lazy_pat :: "('a \ 'b::pcpo maybe) \ ('a \ 'b maybe)" "lazy_pat p \ \ x. return\(run\(p\x))" diff -r a6a15d3446ad -r e32cf29f01fc src/HOLCF/domain/library.ML --- a/src/HOLCF/domain/library.ML Thu Feb 16 23:40:32 2006 +0100 +++ b/src/HOLCF/domain/library.ML Fri Feb 17 01:46:38 2006 +0100 @@ -207,6 +207,7 @@ | mk_stuple ts = foldr1 spair ts; fun mk_ctupleT [] = HOLogic.unitT (* used in match_defs *) | mk_ctupleT Ts = foldr1 HOLogic.mk_prodT Ts; +fun mk_maybeT T = Type ("Fixrec.maybe",[T]); fun cpair_pat (p1,p2) = %%:cpair_patN $ p1 $ p2; fun lift_defined f = lift (fn x => defined (f x)); fun bound_arg vns v = Bound(length vns -find_index_eq v vns -1); diff -r a6a15d3446ad -r e32cf29f01fc src/HOLCF/domain/syntax.ML --- a/src/HOLCF/domain/syntax.ML Thu Feb 16 23:40:32 2006 +0100 +++ b/src/HOLCF/domain/syntax.ML Fri Feb 17 01:46:38 2006 +0100 @@ -47,13 +47,13 @@ (* strictly speaking, these constants have one argument, but the mixfix (without arguments) is introduced only to generate parse rules for non-alphanumeric names*) - fun mat (con ,s,args) = (mat_name_ con, dtype->>mk_ssumT(oneT,mk_uT(mk_ctupleT(map third args))), + fun mat (con ,s,args) = (mat_name_ con, dtype->>mk_maybeT(mk_ctupleT(map third args)), Mixfix(escape ("match_" ^ con), [], Syntax.max_pri)); fun sel1 (_,sel,typ) = Option.map (fn s => (s,dtype ->> typ,NoSyn)) sel; fun sel (_ ,_,args) = List.mapPartial sel1 args; fun freetvar s n = let val tvar = mk_TFree (s ^ string_of_int n) in if tvar mem typevars then freetvar ("t"^s) n else tvar end; - fun mk_patT (a,b) = a ->> mk_ssumT (oneT, mk_uT b); + fun mk_patT (a,b) = a ->> mk_maybeT b; fun pat_arg_typ n arg = mk_patT (third arg, freetvar "t" n); fun pat (con ,s,args) = (pat_name_ con, (mapn pat_arg_typ 1 args) ---> mk_patT (dtype, mk_ctupleT (map (freetvar "t") (1 upto length args))),