| author | himmelma | 
| Thu, 28 May 2009 17:03:14 +0200 | |
| changeset 31282 | b98cbfabe824 | 
| parent 30343 | 79f022df8527 | 
| child 31946 | 99ac0321cd47 | 
| permissions | -rw-r--r-- | 
| 256 | 1  | 
(* Title: Pure/type.ML  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
2  | 
Author: Tobias Nipkow, Lawrence C Paulson, and Markus Wenzel  | 
| 0 | 3  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
4  | 
Type signatures and certified types, special treatment of type vars,  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
5  | 
matching and unification of types, extend and merge type signatures.  | 
| 0 | 6  | 
*)  | 
7  | 
||
8  | 
signature TYPE =  | 
|
| 2964 | 9  | 
sig  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
10  | 
(*type signatures and certified types*)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
11  | 
datatype decl =  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
12  | 
LogicalType of int |  | 
| 14989 | 13  | 
Abbreviation of string list * typ * bool |  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
14  | 
Nonterminal  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
15  | 
type tsig  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
16  | 
val rep_tsig: tsig ->  | 
| 19642 | 17  | 
   {classes: NameSpace.T * Sorts.algebra,
 | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
18  | 
default: sort,  | 
| 28017 | 19  | 
types: ((decl * Properties.T) * serial) NameSpace.table,  | 
| 26641 | 20  | 
log_types: string list}  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
21  | 
val empty_tsig: tsig  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
22  | 
val defaultS: tsig -> sort  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
23  | 
val logical_types: tsig -> string list  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
24  | 
val eq_sort: tsig -> sort * sort -> bool  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
25  | 
val subsort: tsig -> sort * sort -> bool  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
26  | 
val of_sort: tsig -> typ * sort -> bool  | 
| 19464 | 27  | 
val inter_sort: tsig -> sort * sort -> sort  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
28  | 
val cert_class: tsig -> class -> class  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
29  | 
val cert_sort: tsig -> sort -> sort  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
30  | 
val witness_sorts: tsig -> sort list -> sort list -> (typ * sort) list  | 
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
31  | 
type mode  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
32  | 
val mode_default: mode  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
33  | 
val mode_syntax: mode  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
34  | 
val mode_abbrev: mode  | 
| 24484 | 35  | 
val get_mode: Proof.context -> mode  | 
36  | 
val set_mode: mode -> Proof.context -> Proof.context  | 
|
37  | 
val restore_mode: Proof.context -> Proof.context -> Proof.context  | 
|
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
38  | 
val cert_typ_mode: mode -> tsig -> typ -> typ  | 
| 14993 | 39  | 
val cert_typ: tsig -> typ -> typ  | 
| 19464 | 40  | 
val arity_number: tsig -> string -> int  | 
41  | 
val arity_sorts: Pretty.pp -> tsig -> string -> sort -> sort list  | 
|
| 28017 | 42  | 
val the_tags: tsig -> string -> Properties.T  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
43  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
44  | 
(*special treatment of type vars*)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
45  | 
val strip_sorts: typ -> typ  | 
| 24982 | 46  | 
val similar_types: term * term -> bool  | 
| 621 | 47  | 
val no_tvars: typ -> typ  | 
| 21116 | 48  | 
val varify: (string * sort) list -> term -> ((string * sort) * indexname) list * term  | 
| 16289 | 49  | 
val freeze_thaw_type: typ -> typ * (typ -> typ)  | 
50  | 
val freeze_type: typ -> typ  | 
|
51  | 
val freeze_thaw: term -> term * (term -> term)  | 
|
52  | 
val freeze: term -> term  | 
|
| 2964 | 53  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
54  | 
(*matching and unification*)  | 
| 2964 | 55  | 
exception TYPE_MATCH  | 
| 25324 | 56  | 
type tyenv = (sort * typ) Vartab.table  | 
| 26327 | 57  | 
val lookup: tyenv -> indexname * sort -> typ option  | 
| 16946 | 58  | 
val typ_match: tsig -> typ * typ -> tyenv -> tyenv  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
59  | 
val typ_instance: tsig -> typ * typ -> bool  | 
| 16946 | 60  | 
val raw_match: typ * typ -> tyenv -> tyenv  | 
| 19694 | 61  | 
val raw_matches: typ list * typ list -> tyenv -> tyenv  | 
| 16946 | 62  | 
val raw_instance: typ * typ -> bool  | 
| 2964 | 63  | 
exception TUNIFY  | 
| 16946 | 64  | 
val unify: tsig -> typ * typ -> tyenv * int -> tyenv * int  | 
65  | 
val raw_unify: typ * typ -> tyenv -> tyenv  | 
|
| 19696 | 66  | 
val raw_unifys: typ list * typ list -> tyenv -> tyenv  | 
| 17804 | 67  | 
val could_unify: typ * typ -> bool  | 
| 19696 | 68  | 
val could_unifys: typ list * typ list -> bool  | 
| 16650 | 69  | 
val eq_type: tyenv -> typ * typ -> bool  | 
| 0 | 70  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
71  | 
(*extend and merge type signatures*)  | 
| 
30343
 
79f022df8527
replace old bstring by binding for logical primitives: class, type, const etc.;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
72  | 
val add_class: Pretty.pp -> NameSpace.naming -> binding * class list -> tsig -> tsig  | 
| 26669 | 73  | 
val hide_class: bool -> string -> tsig -> tsig  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
74  | 
val set_defsort: sort -> tsig -> tsig  | 
| 
30343
 
79f022df8527
replace old bstring by binding for logical primitives: class, type, const etc.;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
75  | 
val add_type: NameSpace.naming -> Properties.T -> binding * int -> tsig -> tsig  | 
| 
 
79f022df8527
replace old bstring by binding for logical primitives: class, type, const etc.;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
76  | 
val add_abbrev: NameSpace.naming -> Properties.T -> binding * string list * typ -> tsig -> tsig  | 
| 
 
