| author | wenzelm | 
| Sun, 03 Jan 2016 21:31:57 +0100 | |
| changeset 62047 | 1ae53588dcbb | 
| parent 61877 | 276ad4354069 | 
| child 62179 | e089e5b02443 | 
| permissions | -rw-r--r-- | 
| 17707 
bc0270e9d27f
back to simple 'defs' (cf. revision 1.79 of theory.ML);
 wenzelm parents: 
17670diff
changeset | 1 | (* Title: Pure/defs.ML | 
| 
bc0270e9d27f
back to simple 'defs' (cf. revision 1.79 of theory.ML);
 wenzelm parents: 
17670diff
changeset | 2 | Author: Makarius | 
| 16108 
cf468b93a02e
Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
 obua parents: diff
changeset | 3 | |
| 19692 | 4 | Global well-formedness checks for constant definitions. Covers plain | 
| 19701 | 5 | definitions and simple sub-structural overloading. | 
| 16108 
cf468b93a02e
Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
 obua parents: diff
changeset | 6 | *) | 
| 
cf468b93a02e
Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
 obua parents: diff
changeset | 7 | |
| 16877 
e92cba1d4842
tuned interfaces declare, define, finalize, merge:
 wenzelm parents: 
16838diff
changeset | 8 | signature DEFS = | 
| 
e92cba1d4842
tuned interfaces declare, define, finalize, merge:
 wenzelm parents: 
16838diff
changeset | 9 | sig | 
| 61256 | 10 | datatype item_kind = Const | Type | 
| 61249 | 11 | type item = item_kind * string | 
| 61254 | 12 | type entry = item * typ list | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 13 | val item_kind_ord: item_kind * item_kind -> order | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 14 | val plain_args: typ list -> bool | 
| 61265 | 15 | type context = Proof.context * (Name_Space.T * Name_Space.T) | 
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61261diff
changeset | 16 | val global_context: theory -> context | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 17 | val space: context -> item_kind -> Name_Space.T | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 18 | val pretty_item: context -> item -> Pretty.T | 
| 61253 | 19 | val pretty_args: Proof.context -> typ list -> Pretty.T list | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 20 | val pretty_entry: context -> entry -> Pretty.T | 
| 17707 
bc0270e9d27f
back to simple 'defs' (cf. revision 1.79 of theory.ML);
 wenzelm parents: 
17670diff
changeset | 21 | type T | 
| 33712 | 22 | type spec = | 
| 55544 | 23 |    {def: string option,
 | 
| 24 | description: string, | |
| 25 | pos: Position.T, | |
| 26 | lhs: typ list, | |
| 61254 | 27 | rhs: entry list} | 
| 61249 | 28 | val all_specifications_of: T -> (item * spec list) list | 
| 29 | val specifications_of: T -> item -> spec list | |
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 30 | val dest: T -> | 
| 61254 | 31 |    {restricts: (entry * string) list,
 | 
| 32 | reducts: (entry * entry list) list} | |
| 19590 
12af4942923d
simple substructure test for typargs (independent type constructors);
 wenzelm parents: 
19569diff
changeset | 33 | val empty: T | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 34 | val merge: context -> T * T -> T | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 35 | val define: context -> bool -> string option -> string -> entry -> entry list -> T -> T | 
| 61260 | 36 | val get_deps: T -> item -> (typ list * entry list) list | 
| 61249 | 37 | end; | 
| 61246 
077b88f9ec16
HOL typedef with explicit dependency checks according to Ondrey Kuncar, 07-Jul-2015, 16-Jul-2015, 30-Jul-2015;
 wenzelm parents: 
59050diff
changeset | 38 | |
| 17711 | 39 | structure Defs: DEFS = | 
| 17707 
bc0270e9d27f
back to simple 'defs' (cf. revision 1.79 of theory.ML);
 wenzelm parents: 
17670diff
changeset | 40 | struct | 
| 16108 
cf468b93a02e
Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
 obua parents: diff
