| author | haftmann | 
| Mon, 13 Sep 2010 16:43:23 +0200 | |
| changeset 39379 | ab1b070aa412 | 
| parent 38348 | cf7b2121ad9d | 
| child 39557 | fe5722fce758 | 
| permissions | -rw-r--r-- | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 1 | (* Title: HOLCF/Tools/repdef.ML | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 2 | Author: Brian Huffman | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 3 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 4 | Defining representable domains using algebraic deflations. | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 5 | *) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 6 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 7 | signature REPDEF = | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 8 | sig | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 9 | type rep_info = | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 10 |     { emb_def: thm, prj_def: thm, approx_def: thm, REP: thm }
 | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 11 | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 12 | val add_repdef: bool -> binding option -> binding * (string * sort) list * mixfix -> | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 13 | term -> (binding * binding) option -> theory -> | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 14 | (Typedef.info * Pcpodef.cpo_info * Pcpodef.pcpo_info * rep_info) * theory | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 15 | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 16 | val repdef_cmd: (bool * binding) * (binding * (string * string option) list * mixfix) * string | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 17 | * (binding * binding) option -> theory -> theory | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 18 | end; | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 19 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 20 | structure Repdef :> REPDEF = | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 21 | struct | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 22 | |
| 35527 | 23 | open HOLCF_Library; | 
| 24 | ||
| 25 | infixr 6 ->>; | |
| 26 | infix -->>; | |
| 27 | ||
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 28 | (** type definitions **) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 29 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 30 | type rep_info = | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 31 |   { emb_def: thm, prj_def: thm, approx_def: thm, REP: thm };
 | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 32 | |
| 35527 | 33 | (* building types and terms *) | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 34 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 35 | val udomT = @{typ udom};
 | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 36 | fun alg_deflT T = Type (@{type_name alg_defl}, [T]);
 | 
| 35527 | 37 | fun emb_const T = Const (@{const_name emb}, T ->> udomT);
 | 
| 38 | fun prj_const T = Const (@{const_name prj}, udomT ->> T);
 | |
| 39 | fun approx_const T = Const (@{const_name approx}, natT --> (T ->> T));
 | |
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 40 | |
| 35527 | 41 | fun cast_const T = Const (@{const_name cast}, alg_deflT T ->> T ->> T);
 | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 42 | fun mk_cast (t, x) = | 
| 35527 | 43 | capply_const (udomT, udomT) | 
| 44 | $ (capply_const (alg_deflT udomT, udomT ->> udomT) $ cast_const udomT $ t) | |
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 45 | $ x; | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 46 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 47 | (* manipulating theorems *) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 48 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 49 | (* proving class instances *) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 50 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 51 | fun declare_type_name a = | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 52 | Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS))); | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 53 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 54 | fun gen_add_repdef | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 55 | (prep_term: Proof.context -> 'a -> term) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 56 | (def: bool) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 57 | (name: binding) | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 58 | (typ as (tname, raw_args, mx) : binding * (string * sort) list * mixfix) | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 59 | (raw_defl: 'a) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 60 | (opt_morphs: (binding * binding) option) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 61 | (thy: theory) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 62 | : (Typedef.info * Pcpodef.cpo_info * Pcpodef.pcpo_info * rep_info) * theory = | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 63 | let | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 64 | val _ = Theory.requires thy "Representable" "repdefs"; | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 65 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 66 | (*rhs*) | 
| 36153 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
35994diff
changeset | 67 | val tmp_ctxt = | 
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36241diff
changeset | 68 | ProofContext.init_global thy | 
| 36153 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
35994diff
changeset | 69 | |> fold (Variable.declare_typ o TFree) raw_args; | 
| 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
35994diff
changeset | 70 | val defl = prep_term tmp_ctxt raw_defl; | 
| 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
35994diff
changeset | 71 | val tmp_ctxt = tmp_ctxt |> Variable.declare_constraints defl; | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 72 | |
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 73 | val deflT = Term.fastype_of defl; | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 74 |     val _ = if deflT = @{typ "udom alg_defl"} then ()
 | 
| 36153 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
35994diff
changeset | 75 |             else error ("Not type udom alg_defl: " ^ quote (Syntax.string_of_typ tmp_ctxt deflT));
 | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 76 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 77 | (*lhs*) | 