79f022df8527
replace old bstring by binding for logical primitives: class, type, const etc.;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
77  | 
val add_nonterminal: NameSpace.naming -> Properties.T -> binding -> tsig -> tsig  | 
| 26669 | 78  | 
val hide_type: bool -> string -> tsig -> tsig  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
79  | 
val add_arity: Pretty.pp -> arity -> tsig -> tsig  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
80  | 
val add_classrel: Pretty.pp -> class * class -> tsig -> tsig  | 
| 
14830
 
faa4865ba1ce
removed norm_typ; improved output; refer to Pretty.pp;
 
wenzelm 
parents: 
14790 
diff
changeset
 | 
81  | 
val merge_tsigs: Pretty.pp -> tsig * tsig -> tsig  | 
| 2964 | 82  | 
end;  | 
83  | 
||
84  | 
structure Type: TYPE =  | 
|
| 0 | 85  | 
struct  | 
86  | 
||
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
87  | 
(** type signatures and certified types **)  | 
| 2964 | 88  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
89  | 
(* type declarations *)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
90  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
91  | 
datatype decl =  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
92  | 
LogicalType of int |  | 
| 14989 | 93  | 
Abbreviation of string list * typ * bool |  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
94  | 
Nonterminal;  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
95  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
96  | 
fun str_of_decl (LogicalType _) = "logical type constructor"  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
97  | 
| str_of_decl (Abbreviation _) = "type abbreviation"  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
98  | 
| str_of_decl Nonterminal = "syntactic type";  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
99  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
100  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
101  | 
(* type tsig *)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
102  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
103  | 
datatype tsig =  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
104  | 
  TSig of {
 | 
| 19642 | 105  | 
classes: NameSpace.T * Sorts.algebra, (*order-sorted algebra of type classes*)  | 
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
106  | 
default: sort, (*default sort on input*)  | 
| 28017 | 107  | 
types: ((decl * Properties.T) * serial) NameSpace.table, (*declared types*)  | 
| 26641 | 108  | 
log_types: string list}; (*logical types sorted by number of arguments*)  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
109  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
110  | 
fun rep_tsig (TSig comps) = comps;  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
111  | 
|
| 26641 | 112  | 
fun make_tsig (classes, default, types, log_types) =  | 
113  | 
  TSig {classes = classes, default = default, types = types, log_types = log_types};
 | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
114  | 
|
| 19642 | 115  | 
fun build_tsig ((space, classes), default, types) =  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
116  | 
let  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
117  | 
val log_types =  | 
| 27302 | 118  | 
Symtab.fold (fn (c, ((LogicalType n, _), _)) => cons (c, n) | _ => I) (snd types) []  | 
| 17756 | 119  | 
|> Library.sort (Library.int_ord o pairself snd) |> map fst;  | 
| 26641 | 120  | 
in make_tsig ((space, classes), default, types, log_types) end;  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
121  | 
|
| 26641 | 122  | 
fun map_tsig f (TSig {classes, default, types, log_types = _}) =
 | 
| 19642 | 123  | 
build_tsig (f (classes, default, types));  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
124  | 
|
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
125  | 
val empty_tsig =  | 
| 19642 | 126  | 
build_tsig ((NameSpace.empty, Sorts.empty_algebra), [], NameSpace.empty_table);  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
127  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
128  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
129  | 
(* classes and sorts *)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
130  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
131  | 
fun defaultS (TSig {default, ...}) = default;
 | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
132  | 
fun logical_types (TSig {log_types, ...}) = log_types;
 | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
