# HG changeset patch # User huffman # Date 1267399844 28800 # Node ID 979019ab92eb9a03b18f5c8cd1fce9d2f0f1266b # Parent 7675a41e755a1108d8993407763d88710db9d82f move common functions into new file holcf_library.ML diff -r 7675a41e755a -r 979019ab92eb src/HOLCF/IsaMakefile --- a/src/HOLCF/IsaMakefile Sun Feb 28 15:09:09 2010 -0800 +++ b/src/HOLCF/IsaMakefile Sun Feb 28 15:30:44 2010 -0800 @@ -64,6 +64,7 @@ Tools/adm_tac.ML \ Tools/cont_consts.ML \ Tools/cont_proc.ML \ + Tools/holcf_library.ML \ Tools/Domain/domain_extender.ML \ Tools/Domain/domain_axioms.ML \ Tools/Domain/domain_constructors.ML \ diff -r 7675a41e755a -r 979019ab92eb src/HOLCF/Representable.thy --- a/src/HOLCF/Representable.thy Sun Feb 28 15:09:09 2010 -0800 +++ b/src/HOLCF/Representable.thy Sun Feb 28 15:30:44 2010 -0800 @@ -8,6 +8,7 @@ imports Algebraic Universal Ssum Sprod One Fixrec uses ("Tools/repdef.ML") + ("Tools/holcf_library.ML") ("Tools/Domain/domain_isomorphism.ML") begin @@ -749,6 +750,7 @@ subsection {* Constructing Domain Isomorphisms *} +use "Tools/holcf_library.ML" use "Tools/Domain/domain_isomorphism.ML" setup {* diff -r 7675a41e755a -r 979019ab92eb src/HOLCF/Tools/Domain/domain_constructors.ML --- a/src/HOLCF/Tools/Domain/domain_constructors.ML Sun Feb 28 15:09:09 2010 -0800 +++ b/src/HOLCF/Tools/Domain/domain_constructors.ML Sun Feb 28 15:30:44 2010 -0800 @@ -36,256 +36,7 @@ structure Domain_Constructors :> DOMAIN_CONSTRUCTORS = struct -(******************************************************************************) -(************************** building types and terms **************************) -(******************************************************************************) - - -(*** Operations from Isabelle/HOL ***) - -val boolT = HOLogic.boolT; - -val mk_equals = Logic.mk_equals; -val mk_eq = HOLogic.mk_eq; -val mk_trp = HOLogic.mk_Trueprop; -val mk_fst = HOLogic.mk_fst; -val mk_snd = HOLogic.mk_snd; -val mk_not = HOLogic.mk_not; -val mk_conj = HOLogic.mk_conj; -val mk_disj = HOLogic.mk_disj; - -fun mk_ex (x, t) = HOLogic.exists_const (fastype_of x) $ Term.lambda x t; - -(*** Basic HOLCF concepts ***) - -fun mk_bottom T = Const (@{const_name UU}, T); - -fun below_const T = Const (@{const_name below}, [T, T] ---> boolT); -fun mk_below (t, u) = below_const (fastype_of t) $ t $ u; - -fun mk_undef t = mk_eq (t, mk_bottom (fastype_of t)); - -fun mk_defined t = mk_not (mk_undef t); - -fun mk_compact t = - Const (@{const_name compact}, fastype_of t --> boolT) $ t; - -fun mk_cont t = - Const (@{const_name cont}, fastype_of t --> boolT) $ t; - - -(*** Continuous function space ***) - -(* ->> is taken from holcf_logic.ML *) -fun mk_cfunT (T, U) = Type(@{type_name "->"}, [T, U]); - -infixr 6 ->>; val (op ->>) = mk_cfunT; -infix -->>; val (op -->>) = Library.foldr mk_cfunT; - -fun dest_cfunT (Type(@{type_name "->"}, [T, U])) = (T, U) - | dest_cfunT T = raise TYPE ("dest_cfunT", [T], []); - -fun capply_const (S, T) = - Const(@{const_name Rep_CFun}, (S ->> T) --> (S --> T)); - -fun cabs_const (S, T) = - Const(@{const_name Abs_CFun}, (S --> T) --> (S ->> T)); - -fun mk_cabs t = - let val T = fastype_of t - in cabs_const (Term.domain_type T, Term.range_type T) $ t end - -(* builds the expression (% v1 v2 .. vn. rhs) *) -fun lambdas [] rhs = rhs - | lambdas (v::vs) rhs = Term.lambda v (lambdas vs rhs); - -(* builds the expression (LAM v. rhs) *) -fun big_lambda v rhs = - cabs_const (fastype_of v, fastype_of rhs) $ Term.lambda v rhs; - -(* builds the expression (LAM v1 v2 .. vn. rhs) *) -fun big_lambdas [] rhs = rhs - | big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs); - -fun mk_capply (t, u) = - let val (S, T) = - case fastype_of t of - Type(@{type_name "->"}, [S, T]) => (S, T) - | _ => raise TERM ("mk_capply " ^ ML_Syntax.print_list ML_Syntax.print_term [t, u], [t, u]); - in capply_const (S, T) $ t $ u end; - -infix 9 ` ; val (op `) = mk_capply; - -val list_ccomb : term * term list -> term = Library.foldl mk_capply; - -fun mk_ID T = Const (@{const_name ID}, T ->> T); - -fun cfcomp_const (T, U, V) = - Const (@{const_name cfcomp}, (U ->> V) ->> (T ->> U) ->> (T ->> V)); - -fun mk_cfcomp (f, g) = - let - val (U, V) = dest_cfunT (fastype_of f); - val (T, U') = dest_cfunT (fastype_of g); - in - if U = U' - then mk_capply (mk_capply (cfcomp_const (T, U, V), f), g) - else raise TYPE ("mk_cfcomp", [U, U'], [f, g]) - end; - - -(*** Product type ***) - -val mk_prodT = HOLogic.mk_prodT - -fun mk_tupleT [] = HOLogic.unitT - | mk_tupleT [T] = T - | mk_tupleT (T :: Ts) = mk_prodT (T, mk_tupleT Ts); - -(* builds the expression (v1,v2,..,vn) *) -fun mk_tuple [] = HOLogic.unit - | mk_tuple (t::[]) = t - | mk_tuple (t::ts) = HOLogic.mk_prod (t, mk_tuple ts); - -(* builds the expression (%(v1,v2,..,vn). rhs) *) -fun lambda_tuple [] rhs = Term.lambda (Free("unit", HOLogic.unitT)) rhs - | lambda_tuple (v::[]) rhs = Term.lambda v rhs - | lambda_tuple (v::vs) rhs = - HOLogic.mk_split (Term.lambda v (lambda_tuple vs rhs)); - - -(*** Lifted cpo type ***) - -fun mk_upT T = Type(@{type_name "u"}, [T]); - -fun dest_upT (Type(@{type_name "u"}, [T])) = T - | dest_upT T = raise TYPE ("dest_upT", [T], []); - -fun up_const T = Const(@{const_name up}, T ->> mk_upT T); - -fun mk_up t = up_const (fastype_of t) ` t; - -fun fup_const (T, U) = - Const(@{const_name fup}, (T ->> U) ->> mk_upT T ->> U); - -fun from_up T = fup_const (T, T) ` mk_ID T; - - -(*** Strict product type ***) - -val oneT = @{typ "one"}; - -fun mk_sprodT (T, U) = Type(@{type_name "**"}, [T, U]); - -fun dest_sprodT (Type(@{type_name "**"}, [T, U])) = (T, U) - | dest_sprodT T = raise TYPE ("dest_sprodT", [T], []); - -fun spair_const (T, U) = - Const(@{const_name spair}, T ->> U ->> mk_sprodT (T, U)); - -(* builds the expression (:t, u:) *) -fun mk_spair (t, u) = - spair_const (fastype_of t, fastype_of u) ` t ` u; - -(* builds the expression (:t1,t2,..,tn:) *) -fun mk_stuple [] = @{term "ONE"} - | mk_stuple (t::[]) = t - | mk_stuple (t::ts) = mk_spair (t, mk_stuple ts); - -fun sfst_const (T, U) = - Const(@{const_name sfst}, mk_sprodT (T, U) ->> T); - -fun ssnd_const (T, U) = - Const(@{const_name ssnd}, mk_sprodT (T, U) ->> U); - - -(*** Strict sum type ***) - -fun mk_ssumT (T, U) = Type(@{type_name "++"}, [T, U]); - -fun dest_ssumT (Type(@{type_name "++"}, [T, U])) = (T, U) - | dest_ssumT T = raise TYPE ("dest_ssumT", [T], []); - -fun sinl_const (T, U) = Const(@{const_name sinl}, T ->> mk_ssumT (T, U)); -fun sinr_const (T, U) = Const(@{const_name sinr}, U ->> mk_ssumT (T, U)); - -(* builds the list [sinl(t1), sinl(sinr(t2)), ... sinr(...sinr(tn))] *) -fun mk_sinjects ts = - let - val Ts = map fastype_of ts; - fun combine (t, T) (us, U) = - let - val v = sinl_const (T, U) ` t; - val vs = map (fn u => sinr_const (T, U) ` u) us; - in - (v::vs, mk_ssumT (T, U)) - end - fun inj [] = error "mk_sinjects: empty list" - | inj ((t, T)::[]) = ([t], T) - | inj ((t, T)::ts) = combine (t, T) (inj ts); - in - fst (inj (ts ~~ Ts)) - end; - -fun sscase_const (T, U, V) = - Const(@{const_name sscase}, - (T ->> V) ->> (U ->> V) ->> mk_ssumT (T, U) ->> V); - -fun from_sinl (T, U) = - sscase_const (T, U, T) ` mk_ID T ` mk_bottom (U ->> T); - -fun from_sinr (T, U) = - sscase_const (T, U, U) ` mk_bottom (T ->> U) ` mk_ID U; - - -(*** pattern match monad type ***) - -fun mk_matchT T = Type (@{type_name "maybe"}, [T]); - -fun dest_matchT (Type(@{type_name "maybe"}, [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; - - -(*** miscellaneous constructions ***) - -val trT = @{typ "tr"}; - -val deflT = @{typ "udom alg_defl"}; - -fun mapT T = - let - fun argTs (Type (_, Ts)) = Ts | argTs _ = []; - fun auto T = T ->> T; - in - map auto (argTs T) -->> auto T - end; - -fun mk_strict t = - let val (T, U) = dest_cfunT (fastype_of t); - in mk_eq (t ` mk_bottom T, mk_bottom U) end; - -fun mk_fix t = - let val (T, _) = dest_cfunT (fastype_of t) - in mk_capply (Const(@{const_name fix}, (T ->> T) ->> T), t) end; - -fun mk_Rep_of T = - Const (@{const_name Rep_of}, Term.itselfT T --> deflT) $ Logic.mk_type T; - -fun coerce_const T = Const (@{const_name coerce}, T); - -fun isodefl_const T = - Const (@{const_name isodefl}, (T ->> T) --> deflT --> boolT); - -(* splits a cterm into the right and lefthand sides of equality *) -fun dest_eqs t = HOLogic.dest_eq (HOLogic.dest_Trueprop t); - -fun mk_eqs (t, u) = HOLogic.mk_Trueprop (HOLogic.mk_eq (t, u)); - +open HOLCF_Library; (************************** miscellaneous functions ***************************) diff -r 7675a41e755a -r 979019ab92eb src/HOLCF/Tools/Domain/domain_isomorphism.ML --- a/src/HOLCF/Tools/Domain/domain_isomorphism.ML Sun Feb 28 15:09:09 2010 -0800 +++ b/src/HOLCF/Tools/Domain/domain_isomorphism.ML Sun Feb 28 15:30:44 2010 -0800 @@ -94,99 +94,20 @@ (******************************************************************************) -(******************************* building types *******************************) +(************************** building types and terms **************************) (******************************************************************************) -(* ->> is taken from holcf_logic.ML *) -fun cfunT (T, U) = Type(@{type_name "->"}, [T, U]); - -infixr 6 ->>; val (op ->>) = cfunT; +open HOLCF_Library; -fun dest_cfunT (Type(@{type_name "->"}, [T, U])) = (T, U) - | dest_cfunT T = raise TYPE ("dest_cfunT", [T], []); - -fun tupleT [] = HOLogic.unitT - | tupleT [T] = T - | tupleT (T :: Ts) = HOLogic.mk_prodT (T, tupleT Ts); +infixr 6 ->>; +infix -->>; val deflT = @{typ "udom alg_defl"}; fun mapT (T as Type (_, Ts)) = - Library.foldr cfunT (map (fn T => T ->> T) Ts, T ->> T) + (map (fn T => T ->> T) Ts) -->> (T ->> T) | mapT T = T ->> T; -(******************************************************************************) -(******************************* building terms *******************************) -(******************************************************************************) - -(* builds the expression (v1,v2,..,vn) *) -fun mk_tuple [] = HOLogic.unit -| mk_tuple (t::[]) = t -| mk_tuple (t::ts) = HOLogic.mk_prod (t, mk_tuple ts); - -(* builds the expression (%(v1,v2,..,vn). rhs) *) -fun lambda_tuple [] rhs = Term.lambda (Free("unit", HOLogic.unitT)) rhs - | lambda_tuple (v::[]) rhs = Term.lambda v rhs - | lambda_tuple (v::vs) rhs = - HOLogic.mk_split (Term.lambda v (lambda_tuple vs rhs)); - -(* continuous application and abstraction *) - -fun capply_const (S, T) = - Const(@{const_name Rep_CFun}, (S ->> T) --> (S --> T)); - -fun cabs_const (S, T) = - Const(@{const_name Abs_CFun}, (S --> T) --> (S ->> T)); - -fun mk_cabs t = - let val T = Term.fastype_of t - in cabs_const (Term.domain_type T, Term.range_type T) $ t end - -(* builds the expression (LAM v. rhs) *) -fun big_lambda v rhs = - cabs_const (Term.fastype_of v, Term.fastype_of rhs) $ Term.lambda v rhs; - -(* builds the expression (LAM v1 v2 .. vn. rhs) *) -fun big_lambdas [] rhs = rhs - | big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs); - -fun mk_capply (t, u) = - let val (S, T) = - case Term.fastype_of t of - Type(@{type_name "->"}, [S, T]) => (S, T) - | _ => raise TERM ("mk_capply " ^ ML_Syntax.print_list ML_Syntax.print_term [t, u], [t, u]); - in capply_const (S, T) $ t $ u end; - -(* miscellaneous term constructions *) - -val mk_trp = HOLogic.mk_Trueprop; - -val mk_fst = HOLogic.mk_fst; -val mk_snd = HOLogic.mk_snd; - -fun mk_cont t = - let val T = Term.fastype_of t - in Const(@{const_name cont}, T --> HOLogic.boolT) $ t end; - -fun mk_fix t = - let val (T, _) = dest_cfunT (Term.fastype_of t) - in mk_capply (Const(@{const_name fix}, (T ->> T) ->> T), t) end; - -fun ID_const T = Const (@{const_name ID}, cfunT (T, T)); - -fun cfcomp_const (T, U, V) = - Const (@{const_name cfcomp}, (U ->> V) ->> (T ->> U) ->> (T ->> V)); - -fun mk_cfcomp (f, g) = - let - val (U, V) = dest_cfunT (Term.fastype_of f); - val (T, U') = dest_cfunT (Term.fastype_of g); - in - if U = U' - then mk_capply (mk_capply (cfcomp_const (T, U, V), f), g) - else raise TYPE ("mk_cfcomp", [U, U'], [f, g]) - end; - fun mk_Rep_of T = Const (@{const_name Rep_of}, Term.itselfT T --> deflT) $ Logic.mk_type T; @@ -301,7 +222,7 @@ case Symtab.lookup tab c of SOME t => Library.foldl mk_capply (Const (t, mapT T), map map_of Ts) | NONE => if is_closed_typ T - then ID_const T + then mk_ID T else error ("map_of_typ: type variable under unsupported type constructor " ^ c); in map_of T end; @@ -383,7 +304,7 @@ (* declare deflation combinator constants *) fun declare_defl_const (vs, tbind, mx, rhs, morphs) thy = let - val defl_type = Library.foldr cfunT (map (K deflT) vs, deflT); + val defl_type = map (K deflT) vs -->> deflT; val defl_bind = Binding.suffix_name "_defl" tbind; in Sign.declare_const ((defl_bind, defl_type), NoSyn) thy @@ -442,8 +363,8 @@ (* define rep/abs functions *) fun mk_rep_abs ((tbind, morphs), (lhsT, rhsT)) thy = let - val rep_type = cfunT (lhsT, rhsT); - val abs_type = cfunT (rhsT, lhsT); + val rep_type = lhsT ->> rhsT; + val abs_type = rhsT ->> lhsT; val rep_bind = Binding.suffix_name "_rep" tbind; val abs_bind = Binding.suffix_name "_abs" tbind; val (rep_bind, abs_bind) = the_default (rep_bind, abs_bind) morphs; @@ -594,8 +515,8 @@ (((map_const, (lhsT, _)), REP_thm), isodefl_thm) = let val Ts = snd (dest_Type lhsT); - val lhs = Library.foldl mk_capply (map_const, map ID_const Ts); - val goal = mk_eqs (lhs, ID_const lhsT); + val lhs = Library.foldl mk_capply (map_const, map mk_ID Ts); + val goal = mk_eqs (lhs, mk_ID lhsT); val tac = EVERY [rtac @{thm isodefl_REP_imp_ID} 1, stac REP_thm 1, @@ -616,7 +537,7 @@ (* define copy combinators *) val new_dts = map (apsnd (map (fst o dest_TFree)) o dest_Type o fst) dom_eqns; - val copy_arg_type = tupleT (map (fn (T, _) => T ->> T) dom_eqns); + val copy_arg_type = mk_tupleT (map (fn (T, _) => T ->> T) dom_eqns); val copy_arg = Free ("f", copy_arg_type); val copy_args = let fun mk_copy_args [] t = [] @@ -627,16 +548,16 @@ fun copy_of_dtyp (T, dt) = if Datatype_Aux.is_rec_type dt then copy_of_dtyp' (T, dt) - else ID_const T + else mk_ID T and copy_of_dtyp' (T, Datatype_Aux.DtRec i) = nth copy_args i - | copy_of_dtyp' (T, Datatype_Aux.DtTFree a) = ID_const T + | copy_of_dtyp' (T, Datatype_Aux.DtTFree a) = mk_ID T | copy_of_dtyp' (T, Datatype_Aux.DtType (c, ds)) = case Symtab.lookup map_tab' c of SOME f => Library.foldl mk_capply (Const (f, mapT T), map copy_of_dtyp (snd (dest_Type T) ~~ ds)) | NONE => - (warning ("copy_of_dtyp: unknown type constructor " ^ c); ID_const T); + (warning ("copy_of_dtyp: unknown type constructor " ^ c); mk_ID T); fun define_copy ((tbind, (rep_const, abs_const)), (lhsT, rhsT)) thy = let val copy_type = copy_arg_type ->> (lhsT ->> lhsT); @@ -686,7 +607,7 @@ val fix_copy_lemma = let fun mk_map_ID (map_const, (T, rhsT)) = - Library.foldl mk_capply (map_const, map ID_const (snd (dest_Type T))); + Library.foldl mk_capply (map_const, map mk_ID (snd (dest_Type T))); val rhs = mk_tuple (map mk_map_ID (map_consts ~~ dom_eqns)); val goal = mk_eqs (mk_fix c_const, rhs); val rules = diff -r 7675a41e755a -r 979019ab92eb src/HOLCF/Tools/holcf_library.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/Tools/holcf_library.ML Sun Feb 28 15:30:44 2010 -0800 @@ -0,0 +1,237 @@ +(* Title: HOLCF/Tools/holcf_library.ML + Author: Brian Huffman + +Functions for constructing HOLCF types and terms. +*) + +structure HOLCF_Library = +struct + +(*** Operations from Isabelle/HOL ***) + +val boolT = HOLogic.boolT; + +val mk_equals = Logic.mk_equals; +val mk_eq = HOLogic.mk_eq; +val mk_trp = HOLogic.mk_Trueprop; +val mk_fst = HOLogic.mk_fst; +val mk_snd = HOLogic.mk_snd; +val mk_not = HOLogic.mk_not; +val mk_conj = HOLogic.mk_conj; +val mk_disj = HOLogic.mk_disj; + +fun mk_ex (x, t) = HOLogic.exists_const (fastype_of x) $ Term.lambda x t; + + +(*** Basic HOLCF concepts ***) + +fun mk_bottom T = Const (@{const_name UU}, T); + +fun below_const T = Const (@{const_name below}, [T, T] ---> boolT); +fun mk_below (t, u) = below_const (fastype_of t) $ t $ u; + +fun mk_undef t = mk_eq (t, mk_bottom (fastype_of t)); + +fun mk_defined t = mk_not (mk_undef t); + +fun mk_compact t = + Const (@{const_name compact}, fastype_of t --> boolT) $ t; + +fun mk_cont t = + Const (@{const_name cont}, fastype_of t --> boolT) $ t; + + +(*** Continuous function space ***) + +(* ->> is taken from holcf_logic.ML *) +fun mk_cfunT (T, U) = Type(@{type_name "->"}, [T, U]); + +infixr 6 ->>; val (op ->>) = mk_cfunT; +infix -->>; val (op -->>) = Library.foldr mk_cfunT; + +fun dest_cfunT (Type(@{type_name "->"}, [T, U])) = (T, U) + | dest_cfunT T = raise TYPE ("dest_cfunT", [T], []); + +fun capply_const (S, T) = + Const(@{const_name Rep_CFun}, (S ->> T) --> (S --> T)); + +fun cabs_const (S, T) = + Const(@{const_name Abs_CFun}, (S --> T) --> (S ->> T)); + +fun mk_cabs t = + let val T = fastype_of t + in cabs_const (Term.domain_type T, Term.range_type T) $ t end + +(* builds the expression (% v1 v2 .. vn. rhs) *) +fun lambdas [] rhs = rhs + | lambdas (v::vs) rhs = Term.lambda v (lambdas vs rhs); + +(* builds the expression (LAM v. rhs) *) +fun big_lambda v rhs = + cabs_const (fastype_of v, fastype_of rhs) $ Term.lambda v rhs; + +(* builds the expression (LAM v1 v2 .. vn. rhs) *) +fun big_lambdas [] rhs = rhs + | big_lambdas (v::vs) rhs = big_lambda v (big_lambdas vs rhs); + +fun mk_capply (t, u) = + let val (S, T) = + case fastype_of t of + Type(@{type_name "->"}, [S, T]) => (S, T) + | _ => raise TERM ("mk_capply " ^ ML_Syntax.print_list ML_Syntax.print_term [t, u], [t, u]); + in capply_const (S, T) $ t $ u end; + +infix 9 ` ; val (op `) = mk_capply; + +val list_ccomb : term * term list -> term = Library.foldl mk_capply; + +fun mk_ID T = Const (@{const_name ID}, T ->> T); + +fun cfcomp_const (T, U, V) = + Const (@{const_name cfcomp}, (U ->> V) ->> (T ->> U) ->> (T ->> V)); + +fun mk_cfcomp (f, g) = + let + val (U, V) = dest_cfunT (fastype_of f); + val (T, U') = dest_cfunT (fastype_of g); + in + if U = U' + then mk_capply (mk_capply (cfcomp_const (T, U, V), f), g) + else raise TYPE ("mk_cfcomp", [U, U'], [f, g]) + end; + +fun mk_strict t = + let val (T, U) = dest_cfunT (fastype_of t); + in mk_eq (t ` mk_bottom T, mk_bottom U) end; + + +(*** Product type ***) + +val mk_prodT = HOLogic.mk_prodT + +fun mk_tupleT [] = HOLogic.unitT + | mk_tupleT [T] = T + | mk_tupleT (T :: Ts) = mk_prodT (T, mk_tupleT Ts); + +(* builds the expression (v1,v2,..,vn) *) +fun mk_tuple [] = HOLogic.unit + | mk_tuple (t::[]) = t + | mk_tuple (t::ts) = HOLogic.mk_prod (t, mk_tuple ts); + +(* builds the expression (%(v1,v2,..,vn). rhs) *) +fun lambda_tuple [] rhs = Term.lambda (Free("unit", HOLogic.unitT)) rhs + | lambda_tuple (v::[]) rhs = Term.lambda v rhs + | lambda_tuple (v::vs) rhs = + HOLogic.mk_split (Term.lambda v (lambda_tuple vs rhs)); + + +(*** Lifted cpo type ***) + +fun mk_upT T = Type(@{type_name "u"}, [T]); + +fun dest_upT (Type(@{type_name "u"}, [T])) = T + | dest_upT T = raise TYPE ("dest_upT", [T], []); + +fun up_const T = Const(@{const_name up}, T ->> mk_upT T); + +fun mk_up t = up_const (fastype_of t) ` t; + +fun fup_const (T, U) = + Const(@{const_name fup}, (T ->> U) ->> mk_upT T ->> U); + +fun from_up T = fup_const (T, T) ` mk_ID T; + + +(*** Strict product type ***) + +val oneT = @{typ "one"}; + +fun mk_sprodT (T, U) = Type(@{type_name "**"}, [T, U]); + +fun dest_sprodT (Type(@{type_name "**"}, [T, U])) = (T, U) + | dest_sprodT T = raise TYPE ("dest_sprodT", [T], []); + +fun spair_const (T, U) = + Const(@{const_name spair}, T ->> U ->> mk_sprodT (T, U)); + +(* builds the expression (:t, u:) *) +fun mk_spair (t, u) = + spair_const (fastype_of t, fastype_of u) ` t ` u; + +(* builds the expression (:t1,t2,..,tn:) *) +fun mk_stuple [] = @{term "ONE"} + | mk_stuple (t::[]) = t + | mk_stuple (t::ts) = mk_spair (t, mk_stuple ts); + +fun sfst_const (T, U) = + Const(@{const_name sfst}, mk_sprodT (T, U) ->> T); + +fun ssnd_const (T, U) = + Const(@{const_name ssnd}, mk_sprodT (T, U) ->> U); + + +(*** Strict sum type ***) + +fun mk_ssumT (T, U) = Type(@{type_name "++"}, [T, U]); + +fun dest_ssumT (Type(@{type_name "++"}, [T, U])) = (T, U) + | dest_ssumT T = raise TYPE ("dest_ssumT", [T], []); + +fun sinl_const (T, U) = Const(@{const_name sinl}, T ->> mk_ssumT (T, U)); +fun sinr_const (T, U) = Const(@{const_name sinr}, U ->> mk_ssumT (T, U)); + +(* builds the list [sinl(t1), sinl(sinr(t2)), ... sinr(...sinr(tn))] *) +fun mk_sinjects ts = + let + val Ts = map fastype_of ts; + fun combine (t, T) (us, U) = + let + val v = sinl_const (T, U) ` t; + val vs = map (fn u => sinr_const (T, U) ` u) us; + in + (v::vs, mk_ssumT (T, U)) + end + fun inj [] = error "mk_sinjects: empty list" + | inj ((t, T)::[]) = ([t], T) + | inj ((t, T)::ts) = combine (t, T) (inj ts); + in + fst (inj (ts ~~ Ts)) + end; + +fun sscase_const (T, U, V) = + Const(@{const_name sscase}, + (T ->> V) ->> (U ->> V) ->> mk_ssumT (T, U) ->> V); + +fun from_sinl (T, U) = + sscase_const (T, U, T) ` mk_ID T ` mk_bottom (U ->> T); + +fun from_sinr (T, U) = + sscase_const (T, U, U) ` mk_bottom (T ->> U) ` mk_ID U; + + +(*** pattern match monad type ***) + +fun mk_matchT T = Type (@{type_name "maybe"}, [T]); + +fun dest_matchT (Type(@{type_name "maybe"}, [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; + + +(*** lifted boolean type ***) + +val trT = @{typ "tr"}; + + +(*** theory of fixed points ***) + +fun mk_fix t = + let val (T, _) = dest_cfunT (fastype_of t) + in mk_capply (Const(@{const_name fix}, (T ->> T) ->> T), t) end; + + +end;