changeset | 41 | |
| 61249 | 42 | (* specification items *) | 
| 43 | ||
| 61256 | 44 | datatype item_kind = Const | Type; | 
| 61249 | 45 | type item = item_kind * string; | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 46 | type entry = item * typ list; | 
| 61249 | 47 | |
| 61256 | 48 | fun item_kind_ord (Const, Type) = LESS | 
| 49 | | item_kind_ord (Type, Const) = GREATER | |
| 61249 | 50 | | item_kind_ord _ = EQUAL; | 
| 61246 
077b88f9ec16
HOL typedef with explicit dependency checks according to Ondrey Kuncar, 07-Jul-2015, 16-Jul-2015, 30-Jul-2015;
 wenzelm parents: 
59050diff
changeset | 51 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 52 | structure Itemtab = Table(type key = item val ord = prod_ord item_kind_ord fast_string_ord); | 
| 61249 | 53 | |
| 61246 
077b88f9ec16
HOL typedef with explicit dependency checks according to Ondrey Kuncar, 07-Jul-2015, 16-Jul-2015, 30-Jul-2015;
 wenzelm parents: 
59050diff
changeset | 54 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 55 | (* pretty printing *) | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 56 | |
| 61265 | 57 | type context = Proof.context * (Name_Space.T * Name_Space.T) | 
| 61262 
7bd1eb4b056e
tuned signature: eliminated pointless type Context.pretty;
 wenzelm parents: 
61261diff
changeset | 58 | |
| 61265 | 59 | fun global_context thy = | 
| 60 | (Syntax.init_pretty_global thy, (Sign.const_space thy, Sign.type_space thy)); | |
| 61 | ||
| 62 | fun space ((ctxt, spaces): context) kind = | |
| 63 | if kind = Const then #1 spaces else #2 spaces; | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 64 | |
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 65 | fun pretty_item (context as (ctxt, _)) (kind, name) = | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 66 | let val prt_name = Name_Space.pretty ctxt (space context kind) name in | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 67 | if kind = Const then prt_name | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 68 | else Pretty.block [Pretty.keyword1 "type", Pretty.brk 1, prt_name] | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 69 | end; | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 70 | |
| 61253 | 71 | fun pretty_args ctxt args = | 
| 72 | if null args then [] | |
| 73 |   else [Pretty.list "(" ")" (map (Syntax.pretty_typ ctxt o Logic.unvarifyT_global) args)];
 | |