133  | 
|
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
134  | 
fun eq_sort (TSig {classes, ...}) = Sorts.sort_eq (#2 classes);
 | 
| 
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
135  | 
fun subsort (TSig {classes, ...}) = Sorts.sort_le (#2 classes);
 | 
| 19642 | 136  | 
fun of_sort (TSig {classes, ...}) = Sorts.of_sort (#2 classes);
 | 
| 19464 | 137  | 
fun inter_sort (TSig {classes, ...}) = Sorts.inter_sort (#2 classes);
 | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
138  | 
|
| 19642 | 139  | 
fun cert_class (TSig {classes, ...}) = Sorts.certify_class (#2 classes);
 | 
140  | 
fun cert_sort (TSig {classes, ...}) = Sorts.certify_sort (#2 classes);
 | 
|
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
141  | 
|
| 19642 | 142  | 
fun witness_sorts (tsig as TSig {classes, log_types, ...}) =
 | 
143  | 
Sorts.witness_sorts (#2 classes) log_types;  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
144  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
145  | 
|
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
146  | 
(* certification mode *)  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
147  | 
|
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
148  | 
datatype mode = Mode of {normalize: bool, logical: bool};
 | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
149  | 
|
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
150  | 
val mode_default = Mode {normalize = true, logical = true};
 | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
151  | 
val mode_syntax = Mode {normalize = true, logical = false};
 | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
152  | 
val mode_abbrev = Mode {normalize = false, logical = false};
 | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
153  | 
|
| 24484 | 154  | 
structure Mode = ProofDataFun  | 
155  | 
(  | 
|
156  | 
type T = mode;  | 
|
157  | 
fun init _ = mode_default;  | 
|
158  | 
);  | 
|
159  | 
||
160  | 
val get_mode = Mode.get;  | 
|
161  | 
fun set_mode mode = Mode.map (K mode);  | 
|
162  | 
fun restore_mode ctxt = set_mode (get_mode ctxt);  | 
|
163  | 
||
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
164  | 
|
| 27313 | 165  | 
(* lookup types *)  | 
166  | 
||
167  | 
fun undecl_type c = "Undeclared type constructor: " ^ quote c;  | 
|
168  | 
||
169  | 
fun lookup_type (TSig {types, ...}) = Option.map fst o Symtab.lookup (snd types);
 | 
|
170  | 
||
171  | 
fun the_tags tsig c =  | 
|
172  | 
(case lookup_type tsig c of  | 
|
173  | 
SOME (_, tags) => tags  | 
|
174  | 
| NONE => error (undecl_type c));  | 
|
175  | 
||
176  | 
||
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
177  | 
(* certified types *)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
178  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
179  | 
fun bad_nargs t = "Bad number of arguments for type constructor: " ^ quote t;  | 
| 
14830
 
faa4865ba1ce
removed norm_typ; improved output; refer to Pretty.pp;
 
wenzelm 
parents: 
14790 
diff
changeset
 | 
180  | 
|
| 14998 | 181  | 
local  | 
182  | 
||
183  | 
fun inst_typ env (Type (c, Ts)) = Type (c, map (inst_typ env) Ts)  | 
|
| 18957 | 184  | 
| inst_typ env (T as TFree (x, _)) = the_default T (AList.lookup (op =) env x)  | 
| 14998 | 185  | 
| inst_typ _ T = T;  | 
186  | 
||
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
187  | 
in  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
188  | 
|
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
189  | 
fun cert_typ_mode (Mode {normalize, logical}) tsig ty =
 | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
190  | 
let  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
191  | 
fun err msg = raise TYPE (msg, [ty], []);  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
192  | 
|
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
193  | 
val check_logical =  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
194  | 
      if logical then fn c => err ("Illegal occurrence of syntactic type: " ^ quote c)
 | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
195  | 
else fn _ => ();  | 
| 14989 | 196  | 
|
197  | 
fun cert (T as Type (c, Ts)) =  | 
|
198  | 
let  | 
|
199  | 
val Ts' = map cert Ts;  | 
|
200  | 
fun nargs n = if length Ts <> n then err (bad_nargs c) else ();  | 
|
201  | 
in  | 
|
| 27302 | 202  | 
(case lookup_type tsig c of  | 
| 15531 | 203  | 
SOME (LogicalType n, _) => (nargs n; Type (c, Ts'))  | 
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
204  | 
| SOME (Abbreviation (vs, U, syn), _) =>  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
205  | 
(nargs (length vs);  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
206  | 
if syn then check_logical c else ();  | 
| 14993 | 207  | 
if normalize then inst_typ (vs ~~ Ts') U  | 
| 14989 | 208  | 
else Type (c, Ts'))  | 
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
209  | 
| SOME (Nonterminal, _) => (nargs 0; check_logical c; T)  | 
| 15531 | 210  | 
| NONE => err (undecl_type c))  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
211  | 
end  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
212  | 
| cert (TFree (x, S)) = TFree (x, cert_sort tsig S)  | 
| 14989 | 213  | 
| cert (TVar (xi as (_, i), S)) =  | 
| 14993 | 214  | 
if i < 0 then  | 
215  | 
            err ("Malformed type variable: " ^ quote (Term.string_of_vname xi))
 | 
|
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
216  | 
else TVar (xi, cert_sort tsig S);  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
217  | 
|
| 14989 | 218  | 
val ty' = cert ty;  | 
| 14993 | 219  | 
in if ty = ty' then ty else ty' end; (*avoid copying of already normal type*)  | 
| 
14830
 
faa4865ba1ce
removed norm_typ; improved output; refer to Pretty.pp;
 
wenzelm 
parents: 
14790 
diff
changeset
 | 
220  | 
|
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
221  | 
val cert_typ = cert_typ_mode mode_default;  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
222  | 
|
| 14998 | 223  | 
end;  | 
224  | 
||
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
225  | 
|
| 19464 | 226  | 
(* type arities *)  | 
227  | 
||
| 27302 | 228  | 
fun arity_number tsig a =  | 
229  | 
(case lookup_type tsig a of  | 
|
| 19464 | 230  | 
SOME (LogicalType n, _) => n  | 
231  | 
| _ => error (undecl_type a));  | 
|
232  | 
||
233  | 
fun arity_sorts _ tsig a [] = replicate (arity_number tsig a) []  | 
|
| 19642 | 234  | 
  | arity_sorts pp (TSig {classes, ...}) a S = Sorts.mg_domain (#2 classes) a S
 | 
| 26641 | 235  | 
handle Sorts.CLASS_ERROR err => error (Sorts.class_error pp err);  | 
| 19464 | 236  | 
|
237  | 
||
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
238  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
239  | 
(** special treatment of type vars **)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
240  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
241  | 
(* strip_sorts *)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
242  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
243  | 
fun strip_sorts (Type (a, Ts)) = Type (a, map strip_sorts Ts)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
244  | 
| strip_sorts (TFree (x, _)) = TFree (x, [])  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
245  | 
| strip_sorts (TVar (xi, _)) = TVar (xi, []);  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
246  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
247  | 
|
| 
25384
 
7b9dfd4774a6
similar_types: uniform treatment of TFrees/TVars;
 
wenzelm 
parents: 
25324 
diff
changeset
 | 
248  | 
(* equivalence up to renaming of atomic types *)  | 
| 24982 | 249  | 
|
250  | 
local  | 
|
251  | 
||
252  | 
fun standard_types t =  | 
|
253  | 
let  | 
|
| 
25384
 
7b9dfd4774a6
similar_types: uniform treatment of TFrees/TVars;
 
wenzelm 
parents: 
25324 
diff
changeset
 | 
254  | 
val Ts = fold_types (fold_atyps (insert (op =))) t [];  | 
| 
 
7b9dfd4774a6
similar_types: uniform treatment of TFrees/TVars;
 
wenzelm 
parents: 
25324 
diff
changeset
 | 
255  | 
val Ts' = map (fn a => TFree (a, [])) (Name.invents Name.context Name.aT (length Ts));  | 
| 
 
7b9dfd4774a6
similar_types: uniform treatment of TFrees/TVars;
 
wenzelm 
parents: 
25324 
diff
changeset
 | 
256  | 
in map_types (map_atyps (perhaps (AList.lookup (op =) (Ts ~~ Ts')))) t end;  | 
| 24982 | 257  | 
|
258  | 
in  | 
|
259  | 
||
260  | 
val similar_types = op aconv o pairself (Term.map_types strip_sorts o standard_types);  | 
|
261  | 
||
262  | 
end;  | 
|
263  | 
||
264  | 
||
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
265  | 
(* no_tvars *)  | 
| 621 | 266  | 
|
267  | 
fun no_tvars T =  | 
|
| 
29275
 
9fa69e3858d6
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
29269 
diff
changeset
 | 
268  | 
(case Term.add_tvarsT T [] of [] => T  | 
| 12501 | 269  | 
  | vs => raise TYPE ("Illegal schematic type variable(s): " ^
 | 
| 
29275
 
9fa69e3858d6
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
29269 
diff
changeset
 | 
270  | 
commas_quote (map (Term.string_of_vname o #1) (rev vs)), [T], []));  | 
| 621 | 271  | 
|
| 7641 | 272  | 
|
| 19806 | 273  | 
(* varify *)  | 
| 621 | 274  | 
|
| 21116 | 275  | 
fun varify fixed t =  | 
| 621 | 276  | 
let  | 
| 19305 | 277  | 
val fs = Term.fold_types (Term.fold_atyps  | 
278  | 
(fn TFree v => if member (op =) fixed v then I else insert (op =) v | _ => I)) t [];  | 
|
| 29260 | 279  | 
val used = Name.context  | 
280  | 
|> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t;  | 
|
281  | 
val fmap = fs ~~ map (rpair 0) (#1 (Name.variants (map fst fs) used));  | 
|
| 2964 | 282  | 
fun thaw (f as (a, S)) =  | 
| 17184 | 283  | 
(case AList.lookup (op =) fmap f of  | 
| 15531 | 284  | 
NONE => TFree f  | 
| 16946 | 285  | 
| SOME xi => TVar (xi, S));  | 
| 21116 | 286  | 
in (fmap, map_types (map_type_tfree thaw) t) end;  | 
| 2964 | 287  | 
|
288  | 
||
| 7641 | 289  | 
(* freeze_thaw: freeze TVars in a term; return the "thaw" inverse *)  | 
| 
3411
 
163f8f4a42d7
Removal of freeze_vars and thaw_vars.  New freeze_thaw
 
paulson 
parents: 
3175 
diff
changeset
 | 
290  | 
|
| 7641 | 291  | 
local  | 
292  | 
||
| 16289 | 293  | 
fun new_name (ix, (pairs, used)) =  | 
| 
20071
 
8f3e1ddb50e6
replaced Term.variant(list) by Name.variant(_list);
 
wenzelm 
parents: 
19806 
diff
changeset
 | 
294  | 
let val v = Name.variant used (string_of_indexname ix)  | 
| 16289 | 295  | 
in ((ix, v) :: pairs, v :: used) end;  | 
| 621 | 296  | 
|
| 16289 | 297  | 
fun freeze_one alist (ix, sort) =  | 
| 17184 | 298  | 
TFree (the (AList.lookup (op =) alist ix), sort)  | 
| 15531 | 299  | 
handle Option =>  | 
| 3790 | 300  | 
      raise TYPE ("Failure during freezing of ?" ^ string_of_indexname ix, [], []);
 | 
| 2964 | 301  | 
|
| 17184 | 302  | 
fun thaw_one alist (a, sort) = TVar (the (AList.lookup (op =) alist a), sort)  | 
| 16289 | 303  | 
handle Option => TFree (a, sort);  | 
| 416 | 304  | 
|
| 10495 | 305  | 
in  | 
306  | 
||
307  | 
(*this sort of code could replace unvarifyT*)  | 
|
| 7641 | 308  | 
fun freeze_thaw_type T =  | 
309  | 
let  | 
|
| 
29275
 
9fa69e3858d6
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
29269 
diff
changeset
 | 
310  | 
val used = OldTerm.add_typ_tfree_names (T, [])  | 
| 
 
9fa69e3858d6
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
29269 
diff
changeset
 | 
311  | 
and tvars = map #1 (OldTerm.add_typ_tvars (T, []));  | 
| 23178 | 312  | 
val (alist, _) = List.foldr new_name ([], used) tvars;  | 
| 7641 | 313  | 
in (map_type_tvar (freeze_one alist) T, map_type_tfree (thaw_one (map swap alist))) end;  | 
314  | 
||
| 16289 | 315  | 
val freeze_type = #1 o freeze_thaw_type;  | 
316  | 
||
| 
3411
 
163f8f4a42d7
Removal of freeze_vars and thaw_vars.  New freeze_thaw
 
paulson 
parents: 
3175 
diff
changeset
 | 
317  | 
fun freeze_thaw t =  | 
| 7641 | 318  | 
let  | 
| 
29275
 
9fa69e3858d6
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
29269 
diff
changeset
 | 
319  | 
val used = OldTerm.it_term_types OldTerm.add_typ_tfree_names (t, [])  | 
| 
 
9fa69e3858d6
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
29269 
diff
changeset
 | 
320  | 
and tvars = map #1 (OldTerm.it_term_types OldTerm.add_typ_tvars (t, []));  | 
| 23178 | 321  | 
val (alist, _) = List.foldr new_name ([], used) tvars;  | 
| 7641 | 322  | 
in  | 
323  | 
(case alist of  | 
|
324  | 
[] => (t, fn x => x) (*nothing to do!*)  | 
|
| 
20548
 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 
wenzelm 
parents: 
20071 
diff
changeset
 | 
325  | 
| _ => (map_types (map_type_tvar (freeze_one alist)) t,  | 
| 
 
8ef25fe585a8
renamed Term.map_term_types to Term.map_types (cf. Term.fold_types);
 
wenzelm 
parents: 
20071 
diff
changeset
 | 
326  | 
map_types (map_type_tfree (thaw_one (map swap alist)))))  | 
| 7641 | 327  | 
end;  | 
328  | 
||
| 16289 | 329  | 
val freeze = #1 o freeze_thaw;  | 
330  | 
||
| 7641 | 331  | 
end;  | 
332  | 
||
| 256 | 333  | 
|
334  | 
||
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
335  | 
(** matching and unification of types **)  | 
| 8899 | 336  | 
|
| 15797 | 337  | 
type tyenv = (sort * typ) Vartab.table;  | 
| 256 | 338  | 
|
| 15797 | 339  | 
fun tvar_clash ixn S S' = raise TYPE ("Type variable " ^
 | 
340  | 
quote (Term.string_of_vname ixn) ^ " has two distinct sorts",  | 
|
341  | 
[TVar (ixn, S), TVar (ixn, S')], []);  | 
|
| 0 | 342  | 
|
| 26327 | 343  | 
fun lookup tye (ixn, S) =  | 
| 17412 | 344  | 
(case Vartab.lookup tye ixn of  | 
| 15797 | 345  | 
NONE => NONE  | 
| 16289 | 346  | 
| SOME (S', T) => if S = S' then SOME T else tvar_clash ixn S S');  | 
347  | 
||
| 0 | 348  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
349  | 
(* matching *)  | 
| 0 | 350  | 
|
| 2964 | 351  | 
exception TYPE_MATCH;  | 
| 0 | 352  | 
|
| 16946 | 353  | 
fun typ_match tsig =  | 
| 2964 | 354  | 
let  | 
| 16340 | 355  | 
fun match (TVar (v, S), T) subs =  | 
| 26327 | 356  | 
(case lookup subs (v, S) of  | 
| 15531 | 357  | 
NONE =>  | 
| 17412 | 358  | 
if of_sort tsig (T, S) then Vartab.update_new (v, (S, T)) subs  | 
| 14993 | 359  | 
else raise TYPE_MATCH  | 
| 15531 | 360  | 
| SOME U => if U = T then subs else raise TYPE_MATCH)  | 
| 16340 | 361  | 
| match (Type (a, Ts), Type (b, Us)) subs =  | 
| 2964 | 362  | 
if a <> b then raise TYPE_MATCH  | 
| 16885 | 363  | 
else matches (Ts, Us) subs  | 
| 16340 | 364  | 
| match (TFree x, TFree y) subs =  | 
| 2964 | 365  | 
if x = y then subs else raise TYPE_MATCH  | 
| 16885 | 366  | 
| match _ _ = raise TYPE_MATCH  | 
367  | 
and matches (T :: Ts, U :: Us) subs = matches (Ts, Us) (match (T, U) subs)  | 
|
368  | 
| matches _ subs = subs;  | 
|
| 16946 | 369  | 
in match end;  | 
| 0 | 370  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
371  | 
fun typ_instance tsig (T, U) =  | 
| 16946 | 372  | 
(typ_match tsig (U, T) Vartab.empty; true) handle TYPE_MATCH => false;  | 
373  | 
||
374  | 
(*purely structural matching*)  | 
|
375  | 
fun raw_match (TVar (v, S), T) subs =  | 
|
| 26327 | 376  | 
(case lookup subs (v, S) of  | 
| 17412 | 377  | 
NONE => Vartab.update_new (v, (S, T)) subs  | 
| 16946 | 378  | 
| SOME U => if U = T then subs else raise TYPE_MATCH)  | 
379  | 
| raw_match (Type (a, Ts), Type (b, Us)) subs =  | 
|
380  | 
if a <> b then raise TYPE_MATCH  | 
|
381  | 
else raw_matches (Ts, Us) subs  | 
|
382  | 
| raw_match (TFree x, TFree y) subs =  | 
|
383  | 
if x = y then subs else raise TYPE_MATCH  | 
|
384  | 
| raw_match _ _ = raise TYPE_MATCH  | 
|
385  | 
and raw_matches (T :: Ts, U :: Us) subs = raw_matches (Ts, Us) (raw_match (T, U) subs)  | 
|
| 19694 | 386  | 
| raw_matches ([], []) subs = subs  | 
387  | 
| raw_matches _ _ = raise TYPE_MATCH;  | 
|
| 16946 | 388  | 
|
389  | 
fun raw_instance (T, U) =  | 
|
390  | 
(raw_match (U, T) Vartab.empty; true) handle TYPE_MATCH => false;  | 
|
| 2964 | 391  | 
|
| 0 | 392  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
393  | 
(* unification *)  | 
| 2964 | 394  | 
|
| 0 | 395  | 
exception TUNIFY;  | 
396  | 
||
| 26641 | 397  | 
(*occurs check*)  | 
| 2964 | 398  | 
fun occurs v tye =  | 
399  | 
let  | 
|
400  | 
fun occ (Type (_, Ts)) = exists occ Ts  | 
|
401  | 
| occ (TFree _) = false  | 
|
| 15797 | 402  | 
| occ (TVar (w, S)) =  | 
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
29260 
diff
changeset
 | 
403  | 
Term.eq_ix (v, w) orelse  | 
| 26327 | 404  | 
(case lookup tye (w, S) of  | 
| 15531 | 405  | 
NONE => false  | 
406  | 
| SOME U => occ U);  | 
|
| 0 | 407  | 
in occ end;  | 
408  | 
||
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
409  | 
(*chase variable assignments; if devar returns a type var then it must be unassigned*)  | 
| 16885 | 410  | 
fun devar tye (T as TVar v) =  | 
| 26327 | 411  | 
(case lookup tye v of  | 
| 16885 | 412  | 
SOME U => devar tye U  | 
| 15531 | 413  | 
| NONE => T)  | 
| 16885 | 414  | 
| devar tye T = T;  | 
| 0 | 415  | 
|
| 17804 | 416  | 
(*order-sorted unification*)  | 
| 19642 | 417  | 
fun unify (tsig as TSig {classes = (_, classes), ...}) TU (tyenv, maxidx) =
 | 
| 2964 | 418  | 
let  | 
419  | 
val tyvar_count = ref maxidx;  | 
|
| 24848 | 420  | 
fun gen_tyvar S = TVar ((Name.aT, inc tyvar_count), S);  | 
| 2964 | 421  | 
|
| 19642 | 422  | 
fun mg_domain a S = Sorts.mg_domain classes a S  | 
423  | 
handle Sorts.CLASS_ERROR _ => raise TUNIFY;  | 
|
| 2964 | 424  | 
|
| 16885 | 425  | 
fun meet (_, []) tye = tye  | 
426  | 
| meet (TVar (xi, S'), S) tye =  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
427  | 
if Sorts.sort_le classes (S', S) then tye  | 
| 17412 | 428  | 
else Vartab.update_new  | 
| 17221 | 429  | 
(xi, (S', gen_tyvar (Sorts.inter_sort classes (S', S)))) tye  | 
| 16885 | 430  | 
| meet (TFree (_, S'), S) tye =  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
431  | 
if Sorts.sort_le classes (S', S) then tye  | 
| 2964 | 432  | 
else raise TUNIFY  | 
| 16885 | 433  | 
| meet (Type (a, Ts), S) tye = meets (Ts, mg_domain a S) tye  | 
434  | 
and meets (T :: Ts, S :: Ss) tye = meets (Ts, Ss) (meet (devar tye T, S) tye)  | 
|
435  | 
| meets _ tye = tye;  | 
|
| 2964 | 436  | 
|
| 16885 | 437  | 
fun unif (ty1, ty2) tye =  | 
438  | 
(case (devar tye ty1, devar tye ty2) of  | 
|
| 2964 | 439  | 
(T as TVar (v, S1), U as TVar (w, S2)) =>  | 
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
29260 
diff
changeset
 | 
440  | 
if Term.eq_ix (v, w) then  | 
| 15797 | 441  | 
if S1 = S2 then tye else tvar_clash v S1 S2  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
442  | 
else if Sorts.sort_le classes (S1, S2) then  | 
| 17412 | 443  | 
Vartab.update_new (w, (S2, T)) tye  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
444  | 
else if Sorts.sort_le classes (S2, S1) then  | 
| 17412 | 445  | 
Vartab.update_new (v, (S1, U)) tye  | 
| 2964 | 446  | 
else  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
447  | 
let val S = gen_tyvar (Sorts.inter_sort classes (S1, S2)) in  | 
| 17412 | 448  | 
Vartab.update_new (v, (S1, S)) (Vartab.update_new (w, (S2, S)) tye)  | 
| 2964 | 449  | 
end  | 
450  | 
| (TVar (v, S), T) =>  | 
|
451  | 
if occurs v tye T then raise TUNIFY  | 
|
| 17412 | 452  | 
else meet (T, S) (Vartab.update_new (v, (S, T)) tye)  | 
| 2964 | 453  | 
| (T, TVar (v, S)) =>  | 
454  | 
if occurs v tye T then raise TUNIFY  | 
|
| 17412 | 455  | 
else meet (T, S) (Vartab.update_new (v, (S, T)) tye)  | 
| 2964 | 456  | 
| (Type (a, Ts), Type (b, Us)) =>  | 
457  | 
if a <> b then raise TUNIFY  | 
|
| 16885 | 458  | 
else unifs (Ts, Us) tye  | 
459  | 
| (T, U) => if T = U then tye else raise TUNIFY)  | 
|
460  | 
and unifs (T :: Ts, U :: Us) tye = unifs (Ts, Us) (unif (T, U) tye)  | 
|
461  | 
| unifs _ tye = tye;  | 
|
462  | 
in (unif TU tyenv, ! tyvar_count) end;  | 
|
| 0 | 463  | 
|
| 16946 | 464  | 
(*purely structural unification*)  | 
465  | 
fun raw_unify (ty1, ty2) tye =  | 
|
466  | 
(case (devar tye ty1, devar tye ty2) of  | 
|
467  | 
(T as TVar (v, S1), U as TVar (w, S2)) =>  | 
|
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
29260 
diff
changeset
 | 
468  | 
if Term.eq_ix (v, w) then  | 
| 16946 | 469  | 
if S1 = S2 then tye else tvar_clash v S1 S2  | 
| 17412 | 470  | 
else Vartab.update_new (w, (S2, T)) tye  | 
| 16946 | 471  | 
| (TVar (v, S), T) =>  | 
472  | 
if occurs v tye T then raise TUNIFY  | 
|
| 17412 | 473  | 
else Vartab.update_new (v, (S, T)) tye  | 
| 16946 | 474  | 
| (T, TVar (v, S)) =>  | 
475  | 
if occurs v tye T then raise TUNIFY  | 
|
| 17412 | 476  | 
else Vartab.update_new (v, (S, T)) tye  | 
| 16946 | 477  | 
| (Type (a, Ts), Type (b, Us)) =>  | 
478  | 
if a <> b then raise TUNIFY  | 
|
479  | 
else raw_unifys (Ts, Us) tye  | 
|
480  | 
| (T, U) => if T = U then tye else raise TUNIFY)  | 
|
481  | 
and raw_unifys (T :: Ts, U :: Us) tye = raw_unifys (Ts, Us) (raw_unify (T, U) tye)  | 
|
| 19696 | 482  | 
| raw_unifys ([], []) tye = tye  | 
483  | 
| raw_unifys _ _ = raise TUNIFY;  | 
|
| 16946 | 484  | 
|
| 17804 | 485  | 
(*fast unification filter*)  | 
486  | 
fun could_unify (Type (a, Ts), Type (b, Us)) = a = b andalso could_unifys (Ts, Us)  | 
|
487  | 
| could_unify (TFree (a, _), TFree (b, _)) = a = b  | 
|
488  | 
| could_unify (TVar _, _) = true  | 
|
489  | 
| could_unify (_, TVar _) = true  | 
|
490  | 
| could_unify _ = false  | 
|
491  | 
and could_unifys (T :: Ts, U :: Us) = could_unify (T, U) andalso could_unifys (Ts, Us)  | 
|
| 19696 | 492  | 
| could_unifys ([], []) = true  | 
493  | 
| could_unifys _ = false;  | 
|
| 450 | 494  | 
|
| 17804 | 495  | 
|
496  | 
(*equality with respect to a type environment*)  | 
|
| 16650 | 497  | 
fun eq_type tye (T, T') =  | 
| 16885 | 498  | 
(case (devar tye T, devar tye T') of  | 
| 16650 | 499  | 
(Type (s, Ts), Type (s', Ts')) =>  | 
500  | 
s = s' andalso ListPair.all (eq_type tye) (Ts, Ts')  | 
|
501  | 
| (U, U') => U = U');  | 
|
502  | 
||
| 450 | 503  | 
|
| 0 | 504  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
505  | 
(** extend and merge type signatures **)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
506  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
507  | 
(* classes *)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
508  | 
|
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
509  | 
fun add_class pp naming (c, cs) tsig =  | 
| 19642 | 510  | 
tsig |> map_tsig (fn ((space, classes), default, types) =>  | 
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
511  | 
let  | 
| 
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
512  | 
val cs' = map (cert_class tsig) cs  | 
| 
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
513  | 
handle TYPE (msg, _, _) => error msg;  | 
| 
30343
 
79f022df8527
replace old bstring by binding for logical primitives: class, type, const etc.;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
514  | 
val (c', space') = space |> NameSpace.declare naming c;  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
515  | 
val classes' = classes |> Sorts.add_class pp (c', cs');  | 
| 19642 | 516  | 
in ((space', classes'), default, types) end);  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
517  | 
|
| 26669 | 518  | 
fun hide_class fully c = map_tsig (fn ((space, classes), default, types) =>  | 
519  | 
((NameSpace.hide fully c space, classes), default, types));  | 
|
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
520  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
521  | 
|
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
522  | 
(* arities *)  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
523  | 
|
| 19642 | 524  | 
fun add_arity pp (t, Ss, S) tsig = tsig |> map_tsig (fn ((space, classes), default, types) =>  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
525  | 
let  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
526  | 
val _ =  | 
| 27302 | 527  | 
(case lookup_type tsig t of  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
528  | 
SOME (LogicalType n, _) => if length Ss <> n then error (bad_nargs t) else ()  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
529  | 
      | SOME (decl, _) => error ("Illegal " ^ str_of_decl decl ^ ": " ^ quote t)
 | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
530  | 
| NONE => error (undecl_type t));  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
531  | 
val (Ss', S') = (map (cert_sort tsig) Ss, cert_sort tsig S)  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
532  | 
handle TYPE (msg, _, _) => error msg;  | 
| 19642 | 533  | 
val classes' = classes |> Sorts.add_arities pp ((t, map (fn c' => (c', Ss')) S'));  | 
534  | 
in ((space, classes'), default, types) end);  | 
|
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
535  | 
|
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
536  | 
|
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
537  | 
(* classrel *)  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
538  | 
|
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
539  | 
fun add_classrel pp rel tsig =  | 
| 19642 | 540  | 
tsig |> map_tsig (fn ((space, classes), default, types) =>  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
541  | 
let  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
542  | 
val rel' = pairself (cert_class tsig) rel  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
543  | 
handle TYPE (msg, _, _) => error msg;  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
544  | 
val classes' = classes |> Sorts.add_classrel pp rel;  | 
| 19642 | 545  | 
in ((space, classes'), default, types) end);  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
546  | 
|
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
547  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
548  | 
(* default sort *)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
549  | 
|
| 19642 | 550  | 
fun set_defsort S tsig = tsig |> map_tsig (fn (classes, _, types) =>  | 
551  | 
(classes, cert_sort tsig S handle TYPE (msg, _, _) => error msg, types));  | 
|
| 0 | 552  | 
|
553  | 
||
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
554  | 
(* types *)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
555  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
556  | 
local  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
557  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
558  | 
fun err_in_decls c decl decl' =  | 
| 14906 | 559  | 
let val s = str_of_decl decl and s' = str_of_decl decl' in  | 
560  | 
    if s = s' then error ("Duplicate declaration of " ^ s ^ ": " ^ quote c)
 | 
|
| 
14830
 
faa4865ba1ce
removed norm_typ; improved output; refer to Pretty.pp;
 
wenzelm 
parents: 
14790 
diff
changeset
 | 
561  | 
    else error ("Conflict of " ^ s ^ " with " ^ s' ^ ": " ^ quote c)
 | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
562  | 
end;  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
563  | 
|
| 27302 | 564  | 
fun new_decl naming tags (c, decl) (space, types) =  | 
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
565  | 
let  | 
| 27313 | 566  | 
val tags' = tags |> Position.default_properties (Position.thread_data ());  | 
| 
30343
 
79f022df8527
replace old bstring by binding for logical primitives: class, type, const etc.;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
567  | 
val (c', space') = NameSpace.declare naming c space;  | 
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
568  | 
val types' =  | 
| 17412 | 569  | 
(case Symtab.lookup types c' of  | 
| 27302 | 570  | 
SOME ((decl', _), _) => err_in_decls c' decl decl'  | 
| 27313 | 571  | 
| NONE => Symtab.update (c', ((decl, tags'), serial ())) types);  | 
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
572  | 
in (space', types') end;  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
573  | 
|
| 27302 | 574  | 
fun the_decl (_, types) = fst o fst o the o Symtab.lookup types;  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
575  | 
|
| 19642 | 576  | 
fun map_types f = map_tsig (fn (classes, default, types) =>  | 
| 
19579
 
b802d1804b77
replaced Sorts.DOMAIN by general Sorts.CLASS_ERROR;
 
wenzelm 
parents: 
19530 
diff
changeset
 | 
577  | 
let  | 
| 
 
b802d1804b77
replaced Sorts.DOMAIN by general Sorts.CLASS_ERROR;
 
wenzelm 
parents: 
19530 
diff
changeset
 | 
578  | 
val (space', tab') = f types;  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21116 
diff
changeset
 | 
579  | 
val _ = NameSpace.intern space' "dummy" = "dummy" orelse  | 
| 
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21116 
diff
changeset
 | 
580  | 
error "Illegal declaration of dummy type";  | 
| 19642 | 581  | 
in (classes, default, (space', tab')) end);  | 
| 2964 | 582  | 
|
| 14989 | 583  | 
fun syntactic types (Type (c, Ts)) =  | 
| 27302 | 584  | 
(case Symtab.lookup types c of SOME ((Nonterminal, _), _) => true | _ => false)  | 
| 14989 | 585  | 
orelse exists (syntactic types) Ts  | 
586  | 
| syntactic _ _ = false;  | 
|
587  | 
||
| 27302 | 588  | 
in  | 
589  | 
||
| 
30343
 
79f022df8527
replace old bstring by binding for logical primitives: class, type, const etc.;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
590  | 
fun add_type naming tags (c, n) =  | 
| 
 
79f022df8527
replace old bstring by binding for logical primitives: class, type, const etc.;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
591  | 
  if n < 0 then error ("Bad type constructor declaration: " ^ quote (Binding.str_of c))
 | 
| 
 
79f022df8527
replace old bstring by binding for logical primitives: class, type, const etc.;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
592  | 
else map_types (new_decl naming tags (c, LogicalType n));  | 
| 27302 | 593  | 
|
594  | 
fun add_abbrev naming tags (a, vs, rhs) tsig = tsig |> map_types (fn types =>  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
595  | 
let  | 
| 
30343
 
79f022df8527
replace old bstring by binding for logical primitives: class, type, const etc.;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
596  | 
fun err msg =  | 
| 
 
79f022df8527
replace old bstring by binding for logical primitives: class, type, const etc.;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
597  | 
      cat_error msg ("The error(s) above occurred in type abbreviation: " ^ quote (Binding.str_of a));
 | 
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
598  | 
val rhs' = strip_sorts (no_tvars (cert_typ_mode mode_syntax tsig rhs))  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
599  | 
handle TYPE (msg, _, _) => err msg;  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
600  | 
in  | 
| 18964 | 601  | 
(case duplicates (op =) vs of  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
602  | 
[] => []  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
603  | 
    | dups => err ("Duplicate variables on lhs: " ^ commas_quote dups));
 | 
| 
27273
 
d54ae0bdad80
add_abbrev: check tfrees of rhs, not tvars (addresses a lapse introduced in 1.65);
 
wenzelm 
parents: 
26669 
diff
changeset
 | 
604  | 
(case subtract (op =) vs (map #1 (Term.add_tfreesT rhs' [])) of  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
605  | 
[] => []  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
606  | 
    | extras => err ("Extra variables on rhs: " ^ commas_quote extras));
 | 
| 27302 | 607  | 
types |> new_decl naming tags (a, Abbreviation (vs, rhs', syntactic (#2 types) rhs'))  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
608  | 
end);  | 
| 0 | 609  | 
|
| 27302 | 610  | 
fun add_nonterminal naming tags = map_types o new_decl naming tags o rpair Nonterminal;  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
611  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
612  | 
fun merge_types (types1, types2) =  | 
| 20674 | 613  | 
NameSpace.merge_tables (Library.eq_snd (op = : serial * serial -> bool)) (types1, types2)  | 
| 
23655
 
d2d1138e0ddc
replaced exception TableFun/GraphFun.DUPS by TableFun/GraphFun.DUP;
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
614  | 
handle Symtab.DUP d => err_in_decls d (the_decl types1 d) (the_decl types2 d);  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
615  | 
|
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
616  | 
end;  | 
| 0 | 617  | 
|
| 26669 | 618  | 
fun hide_type fully c = map_tsig (fn (classes, default, (space, types)) =>  | 
619  | 
(classes, default, (NameSpace.hide fully c space, types)));  | 
|
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
620  | 
|
| 3790 | 621  | 
|
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
622  | 
(* merge type signatures *)  | 
| 
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
623  | 
|
| 
14830
 
faa4865ba1ce
removed norm_typ; improved output; refer to Pretty.pp;
 
wenzelm 
parents: 
14790 
diff
changeset
 | 
624  | 
fun merge_tsigs pp (tsig1, tsig2) =  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
625  | 
let  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
626  | 
    val (TSig {classes = (space1, classes1), default = default1, types = types1,
 | 
| 26641 | 627  | 
log_types = _}) = tsig1;  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
628  | 
    val (TSig {classes = (space2, classes2), default = default2, types = types2,
 | 
| 26641 | 629  | 
log_types = _}) = tsig2;  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
630  | 
|
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
631  | 
val space' = NameSpace.merge (space1, space2);  | 
| 19642 | 632  | 
val classes' = Sorts.merge_algebra pp (classes1, classes2);  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
633  | 
val default' = Sorts.inter_sort classes' (default1, default2);  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
634  | 
val types' = merge_types (types1, types2);  | 
| 19642 | 635  | 
in build_tsig ((space', classes'), default', types') end;  | 
| 
14790
 
0d984ee030a1
major cleanup of tsig datastructures and extend/merge operations; fixes old bugs in classes/arities code; proper treatment of nonterminals and syntax-only types;
 
wenzelm 
parents: 
13666 
diff
changeset
 | 
636  | 
|
| 0 | 637  | 
end;  |