| author | wenzelm | 
| Fri, 17 May 2013 19:11:03 +0200 | |
| changeset 52057 | 69137d20ab0b | 
| parent 42389 | b2c6033fc7e4 | 
| child 55544 | cf1baba89a27 | 
| 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  | 
|
| 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: 
16838 
diff
changeset
 | 
8  | 
signature DEFS =  | 
| 
 
e92cba1d4842
tuned interfaces declare, define, finalize, merge:
 
wenzelm 
parents: 
16838 
diff
changeset
 | 
9  | 
sig  | 
| 42384 | 10  | 
val pretty_const: Proof.context -> string * typ list -> Pretty.T  | 
| 19701 | 11  | 
val plain_args: typ list -> bool  | 
| 
17707
 
bc0270e9d27f
back to simple 'defs' (cf. revision 1.79 of theory.ML);
 
wenzelm 
parents: 
17670 
diff
changeset
 | 
12  | 
type T  | 
| 33712 | 13  | 
type spec =  | 
14  | 
    {def: string option, description: string, lhs: typ list, rhs: (string * typ list) list}
 | 
|
15  | 
val all_specifications_of: T -> (string * spec list) list  | 
|
16  | 
val specifications_of: T -> string -> spec list  | 
|
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
17  | 
val dest: T ->  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
18  | 
   {restricts: ((string * typ list) * string) list,
 | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
19  | 
reducts: ((string * typ list) * (string * typ list) list) list}  | 
| 
19590
 
12af4942923d
simple substructure test for typargs (independent type constructors);
 
wenzelm 
parents: 
19569 
diff
changeset
 | 
20  | 
val empty: T  | 
| 42389 | 21  | 
val merge: Proof.context -> T * T -> T  | 
| 42384 | 22  | 
val define: Proof.context -> bool -> string option -> string ->  | 
| 19727 | 23  | 
string * typ list -> (string * typ list) list -> T -> T  | 
| 
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
 | 
24  | 
end  | 
| 
 
cf468b93a02e
Implement cycle-free overloading, so that definitions cannot harm consistency any more (except of course via interaction with axioms).
 
obua 
parents:  
diff
changeset
 | 
25  | 
|
| 17711 | 26  | 
structure Defs: DEFS =  | 
| 
17707
 
bc0270e9d27f
back to simple 'defs' (cf. revision 1.79 of theory.ML);
 
wenzelm 
parents: 
17670 
diff
changeset
 | 
27  | 
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
 | 
28  | 
|
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
29  | 
(* type arguments *)  | 
| 
19613
 
9bf274ec94cf
allow dependencies of disjoint collections of instances;
 
wenzelm 
parents: 
19590 
diff
changeset
 | 
30  | 
|
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
31  | 
type args = typ list;  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
32  | 
|
| 42384 | 33  | 
fun pretty_const ctxt (c, args) =  | 
| 
19613
 
9bf274ec94cf
allow dependencies of disjoint collections of instances;
 
wenzelm 
parents: 
19590 
diff
changeset
 | 
34  | 
let  | 
| 19692 | 35  | 
val prt_args =  | 
36  | 
if null args then []  | 
|
| 42384 | 37  | 
      else [Pretty.list "(" ")" (map (Syntax.pretty_typ ctxt o Logic.unvarifyT_global) args)];
 | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
