# HG changeset patch # User huffman # Date 1267056903 28800 # Node ID 73f645fdd4fff01f7a5806169a08c7c28cfed7c4 # Parent 2e0f9516947e3a502b65e097741ac1868beb8e2f reorganizing domain package code (in progress) diff -r 2e0f9516947e -r 73f645fdd4ff src/HOLCF/Domain.thy --- a/src/HOLCF/Domain.thy Wed Feb 24 14:20:07 2010 -0800 +++ b/src/HOLCF/Domain.thy Wed Feb 24 16:15:03 2010 -0800 @@ -9,6 +9,7 @@ uses ("Tools/cont_consts.ML") ("Tools/cont_proc.ML") + ("Tools/Domain/domain_constructors.ML") ("Tools/Domain/domain_library.ML") ("Tools/Domain/domain_syntax.ML") ("Tools/Domain/domain_axioms.ML") @@ -230,6 +231,7 @@ use "Tools/cont_consts.ML" use "Tools/cont_proc.ML" +use "Tools/Domain/domain_constructors.ML" use "Tools/Domain/domain_library.ML" use "Tools/Domain/domain_syntax.ML" use "Tools/Domain/domain_axioms.ML" diff -r 2e0f9516947e -r 73f645fdd4ff src/HOLCF/Tools/Domain/domain_axioms.ML --- a/src/HOLCF/Tools/Domain/domain_axioms.ML Wed Feb 24 14:20:07 2010 -0800 +++ b/src/HOLCF/Tools/Domain/domain_axioms.ML Wed Feb 24 16:15:03 2010 -0800 @@ -78,19 +78,8 @@ (dc_abs oo (copy_of_dtyp map_tab r (dtyp_of_eq eqn)) oo dc_rep)) end; -(* -- definitions concerning the constructors, discriminators and selectors - *) +(* -- definitions concerning the discriminators and selectors - *) - fun con_def m n (_,args) = let - fun idxs z x arg = (if is_lazy arg then mk_up else I) (Bound(z-x)); - fun parms vs = mk_stuple (mapn (idxs(length vs)) 1 vs); - fun inj y 1 _ = y - | inj y _ 0 = mk_sinl y - | inj y i j = mk_sinr (inj y (i-1) (j-1)); - in List.foldr /\# (dc_abs`(inj (parms args) m n)) args end; - - val con_defs = mapn (fn n => fn (con, _, args) => - (extern_name con ^"_def", %%:con == con_def (length cons) n (con,args))) 0 cons; - val dis_defs = let fun ddef (con,_,_) = (dis_name con ^"_def",%%:(dis_name con) == list_ccomb(%%:(dname^"_when"),map @@ -152,7 +141,7 @@ in (dnam, (if definitional then [] else [abs_iso_ax, rep_iso_ax, reach_ax]), (if definitional then [when_def] else [when_def, copy_def]) @ - con_defs @ dis_defs @ mat_defs @ pat_defs @ sel_defs @ + dis_defs @ mat_defs @ pat_defs @ sel_defs @ [take_def, finite_def]) end; (* let (calc_axioms) *) @@ -226,9 +215,11 @@ mk_imp (proj (Bound 2) eqs n $ Bound 1 $ Bound 0, foldr1 mk_disj (mk_conj(Bound 1 === UU,Bound 0 === UU) ::map one_con cons)))); +(* TEMPORARILY DISABLED val bisim_def = ("bisim_def", %%:(comp_dname^"_bisim") == mk_lam("R", foldr1 mk_conj (mapn one_comp 0 eqs))); +TEMPORARILY DISABLED *) fun add_one (dnam, axs, dfs) = Sign.add_path dnam @@ -245,7 +236,7 @@ in thy |> Sign.add_path comp_dnam - |> add_defs_infer (bisim_def::(if use_copy_def then [copy_def] else [])) + |> add_defs_infer ((*bisim_def::*)(if use_copy_def then [copy_def] else [])) |> Sign.parent_path |> fold add_matchers eqs end; (* let (add_axioms) *) diff -r 2e0f9516947e -r 73f645fdd4ff src/HOLCF/Tools/Domain/domain_constructors.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOLCF/Tools/Domain/domain_constructors.ML Wed Feb 24 16:15:03 2010 -0800 @@ -0,0 +1,278 @@ +(* Title: HOLCF/Tools/domain/domain_constructors.ML + Author: Brian Huffman + +Defines constructor functions for a given domain isomorphism +and proves related theorems. +*) + +signature DOMAIN_CONSTRUCTORS = +sig + val add_domain_constructors : + string + -> typ * (binding * (bool * binding option * typ) list * mixfix) list + -> term * term + -> thm * thm + -> theory + -> { con_consts : term list, + con_defs : thm list } + * theory; +end; + + +structure Domain_Constructors :> DOMAIN_CONSTRUCTORS = +struct + +(******************************************************************************) +(************************** building types and terms **************************) +(******************************************************************************) + + +(*** 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; + +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 = 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; + +infix 9 ` ; val (op `) = mk_capply; + + +(*** Product type ***) + +fun mk_tupleT [] = HOLogic.unitT + | mk_tupleT [T] = T + | mk_tupleT (T :: Ts) = HOLogic.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 up_const T = Const(@{const_name up}, T ->> mk_upT T); + +fun mk_up t = up_const (Term.fastype_of t) ` t; + + +(*** Strict product type ***) + +val oneT = @{typ "one"}; + +fun mk_sprodT (T, U) = Type(@{type_name "**"}, [T, U]); + +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 (Term.fastype_of t, Term.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); + + +(*** 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 Term.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; + + +(*** 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 + Library.foldr mk_cfunT (map auto (argTs T), auto T) + end; + +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; + +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}, 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; + +fun coerce_const T = Const (@{const_name coerce}, T); + +fun isodefl_const T = + Const (@{const_name isodefl}, (T ->> T) --> deflT --> HOLogic.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)); + +(*** miscellaneous ***) + +fun declare_consts + (decls : (binding * typ * mixfix) list) + (thy : theory) + : term list * theory = + let + fun con (b, T, mx) = Const (Sign.full_name thy b, T); + val thy = Cont_Consts.add_consts decls thy; + in + (map con decls, thy) + end; + +fun define_consts + (specs : (binding * term * mixfix) list) + (thy : theory) + : (term list * thm list) * theory = + let + fun mk_decl (b, t, mx) = (b, Term.fastype_of t, mx); + val decls = map mk_decl specs; + val thy = Cont_Consts.add_consts decls thy; + fun mk_const (b, T, mx) = Const (Sign.full_name thy b, T); + val consts = map mk_const decls; + fun mk_def c (b, t, mx) = + (Binding.suffix_name "_def" b, Logic.mk_equals (c, t)); + val defs = map2 mk_def consts specs; + val (def_thms, thy) = + PureThy.add_defs false (map Thm.no_attributes defs) thy; + in + ((consts, def_thms), thy) + end; + +(*** argument preprocessing ***) + +type arg = (bool * binding option * typ) * string; + + + +(******************************* main function ********************************) + +fun add_domain_constructors + (dname : string) + (lhsT : typ, + cons : (binding * (bool * binding option * typ) list * mixfix) list) + (rep_const : term, abs_const : term) + (rep_iso_thm : thm, abs_iso_thm : thm) + (thy : theory) = + let + (* TODO: re-enable this *) + (* val thy = Sign.add_path dname thy; *) + + (* define constructor functions *) + val ((con_consts, con_def_thms), thy) = + let + fun prep_con (bind, args, mx) = + (bind, args ~~ Datatype_Prop.make_tnames (map #3 args), mx); + fun var_of ((lazy, sel, T), name) = Free (name, T); + fun is_lazy ((lazy, sel, T), name) = lazy; + val cons' = map prep_con cons; + fun one_arg arg = (if is_lazy arg then mk_up else I) (var_of arg); + fun one_con (bind, args, mx) = mk_stuple (map one_arg args); + fun mk_abs t = abs_const ` t; + val rhss = map mk_abs (mk_sinjects (map one_con cons')); + fun mk_def (bind, args, mx) rhs = + (bind, big_lambdas (map var_of args) rhs, mx); + in + define_consts (map2 mk_def cons' rhss) thy + end; + + (* TODO: re-enable this *) + (* val thy = Sign.parent_path thy; *) + + val result = + { con_consts = con_consts, + con_defs = con_def_thms }; + in + (result, thy) + end; + +end; diff -r 2e0f9516947e -r 73f645fdd4ff src/HOLCF/Tools/Domain/domain_extender.ML --- a/src/HOLCF/Tools/Domain/domain_extender.ML Wed Feb 24 14:20:07 2010 -0800 +++ b/src/HOLCF/Tools/Domain/domain_extender.ML Wed Feb 24 16:15:03 2010 -0800 @@ -167,7 +167,9 @@ val thy = thy' |> Domain_Axioms.add_axioms false comp_dnam eqs; val ((rewss, take_rews), theorems_thy) = thy - |> fold_map (fn eq => Domain_Theorems.theorems (eq, eqs)) eqs + |> fold_map (fn (eq, (x,cs)) => + Domain_Theorems.theorems (eq, eqs) (Type x, cs)) + (eqs ~~ eqs') ||>> Domain_Theorems.comp_theorems (comp_dnam, eqs); in theorems_thy @@ -238,7 +240,9 @@ val thy = thy' |> Domain_Axioms.add_axioms true comp_dnam eqs; val ((rewss, take_rews), theorems_thy) = thy - |> fold_map (fn eq => Domain_Theorems.theorems (eq, eqs)) eqs + |> fold_map (fn (eq, (x,cs)) => + Domain_Theorems.theorems (eq, eqs) (Type x, cs)) + (eqs ~~ eqs') ||>> Domain_Theorems.comp_theorems (comp_dnam, eqs); in theorems_thy diff -r 2e0f9516947e -r 73f645fdd4ff src/HOLCF/Tools/Domain/domain_syntax.ML --- a/src/HOLCF/Tools/Domain/domain_syntax.ML Wed Feb 24 14:20:07 2010 -0800 +++ b/src/HOLCF/Tools/Domain/domain_syntax.ML Wed Feb 24 16:15:03 2010 -0800 @@ -70,8 +70,6 @@ Binding.name ("match_" ^ strip_esc (Binding.name_of con)); fun pat_name_ con = Binding.name (strip_esc (Binding.name_of con) ^ "_pat"); - fun con (name,args,mx) = - (name, List.foldr (op ->>) dtype (map third args), mx); fun dis (con,args,mx) = (dis_name_ con, dtype->>trT, Mixfix(escape ("is_" ^ Binding.name_of con), [], Syntax.max_pri)); @@ -100,7 +98,6 @@ mk_patT (dtype, mk_ctupleT (map (freetvar "t") (1 upto length args))), Mixfix(escape (Binding.name_of con ^ "_pat"), [], Syntax.max_pri)); in - val consts_con = map con cons'; val consts_dis = map dis cons'; val consts_mat = map mat cons'; val consts_pat = map pat cons'; @@ -172,7 +169,7 @@ if definitional then [] else [const_rep, const_abs, const_copy]; in (optional_consts @ [const_when] @ - consts_con @ consts_dis @ consts_mat @ consts_pat @ consts_sel @ + consts_dis @ consts_mat @ consts_pat @ consts_sel @ [const_take, const_finite], (case_trans false :: case_trans true :: (abscon_trans false @ abscon_trans true @ Case_trans))) end; (* let *) diff -r 2e0f9516947e -r 73f645fdd4ff src/HOLCF/Tools/Domain/domain_theorems.ML --- a/src/HOLCF/Tools/Domain/domain_theorems.ML Wed Feb 24 14:20:07 2010 -0800 +++ b/src/HOLCF/Tools/Domain/domain_theorems.ML Wed Feb 24 16:15:03 2010 -0800 @@ -9,7 +9,11 @@ signature DOMAIN_THEOREMS = sig - val theorems: Domain_Library.eq * Domain_Library.eq list -> theory -> thm list * theory; + val theorems: + Domain_Library.eq * Domain_Library.eq list + -> typ * (binding * (bool * binding option * typ) list * mixfix) list + -> theory -> thm list * theory; + val comp_theorems: bstring * Domain_Library.eq list -> theory -> thm list * theory; val quiet_mode: bool Unsynchronized.ref; val trace_domain: bool Unsynchronized.ref; @@ -135,11 +139,13 @@ val dist_eqI = @{lemma "!!x::'a::po. ~ x << y ==> x ~= y" by (blast dest!: below_antisym_inverse)} -fun theorems (((dname, _), cons) : eq, eqs : eq list) thy = +fun theorems + (((dname, _), cons) : eq, eqs : eq list) + (dom_eqn : typ * (binding * (bool * binding option * typ) list * mixfix) list) + (thy : theory) = let val _ = message ("Proving isomorphism properties of domain "^dname^" ..."); -val pg = pg' thy; val map_tab = Domain_Isomorphism.get_map_tab thy; @@ -152,7 +158,6 @@ val ax_rep_iso = ga "rep_iso" dname; val ax_when_def = ga "when_def" dname; fun get_def mk_name (con, _, _) = ga (mk_name con^"_def") dname; - val axs_con_def = map (get_def extern_name) cons; val axs_dis_def = map (get_def dis_name) cons; val axs_mat_def = map (get_def mat_name) cons; val axs_pat_def = map (get_def pat_name) cons; @@ -167,8 +172,35 @@ val ax_copy_def = ga "copy_def" dname; end; (* local *) +(* ----- define constructors ------------------------------------------------ *) + +val lhsT = fst dom_eqn; + +val rhsT = + let + fun mk_arg_typ (lazy, sel, T) = if lazy then mk_uT T else T; + fun mk_con_typ (bind, args, mx) = + if null args then oneT else foldr1 mk_sprodT (map mk_arg_typ args); + fun mk_eq_typ (_, cons) = foldr1 mk_ssumT (map mk_con_typ cons); + in + mk_eq_typ dom_eqn + end; + +val rep_const = Const(dname^"_rep", lhsT ->> rhsT); + +val abs_const = Const(dname^"_abs", rhsT ->> lhsT); + +val (result, thy) = + Domain_Constructors.add_domain_constructors + (Long_Name.base_name dname) dom_eqn + (rep_const, abs_const) (ax_rep_iso, ax_abs_iso) thy; + +val axs_con_def = #con_defs result; + (* ----- theorems concerning the isomorphism -------------------------------- *) +val pg = pg' thy; + val dc_abs = %%:(dname^"_abs"); val dc_rep = %%:(dname^"_rep"); val dc_copy = %%:(dname^"_copy"); @@ -711,7 +743,9 @@ val axs_take_def = map (ga "take_def" ) dnames; val axs_finite_def = map (ga "finite_def") dnames; val ax_copy2_def = ga "copy_def" comp_dnam; +(* TEMPORARILY DISABLED val ax_bisim_def = ga "bisim_def" comp_dnam; +TEMPORARILY DISABLED *) end; local @@ -1031,6 +1065,7 @@ (* ----- theorem concerning coinduction ------------------------------------- *) +(* COINDUCTION TEMPORARILY DISABLED local val xs = mapn (fn n => K (x_name n)) 1 dnames; fun bnd_arg n i = Bound(2*(n_eqs - n)-i-1); @@ -1081,6 +1116,7 @@ take_lemmas; in pg [] goal (K tacs) end; end; (* local *) +COINDUCTION TEMPORARILY DISABLED *) val inducts = Project_Rule.projections (ProofContext.init thy) ind; fun ind_rule (dname, rule) = ((Binding.empty, [rule]), [Induct.induct_type dname]); @@ -1092,8 +1128,8 @@ ((Binding.name "take_lemmas", take_lemmas ), []), ((Binding.name "finites" , finites ), []), ((Binding.name "finite_ind" , [finite_ind]), []), - ((Binding.name "ind" , [ind] ), []), - ((Binding.name "coind" , [coind] ), [])] + ((Binding.name "ind" , [ind] ), [])(*, + ((Binding.name "coind" , [coind] ), [])*)] |> (if induct_failed then I else snd o PureThy.add_thmss (map ind_rule (dnames ~~ inducts))) |> Sign.parent_path |> pair take_rews diff -r 2e0f9516947e -r 73f645fdd4ff src/HOLCF/ex/Domain_ex.thy --- a/src/HOLCF/ex/Domain_ex.thy Wed Feb 24 14:20:07 2010 -0800 +++ b/src/HOLCF/ex/Domain_ex.thy Wed Feb 24 16:15:03 2010 -0800 @@ -122,7 +122,7 @@ text {* Rules about constructors *} term Leaf term Node -thm tree.Leaf_def tree.Node_def +thm Leaf_def Node_def thm tree.exhaust thm tree.casedist thm tree.compacts @@ -175,9 +175,11 @@ thm tree.finites text {* Rules about bisimulation predicate *} +(* COINDUCTION TEMPORARILY DISABLED term tree_bisim thm tree.bisim_def thm tree.coind +COINDUCTION TEMPORARILY DISABLED *) text {* Induction rule *} thm tree.ind diff -r 2e0f9516947e -r 73f645fdd4ff src/HOLCF/ex/Stream.thy --- a/src/HOLCF/ex/Stream.thy Wed Feb 24 14:20:07 2010 -0800 +++ b/src/HOLCF/ex/Stream.thy Wed Feb 24 16:15:03 2010 -0800 @@ -273,6 +273,7 @@ section "coinduction" +(* COINDUCTION TEMPORARILY DISABLED lemma stream_coind_lemma2: "!s1 s2. R s1 s2 --> ft$s1 = ft$s2 & R (rt$s1) (rt$s2) ==> stream_bisim R" apply (simp add: stream.bisim_def,clarsimp) apply (case_tac "x=UU",clarsimp) @@ -286,6 +287,7 @@ apply (erule_tac x="a && y" in allE) apply (erule_tac x="aa && ya" in allE) back by auto +COINDUCTION TEMPORARILY DISABLED *)