| author | noschinl | 
| Wed, 23 Feb 2011 11:23:26 +0100 | |
| changeset 41827 | 98eda7ffde79 | 
| parent 41436 | 480978f80eae | 
| child 42151 | 4da4fc77664b | 
| 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 | |
| 40575 | 7 | signature DOMAINDEF = | 
| 33679 
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 = | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 10 |     {
 | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 11 | emb_def : thm, | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 12 | prj_def : thm, | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 13 | defl_def : thm, | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 14 | liftemb_def : thm, | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 15 | liftprj_def : thm, | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 16 | liftdefl_def : thm, | 
| 40494 
db8a09daba7b
add class liftdomain, for bifinite domains where DEFL('a u) = u_defl('a)
 huffman parents: 
40491diff
changeset | 17 | DEFL : thm | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 18 | } | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 19 | |
| 40575 | 20 | val add_domaindef: bool -> binding option -> binding * (string * sort) list * mixfix -> | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 21 | term -> (binding * binding) option -> theory -> | 
| 40772 | 22 | (Typedef.info * Cpodef.cpo_info * Cpodef.pcpo_info * rep_info) * theory | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 23 | |
| 40575 | 24 | val domaindef_cmd: (bool * binding) * (binding * (string * string option) list * mixfix) * string | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 25 | * (binding * binding) option -> theory -> theory | 
| 40832 | 26 | end | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 27 | |
| 41296 
6aaf80ea9715
switch to transparent ascription, to avoid warning messages
 huffman parents: 
41292diff
changeset | 28 | structure Domaindef : DOMAINDEF = | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 29 | struct | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 30 | |
| 40832 | 31 | open HOLCF_Library | 
| 35527 | 32 | |
| 40832 | 33 | infixr 6 ->> | 
| 34 | infix -->> | |
| 35527 | 35 | |
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 36 | (** type definitions **) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 37 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 38 | type rep_info = | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 39 |   {
 | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 40 | emb_def : thm, | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 41 | prj_def : thm, | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 42 | defl_def : thm, | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 43 | liftemb_def : thm, | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 44 | liftprj_def : thm, | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 45 | liftdefl_def : thm, | 
| 40494 
db8a09daba7b
add class liftdomain, for bifinite domains where DEFL('a u) = u_defl('a)
 huffman parents: 
40491diff
changeset | 46 | DEFL : thm | 
| 40832 | 47 | } | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 48 | |
| 35527 | 49 | (* building types and terms *) | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 50 | |
| 40832 | 51 | val udomT = @{typ udom}
 | 
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41228diff
changeset | 52 | val deflT = @{typ "udom defl"}
 | 
| 41292 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 huffman parents: 
41290diff
changeset | 53 | val udeflT = @{typ "udom u defl"}
 | 
| 40832 | 54 | fun emb_const T = Const (@{const_name emb}, T ->> udomT)
 | 
| 55 | fun prj_const T = Const (@{const_name prj}, udomT ->> T)
 | |
| 56 | fun defl_const T = Const (@{const_name defl}, Term.itselfT T --> deflT)
 | |
| 41292 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 huffman parents: 
41290diff
changeset | 57 | fun liftemb_const T = Const (@{const_name liftemb}, mk_upT T ->> mk_upT udomT)
 | 
| 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 huffman parents: 
41290diff
changeset | 58 | fun liftprj_const T = Const (@{const_name liftprj}, mk_upT udomT ->> mk_upT T)
 | 
| 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 huffman parents: 
41290diff
changeset | 59 | fun liftdefl_const T = Const (@{const_name liftdefl}, Term.itselfT T --> udeflT)
 | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 60 | |
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 61 | fun mk_u_map t = | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 62 | let | 
| 40832 | 63 | val (T, U) = dest_cfunT (fastype_of t) | 
| 64 | val u_map_type = (T ->> U) ->> (mk_upT T ->> mk_upT U) | |
| 65 |     val u_map_const = Const (@{const_name u_map}, u_map_type)
 | |
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 66 | in | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 67 | mk_capply (u_map_const, t) | 
| 40832 | 68 | end | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 69 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 70 | fun mk_cast (t, x) = | 
| 35527 | 71 | capply_const (udomT, udomT) | 
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41228diff
changeset | 72 |   $ (capply_const (deflT, udomT ->> udomT) $ @{term "cast :: udom defl -> udom -> udom"} $ t)
 | 