| 74 | ||
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 75 | fun pretty_entry context (c, args) = | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 76 | Pretty.block (pretty_item context c :: pretty_args (#1 context) args); | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 77 | |
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 78 | |
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 79 | (* type arguments *) | 
| 19624 | 80 | |
| 19707 | 81 | fun plain_args args = | 
| 82 | forall Term.is_TVar args andalso not (has_duplicates (op =) args); | |
| 83 | ||
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 84 | fun disjoint_args (Ts, Us) = | 
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 85 | not (Type.could_unifys (Ts, Us)) orelse | 
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 86 | ((Type.raw_unifys (Ts, map (Logic.incr_tvar (maxidx_of_typs Ts + 1)) Us) Vartab.empty; false) | 
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 87 | handle Type.TUNIFY => true); | 
| 19692 | 88 | |
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 89 | fun match_args (Ts, Us) = | 
| 56050 
fdccbb97915a
minor performance tuning via fast matching filter;
 wenzelm parents: 
55544diff
changeset | 90 | if Type.could_matches (Ts, Us) then | 
| 
fdccbb97915a
minor performance tuning via fast matching filter;
 wenzelm parents: 
55544diff
changeset | 91 | Option.map Envir.subst_type | 
| 
fdccbb97915a
minor performance tuning via fast matching filter;
 wenzelm parents: 
55544diff
changeset | 92 | (SOME (Type.raw_matches (Ts, Us) Vartab.empty) handle Type.TYPE_MATCH => NONE) | 
| 
fdccbb97915a
minor performance tuning via fast matching filter;
 wenzelm parents: 
55544diff
changeset | 93 | else NONE; | 
| 19692 | 94 | |
| 95 | ||
| 96 | (* datatype defs *) | |
| 97 | ||
| 33701 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 wenzelm parents: 
32785diff
changeset | 98 | type spec = | 
| 55544 | 99 |  {def: string option,
 | 
| 100 | description: string, | |
| 101 | pos: Position.T, | |
| 61254 | 102 | lhs: typ list, | 
| 103 | rhs: entry list}; | |
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 104 | |
| 19692 | 105 | type def = | 
| 55544 | 106 |  {specs: spec Inttab.table,  (*source specifications*)
 | 
| 61254 | 107 | restricts: (typ list * string) list, (*global restrictions imposed by incomplete patterns*) | 
| 108 | reducts: (typ list * entry list) list}; (*specifications as reduction system*) | |
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 109 | |
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 110 | fun make_def (specs, restricts, reducts) = | 
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 111 |   {specs = specs, restricts = restricts, reducts = reducts}: def;
 | 
| 19692 | 112 | |
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 113 | fun map_def c f = | 
| 61249 | 114 | Itemtab.default (c, make_def (Inttab.empty, [], [])) #> | 
| 115 |   Itemtab.map_entry c (fn {specs, restricts, reducts}: def =>
 | |
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 116 | make_def (f (specs, restricts, reducts))); | 
| 19692 | 117 | |
| 118 | ||
| 61249 | 119 | datatype T = Defs of def Itemtab.table; | 
| 19692 | 120 | |
| 19712 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 wenzelm parents: 
19707diff
changeset | 121 | fun lookup_list which defs c = | 
| 61249 | 122 | (case Itemtab.lookup defs c of | 
| 19713 | 123 | SOME (def: def) => which def | 
| 19692 | 124 | | NONE => []); | 
| 125 | ||
| 32050 | 126 | fun all_specifications_of (Defs defs) = | 
| 61249 | 127 | (map o apsnd) (map snd o Inttab.dest o #specs) (Itemtab.dest defs); | 
| 32050 | 128 | |
| 24199 | 129 | fun specifications_of (Defs defs) = lookup_list (map snd o Inttab.dest o #specs) defs; | 
| 32050 | 130 | |
| 19692 | 131 | val restricts_of = lookup_list #restricts; | 
| 132 | val reducts_of = lookup_list #reducts; | |
| 133 | ||
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 134 | fun dest (Defs defs) = | 
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 135 | let | 
| 61249 | 136 |     val restricts = Itemtab.fold (fn (c, {restricts, ...}) =>
 | 
| 33701 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 wenzelm parents: 
32785diff
changeset | 137 | fold (fn (args, description) => cons ((c, args), description)) restricts) defs []; | 
| 61249 | 138 |     val reducts = Itemtab.fold (fn (c, {reducts, ...}) =>
 | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 139 | fold (fn (args, deps) => cons ((c, args), deps)) reducts) defs []; | 
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 140 |   in {restricts = restricts, reducts = reducts} end;
 | 
| 19692 | 141 | |
| 61249 | 142 | val empty = Defs Itemtab.empty; | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 143 | |
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 144 | |
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 145 | (* specifications *) | 
| 19692 | 146 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 147 | fun disjoint_specs context c (i, {description = a, pos = pos_a, lhs = Ts, ...}: spec) =
 | 
| 55544 | 148 |   Inttab.forall (fn (j, {description = b, pos = pos_b, lhs = Us, ...}: spec) =>
 | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 149 | i = j orelse disjoint_args (Ts, Us) orelse | 
| 61877 
276ad4354069
renamed Pretty.str_of to Pretty.unformatted_string_of to emphasize its meaning;
 wenzelm parents: 
61265diff
changeset | 150 |       error ("Clash of specifications for " ^
 | 
| 
276ad4354069
renamed Pretty.str_of to Pretty.unformatted_string_of to emphasize its meaning;
 wenzelm parents: 
61265diff
changeset | 151 | Pretty.unformatted_string_of (pretty_item context c) ^ ":\n" ^ | 
| 55544 | 152 | " " ^ quote a ^ Position.here pos_a ^ "\n" ^ | 
| 153 | " " ^ quote b ^ Position.here pos_b)); | |
| 19692 | 154 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 155 | fun join_specs context c ({specs = specs1, restricts, reducts}, {specs = specs2, ...}: def) =
 | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 156 | let | 
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 157 | val specs' = | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 158 | Inttab.fold (fn spec2 => (disjoint_specs context c spec2 specs1; Inttab.update spec2)) | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 159 | specs2 specs1; | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 160 | in make_def (specs', restricts, reducts) end; | 
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 161 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 162 | fun update_specs context c spec = map_def c (fn (specs, restricts, reducts) => | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 163 | (disjoint_specs context c spec specs; (Inttab.update spec specs, restricts, reducts))); | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 164 | |
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 165 | |
| 19701 | 166 | (* normalized dependencies: reduction with well-formedness check *) | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 167 | |
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 168 | local | 
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 169 | |
| 61253 | 170 | val prt = Pretty.string_of oo pretty_entry; | 
| 19729 | 171 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 172 | fun err context (c, args) (d, Us) s1 s2 = | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 173 | error (s1 ^ " dependency of " ^ prt context (c, args) ^ " -> " ^ prt context (d, Us) ^ s2); | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 174 | |
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 175 | fun acyclic context (c, args) (d, Us) = | 
| 19729 | 176 | c <> d orelse | 
| 177 | is_none (match_args (args, Us)) orelse | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 178 | err context (c, args) (d, Us) "Circular" ""; | 
| 19729 | 179 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 180 | fun wellformed context defs (c, args) (d, Us) = | 
| 57520 
3ad1b289f21b
proper plain_args to ensure that multi-argument overloading cannot escape pattern restriction (despite more liberal structural containment before 3ae3cc4b1eac);
 wenzelm parents: 
56050diff
changeset | 181 | plain_args Us orelse | 
| 19729 | 182 | (case find_first (fn (Ts, _) => not (disjoint_args (Ts, Us))) (restricts_of defs d) of | 
| 33701 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 wenzelm parents: 
32785diff
changeset | 183 | SOME (Ts, description) => | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 184 | err context (c, args) (d, Us) "Malformed" | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 185 |         ("\n(restriction " ^ prt context (d, Ts) ^ " from " ^ quote description ^ ")")
 | 
| 19729 | 186 | | NONE => true); | 
| 19692 | 187 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 188 | fun reduction context defs const deps = | 
| 19692 | 189 | let | 
| 19701 | 190 | fun reduct Us (Ts, rhs) = | 
| 191 | (case match_args (Ts, Us) of | |
| 192 | NONE => NONE | |
| 193 | | SOME subst => SOME (map (apsnd (map subst)) rhs)); | |
| 194 | fun reducts (d, Us) = get_first (reduct Us) (reducts_of defs d); | |
| 195 | ||
| 196 | val reds = map (`reducts) deps; | |
| 197 | val deps' = | |
| 198 | if forall (is_none o #1) reds then NONE | |
| 20668 | 199 | else SOME (fold_rev | 
| 200 | (fn (NONE, dp) => insert (op =) dp | (SOME dps, _) => fold (insert (op =)) dps) reds []); | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 201 | val _ = forall (acyclic context const) (the_default deps deps'); | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 202 | in deps' end; | 
| 19692 | 203 | |
| 19760 | 204 | in | 
| 205 | ||
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 206 | fun normalize context = | 
| 19692 | 207 | let | 
| 19701 | 208 |     fun norm_update (c, {reducts, ...}: def) (changed, defs) =
 | 
| 209 | let | |
| 210 | val reducts' = reducts |> map (fn (args, deps) => | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 211 | (args, perhaps (reduction context defs (c, args)) deps)); | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 212 | in | 
| 19701 | 213 | if reducts = reducts' then (changed, defs) | 
| 32785 | 214 | else (true, defs |> map_def c (fn (specs, restricts, _) => (specs, restricts, reducts'))) | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 215 | end; | 
| 19701 | 216 | fun norm_all defs = | 
| 61249 | 217 | (case Itemtab.fold norm_update defs (false, defs) of | 
| 19701 | 218 | (true, defs') => norm_all defs' | 
| 219 | | (false, _) => defs); | |
| 19729 | 220 |     fun check defs (c, {reducts, ...}: def) =
 | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 221 | reducts |> forall (fn (args, deps) => forall (wellformed context defs (c, args)) deps); | 
| 61249 | 222 | in norm_all #> (fn defs => tap (Itemtab.forall (check defs)) defs) end; | 
| 19701 | 223 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 224 | fun dependencies context (c, args) restr deps = | 
| 19712 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 wenzelm parents: 
19707diff
changeset | 225 | map_def c (fn (specs, restricts, reducts) => | 
| 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 wenzelm parents: 
19707diff
changeset | 226 | let | 
| 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 wenzelm parents: 
19707diff
changeset | 227 | val restricts' = Library.merge (op =) (restricts, restr); | 
| 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 wenzelm parents: 
19707diff
changeset | 228 | val reducts' = insert (op =) (args, deps) reducts; | 
| 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 wenzelm parents: 
19707diff
changeset | 229 | in (specs, restricts', reducts') end) | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 230 | #> normalize context; | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 231 | |
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 232 | end; | 
| 19692 | 233 | |
| 234 | ||
| 19624 | 235 | (* merge *) | 
| 236 | ||
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 237 | fun merge context (Defs defs1, Defs defs2) = | 
| 19613 
9bf274ec94cf
allow dependencies of disjoint collections of instances;
 wenzelm parents: 
19590diff
changeset | 238 | let | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 239 | fun add_deps (c, args) restr deps defs = | 
| 19692 | 240 | if AList.defined (op =) (reducts_of defs c) args then defs | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 241 | else dependencies context (c, args) restr deps defs; | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 242 |     fun add_def (c, {restricts, reducts, ...}: def) =
 | 
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 243 | fold (fn (args, deps) => add_deps (c, args) restricts deps) reducts; | 
| 19760 | 244 | in | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 245 | Defs (Itemtab.join (join_specs context) (defs1, defs2) | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 246 | |> normalize context |> Itemtab.fold add_def defs2) | 
| 19760 | 247 | end; | 
| 19613 
9bf274ec94cf
allow dependencies of disjoint collections of instances;
 wenzelm parents: 
19590diff
changeset | 248 | |
| 
9bf274ec94cf
allow dependencies of disjoint collections of instances;
 wenzelm parents: 
19590diff
changeset | 249 | |
| 
9bf274ec94cf
allow dependencies of disjoint collections of instances;
 wenzelm parents: 
19590diff
changeset | 250 | (* define *) | 
| 19590 
12af4942923d
simple substructure test for typargs (independent type constructors);
 wenzelm parents: 
19569diff
changeset | 251 | |
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 252 | fun define context unchecked def description (c, args) deps (Defs defs) = | 
| 17707 
bc0270e9d27f
back to simple 'defs' (cf. revision 1.79 of theory.ML);
 wenzelm parents: 
17670diff
changeset | 253 | let | 
| 55544 | 254 | val pos = Position.thread_data (); | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 255 | val restr = | 
| 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 256 | if plain_args args orelse | 
| 61249 | 257 | (case args of [Term.Type (_, rec_args)] => plain_args rec_args | _ => false) | 
| 33701 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 wenzelm parents: 
32785diff
changeset | 258 | then [] else [(args, description)]; | 
| 19692 | 259 | val spec = | 
| 55544 | 260 |       (serial (), {def = def, description = description, pos = pos, lhs = args, rhs = deps});
 | 
| 61261 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 261 | val defs' = defs |> update_specs context c spec; | 
| 
ddb2da7cb2e4
more explicit Defs.context: use proper name spaces as far as possible;
 wenzelm parents: 
61260diff
changeset | 262 | in Defs (defs' |> (if unchecked then I else dependencies context (c, args) restr deps)) end; | 
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 263 | |
| 61260 | 264 | fun get_deps (Defs defs) c = reducts_of defs c; | 
| 265 | ||
| 19697 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 wenzelm parents: 
19695diff
changeset | 266 | end; |