38  | 
in Pretty.block (Pretty.str c :: prt_args) end;  | 
| 19624 | 39  | 
|
| 19707 | 40  | 
fun plain_args args =  | 
41  | 
forall Term.is_TVar args andalso not (has_duplicates (op =) args);  | 
|
42  | 
||
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
43  | 
fun disjoint_args (Ts, Us) =  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
44  | 
not (Type.could_unifys (Ts, Us)) orelse  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
45  | 
((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: 
19695 
diff
changeset
 | 
46  | 
handle Type.TUNIFY => true);  | 
| 19692 | 47  | 
|
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
48  | 
fun match_args (Ts, Us) =  | 
| 32035 | 49  | 
Option.map Envir.subst_type  | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
50  | 
(SOME (Type.raw_matches (Ts, Us) Vartab.empty) handle Type.TYPE_MATCH => NONE);  | 
| 19692 | 51  | 
|
52  | 
||
53  | 
(* datatype defs *)  | 
|
54  | 
||
| 
33701
 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 
wenzelm 
parents: 
32785 
diff
changeset
 | 
55  | 
type spec =  | 
| 
 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 
wenzelm 
parents: 
32785 
diff
changeset
 | 
56  | 
  {def: string option, description: string, lhs: args, rhs: (string * args) list};
 | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
57  | 
|
| 19692 | 58  | 
type def =  | 
| 
19712
 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 
wenzelm 
parents: 
19707 
diff
changeset
 | 
59  | 
 {specs: spec Inttab.table,                 (*source specifications*)
 | 
| 
 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 
wenzelm 
parents: 
19707 
diff
changeset
 | 
60  | 
restricts: (args * string) list, (*global restrictions imposed by incomplete patterns*)  | 
| 
 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 
wenzelm 
parents: 
19707 
diff
changeset
 | 
61  | 
reducts: (args * (string * args) list) list}; (*specifications as reduction system*)  | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
62  | 
|
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
63  | 
fun make_def (specs, restricts, reducts) =  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
64  | 
  {specs = specs, restricts = restricts, reducts = reducts}: def;
 | 
| 19692 | 65  | 
|
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
66  | 
fun map_def c f =  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
67  | 
Symtab.default (c, make_def (Inttab.empty, [], [])) #>  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
68  | 
  Symtab.map_entry c (fn {specs, restricts, reducts}: def =>
 | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
69  | 
make_def (f (specs, restricts, reducts)));  | 
| 19692 | 70  | 
|
71  | 
||
72  | 
datatype T = Defs of def Symtab.table;  | 
|
73  | 
||
| 
19712
 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 
wenzelm 
parents: 
19707 
diff
changeset
 | 
74  | 
fun lookup_list which defs c =  | 
| 19692 | 75  | 
(case Symtab.lookup defs c of  | 
| 19713 | 76  | 
SOME (def: def) => which def  | 
| 19692 | 77  | 
| NONE => []);  | 
78  | 
||
| 32050 | 79  | 
fun all_specifications_of (Defs defs) =  | 
80  | 
(map o apsnd) (map snd o Inttab.dest o #specs) (Symtab.dest defs);  | 
|
81  | 
||
| 24199 | 82  | 
fun specifications_of (Defs defs) = lookup_list (map snd o Inttab.dest o #specs) defs;  | 
| 32050 | 83  | 
|
| 19692 | 84  | 
val restricts_of = lookup_list #restricts;  | 
85  | 
val reducts_of = lookup_list #reducts;  | 
|
86  | 
||
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
87  | 
fun dest (Defs defs) =  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
88  | 
let  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
89  | 
    val restricts = Symtab.fold (fn (c, {restricts, ...}) =>
 | 
| 
33701
 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 
wenzelm 
parents: 
32785 
diff
changeset
 | 
90  | 
fold (fn (args, description) => cons ((c, args), description)) restricts) defs [];  | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
91  | 
    val reducts = Symtab.fold (fn (c, {reducts, ...}) =>
 | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
92  | 
fold (fn (args, deps) => cons ((c, args), deps)) reducts) defs [];  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
93  | 
  in {restricts = restricts, reducts = reducts} end;
 | 
| 19692 | 94  | 
|
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
95  | 
val empty = Defs Symtab.empty;  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
96  | 
|
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
97  | 
|
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
98  | 
(* specifications *)  | 
| 19692 | 99  | 
|
| 
33701
 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 
wenzelm 
parents: 
32785 
diff
changeset
 | 
100  | 
fun disjoint_specs c (i, {lhs = Ts, description = a, ...}: spec) =
 | 
| 
 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 
wenzelm 
parents: 
32785 
diff
changeset
 | 
101  | 
  Inttab.forall (fn (j, {lhs = Us, description = b, ...}: spec) =>
 | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
102  | 
i = j orelse disjoint_args (Ts, Us) orelse  | 
| 24792 | 103  | 
      error ("Clash of specifications " ^ quote a ^ " and " ^ quote b ^
 | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
104  | 
" for constant " ^ quote c));  | 
| 19692 | 105  | 
|
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
106  | 
fun join_specs c ({specs = specs1, restricts, reducts}, {specs = specs2, ...}: def) =
 | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
107  | 
let  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
108  | 
val specs' =  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
109  | 
Inttab.fold (fn spec2 => (disjoint_specs c spec2 specs1; Inttab.update spec2)) specs2 specs1;  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
110  | 
in make_def (specs', restricts, reducts) end;  | 
| 
 
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 update_specs c spec = map_def c (fn (specs, restricts, reducts) =>  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
113  | 
(disjoint_specs c spec specs; (Inttab.update spec specs, restricts, reducts)));  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
114  | 
|
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
115  | 
|
| 19701 | 116  | 
(* normalized dependencies: reduction with well-formedness check *)  | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
117  | 
|
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
118  | 
local  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
119  | 
|
| 19729 | 120  | 
val prt = Pretty.string_of oo pretty_const;  | 
| 42384 | 121  | 
fun err ctxt (c, args) (d, Us) s1 s2 =  | 
122  | 
error (s1 ^ " dependency of constant " ^ prt ctxt (c, args) ^ " -> " ^ prt ctxt (d, Us) ^ s2);  | 
|
| 19729 | 123  | 
|
| 
19712
 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 
wenzelm 
parents: 
19707 
diff
changeset
 | 
124  | 
fun contained (U as TVar _) (Type (_, Ts)) = exists (fn T => T = U orelse contained U T) Ts  | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
125  | 
| contained _ _ = false;  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
126  | 
|
| 42384 | 127  | 
fun acyclic ctxt (c, args) (d, Us) =  | 
| 19729 | 128  | 
c <> d orelse  | 
129  | 
exists (fn U => exists (contained U) args) Us orelse  | 
|
130  | 
is_none (match_args (args, Us)) orelse  | 
|
| 42384 | 131  | 
err ctxt (c, args) (d, Us) "Circular" "";  | 
| 19729 | 132  | 
|
| 42384 | 133  | 
fun wellformed ctxt defs (c, args) (d, Us) =  | 
| 19729 | 134  | 
forall is_TVar Us orelse  | 
135  | 
(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: 
32785 
diff
changeset
 | 
136  | 
SOME (Ts, description) =>  | 
| 42384 | 137  | 
err ctxt (c, args) (d, Us) "Malformed"  | 
138  | 
        ("\n(restriction " ^ prt ctxt (d, Ts) ^ " from " ^ quote description ^ ")")
 | 
|
| 19729 | 139  | 
| NONE => true);  | 
| 19692 | 140  | 
|
| 42384 | 141  | 
fun reduction ctxt defs const deps =  | 
| 19692 | 142  | 
let  | 
| 19701 | 143  | 
fun reduct Us (Ts, rhs) =  | 
144  | 
(case match_args (Ts, Us) of  | 
|
145  | 
NONE => NONE  | 
|
146  | 
| SOME subst => SOME (map (apsnd (map subst)) rhs));  | 
|
147  | 
fun reducts (d, Us) = get_first (reduct Us) (reducts_of defs d);  | 
|
148  | 
||
149  | 
val reds = map (`reducts) deps;  | 
|
150  | 
val deps' =  | 
|
151  | 
if forall (is_none o #1) reds then NONE  | 
|
| 20668 | 152  | 
else SOME (fold_rev  | 
153  | 
(fn (NONE, dp) => insert (op =) dp | (SOME dps, _) => fold (insert (op =)) dps) reds []);  | 
|
| 42384 | 154  | 
val _ = forall (acyclic ctxt const) (the_default deps deps');  | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
155  | 
in deps' end;  | 
| 19692 | 156  | 
|
| 19760 | 157  | 
in  | 
158  | 
||
| 42384 | 159  | 
fun normalize ctxt =  | 
| 19692 | 160  | 
let  | 
| 19701 | 161  | 
    fun norm_update (c, {reducts, ...}: def) (changed, defs) =
 | 
162  | 
let  | 
|
163  | 
val reducts' = reducts |> map (fn (args, deps) =>  | 
|
| 42384 | 164  | 
(args, perhaps (reduction ctxt defs (c, args)) deps));  | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
165  | 
in  | 
| 19701 | 166  | 
if reducts = reducts' then (changed, defs)  | 
| 32785 | 167  | 
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
 | 
168  | 
end;  | 
| 19701 | 169  | 
fun norm_all defs =  | 
170  | 
(case Symtab.fold norm_update defs (false, defs) of  | 
|
171  | 
(true, defs') => norm_all defs'  | 
|
172  | 
| (false, _) => defs);  | 
|
| 19729 | 173  | 
    fun check defs (c, {reducts, ...}: def) =
 | 
| 42384 | 174  | 
reducts |> forall (fn (args, deps) => forall (wellformed ctxt defs (c, args)) deps);  | 
| 19729 | 175  | 
in norm_all #> (fn defs => tap (Symtab.forall (check defs)) defs) end;  | 
| 19701 | 176  | 
|
| 42384 | 177  | 
fun dependencies ctxt (c, args) restr deps =  | 
| 
19712
 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 
wenzelm 
parents: 
19707 
diff
changeset
 | 
178  | 
map_def c (fn (specs, restricts, reducts) =>  | 
| 
 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 
wenzelm 
parents: 
19707 
diff
changeset
 | 
179  | 
let  | 
| 
 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 
wenzelm 
parents: 
19707 
diff
changeset
 | 
180  | 
val restricts' = Library.merge (op =) (restricts, restr);  | 
| 
 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 
wenzelm 
parents: 
19707 
diff
changeset
 | 
181  | 
val reducts' = insert (op =) (args, deps) reducts;  | 
| 
 
3ae3cc4b1eac
wellformed: be less ambitious about structural containment;
 
wenzelm 
parents: 
19707 
diff
changeset
 | 
182  | 
in (specs, restricts', reducts') end)  | 
| 42384 | 183  | 
#> normalize ctxt;  | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
184  | 
|
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
185  | 
end;  | 
| 19692 | 186  | 
|
187  | 
||
| 19624 | 188  | 
(* merge *)  | 
189  | 
||
| 42389 | 190  | 
fun merge ctxt (Defs defs1, Defs defs2) =  | 
| 
19613
 
9bf274ec94cf
allow dependencies of disjoint collections of instances;
 
wenzelm 
parents: 
19590 
diff
changeset
 | 
191  | 
let  | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
192  | 
fun add_deps (c, args) restr deps defs =  | 
| 19692 | 193  | 
if AList.defined (op =) (reducts_of defs c) args then defs  | 
| 42384 | 194  | 
else dependencies ctxt (c, args) restr deps defs;  | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
195  | 
    fun add_def (c, {restricts, reducts, ...}: def) =
 | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
196  | 
fold (fn (args, deps) => add_deps (c, args) restricts deps) reducts;  | 
| 19760 | 197  | 
in  | 
198  | 
Defs (Symtab.join join_specs (defs1, defs2)  | 
|
| 42384 | 199  | 
|> normalize ctxt |> Symtab.fold add_def defs2)  | 
| 19760 | 200  | 
end;  | 
| 
19613
 
9bf274ec94cf
allow dependencies of disjoint collections of instances;
 
wenzelm 
parents: 
19590 
diff
changeset
 | 
201  | 
|
| 
 
9bf274ec94cf
allow dependencies of disjoint collections of instances;
 
wenzelm 
parents: 
19590 
diff
changeset
 | 
202  | 
|
| 
 
9bf274ec94cf
allow dependencies of disjoint collections of instances;
 
wenzelm 
parents: 
19590 
diff
changeset
 | 
203  | 
(* define *)  | 
| 
19590
 
12af4942923d
simple substructure test for typargs (independent type constructors);
 
wenzelm 
parents: 
19569 
diff
changeset
 | 
204  | 
|
| 42384 | 205  | 
fun define ctxt 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
 | 
206  | 
let  | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
207  | 
val restr =  | 
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
208  | 
if plain_args args orelse  | 
| 32785 | 209  | 
(case args of [Type (_, rec_args)] => plain_args rec_args | _ => false)  | 
| 
33701
 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 
wenzelm 
parents: 
32785 
diff
changeset
 | 
210  | 
then [] else [(args, description)];  | 
| 19692 | 211  | 
val spec =  | 
| 
33701
 
9dd1079cec3a
primitive defs: clarified def (axiom name) vs. description;
 
wenzelm 
parents: 
32785 
diff
changeset
 | 
212  | 
      (serial (), {def = def, description = description, lhs = args, rhs = deps});
 | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
213  | 
val defs' = defs |> update_specs c spec;  | 
| 42384 | 214  | 
in Defs (defs' |> (if unchecked then I else dependencies ctxt (c, args) restr deps)) end;  | 
| 
19697
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
215  | 
|
| 
 
423af2e013b8
specifications_of: lhs/rhs represented as typargs;
 
wenzelm 
parents: 
19695 
diff
changeset
 | 
216  | 
end;  |