| 40832 | 73 | $ x | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 74 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 75 | (* manipulating theorems *) | 
| 
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 | (* proving class instances *) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 78 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 79 | fun declare_type_name a = | 
| 40832 | 80 | Variable.declare_constraints (Logic.mk_type (TFree (a, dummyS))) | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 81 | |
| 40575 | 82 | fun gen_add_domaindef | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 83 | (prep_term: Proof.context -> 'a -> term) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 84 | (def: bool) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 85 | (name: binding) | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 86 | (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 | 87 | (raw_defl: 'a) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 88 | (opt_morphs: (binding * binding) option) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 89 | (thy: theory) | 
| 40772 | 90 | : (Typedef.info * Cpodef.cpo_info * Cpodef.pcpo_info * rep_info) * theory = | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 91 | let | 
| 40832 | 92 | val _ = Theory.requires thy "Domain" "domaindefs" | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 93 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 94 | (*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 | 95 | 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 | 96 | ProofContext.init_global thy | 
| 40832 | 97 | |> fold (Variable.declare_typ o TFree) raw_args | 
| 98 | val defl = prep_term tmp_ctxt raw_defl | |
| 99 | 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 | 100 | |
| 40832 | 101 | val deflT = Term.fastype_of defl | 
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41228diff
changeset | 102 |     val _ = if deflT = @{typ "udom defl"} then ()
 | 
| 40832 | 103 |             else error ("Not type defl: " ^ quote (Syntax.string_of_typ tmp_ctxt deflT))
 | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 104 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 105 | (*lhs*) | 
| 40832 | 106 | val lhs_tfrees = map (ProofContext.check_tfree tmp_ctxt) raw_args | 
| 107 | val lhs_sorts = map snd lhs_tfrees | |
| 108 | val full_tname = Sign.full_name thy tname | |
| 109 | val newT = Type (full_tname, map TFree lhs_tfrees) | |
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 110 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 111 | (*morphisms*) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 112 | val morphs = opt_morphs | 
| 40832 | 113 | |> the_default (Binding.prefix_name "Rep_" name, Binding.prefix_name "Abs_" name) | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 114 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 115 | (*set*) | 
| 41287 
029a6fc1bfb8
type 'defl' takes a type parameter again (cf. b525988432e9)
 huffman parents: 
41228diff
changeset | 116 |     val set = @{term "defl_set :: udom defl => udom => bool"} $ defl
 | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 117 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 118 | (*pcpodef*) | 
| 40832 | 119 |     val tac1 = rtac @{thm defl_set_bottom} 1
 | 
| 120 |     val tac2 = rtac @{thm adm_defl_set} 1
 | |
| 35904 | 121 | val ((info, cpo_info, pcpo_info), thy) = thy | 
| 40832 | 122 | |> Cpodef.add_pcpodef def (SOME name) typ set (SOME morphs) (tac1, tac2) | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 123 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 124 | (*definitions*) | 
| 40832 | 125 | val Rep_const = Const (#Rep_name (#1 info), newT --> udomT) | 
| 126 | val Abs_const = Const (#Abs_name (#1 info), udomT --> newT) | |
| 127 | val emb_eqn = Logic.mk_equals (emb_const newT, cabs_const (newT, udomT) $ Rep_const) | |
| 35527 | 128 | val prj_eqn = Logic.mk_equals (prj_const newT, cabs_const (udomT, newT) $ | 
| 40832 | 129 |       Abs ("x", udomT, Abs_const $ mk_cast (defl, Bound 0)))
 | 
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39986diff
changeset | 130 | val defl_eqn = Logic.mk_equals (defl_const newT, | 
| 40832 | 131 |       Abs ("x", Term.itselfT newT, defl))
 | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 132 | val liftemb_eqn = | 
| 41292 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 huffman parents: 
41290diff
changeset | 133 | Logic.mk_equals (liftemb_const newT, mk_u_map (emb_const newT)) | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 134 | val liftprj_eqn = | 
| 41292 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 huffman parents: 
41290diff
changeset | 135 | Logic.mk_equals (liftprj_const newT, mk_u_map (prj_const newT)) | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 136 | val liftdefl_eqn = | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 137 | Logic.mk_equals (liftdefl_const newT, | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 138 |         Abs ("t", Term.itselfT newT,
 | 
| 41436 | 139 |           mk_capply (@{const liftdefl_of}, defl_const newT $ Logic.mk_type newT)))
 | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 140 | |
| 40832 | 141 | val name_def = Binding.suffix_name "_def" name | 
| 142 | val emb_bind = (Binding.prefix_name "emb_" name_def, []) | |
| 143 | val prj_bind = (Binding.prefix_name "prj_" name_def, []) | |
| 144 | val defl_bind = (Binding.prefix_name "defl_" name_def, []) | |
| 145 | val liftemb_bind = (Binding.prefix_name "liftemb_" name_def, []) | |
| 146 | val liftprj_bind = (Binding.prefix_name "liftprj_" name_def, []) | |
| 147 | val liftdefl_bind = (Binding.prefix_name "liftdefl_" name_def, []) | |
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 148 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 149 | (*instantiate class rep*) | 
| 35904 | 150 | val lthy = thy | 
| 41292 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 huffman parents: 
41290diff
changeset | 151 |       |> Class.instantiation ([full_tname], lhs_tfrees, @{sort domain})
 | 
| 35904 | 152 | val ((_, (_, emb_ldef)), lthy) = | 
| 40832 | 153 | Specification.definition (NONE, (emb_bind, emb_eqn)) lthy | 
| 35904 | 154 | val ((_, (_, prj_ldef)), lthy) = | 
| 40832 | 155 | Specification.definition (NONE, (prj_bind, prj_eqn)) lthy | 
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39986diff
changeset | 156 | val ((_, (_, defl_ldef)), lthy) = | 
| 40832 | 157 | Specification.definition (NONE, (defl_bind, defl_eqn)) lthy | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 158 | val ((_, (_, liftemb_ldef)), lthy) = | 
| 40832 | 159 | Specification.definition (NONE, (liftemb_bind, liftemb_eqn)) lthy | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 160 | val ((_, (_, liftprj_ldef)), lthy) = | 
| 40832 | 161 | Specification.definition (NONE, (liftprj_bind, liftprj_eqn)) lthy | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 162 | val ((_, (_, liftdefl_ldef)), lthy) = | 
| 40832 | 163 | Specification.definition (NONE, (liftdefl_bind, liftdefl_eqn)) lthy | 
| 164 | val ctxt_thy = ProofContext.init_global (ProofContext.theory_of lthy) | |
| 165 | val emb_def = singleton (ProofContext.export lthy ctxt_thy) emb_ldef | |
| 166 | val prj_def = singleton (ProofContext.export lthy ctxt_thy) prj_ldef | |
| 167 | val defl_def = singleton (ProofContext.export lthy ctxt_thy) defl_ldef | |
| 168 | val liftemb_def = singleton (ProofContext.export lthy ctxt_thy) liftemb_ldef | |
| 169 | val liftprj_def = singleton (ProofContext.export lthy ctxt_thy) liftprj_ldef | |
| 170 | val liftdefl_def = singleton (ProofContext.export lthy ctxt_thy) liftdefl_ldef | |
| 33826 | 171 | val type_definition_thm = | 
| 41228 
e1fce873b814
renamed structure MetaSimplifier to raw_Simplifer, to emphasize its meaning;
 wenzelm parents: 
40832diff
changeset | 172 | Raw_Simplifier.rewrite_rule | 
| 35994 
9cc3df9a606e
Typedef.info: separate global and local part, only the latter is transformed by morphisms;
 wenzelm parents: 
35904diff
changeset | 173 | (the_list (#set_def (#2 info))) | 
| 40832 | 174 | (#type_definition (#2 info)) | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 175 | val typedef_thms = | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 176 | [type_definition_thm, #below_def cpo_info, emb_def, prj_def, defl_def, | 
| 40832 | 177 | liftemb_def, liftprj_def, liftdefl_def] | 
| 35904 | 178 | val thy = lthy | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 179 | |> Class.prove_instantiation_instance | 
| 41292 
2b7bc8d9fd6e
use deflations over type 'udom u' to represent predomains;
 huffman parents: 
41290diff
changeset | 180 |           (K (Tactic.rtac (@{thm typedef_domain_class} OF typedef_thms) 1))
 | 
| 40832 | 181 | |> Local_Theory.exit_global | 
| 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 | (*other theorems*) | 
| 40832 | 184 | val defl_thm' = Thm.transfer thy defl_def | 
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39986diff
changeset | 185 | val (DEFL_thm, thy) = thy | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 186 | |> Sign.add_path (Binding.name_of name) | 
| 39557 
fe5722fce758
renamed structure PureThy to Pure_Thy and moved most content to Global_Theory, to emphasize that this is global-only;
 wenzelm parents: 
38348diff
changeset | 187 | |> Global_Theory.add_thm | 
| 39989 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39986diff
changeset | 188 | ((Binding.prefix_name "DEFL_" name, | 
| 
ad60d7311f43
renamed type and constant 'sfp' to 'defl'; replaced syntax SFP('a) with DEFL('a)
 huffman parents: 
39986diff
changeset | 189 |           Drule.zero_var_indexes (@{thm typedef_DEFL} OF [defl_thm'])), [])
 | 
| 40832 | 190 | ||> Sign.restore_naming thy | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 191 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 192 | val rep_info = | 
| 40491 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 193 |       { emb_def = emb_def, prj_def = prj_def, defl_def = defl_def,
 | 
| 
6de5839e2fb3
add 'predomain' class: unpointed version of bifinite
 huffman parents: 
40038diff
changeset | 194 | liftemb_def = liftemb_def, liftprj_def = liftprj_def, | 
| 40832 | 195 | liftdefl_def = liftdefl_def, DEFL = DEFL_thm } | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 196 | in | 
| 35904 | 197 | ((info, cpo_info, pcpo_info, rep_info), thy) | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 198 | end | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 199 | handle ERROR msg => | 
| 40832 | 200 |     cat_error msg ("The error(s) above occurred in domaindef " ^ quote (Binding.str_of name))
 | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 201 | |
| 40575 | 202 | fun add_domaindef def opt_name typ defl opt_morphs thy = | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 203 | let | 
| 40832 | 204 | val name = the_default (#1 typ) opt_name | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 205 | in | 
| 40575 | 206 | gen_add_domaindef Syntax.check_term def name typ defl opt_morphs thy | 
| 40832 | 207 | end | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 208 | |
| 40575 | 209 | fun domaindef_cmd ((def, name), (b, raw_args, mx), A, morphs) thy = | 
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 210 | let | 
| 40832 | 211 | val ctxt = ProofContext.init_global thy | 
| 212 | val args = map (apsnd (Typedecl.read_constraint ctxt)) raw_args | |
| 213 | in snd (gen_add_domaindef Syntax.read_term def name (b, args, mx) A morphs thy) end | |
| 35840 
01d7c4ba9050
allow sort constraints in HOL/typedef and related HOLCF variants;
 wenzelm parents: 
35527diff
changeset | 214 | |
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 215 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 216 | (** outer syntax **) | 
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 217 | |
| 40575 | 218 | val domaindef_decl = | 
| 36960 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 219 |   Scan.optional (Parse.$$$ "(" |--
 | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 220 | ((Parse.$$$ "open" >> K false) -- Scan.option Parse.binding || | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 221 | Parse.binding >> (fn s => (true, SOME s))) | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 222 | --| Parse.$$$ ")") (true, NONE) -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 223 | (Parse.type_args_constrained -- Parse.binding) -- | 
| 
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
 wenzelm parents: 
36610diff
changeset | 224 | Parse.opt_mixfix -- (Parse.$$$ "=" |-- Parse.term) -- | 
| 40832 | 225 | Scan.option (Parse.$$$ "morphisms" |-- Parse.!!! (Parse.binding -- Parse.binding)) | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 226 | |
| 40575 | 227 | fun mk_domaindef ((((((def, opt_name), (args, t)), mx), A), morphs)) = | 
| 40832 | 228 | domaindef_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 | 229 | |
| 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 230 | val _ = | 
| 40575 | 231 | Outer_Syntax.command "domaindef" "HOLCF definition of domains from deflations" Keyword.thy_decl | 
| 232 | (domaindef_decl >> | |
| 40832 | 233 | (Toplevel.print oo (Toplevel.theory o mk_domaindef))) | 
| 33679 
331712879666
automate definition of representable domains from algebraic deflations
 huffman parents: diff
changeset | 234 | |
| 40832 | 235 | end |