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