| 36153 
1ac501e16a6a
replaced slightly odd Typedecl.predeclare_constraints by plain declaration of type arguments -- also avoid "recursive" declaration of type constructor, which can cause problems with sequential definitions B.foo = A.foo;
 wenzelm parents: 
35994diff
changeset | 78 | val lhs_tfrees = map (ProofContext.check_tfree tmp_ctxt) raw_args; | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 79 | val lhs_sorts = map snd lhs_tfrees; | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 80 | val full_tname = Sign.full_name thy tname; | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 81 | val newT = Type (full_tname, map TFree lhs_tfrees); | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 82 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 83 | (*morphisms*) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 84 | val morphs = opt_morphs | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 85 | |> the_default (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name); | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 86 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 87 | (*set*) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 88 |     val in_defl = @{term "in_deflation :: udom => udom alg_defl => bool"};
 | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 89 |     val set = HOLogic.Collect_const udomT $ Abs ("x", udomT, in_defl $ Bound 0 $ defl);
 | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 90 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 91 | (*pcpodef*) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 92 |     val tac1 = rtac @{thm CollectI} 1 THEN rtac @{thm bottom_in_deflation} 1;
 | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 93 |     val tac2 = rtac @{thm adm_mem_Collect_in_deflation} 1;
 | 
| 35904 | 94 | val ((info, cpo_info, pcpo_info), thy) = thy | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 95 | |> Pcpodef.add_pcpodef def (SOME name) typ set (SOME morphs) (tac1, tac2); | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 96 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 97 | (*definitions*) | 
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35904diff
changeset | 98 | val Rep_const = Const (#Rep_name (#1 info), newT --> udomT); | 
| 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35904diff
changeset | 99 | val Abs_const = Const (#Abs_name (#1 info), udomT --> newT); | 
| 35527 | 100 | val emb_eqn = Logic.mk_equals (emb_const newT, cabs_const (newT, udomT) $ Rep_const); | 
| 101 | val prj_eqn = Logic.mk_equals (prj_const newT, cabs_const (udomT, newT) $ | |
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 102 |       Abs ("x", udomT, Abs_const $ mk_cast (defl, Bound 0)));
 | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 103 | val repdef_approx_const = | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 104 |       Const (@{const_name repdef_approx}, (newT --> udomT) --> (udomT --> newT)
 | 
| 35527 | 105 | --> alg_deflT udomT --> natT --> (newT ->> newT)); | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 106 | val approx_eqn = Logic.mk_equals (approx_const newT, | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 107 | repdef_approx_const $ Rep_const $ Abs_const $ defl); | 
| 35904 | 108 | val name_def = Binding.suffix_name "_def" name; | 
| 109 | val emb_bind = (Binding.prefix_name "emb_" name_def, []); | |
| 110 | val prj_bind = (Binding.prefix_name "prj_" name_def, []); | |
| 111 | val approx_bind = (Binding.prefix_name "approx_" name_def, []); | |
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 112 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 113 | (*instantiate class rep*) | 
| 35904 | 114 | val lthy = thy | 
| 38348 
cf7b2121ad9d
moved instantiation target formally to class_target.ML
 haftmann parents: 
36960diff
changeset | 115 |       |> Class.instantiation ([full_tname], lhs_tfrees, @{sort rep});
 | 
| 35904 | 116 | val ((_, (_, emb_ldef)), lthy) = | 
| 117 | Specification.definition (NONE, (emb_bind, emb_eqn)) lthy; | |
| 118 | val ((_, (_, prj_ldef)), lthy) = | |
| 119 | Specification.definition (NONE, (prj_bind, prj_eqn)) lthy; | |
| 120 | val ((_, (_, approx_ldef)), lthy) = | |
| 121 | Specification.definition (NONE, (approx_bind, approx_eqn)) lthy; | |
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36241diff
changeset | 122 | val ctxt_thy = ProofContext.init_global (ProofContext.theory_of lthy); | 
| 35904 | 123 | val emb_def = singleton (ProofContext.export lthy ctxt_thy) emb_ldef; | 
| 124 | val prj_def = singleton (ProofContext.export lthy ctxt_thy) prj_ldef; | |
| 125 | val approx_def = singleton (ProofContext.export lthy ctxt_thy) approx_ldef; | |
| 33826 | 126 | val type_definition_thm = | 
| 127 | MetaSimplifier.rewrite_rule | |
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35904diff
changeset | 128 | (the_list (#set_def (#2 info))) | 
| 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35904diff
changeset | 129 | (#type_definition (#2 info)); | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 130 | val typedef_thms = | 
| 33826 | 131 | [type_definition_thm, #below_def cpo_info, emb_def, prj_def, approx_def]; | 
| 35904 | 132 | val thy = lthy | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 133 | |> Class.prove_instantiation_instance | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 134 |           (K (Tactic.rtac (@{thm typedef_rep_class} OF typedef_thms) 1))
 | 
| 33681 | 135 | |> Local_Theory.exit_global; | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 136 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 137 | (*other theorems*) | 
| 35904 | 138 | val typedef_thms' = map (Thm.transfer thy) | 
| 33826 | 139 | [type_definition_thm, #below_def cpo_info, emb_def, prj_def]; | 
| 35904 | 140 | val (REP_thm, thy) = thy | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 141 | |> Sign.add_path (Binding.name_of name) | 
| 35904 | 142 | |> PureThy.add_thm | 
| 143 | ((Binding.prefix_name "REP_" name, | |
| 36241 
2a4cec6bcae2
replace many uses of Drule.export_without_context with Drule.zero_var_indexes
 huffman parents: 
36153diff
changeset | 144 |           Drule.zero_var_indexes (@{thm typedef_REP} OF typedef_thms')), [])
 | 
| 35904 | 145 | ||> Sign.restore_naming thy; | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 146 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 147 | val rep_info = | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 148 |       { emb_def = emb_def, prj_def = prj_def, approx_def = approx_def, REP = REP_thm };
 | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 149 | in | 
| 35904 | 150 | ((info, cpo_info, pcpo_info, rep_info), thy) | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 151 | end | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 152 | handle ERROR msg => | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 153 |     cat_error msg ("The error(s) above occurred in repdef " ^ quote (Binding.str_of name));
 | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 154 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 155 | fun add_repdef def opt_name typ defl opt_morphs thy = | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 156 | let | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 157 | val name = the_default (#1 typ) opt_name; | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 158 | in | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 159 | gen_add_repdef Syntax.check_term def name typ defl opt_morphs thy | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 160 | end; | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 161 | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 162 | fun repdef_cmd ((def, name), (b, raw_args, mx), A, morphs) thy = | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 163 | let | 
| 36610 
bafd82950e24
renamed ProofContext.init to ProofContext.init_global to emphasize that this is not the real thing;
 wenzelm parents: 
36241diff
changeset | 164 | val ctxt = ProofContext.init_global thy; | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 165 | val args = map (apsnd (Typedecl.read_constraint ctxt)) raw_args; | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 166 | in snd (gen_add_repdef Syntax.read_term def name (b, args, mx) A morphs thy) end; | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 167 | |
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 168 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 169 | (** outer syntax **) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 170 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 171 | val repdef_decl = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 172 |   Scan.optional (Parse.$$$ "(" |--
 | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 173 | ((Parse.$$$ "open" >> K false) -- Scan.option Parse.binding || | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 174 | Parse.binding >> (fn s => (true, SOME s))) | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 175 | --| Parse.$$$ ")") (true, NONE) -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 176 | (Parse.type_args_constrained -- Parse.binding) -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 177 | Parse.opt_mixfix -- (Parse.$$$ "=" |-- Parse.term) -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 178 | Scan.option (Parse.$$$ "morphisms" |-- Parse.!!! (Parse.binding -- Parse.binding)); | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 179 | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 180 | fun mk_repdef ((((((def, opt_name), (args, t)), mx), A), morphs)) = | 
| 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 181 | repdef_cmd ((def, the_default t opt_name), (t, args, mx), A, morphs); | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 182 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 183 | val _ = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 184 | Outer_Syntax.command "repdef" "HOLCF definition of representable domains" Keyword.thy_decl | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 185 | (repdef_decl >> | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 186 | (Toplevel.print oo (Toplevel.theory o mk_repdef))); | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 187 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 188 | end; |