author | bulwahn |
Thu, 29 Jul 2010 17:27:51 +0200 | |
changeset 38072 | 7b8c295af291 |
parent 36960 | 01594f816e3a |
child 39557 | fe5722fce758 |
permissions | -rw-r--r-- |
32126 | 1 |
(* Title: HOLCF/Tools/Domain/domain_extender.ML |
23152 | 2 |
Author: David von Oheimb |
35794 | 3 |
Author: Brian Huffman |
23152 | 4 |
|
5 |
Theory extender for domain command, including theory syntax. |
|
6 |
*) |
|
7 |
||
8 |
signature DOMAIN_EXTENDER = |
|
9 |
sig |
|
33796 | 10 |
val add_domain_cmd: |
35774 | 11 |
binding -> |
33796 | 12 |
((string * string option) list * binding * mixfix * |
13 |
(binding * (bool * binding option * string) list * mixfix) list) list |
|
14 |
-> theory -> theory |
|
15 |
||
16 |
val add_domain: |
|
35774 | 17 |
binding -> |
33796 | 18 |
((string * string option) list * binding * mixfix * |
19 |
(binding * (bool * binding option * typ) list * mixfix) list) list |
|
20 |
-> theory -> theory |
|
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
21 |
|
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
22 |
val add_new_domain_cmd: |
35774 | 23 |
binding -> |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
24 |
((string * string option) list * binding * mixfix * |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
25 |
(binding * (bool * binding option * string) list * mixfix) list) list |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
26 |
-> theory -> theory |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
27 |
|
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
28 |
val add_new_domain: |
35774 | 29 |
binding -> |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
30 |
((string * string option) list * binding * mixfix * |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
31 |
(binding * (bool * binding option * typ) list * mixfix) list) list |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
32 |
-> theory -> theory |
23152 | 33 |
end; |
34 |
||
31023 | 35 |
structure Domain_Extender :> DOMAIN_EXTENDER = |
23152 | 36 |
struct |
37 |
||
38 |
open Domain_Library; |
|
39 |
||
40 |
(* ----- general testing and preprocessing of constructor list -------------- *) |
|
30919
dcf8a7a66bd1
make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents:
30916
diff
changeset
|
41 |
fun check_and_sort_domain |
33796 | 42 |
(dtnvs : (string * typ list) list) |
43 |
(cons'' : (binding * (bool * binding option * typ) list * mixfix) list list) |
|
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
44 |
(thy : theory) |
31288 | 45 |
: ((string * typ list) * |
46 |
(binding * (bool * binding option * typ) list * mixfix) list) list = |
|
33796 | 47 |
let |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
48 |
val defaultS = Sign.defaultS thy; |
33796 | 49 |
|
50 |
val test_dupl_typs = |
|
51 |
case duplicates (op =) (map fst dtnvs) of |
|
52 |
[] => false | dups => error ("Duplicate types: " ^ commas_quote dups); |
|
53 |
||
54 |
val all_cons = map (Binding.name_of o first) (flat cons''); |
|
55 |
val test_dupl_cons = |
|
56 |
case duplicates (op =) all_cons of |
|
57 |
[] => false | dups => error ("Duplicate constructors: " |
|
58 |
^ commas_quote dups); |
|
59 |
val all_sels = |
|
60 |
(map Binding.name_of o map_filter second o maps second) (flat cons''); |
|
61 |
val test_dupl_sels = |
|
62 |
case duplicates (op =) all_sels of |
|
63 |
[] => false | dups => error("Duplicate selectors: "^commas_quote dups); |
|
64 |
||
65 |
fun test_dupl_tvars s = |
|
66 |
case duplicates (op =) (map(fst o dest_TFree)s) of |
|
67 |
[] => false | dups => error("Duplicate type arguments: " |
|
68 |
^commas_quote dups); |
|
69 |
val test_dupl_tvars' = exists test_dupl_tvars (map snd dtnvs); |
|
70 |
||
71 |
(* test for free type variables, illegal sort constraints on rhs, |
|
72 |
non-pcpo-types and invalid use of recursive type; |
|
73 |
replace sorts in type variables on rhs *) |
|
74 |
fun analyse_equation ((dname,typevars),cons') = |
|
75 |
let |
|
76 |
val tvars = map dest_TFree typevars; |
|
77 |
val distinct_typevars = map TFree tvars; |
|
78 |
fun rm_sorts (TFree(s,_)) = TFree(s,[]) |
|
79 |
| rm_sorts (Type(s,ts)) = Type(s,remove_sorts ts) |
|
80 |
| rm_sorts (TVar(s,_)) = TVar(s,[]) |
|
81 |
and remove_sorts l = map rm_sorts l; |
|
82 |
fun analyse indirect (TFree(v,s)) = |
|
83 |
(case AList.lookup (op =) tvars v of |
|
84 |
NONE => error ("Free type variable " ^ quote v ^ " on rhs.") |
|
85 |
| SOME sort => if eq_set (op =) (s, defaultS) orelse |
|
86 |
eq_set (op =) (s, sort) |
|
87 |
then TFree(v,sort) |
|
88 |
else error ("Inconsistent sort constraint" ^ |
|
89 |
" for type variable " ^ quote v)) |
|
90 |
| analyse indirect (t as Type(s,typl)) = |
|
91 |
(case AList.lookup (op =) dtnvs s of |
|
36119 | 92 |
NONE => Type (s, map (analyse false) typl) |
33796 | 93 |
| SOME typevars => |
94 |
if indirect |
|
95 |
then error ("Indirect recursion of type " ^ |
|
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
96 |
quote (string_of_typ thy t)) |
33796 | 97 |
else if dname <> s orelse |
98 |
(** BUG OR FEATURE?: |
|
99 |
mutual recursion may use different arguments **) |
|
100 |
remove_sorts typevars = remove_sorts typl |
|
101 |
then Type(s,map (analyse true) typl) |
|
102 |
else error ("Direct recursion of type " ^ |
|
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
103 |
quote (string_of_typ thy t) ^ |
33796 | 104 |
" with different arguments")) |
105 |
| analyse indirect (TVar _) = Imposs "extender:analyse"; |
|
106 |
fun check_pcpo lazy T = |
|
107 |
let val ok = if lazy then cpo_type else pcpo_type |
|
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
108 |
in if ok thy T then T |
33796 | 109 |
else error ("Constructor argument type is not of sort pcpo: " ^ |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
110 |
string_of_typ thy T) |
33796 | 111 |
end; |
112 |
fun analyse_arg (lazy, sel, T) = |
|
113 |
(lazy, sel, check_pcpo lazy (analyse false T)); |
|
114 |
fun analyse_con (b, args, mx) = (b, map analyse_arg args, mx); |
|
115 |
in ((dname,distinct_typevars), map analyse_con cons') end; |
|
116 |
in ListPair.map analyse_equation (dtnvs,cons'') |
|
117 |
end; (* let *) |
|
23152 | 118 |
|
119 |
(* ----- calls for building new thy and thms -------------------------------- *) |
|
120 |
||
36118
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
121 |
type info = |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
122 |
Domain_Take_Proofs.iso_info list * Domain_Take_Proofs.take_induct_info; |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
123 |
|
30919
dcf8a7a66bd1
make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents:
30916
diff
changeset
|
124 |
fun gen_add_domain |
33796 | 125 |
(prep_typ : theory -> 'a -> typ) |
36118
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
126 |
(add_isos : (binding * mixfix * (typ * typ)) list -> theory -> info * theory) |
35774 | 127 |
(comp_dbind : binding) |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
128 |
(eqs''' : ((string * string option) list * binding * mixfix * |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
129 |
(binding * (bool * binding option * 'a) list * mixfix) list) list) |
35516 | 130 |
(thy : theory) = |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
131 |
let |
35520 | 132 |
val dtnvs : (binding * typ list * mixfix) list = |
133 |
let |
|
134 |
fun readS (SOME s) = Syntax.read_sort_global thy s |
|
135 |
| readS NONE = Sign.defaultS thy; |
|
136 |
fun readTFree (a, s) = TFree (a, readS s); |
|
137 |
in |
|
138 |
map (fn (vs,dname:binding,mx,_) => |
|
139 |
(dname, map readTFree vs, mx)) eqs''' |
|
140 |
end; |
|
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
141 |
|
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
142 |
fun thy_type (dname,tvars,mx) = (dname, length tvars, mx); |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
143 |
fun thy_arity (dname,tvars,mx) = |
35516 | 144 |
(Sign.full_name thy dname, map (snd o dest_TFree) tvars, @{sort rep}); |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
145 |
|
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
146 |
(* this theory is used just for parsing and error checking *) |
35516 | 147 |
val tmp_thy = thy |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
148 |
|> Theory.copy |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
149 |
|> Sign.add_types (map thy_type dtnvs) |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
150 |
|> fold (AxClass.axiomatize_arity o thy_arity) dtnvs; |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
151 |
|
35776
b0bc15d8ad3b
pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents:
35775
diff
changeset
|
152 |
val dbinds : binding list = |
b0bc15d8ad3b
pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents:
35775
diff
changeset
|
153 |
map (fn (_,dbind,_,_) => dbind) eqs'''; |
35520 | 154 |
val cons''' : |
155 |
(binding * (bool * binding option * 'a) list * mixfix) list list = |
|
156 |
map (fn (_,_,_,cons) => cons) eqs'''; |
|
157 |
val cons'' : |
|
158 |
(binding * (bool * binding option * typ) list * mixfix) list list = |
|
159 |
map (map (upd_second (map (upd_third (prep_typ tmp_thy))))) cons'''; |
|
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
160 |
val dtnvs' : (string * typ list) list = |
35520 | 161 |
map (fn (dname,vs,mx) => (Sign.full_name thy dname,vs)) dtnvs; |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
162 |
val eqs' : ((string * typ list) * |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
163 |
(binding * (bool * binding option * typ) list * mixfix) list) list = |
36119 | 164 |
check_and_sort_domain dtnvs' cons'' tmp_thy; |
36118
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
165 |
val dts : typ list = map (Type o fst) eqs'; |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
166 |
|
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
167 |
fun mk_arg_typ (lazy, dest_opt, T) = if lazy then mk_uT T else T; |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
168 |
fun mk_con_typ (bind, args, mx) = |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
169 |
if null args then oneT else foldr1 mk_sprodT (map mk_arg_typ args); |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
170 |
fun mk_eq_typ (_, cons) = foldr1 mk_ssumT (map mk_con_typ cons); |
36118
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
171 |
val repTs : typ list = map mk_eq_typ eqs'; |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
172 |
|
36118
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
173 |
val iso_spec : (binding * mixfix * (typ * typ)) list = |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
174 |
map (fn ((dbind, _, mx), eq) => (dbind, mx, eq)) |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
175 |
(dtnvs ~~ (dts ~~ repTs)); |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
176 |
|
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
177 |
val ((iso_infos, take_info), thy) = add_isos iso_spec thy; |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
178 |
|
35520 | 179 |
val new_dts : (string * string list) list = |
180 |
map (fn ((s,Ts),_) => (s, map (fst o dest_TFree) Ts)) eqs'; |
|
181 |
fun one_con (con,args,mx) : cons = |
|
36118
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
182 |
(Binding.name_of con, (* FIXME preverse binding (!?) *) |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
183 |
ListPair.map (fn ((lazy,sel,tp),vn) => |
35519 | 184 |
mk_arg ((lazy, Datatype_Aux.dtyp_of_typ new_dts tp), vn)) |
36118
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
185 |
(args, Datatype_Prop.make_tnames (map third args))); |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
186 |
val eqs : eq list = |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
187 |
map (fn (dtnvs,cons') => (dtnvs, map one_con cons')) eqs'; |
36118
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
188 |
|
33796 | 189 |
val ((rewss, take_rews), theorems_thy) = |
190 |
thy |
|
36118
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
191 |
|> fold_map (fn (((dbind, eq), (_,cs)), info) => |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
192 |
Domain_Theorems.theorems (eq, eqs) dbind cs info take_info) |
35776
b0bc15d8ad3b
pass domain binding as argument to Domain_Theorems.theorems; proper qualified bindings for theorem names
huffman
parents:
35775
diff
changeset
|
193 |
(dbinds ~~ eqs ~~ eqs' ~~ iso_infos) |
35774 | 194 |
||>> Domain_Theorems.comp_theorems (comp_dbind, eqs) take_info; |
33796 | 195 |
in |
196 |
theorems_thy |
|
197 |
|> PureThy.add_thmss |
|
35774 | 198 |
[((Binding.qualified true "rews" comp_dbind, |
199 |
flat rewss @ take_rews), [])] |
|
33796 | 200 |
|> snd |
201 |
end; |
|
23152 | 202 |
|
36118
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
203 |
fun define_isos (spec : (binding * mixfix * (typ * typ)) list) = |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
204 |
let |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
205 |
fun prep (dbind, mx, (lhsT, rhsT)) = |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
206 |
let val (dname, vs) = dest_Type lhsT; |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
207 |
in (map (fst o dest_TFree) vs, dbind, mx, rhsT, NONE) end; |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
208 |
in |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
209 |
Domain_Isomorphism.domain_isomorphism (map prep spec) |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
210 |
end; |
23152 | 211 |
|
36118
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
212 |
val add_domain = |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
213 |
gen_add_domain Sign.certify_typ Domain_Axioms.add_axioms; |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
214 |
|
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
215 |
val add_new_domain = |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
216 |
gen_add_domain Sign.certify_typ define_isos; |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
217 |
|
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
218 |
val add_domain_cmd = |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
219 |
gen_add_domain Syntax.read_typ_global Domain_Axioms.add_axioms; |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
220 |
|
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
221 |
val add_new_domain_cmd = |
413d6d1f0f6e
share more code between definitional and axiomatic domain packages
huffman
parents:
36117
diff
changeset
|
222 |
gen_add_domain Syntax.read_typ_global define_isos; |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
223 |
|
23152 | 224 |
|
225 |
(** outer syntax **) |
|
226 |
||
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
227 |
val _ = Keyword.keyword "lazy"; |
24867 | 228 |
|
30919
dcf8a7a66bd1
make domain package ML interface more consistent with datatype package; use binding instead of bstring
huffman
parents:
30916
diff
changeset
|
229 |
val dest_decl : (bool * binding option * string) parser = |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
230 |
Parse.$$$ "(" |-- Scan.optional (Parse.$$$ "lazy" >> K true) false -- |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
231 |
(Parse.binding >> SOME) -- (Parse.$$$ "::" |-- Parse.typ) --| Parse.$$$ ")" >> Parse.triple1 |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
232 |
|| Parse.$$$ "(" |-- Parse.$$$ "lazy" |-- Parse.typ --| Parse.$$$ ")" |
33796 | 233 |
>> (fn t => (true,NONE,t)) |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
234 |
|| Parse.typ >> (fn t => (false,NONE,t)); |
23152 | 235 |
|
236 |
val cons_decl = |
|
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
237 |
Parse.binding -- Scan.repeat dest_decl -- Parse.opt_mixfix; |
30916
a3d2128cac92
allow infix declarations for type constructors defined with domain package
huffman
parents:
30915
diff
changeset
|
238 |
|
a3d2128cac92
allow infix declarations for type constructors defined with domain package
huffman
parents:
30915
diff
changeset
|
239 |
val domain_decl = |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
240 |
(Parse.type_args_constrained -- Parse.binding -- Parse.opt_mixfix) -- |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
241 |
(Parse.$$$ "=" |-- Parse.enum1 "|" cons_decl); |
23152 | 242 |
|
30916
a3d2128cac92
allow infix declarations for type constructors defined with domain package
huffman
parents:
30915
diff
changeset
|
243 |
val domains_decl = |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
244 |
Scan.option (Parse.$$$ "(" |-- Parse.binding --| Parse.$$$ ")") -- |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
245 |
Parse.and_list1 domain_decl; |
23152 | 246 |
|
33796 | 247 |
fun mk_domain |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
248 |
(definitional : bool) |
35774 | 249 |
(opt_name : binding option, |
33796 | 250 |
doms : ((((string * string option) list * binding) * mixfix) * |
251 |
((binding * (bool * binding option * string) list) * mixfix) list) list ) = |
|
252 |
let |
|
253 |
val names = map (fn (((_, t), _), _) => Binding.name_of t) doms; |
|
254 |
val specs : ((string * string option) list * binding * mixfix * |
|
255 |
(binding * (bool * binding option * string) list * mixfix) list) list = |
|
256 |
map (fn (((vs, t), mx), cons) => |
|
257 |
(vs, t, mx, map (fn ((c, ds), mx) => (c, ds, mx)) cons)) doms; |
|
35774 | 258 |
val comp_dbind = |
259 |
case opt_name of NONE => Binding.name (space_implode "_" names) |
|
260 |
| SOME s => s; |
|
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
261 |
in |
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
262 |
if definitional |
35774 | 263 |
then add_new_domain_cmd comp_dbind specs |
264 |
else add_domain_cmd comp_dbind specs |
|
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
265 |
end; |
23152 | 266 |
|
24867 | 267 |
val _ = |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
268 |
Outer_Syntax.command "domain" "define recursive domains (HOLCF)" |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
269 |
Keyword.thy_decl (domains_decl >> (Toplevel.theory o mk_domain false)); |
33798
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
270 |
|
46cbbcbd4e68
clean up indentation; add 'definitional' option flag
huffman
parents:
33796
diff
changeset
|
271 |
val _ = |
36960
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
272 |
Outer_Syntax.command "new_domain" "define recursive domains (HOLCF)" |
01594f816e3a
prefer structure Keyword, Parse, Parse_Spec, Outer_Syntax;
wenzelm
parents:
36119
diff
changeset
|
273 |
Keyword.thy_decl (domains_decl >> (Toplevel.theory o mk_domain true)); |
23152 | 274 |
|
24867 | 275 |
end; |