# HG changeset patch # User huffman # Date 1267206457 28800 # Node ID cf8c5a751a9ad1f3ff3ca5b7afc3c543b943fc59 # Parent a726a033b3135b56c9e4343e35100c83e4d5bce3 move proof of con_rews into domain_constructor.ML diff -r a726a033b313 -r cf8c5a751a9a src/HOLCF/Tools/Domain/domain_constructors.ML --- a/src/HOLCF/Tools/Domain/domain_constructors.ML Fri Feb 26 09:13:29 2010 -0800 +++ b/src/HOLCF/Tools/Domain/domain_constructors.ML Fri Feb 26 09:47:37 2010 -0800 @@ -16,6 +16,7 @@ -> { con_consts : term list, con_betas : thm list, con_compacts : thm list, + con_rews : thm list, sel_rews : thm list } * theory; end; @@ -219,7 +220,9 @@ let val (T, U) = dest_cfunT (Term.fastype_of t); in mk_eq (t ` mk_bottom T, mk_bottom U) end; -fun mk_defined t = mk_not (mk_eq (t, mk_bottom (Term.fastype_of t))); +fun mk_undef t = mk_eq (t, mk_bottom (Term.fastype_of t)); + +fun mk_defined t = mk_not (mk_undef t); fun mk_compact t = Const (@{const_name compact}, Term.fastype_of t --> boolT) $ t; @@ -458,6 +461,9 @@ = let + (* get theorems about rep and abs *) + val abs_strict = iso_locale RS @{thm iso.abs_strict}; + (* define constructor functions *) val ((con_consts, con_defs), thy) = let @@ -511,11 +517,51 @@ map con_compact spec' end; + (* prove strictness rules for constructors *) + local + fun con_strict (con, args) = + let + val rules = abs_strict :: @{thms con_strict_rules}; + val Ts = map snd args; + val ns = Datatype_Prop.make_tnames Ts; + val vs = map Free (ns ~~ Ts); + val nonlazy = map snd (filter_out fst (map fst args ~~ vs)); + fun one_strict v' = + let + val UU = mk_bottom (Term.fastype_of v'); + val vs' = map (fn v => if v = v' then UU else v) vs; + val goal = mk_trp (mk_undef (list_ccomb (con, vs'))); + val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1]; + in prove thy con_betas goal (K tacs) end; + in map one_strict nonlazy end; + + fun con_defin (con, args) = + let + fun iff_disj (t, []) = HOLogic.mk_not t + | iff_disj (t, ts) = mk_eq (t, foldr1 HOLogic.mk_disj ts); + val Ts = map snd args; + val ns = Datatype_Prop.make_tnames Ts; + val vs = map Free (ns ~~ Ts); + val nonlazy = map snd (filter_out fst (map fst args ~~ vs)); + val lhs = mk_undef (list_ccomb (con, vs)); + val rhss = map mk_undef nonlazy; + val goal = mk_trp (iff_disj (lhs, rhss)); + val rule1 = iso_locale RS @{thm iso.abs_defined_iff}; + val rules = rule1 :: @{thms con_defined_iff_rules}; + val tacs = [simp_tac (HOL_ss addsimps rules) 1]; + in prove thy con_betas goal (K tacs) end; + in + val con_stricts = maps con_strict spec'; + val con_defins = map con_defin spec'; + val con_rews = con_stricts @ con_defins; + end; + val result = { con_consts = con_consts, con_betas = con_betas, - con_compacts = con_compacts + con_compacts = con_compacts, + con_rews = con_rews }; in (result, thy) @@ -542,7 +588,7 @@ val abs_defined_iff = iso_locale RS @{thm iso.abs_defined_iff}; (* define constructor functions *) - val ({con_consts, con_betas, con_compacts}, thy) = + val ({con_consts, con_betas, con_compacts, con_rews}, thy) = add_constructors spec abs_const iso_locale thy; (* TODO: enable this earlier *) @@ -564,6 +610,7 @@ { con_consts = con_consts, con_betas = con_betas, con_compacts = con_compacts, + con_rews = con_rews, sel_rews = sel_thms }; in (result, thy) diff -r a726a033b313 -r cf8c5a751a9a src/HOLCF/Tools/Domain/domain_theorems.ML --- a/src/HOLCF/Tools/Domain/domain_theorems.ML Fri Feb 26 09:13:29 2010 -0800 +++ b/src/HOLCF/Tools/Domain/domain_theorems.ML Fri Feb 26 09:47:37 2010 -0800 @@ -199,6 +199,7 @@ val con_appls = #con_betas result; val con_compacts = #con_compacts result; +val con_rews = #con_rews result; val sel_rews = #sel_rews result; (* ----- theorems concerning the isomorphism -------------------------------- *) @@ -427,37 +428,6 @@ val pat_rews = pat_stricts @ pat_apps; end; -local - fun con_strict (con, _, args) = - let - val rules = abs_strict :: @{thms con_strict_rules}; - fun one_strict vn = - let - fun f arg = if vname arg = vn then UU else %# arg; - val goal = mk_trp (con_app2 con f args === UU); - val tacs = [simp_tac (HOL_basic_ss addsimps rules) 1]; - in pg con_appls goal (K tacs) end; - in map one_strict (nonlazy args) end; - - fun con_defin (con, _, args) = - let - fun iff_disj (t, []) = HOLogic.mk_not t - | iff_disj (t, ts) = t === foldr1 HOLogic.mk_disj ts; - val lhs = con_app con args === UU; - val rhss = map (fn x => %:x === UU) (nonlazy args); - val goal = mk_trp (iff_disj (lhs, rhss)); - val rule1 = iso_locale RS @{thm iso.abs_defined_iff}; - val rules = rule1 :: @{thms con_defined_iff_rules}; - val tacs = [simp_tac (HOL_ss addsimps rules) 1]; - in pg con_appls goal (K tacs) end; -in - val _ = trace " Proving con_stricts..."; - val con_stricts = maps con_strict cons; - val _ = trace " Proving con_defins..."; - val con_defins = map con_defin cons; - val con_rews = con_stricts @ con_defins; -end; - val _ = trace " Proving dist_les..."; val dist_les = let