| author | kuncar | 
| Wed, 05 Jun 2013 15:21:52 +0200 | |
| changeset 52307 | 32c433c38ddd | 
| parent 51930 | 52fd62618631 | 
| child 55841 | a232c0ff3c20 | 
| 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  | 
| 39288 | 10  | 
(*constraints*)  | 
| 
39290
 
44e4d8dfd6bf
load type_infer.ML later -- proper context for Type_Infer.infer_types;
 
wenzelm 
parents: 
39289 
diff
changeset
 | 
11  | 
val mark_polymorphic: typ -> typ  | 
| 39288 | 12  | 
val constraint: typ -> term -> term  | 
| 
45445
 
41e641a870de
pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
 
wenzelm 
parents: 
44116 
diff
changeset
 | 
13  | 
val constraint_type: Proof.context -> typ -> typ  | 
| 
39292
 
6f085332c7d3
Type_Infer.preterm: eliminated separate Constraint;
 
wenzelm 
parents: 
39290 
diff
changeset
 | 
14  | 
val strip_constraints: term -> term  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
15  | 
val appl_error: Proof.context -> term -> typ -> term -> typ -> string  | 
| 
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
 | 
16  | 
(*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
 | 
17  | 
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
 | 
18  | 
LogicalType of int |  | 
| 14989 | 19  | 
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
 | 
20  | 
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
 | 
21  | 
type tsig  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
22  | 
val eq_tsig: tsig * tsig -> 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
 | 
23  | 
val rep_tsig: tsig ->  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33094 
diff
changeset
 | 
24  | 
   {classes: Name_Space.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
 | 
25  | 
default: sort,  | 
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33097 
diff
changeset
 | 
26  | 
types: decl Name_Space.table,  | 
| 26641 | 27  | 
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
 | 
28  | 
val empty_tsig: tsig  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
29  | 
val class_space: tsig -> Name_Space.T  | 
| 35680 | 30  | 
val class_alias: Name_Space.naming -> binding -> string -> tsig -> tsig  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
31  | 
val intern_class: tsig -> xstring -> string  | 
| 
42358
 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 
wenzelm 
parents: 
41421 
diff
changeset
 | 
32  | 
val extern_class: Proof.context -> tsig -> string -> xstring  | 
| 
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
 | 
33  | 
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
 | 
34  | 
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
 | 
35  | 
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
 | 
36  | 
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
 | 
37  | 
val of_sort: tsig -> typ * sort -> bool  | 
| 19464 | 38  | 
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
 | 
39  | 
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
 | 
40  | 
val cert_sort: tsig -> sort -> sort  | 
| 36447 | 41  | 
val minimize_sort: tsig -> sort -> sort  | 
| 
31946
 
99ac0321cd47
witness_sorts: proper type witnesses for hyps, not invented "'hyp" variables;
 
wenzelm 
parents: 
30343 
diff
changeset
 | 
42  | 
val witness_sorts: tsig -> (typ * sort) list -> sort list -> (typ * sort) list  | 
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
43  | 
type mode  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
44  | 
val mode_default: mode  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
45  | 
val mode_syntax: mode  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
46  | 
val mode_abbrev: mode  | 
| 24484 | 47  | 
val get_mode: Proof.context -> mode  | 
48  | 
val set_mode: mode -> Proof.context -> Proof.context  | 
|
49  | 
val restore_mode: Proof.context -> Proof.context -> Proof.context  | 
|
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
50  | 
val type_space: tsig -> Name_Space.T  | 
| 35680 | 51  | 
val type_alias: Name_Space.naming -> binding -> string -> tsig -> tsig  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
52  | 
val intern_type: tsig -> xstring -> string  | 
| 
42358
 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 
wenzelm 
parents: 
41421 
diff
changeset
 | 
53  | 
val extern_type: Proof.context -> tsig -> string -> xstring  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
54  | 
val is_logtype: tsig -> string -> bool  | 
| 42468 | 55  | 
val the_decl: tsig -> string * Position.T -> decl  | 
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
56  | 
val cert_typ_mode: mode -> tsig -> typ -> typ  | 
| 14993 | 57  | 
val cert_typ: tsig -> typ -> typ  | 
| 19464 | 58  | 
val arity_number: tsig -> string -> int  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
59  | 
val arity_sorts: Context.pretty -> tsig -> string -> sort -> sort 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
 | 
60  | 
|
| 
 
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
 | 
61  | 
(*special treatment of type vars*)  | 
| 
36621
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36450 
diff
changeset
 | 
62  | 
val sort_of_atyp: typ -> 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
 | 
63  | 
val strip_sorts: typ -> typ  | 
| 
49687
 
4b9034f089eb
use explicit Type.strip_sorts_dummy to suppress sort constraints in output;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
64  | 
val strip_sorts_dummy: typ -> typ  | 
| 621 | 65  | 
val no_tvars: typ -> typ  | 
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35680 
diff
changeset
 | 
66  | 
val varify_global: (string * sort) list -> term -> ((string * sort) * indexname) list * term  | 
| 33832 | 67  | 
val legacy_freeze_thaw_type: typ -> typ * (typ -> typ)  | 
68  | 
val legacy_freeze_type: typ -> typ  | 
|
69  | 
val legacy_freeze_thaw: term -> term * (term -> term)  | 
|
70  | 
val legacy_freeze: term -> term  | 
|
| 2964 | 71  | 
|
| 
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
 | 
72  | 
(*matching and unification*)  | 
| 2964 | 73  | 
exception TYPE_MATCH  | 
| 25324 | 74  | 
type tyenv = (sort * typ) Vartab.table  | 
| 26327 | 75  | 
val lookup: tyenv -> indexname * sort -> typ option  | 
| 
32648
 
143e0b0a6b33
Correct chasing of type variable instantiations during type unification.
 
paulson 
parents: 
32030 
diff
changeset
 | 
76  | 
val devar: tyenv -> typ -> typ  | 
| 16946 | 77  | 
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
 | 
78  | 
val typ_instance: tsig -> typ * typ -> bool  | 
| 16946 | 79  | 
val raw_match: typ * typ -> tyenv -> tyenv  | 
| 19694 | 80  | 
val raw_matches: typ list * typ list -> tyenv -> tyenv  | 
| 16946 | 81  | 
val raw_instance: typ * typ -> bool  | 
| 2964 | 82  | 
exception TUNIFY  | 
| 16946 | 83  | 
val unify: tsig -> typ * typ -> tyenv * int -> tyenv * int  | 
84  | 
val raw_unify: typ * typ -> tyenv -> tyenv  | 
|
| 19696 | 85  | 
val raw_unifys: typ list * typ list -> tyenv -> tyenv  | 
| 17804 | 86  | 
val could_unify: typ * typ -> bool  | 
| 19696 | 87  | 
val could_unifys: typ list * typ list -> bool  | 
| 16650 | 88  | 
val eq_type: tyenv -> typ * typ -> bool  | 
| 0 | 89  | 
|
| 
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
 | 
90  | 
(*extend and merge type signatures*)  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
91  | 
val add_class: Context.generic -> binding * class list -> tsig -> tsig  | 
| 26669 | 92  | 
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
 | 
93  | 
val set_defsort: sort -> tsig -> tsig  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
94  | 
val add_type: Context.generic -> binding * int -> tsig -> tsig  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
95  | 
val add_abbrev: Context.generic -> binding * string list * typ -> tsig -> tsig  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
96  | 
val add_nonterminal: Context.generic -> binding -> tsig -> tsig  | 
| 26669 | 97  | 
val hide_type: bool -> string -> tsig -> tsig  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
98  | 
val add_arity: Context.pretty -> arity -> tsig -> tsig  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
99  | 
val add_classrel: Context.pretty -> class * class -> tsig -> tsig  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
100  | 
val merge_tsig: Context.pretty -> tsig * tsig -> tsig  | 
| 2964 | 101  | 
end;  | 
102  | 
||
103  | 
structure Type: TYPE =  | 
|
| 0 | 104  | 
struct  | 
105  | 
||
| 39288 | 106  | 
(** constraints **)  | 
107  | 
||
| 
39290
 
44e4d8dfd6bf
load type_infer.ML later -- proper context for Type_Infer.infer_types;
 
wenzelm 
parents: 
39289 
diff
changeset
 | 
108  | 
(*indicate polymorphic Vars*)  | 
| 
 
44e4d8dfd6bf
load type_infer.ML later -- proper context for Type_Infer.infer_types;
 
wenzelm 
parents: 
39289 
diff
changeset
 | 
109  | 
fun mark_polymorphic T = Type ("_polymorphic_", [T]);
 | 
| 
 
44e4d8dfd6bf
load type_infer.ML later -- proper context for Type_Infer.infer_types;
 
wenzelm 
parents: 
39289 
diff
changeset
 | 
110  | 
|
| 39288 | 111  | 
fun constraint T t =  | 
112  | 
if T = dummyT then t  | 
|
113  | 
  else Const ("_type_constraint_", T --> T) $ t;
 | 
|
114  | 
||
| 
45445
 
41e641a870de
pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
 
wenzelm 
parents: 
44116 
diff
changeset
 | 
115  | 
fun constraint_type ctxt T =  | 
| 
 
41e641a870de
pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
 
wenzelm 
parents: 
44116 
diff
changeset
 | 
116  | 
  let fun err () = error ("Malformed internal type constraint: " ^ Syntax.string_of_typ ctxt T);
 | 
| 
 
41e641a870de
pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
 
wenzelm 
parents: 
44116 
diff
changeset
 | 
117  | 
  in (case T of Type ("fun", [A, B]) => if A = B then A else err () | _ => err ()) end;
 | 
| 
 
41e641a870de
pass term positions into check phase, where resulting types are reported accordingly, and eventually shown as tooltips;
 
wenzelm 
parents: 
44116 
diff
changeset
 | 
118  | 
|
| 
39292
 
6f085332c7d3
Type_Infer.preterm: eliminated separate Constraint;
 
wenzelm 
parents: 
39290 
diff
changeset
 | 
119  | 
fun strip_constraints (Const ("_type_constraint_", _) $ t) = strip_constraints t
 | 
| 
 
6f085332c7d3
Type_Infer.preterm: eliminated separate Constraint;
 
wenzelm 
parents: 
39290 
diff
changeset
 | 
120  | 
| strip_constraints (t $ u) = strip_constraints t $ strip_constraints u  | 
| 
 
6f085332c7d3
Type_Infer.preterm: eliminated separate Constraint;
 
wenzelm 
parents: 
39290 
diff
changeset
 | 
121  | 
| strip_constraints (Abs (x, T, t)) = Abs (x, T, strip_constraints t)  | 
| 
 
6f085332c7d3
Type_Infer.preterm: eliminated separate Constraint;
 
wenzelm 
parents: 
39290 
diff
changeset
 | 
122  | 
| strip_constraints a = a;  | 
| 
 
6f085332c7d3
Type_Infer.preterm: eliminated separate Constraint;
 
wenzelm 
parents: 
39290 
diff
changeset
 | 
123  | 
|
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
124  | 
fun appl_error ctxt (Const ("_type_constraint_", Type ("fun", [T, _]))) _ u U =
 | 
| 
39289
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
125  | 
cat_lines  | 
| 
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
126  | 
["Failed to meet type constraint:", "",  | 
| 
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
127  | 
Pretty.string_of (Pretty.block  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
128  | 
[Pretty.str "Term:", Pretty.brk 2, Syntax.pretty_term ctxt u,  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
129  | 
Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt U]),  | 
| 
39289
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
130  | 
Pretty.string_of (Pretty.block  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
131  | 
[Pretty.str "Type:", Pretty.brk 2, Syntax.pretty_typ ctxt T])]  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
132  | 
| appl_error ctxt t T u U =  | 
| 
39289
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
133  | 
cat_lines  | 
| 
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
134  | 
["Type error in application: " ^  | 
| 
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
135  | 
(case T of  | 
| 
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
136  | 
            Type ("fun", _) => "incompatible operand type"
 | 
| 
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
137  | 
| _ => "operator not of function type"),  | 
| 
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
138  | 
"",  | 
| 
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
139  | 
Pretty.string_of (Pretty.block  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
140  | 
[Pretty.str "Operator:", Pretty.brk 2, Syntax.pretty_term ctxt t,  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
141  | 
Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt T]),  | 
| 
39289
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
142  | 
Pretty.string_of (Pretty.block  | 
| 
42383
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
143  | 
[Pretty.str "Operand:", Pretty.brk 3, Syntax.pretty_term ctxt u,  | 
| 
 
0ae4ad40d7b5
simplified pretty printing context, which is only required for certain kernel operations;
 
wenzelm 
parents: 
42381 
diff
changeset
 | 
144  | 
Pretty.str " ::", Pretty.brk 1, Syntax.pretty_typ ctxt U])];  | 
| 
39289
 
92b50c8bb67b
common Type.appl_error, which also covers explicit constraints;
 
wenzelm 
parents: 
39288 
diff
changeset
 | 
145  | 
|
| 39288 | 146  | 
|
147  | 
||
| 
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
 | 
148  | 
(** type signatures and certified types **)  | 
| 2964 | 149  | 
|
| 
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
 | 
150  | 
(* 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
 | 
151  | 
|
| 
 
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
 | 
152  | 
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
 | 
153  | 
LogicalType of int |  | 
| 14989 | 154  | 
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
 | 
155  | 
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
 | 
156  | 
|
| 
 
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
 | 
157  | 
|
| 
 
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
 | 
158  | 
(* 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
 | 
159  | 
|
| 
 
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
 | 
160  | 
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
 | 
161  | 
  TSig of {
 | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33094 
diff
changeset
 | 
162  | 
classes: Name_Space.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
 | 
163  | 
default: sort, (*default sort on input*)  | 
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33097 
diff
changeset
 | 
164  | 
types: decl Name_Space.table, (*declared types*)  | 
| 26641 | 165  | 
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
 | 
166  | 
|
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
167  | 
fun eq_tsig  | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
168  | 
   (TSig {classes = classes1, default = default1, types = types1, log_types = _},
 | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
169  | 
    TSig {classes = classes2, default = default2, types = types2, log_types = _}) =
 | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
170  | 
pointer_eq (classes1, classes2) andalso  | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
171  | 
default1 = default2 andalso  | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
172  | 
pointer_eq (types1, types2);  | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
173  | 
|
| 
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
 | 
174  | 
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
 | 
175  | 
|
| 26641 | 176  | 
fun make_tsig (classes, default, types, log_types) =  | 
177  | 
  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
 | 
178  | 
|
| 
33094
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
179  | 
fun build_tsig (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
 | 
180  | 
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
 | 
181  | 
val log_types =  | 
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33097 
diff
changeset
 | 
182  | 
Symtab.fold (fn (c, LogicalType n) => cons (c, n) | _ => I) (snd types) []  | 
| 
33094
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
183  | 
|> Library.sort (int_ord o pairself snd) |> map fst;  | 
| 
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
184  | 
in make_tsig (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
 | 
185  | 
|
| 26641 | 186  | 
fun map_tsig f (TSig {classes, default, types, log_types = _}) =
 | 
| 19642 | 187  | 
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
 | 
188  | 
|
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
189  | 
val empty_tsig =  | 
| 
50201
 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 
wenzelm 
parents: 
49687 
diff
changeset
 | 
190  | 
build_tsig ((Name_Space.empty Markup.classN, Sorts.empty_algebra), [],  | 
| 
 
c26369c9eda6
Isabelle-specific implementation of quasi-abstract markup elements -- back to module arrangement before d83797ef0d2d;
 
wenzelm 
parents: 
49687 
diff
changeset
 | 
191  | 
Name_Space.empty_table Markup.type_nameN);  | 
| 
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
 | 
192  | 
|
| 
 
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
 | 
193  | 
|
| 
 
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
 | 
194  | 
(* 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
 | 
195  | 
|
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
196  | 
val class_space = #1 o #classes o rep_tsig;  | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
197  | 
|
| 35680 | 198  | 
fun class_alias naming binding name = map_tsig (fn ((space, classes), default, types) =>  | 
199  | 
((Name_Space.alias naming binding name space, classes), default, types));  | 
|
200  | 
||
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
201  | 
val intern_class = Name_Space.intern o class_space;  | 
| 
42358
 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 
wenzelm 
parents: 
41421 
diff
changeset
 | 
202  | 
fun extern_class ctxt = Name_Space.extern ctxt o class_space;  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
203  | 
|
| 
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
 | 
204  | 
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
 | 
205  | 
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
 | 
206  | 
|
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
207  | 
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
 | 
208  | 
fun subsort (TSig {classes, ...}) = Sorts.sort_le (#2 classes);
 | 
| 19642 | 209  | 
fun of_sort (TSig {classes, ...}) = Sorts.of_sort (#2 classes);
 | 
| 19464 | 210  | 
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
 | 
211  | 
|
| 45595 | 212  | 
fun cert_class (TSig {classes = (_, algebra), ...}) c =
 | 
213  | 
if can (Graph.get_entry (Sorts.classes_of algebra)) c then c  | 
|
214  | 
  else raise TYPE ("Undeclared class: " ^ quote c, [], []);
 | 
|
215  | 
||
216  | 
val cert_sort = map o cert_class;  | 
|
217  | 
||
| 36447 | 218  | 
fun minimize_sort (TSig {classes, ...}) = Sorts.minimize_sort (#2 classes);
 | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
219  | 
|
| 32784 | 220  | 
fun witness_sorts (TSig {classes, log_types, ...}) =
 | 
| 19642 | 221  | 
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
 | 
222  | 
|
| 
 
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
 | 
223  | 
|
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
224  | 
(* certification mode *)  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
225  | 
|
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
226  | 
datatype mode = Mode of {normalize: bool, logical: bool};
 | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
227  | 
|
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
228  | 
val mode_default = Mode {normalize = true, logical = true};
 | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
229  | 
val mode_syntax = Mode {normalize = true, logical = false};
 | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
230  | 
val mode_abbrev = Mode {normalize = false, logical = false};
 | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
231  | 
|
| 33519 | 232  | 
structure Mode = Proof_Data  | 
| 24484 | 233  | 
(  | 
234  | 
type T = mode;  | 
|
235  | 
fun init _ = mode_default;  | 
|
236  | 
);  | 
|
237  | 
||
238  | 
val get_mode = Mode.get;  | 
|
239  | 
fun set_mode mode = Mode.map (K mode);  | 
|
240  | 
fun restore_mode ctxt = set_mode (get_mode ctxt);  | 
|
241  | 
||
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
242  | 
|
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
243  | 
(* types *)  | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
244  | 
|
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
245  | 
val type_space = #1 o #types o rep_tsig;  | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
246  | 
|
| 35680 | 247  | 
fun type_alias naming binding name = map_tsig (fn (classes, default, (space, types)) =>  | 
248  | 
(classes, default, (Name_Space.alias naming binding name space, types)));  | 
|
249  | 
||
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
250  | 
val intern_type = Name_Space.intern o type_space;  | 
| 
42358
 
b47d41d9f4b5
Name_Space: proper configuration options long_names, short_names, unique_names instead of former unsynchronized references;
 
wenzelm 
parents: 
41421 
diff
changeset
 | 
251  | 
fun extern_type ctxt = Name_Space.extern ctxt o type_space;  | 
| 
35669
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
252  | 
|
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
253  | 
val is_logtype = member (op =) o logical_types;  | 
| 
 
a91c7ed801b8
added ProofContext.tsig_of -- proforma version for local name space only, not logical content;
 
wenzelm 
parents: 
35359 
diff
changeset
 | 
254  | 
|
| 27313 | 255  | 
|
256  | 
fun undecl_type c = "Undeclared type constructor: " ^ quote c;  | 
|
257  | 
||
| 
33094
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
258  | 
fun lookup_type (TSig {types = (_, types), ...}) = Symtab.lookup types;
 | 
| 27313 | 259  | 
|
| 42468 | 260  | 
fun the_decl tsig (c, pos) =  | 
| 
35359
 
3ec03a3cd9d0
provide direct access to the different kinds of type declarations;
 
wenzelm 
parents: 
34272 
diff
changeset
 | 
261  | 
(case lookup_type tsig c of  | 
| 48992 | 262  | 
NONE => error (undecl_type c ^ Position.here pos)  | 
| 
35359
 
3ec03a3cd9d0
provide direct access to the different kinds of type declarations;
 
wenzelm 
parents: 
34272 
diff
changeset
 | 
263  | 
| SOME decl => decl);  | 
| 
 
3ec03a3cd9d0
provide direct access to the different kinds of type declarations;
 
wenzelm 
parents: 
34272 
diff
changeset
 | 
264  | 
|
| 27313 | 265  | 
|
| 
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
 | 
266  | 
(* 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
 | 
267  | 
|
| 
 
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
 | 
268  | 
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
 | 
269  | 
|
| 14998 | 270  | 
local  | 
271  | 
||
272  | 
fun inst_typ env (Type (c, Ts)) = Type (c, map (inst_typ env) Ts)  | 
|
| 18957 | 273  | 
| inst_typ env (T as TFree (x, _)) = the_default T (AList.lookup (op =) env x)  | 
| 14998 | 274  | 
| inst_typ _ T = T;  | 
275  | 
||
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
276  | 
in  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
277  | 
|
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
278  | 
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
 | 
279  | 
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
 | 
280  | 
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
 | 
281  | 
|
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
282  | 
val check_logical =  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
283  | 
      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
 | 
284  | 
else fn _ => ();  | 
| 14989 | 285  | 
|
286  | 
fun cert (T as Type (c, Ts)) =  | 
|
287  | 
let  | 
|
288  | 
val Ts' = map cert Ts;  | 
|
289  | 
fun nargs n = if length Ts <> n then err (bad_nargs c) else ();  | 
|
290  | 
in  | 
|
| 42468 | 291  | 
(case the_decl tsig (c, Position.none) of  | 
| 
35359
 
3ec03a3cd9d0
provide direct access to the different kinds of type declarations;
 
wenzelm 
parents: 
34272 
diff
changeset
 | 
292  | 
LogicalType n => (nargs n; Type (c, Ts'))  | 
| 
 
3ec03a3cd9d0
provide direct access to the different kinds of type declarations;
 
wenzelm 
parents: 
34272 
diff
changeset
 | 
293  | 
| Abbreviation (vs, U, syn) =>  | 
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
294  | 
(nargs (length vs);  | 
| 
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
295  | 
if syn then check_logical c else ();  | 
| 14993 | 296  | 
if normalize then inst_typ (vs ~~ Ts') U  | 
| 14989 | 297  | 
else Type (c, Ts'))  | 
| 
35359
 
3ec03a3cd9d0
provide direct access to the different kinds of type declarations;
 
wenzelm 
parents: 
34272 
diff
changeset
 | 
298  | 
| Nonterminal => (nargs 0; check_logical c; 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
 | 
299  | 
end  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
300  | 
| cert (TFree (x, S)) = TFree (x, cert_sort tsig S)  | 
| 14989 | 301  | 
| cert (TVar (xi as (_, i), S)) =  | 
| 14993 | 302  | 
if i < 0 then  | 
303  | 
            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
 | 
304  | 
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
 | 
305  | 
|
| 14989 | 306  | 
val ty' = cert ty;  | 
| 14993 | 307  | 
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
 | 
308  | 
|
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
309  | 
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
 | 
310  | 
|
| 14998 | 311  | 
end;  | 
312  | 
||
| 
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
 | 
313  | 
|
| 19464 | 314  | 
(* type arities *)  | 
315  | 
||
| 27302 | 316  | 
fun arity_number tsig a =  | 
317  | 
(case lookup_type tsig a of  | 
|
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33097 
diff
changeset
 | 
318  | 
SOME (LogicalType n) => n  | 
| 19464 | 319  | 
| _ => error (undecl_type a));  | 
320  | 
||
321  | 
fun arity_sorts _ tsig a [] = replicate (arity_number tsig a) []  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
322  | 
  | arity_sorts pp (TSig {classes, ...}) a S =
 | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
323  | 
Sorts.mg_domain (#2 classes) a S  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
324  | 
handle Sorts.CLASS_ERROR err => error (Sorts.class_error pp err);  | 
| 19464 | 325  | 
|
326  | 
||
| 
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
 | 
327  | 
|
| 
 
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
 | 
328  | 
(** 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
 | 
329  | 
|
| 
36621
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36450 
diff
changeset
 | 
330  | 
(* sort_of_atyp *)  | 
| 
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36450 
diff
changeset
 | 
331  | 
|
| 
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36450 
diff
changeset
 | 
332  | 
fun sort_of_atyp (TFree (_, S)) = S  | 
| 
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36450 
diff
changeset
 | 
333  | 
| sort_of_atyp (TVar (_, S)) = S  | 
| 
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36450 
diff
changeset
 | 
334  | 
  | sort_of_atyp T = raise TYPE ("sort_of_atyp", [T], []);
 | 
| 
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36450 
diff
changeset
 | 
335  | 
|
| 
 
2fd4e2c76636
proof terms for strip_shyps, based on the version by krauss/schropp with some notable differences:
 
wenzelm 
parents: 
36450 
diff
changeset
 | 
336  | 
|
| 
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
 | 
337  | 
(* 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
 | 
338  | 
|
| 
49687
 
4b9034f089eb
use explicit Type.strip_sorts_dummy to suppress sort constraints in output;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
339  | 
val strip_sorts = map_atyps  | 
| 
 
4b9034f089eb
use explicit Type.strip_sorts_dummy to suppress sort constraints in output;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
340  | 
(fn TFree (x, _) => TFree (x, [])  | 
| 
 
4b9034f089eb
use explicit Type.strip_sorts_dummy to suppress sort constraints in output;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
341  | 
| TVar (xi, _) => TVar (xi, []));  | 
| 
 
4b9034f089eb
use explicit Type.strip_sorts_dummy to suppress sort constraints in output;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
342  | 
|
| 
 
4b9034f089eb
use explicit Type.strip_sorts_dummy to suppress sort constraints in output;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
343  | 
val strip_sorts_dummy = map_atyps  | 
| 
 
4b9034f089eb
use explicit Type.strip_sorts_dummy to suppress sort constraints in output;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
344  | 
(fn TFree (x, _) => TFree (x, dummyS)  | 
| 
 
4b9034f089eb
use explicit Type.strip_sorts_dummy to suppress sort constraints in output;
 
wenzelm 
parents: 
48992 
diff
changeset
 | 
345  | 
| TVar (xi, _) => TVar (xi, dummyS));  | 
| 
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
 | 
346  | 
|
| 
 
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
 | 
347  | 
|
| 
 
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
 | 
348  | 
(* no_tvars *)  | 
| 621 | 349  | 
|
350  | 
fun no_tvars T =  | 
|
| 
29275
 
9fa69e3858d6
moved old add_type_XXX, add_term_XXX etc. to structure OldTerm;
 
wenzelm 
parents: 
29269 
diff
changeset
 | 
351  | 
(case Term.add_tvarsT T [] of [] => T  | 
| 12501 | 352  | 
  | 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
 | 
353  | 
commas_quote (map (Term.string_of_vname o #1) (rev vs)), [T], []));  | 
| 621 | 354  | 
|
| 7641 | 355  | 
|
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35680 
diff
changeset
 | 
356  | 
(* varify_global *)  | 
| 621 | 357  | 
|
| 
35845
 
e5980f0ad025
renamed varify/unvarify operations to varify_global/unvarify_global to emphasize that these only work in a global situation;
 
wenzelm 
parents: 
35680 
diff
changeset
 | 
358  | 
fun varify_global fixed t =  | 
| 621 | 359  | 
let  | 
| 19305 | 360  | 
val fs = Term.fold_types (Term.fold_atyps  | 
361  | 
(fn TFree v => if member (op =) fixed v then I else insert (op =) v | _ => I)) t [];  | 
|
| 29260 | 362  | 
val used = Name.context  | 
363  | 
|> fold_types (fold_atyps (fn TVar ((a, _), _) => Name.declare a | _ => I)) t;  | 
|
| 
43326
 
47cf4bc789aa
simplified Name.variant -- discontinued builtin fold_map;
 
wenzelm 
parents: 
43324 
diff
changeset
 | 
364  | 
val fmap = fs ~~ map (rpair 0) (#1 (fold_map Name.variant (map fst fs) used));  | 
| 32784 | 365  | 
fun thaw (f as (_, S)) =  | 
| 17184 | 366  | 
(case AList.lookup (op =) fmap f of  | 
| 15531 | 367  | 
NONE => TFree f  | 
| 16946 | 368  | 
| SOME xi => TVar (xi, S));  | 
| 21116 | 369  | 
in (fmap, map_types (map_type_tfree thaw) t) end;  | 
| 2964 | 370  | 
|
371  | 
||
| 7641 | 372  | 
(* 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
 | 
373  | 
|
| 7641 | 374  | 
local  | 
375  | 
||
| 
44116
 
c70257b4cb7e
avoid OldTerm operations -- with subtle changes of semantics;
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
376  | 
fun new_name ix (pairs, used) =  | 
| 
43324
 
2b47822868e4
discontinued Name.variant to emphasize that this is old-style / indirect;
 
wenzelm 
parents: 
42468 
diff
changeset
 | 
377  | 
let val v = singleton (Name.variant_list used) (string_of_indexname ix)  | 
| 16289 | 378  | 
in ((ix, v) :: pairs, v :: used) end;  | 
| 621 | 379  | 
|
| 16289 | 380  | 
fun freeze_one alist (ix, sort) =  | 
| 17184 | 381  | 
TFree (the (AList.lookup (op =) alist ix), sort)  | 
| 
51930
 
52fd62618631
prefer explicitly qualified exceptions, which is particular important for robust handlers;
 
wenzelm 
parents: 
51701 
diff
changeset
 | 
382  | 
handle Option.Option =>  | 
| 3790 | 383  | 
      raise TYPE ("Failure during freezing of ?" ^ string_of_indexname ix, [], []);
 | 
| 2964 | 384  | 
|
| 17184 | 385  | 
fun thaw_one alist (a, sort) = TVar (the (AList.lookup (op =) alist a), sort)  | 
| 
51930
 
52fd62618631
prefer explicitly qualified exceptions, which is particular important for robust handlers;
 
wenzelm 
parents: 
51701 
diff
changeset
 | 
386  | 
handle Option.Option => TFree (a, sort);  | 
| 416 | 387  | 
|
| 10495 | 388  | 
in  | 
389  | 
||
| 33832 | 390  | 
fun legacy_freeze_thaw_type T =  | 
| 7641 | 391  | 
let  | 
| 
44116
 
c70257b4cb7e
avoid OldTerm operations -- with subtle changes of semantics;
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
392  | 
val used = Term.add_tfree_namesT T [];  | 
| 
 
c70257b4cb7e
avoid OldTerm operations -- with subtle changes of semantics;
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
393  | 
val (alist, _) = fold_rev new_name (map #1 (Term.add_tvarsT T [])) ([], used);  | 
| 7641 | 394  | 
in (map_type_tvar (freeze_one alist) T, map_type_tfree (thaw_one (map swap alist))) end;  | 
395  | 
||
| 33832 | 396  | 
val legacy_freeze_type = #1 o legacy_freeze_thaw_type;  | 
| 16289 | 397  | 
|
| 33832 | 398  | 
fun legacy_freeze_thaw t =  | 
| 7641 | 399  | 
let  | 
| 
44116
 
c70257b4cb7e
avoid OldTerm operations -- with subtle changes of semantics;
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
400  | 
val used = Term.add_tfree_names t [];  | 
| 
 
c70257b4cb7e
avoid OldTerm operations -- with subtle changes of semantics;
 
wenzelm 
parents: 
43552 
diff
changeset
 | 
401  | 
val (alist, _) = fold_rev new_name (map #1 (Term.add_tvars t [])) ([], used);  | 
| 7641 | 402  | 
in  | 
403  | 
(case alist of  | 
|
404  | 
[] => (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
 | 
405  | 
| _ => (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
 | 
406  | 
map_types (map_type_tfree (thaw_one (map swap alist)))))  | 
| 7641 | 407  | 
end;  | 
408  | 
||
| 33832 | 409  | 
val legacy_freeze = #1 o legacy_freeze_thaw;  | 
| 16289 | 410  | 
|
| 7641 | 411  | 
end;  | 
412  | 
||
| 256 | 413  | 
|
414  | 
||
| 
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
 | 
415  | 
(** matching and unification of types **)  | 
| 8899 | 416  | 
|
| 15797 | 417  | 
type tyenv = (sort * typ) Vartab.table;  | 
| 256 | 418  | 
|
| 
51701
 
1e29891759c4
tuned exceptions -- avoid composing error messages in low-level situations;
 
wenzelm 
parents: 
50201 
diff
changeset
 | 
419  | 
fun tvar_clash ixn S S' =  | 
| 
 
1e29891759c4
tuned exceptions -- avoid composing error messages in low-level situations;
 
wenzelm 
parents: 
50201 
diff
changeset
 | 
420  | 
  raise TYPE ("Type variable has two distinct sorts", [TVar (ixn, S), TVar (ixn, S')], []);
 | 
| 0 | 421  | 
|
| 26327 | 422  | 
fun lookup tye (ixn, S) =  | 
| 17412 | 423  | 
(case Vartab.lookup tye ixn of  | 
| 15797 | 424  | 
NONE => NONE  | 
| 16289 | 425  | 
| SOME (S', T) => if S = S' then SOME T else tvar_clash ixn S S');  | 
426  | 
||
| 0 | 427  | 
|
| 
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
 | 
428  | 
(* matching *)  | 
| 0 | 429  | 
|
| 2964 | 430  | 
exception TYPE_MATCH;  | 
| 0 | 431  | 
|
| 16946 | 432  | 
fun typ_match tsig =  | 
| 2964 | 433  | 
let  | 
| 
41421
 
2db1d3d2ed54
uniform treatment of typ_match and raw_match (cf. b654fa27fbc4);
 
wenzelm 
parents: 
41254 
diff
changeset
 | 
434  | 
fun match (V as TVar (v, S), T) subs =  | 
| 26327 | 435  | 
(case lookup subs (v, S) of  | 
| 15531 | 436  | 
NONE =>  | 
| 
41421
 
2db1d3d2ed54
uniform treatment of typ_match and raw_match (cf. b654fa27fbc4);
 
wenzelm 
parents: 
41254 
diff
changeset
 | 
437  | 
if V = T then subs  | 
| 
 
2db1d3d2ed54
uniform treatment of typ_match and raw_match (cf. b654fa27fbc4);
 
wenzelm 
parents: 
41254 
diff
changeset
 | 
438  | 
else if of_sort tsig (T, S) then Vartab.update_new (v, (S, T)) subs  | 
| 14993 | 439  | 
else raise TYPE_MATCH  | 
| 15531 | 440  | 
| SOME U => if U = T then subs else raise TYPE_MATCH)  | 
| 16340 | 441  | 
| match (Type (a, Ts), Type (b, Us)) subs =  | 
| 2964 | 442  | 
if a <> b then raise TYPE_MATCH  | 
| 16885 | 443  | 
else matches (Ts, Us) subs  | 
| 16340 | 444  | 
| match (TFree x, TFree y) subs =  | 
| 2964 | 445  | 
if x = y then subs else raise TYPE_MATCH  | 
| 16885 | 446  | 
| match _ _ = raise TYPE_MATCH  | 
447  | 
and matches (T :: Ts, U :: Us) subs = matches (Ts, Us) (match (T, U) subs)  | 
|
448  | 
| matches _ subs = subs;  | 
|
| 16946 | 449  | 
in match end;  | 
| 0 | 450  | 
|
| 
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
 | 
451  | 
fun typ_instance tsig (T, U) =  | 
| 16946 | 452  | 
(typ_match tsig (U, T) Vartab.empty; true) handle TYPE_MATCH => false;  | 
453  | 
||
454  | 
(*purely structural matching*)  | 
|
| 
41421
 
2db1d3d2ed54
uniform treatment of typ_match and raw_match (cf. b654fa27fbc4);
 
wenzelm 
parents: 
41254 
diff
changeset
 | 
455  | 
fun raw_match (V as TVar (v, S), T) subs =  | 
| 26327 | 456  | 
(case lookup subs (v, S) of  | 
| 
41421
 
2db1d3d2ed54
uniform treatment of typ_match and raw_match (cf. b654fa27fbc4);
 
wenzelm 
parents: 
41254 
diff
changeset
 | 
457  | 
NONE => if V = T then subs else Vartab.update_new (v, (S, T)) subs  | 
| 16946 | 458  | 
| SOME U => if U = T then subs else raise TYPE_MATCH)  | 
459  | 
| raw_match (Type (a, Ts), Type (b, Us)) subs =  | 
|
460  | 
if a <> b then raise TYPE_MATCH  | 
|
461  | 
else raw_matches (Ts, Us) subs  | 
|
462  | 
| raw_match (TFree x, TFree y) subs =  | 
|
463  | 
if x = y then subs else raise TYPE_MATCH  | 
|
464  | 
| raw_match _ _ = raise TYPE_MATCH  | 
|
465  | 
and raw_matches (T :: Ts, U :: Us) subs = raw_matches (Ts, Us) (raw_match (T, U) subs)  | 
|
| 19694 | 466  | 
| raw_matches ([], []) subs = subs  | 
467  | 
| raw_matches _ _ = raise TYPE_MATCH;  | 
|
| 16946 | 468  | 
|
469  | 
fun raw_instance (T, U) =  | 
|
470  | 
(raw_match (U, T) Vartab.empty; true) handle TYPE_MATCH => false;  | 
|
| 2964 | 471  | 
|
| 0 | 472  | 
|
| 
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
 | 
473  | 
(* unification *)  | 
| 2964 | 474  | 
|
| 0 | 475  | 
exception TUNIFY;  | 
476  | 
||
| 26641 | 477  | 
(*occurs check*)  | 
| 2964 | 478  | 
fun occurs v tye =  | 
479  | 
let  | 
|
480  | 
fun occ (Type (_, Ts)) = exists occ Ts  | 
|
481  | 
| occ (TFree _) = false  | 
|
| 15797 | 482  | 
| occ (TVar (w, S)) =  | 
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
29260 
diff
changeset
 | 
483  | 
Term.eq_ix (v, w) orelse  | 
| 26327 | 484  | 
(case lookup tye (w, S) of  | 
| 15531 | 485  | 
NONE => false  | 
486  | 
| SOME U => occ U);  | 
|
| 0 | 487  | 
in occ end;  | 
488  | 
||
| 
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
 | 
489  | 
(*chase variable assignments; if devar returns a type var then it must be unassigned*)  | 
| 16885 | 490  | 
fun devar tye (T as TVar v) =  | 
| 26327 | 491  | 
(case lookup tye v of  | 
| 16885 | 492  | 
SOME U => devar tye U  | 
| 15531 | 493  | 
| NONE => T)  | 
| 32784 | 494  | 
| devar _ T = T;  | 
| 0 | 495  | 
|
| 17804 | 496  | 
(*order-sorted unification*)  | 
| 32784 | 497  | 
fun unify (TSig {classes = (_, classes), ...}) TU (tyenv, maxidx) =
 | 
| 2964 | 498  | 
let  | 
| 32738 | 499  | 
val tyvar_count = Unsynchronized.ref maxidx;  | 
500  | 
fun gen_tyvar S = TVar ((Name.aT, Unsynchronized.inc tyvar_count), S);  | 
|
| 2964 | 501  | 
|
| 19642 | 502  | 
fun mg_domain a S = Sorts.mg_domain classes a S  | 
503  | 
handle Sorts.CLASS_ERROR _ => raise TUNIFY;  | 
|
| 2964 | 504  | 
|
| 16885 | 505  | 
fun meet (_, []) tye = tye  | 
506  | 
| 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
 | 
507  | 
if Sorts.sort_le classes (S', S) then tye  | 
| 17412 | 508  | 
else Vartab.update_new  | 
| 17221 | 509  | 
(xi, (S', gen_tyvar (Sorts.inter_sort classes (S', S)))) tye  | 
| 16885 | 510  | 
| 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
 | 
511  | 
if Sorts.sort_le classes (S', S) then tye  | 
| 2964 | 512  | 
else raise TUNIFY  | 
| 16885 | 513  | 
| meet (Type (a, Ts), S) tye = meets (Ts, mg_domain a S) tye  | 
514  | 
and meets (T :: Ts, S :: Ss) tye = meets (Ts, Ss) (meet (devar tye T, S) tye)  | 
|
515  | 
| meets _ tye = tye;  | 
|
| 2964 | 516  | 
|
| 16885 | 517  | 
fun unif (ty1, ty2) tye =  | 
518  | 
(case (devar tye ty1, devar tye ty2) of  | 
|
| 2964 | 519  | 
(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
 | 
520  | 
if Term.eq_ix (v, w) then  | 
| 15797 | 521  | 
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
 | 
522  | 
else if Sorts.sort_le classes (S1, S2) then  | 
| 17412 | 523  | 
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
 | 
524  | 
else if Sorts.sort_le classes (S2, S1) then  | 
| 17412 | 525  | 
Vartab.update_new (v, (S1, U)) tye  | 
| 2964 | 526  | 
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
 | 
527  | 
let val S = gen_tyvar (Sorts.inter_sort classes (S1, S2)) in  | 
| 17412 | 528  | 
Vartab.update_new (v, (S1, S)) (Vartab.update_new (w, (S2, S)) tye)  | 
| 2964 | 529  | 
end  | 
530  | 
| (TVar (v, S), T) =>  | 
|
531  | 
if occurs v tye T then raise TUNIFY  | 
|
| 17412 | 532  | 
else meet (T, S) (Vartab.update_new (v, (S, T)) tye)  | 
| 2964 | 533  | 
| (T, TVar (v, S)) =>  | 
534  | 
if occurs v tye T then raise TUNIFY  | 
|
| 17412 | 535  | 
else meet (T, S) (Vartab.update_new (v, (S, T)) tye)  | 
| 2964 | 536  | 
| (Type (a, Ts), Type (b, Us)) =>  | 
537  | 
if a <> b then raise TUNIFY  | 
|
| 16885 | 538  | 
else unifs (Ts, Us) tye  | 
539  | 
| (T, U) => if T = U then tye else raise TUNIFY)  | 
|
540  | 
and unifs (T :: Ts, U :: Us) tye = unifs (Ts, Us) (unif (T, U) tye)  | 
|
541  | 
| unifs _ tye = tye;  | 
|
542  | 
in (unif TU tyenv, ! tyvar_count) end;  | 
|
| 0 | 543  | 
|
| 16946 | 544  | 
(*purely structural unification*)  | 
545  | 
fun raw_unify (ty1, ty2) tye =  | 
|
546  | 
(case (devar tye ty1, devar tye ty2) of  | 
|
| 32784 | 547  | 
(T as TVar (v, S1), TVar (w, S2)) =>  | 
| 
29269
 
5c25a2012975
moved term order operations to structure TermOrd (cf. Pure/term_ord.ML);
 
wenzelm 
parents: 
29260 
diff
changeset
 | 
548  | 
if Term.eq_ix (v, w) then  | 
| 16946 | 549  | 
if S1 = S2 then tye else tvar_clash v S1 S2  | 
| 17412 | 550  | 
else Vartab.update_new (w, (S2, T)) tye  | 
| 16946 | 551  | 
| (TVar (v, S), T) =>  | 
552  | 
if occurs v tye T then raise TUNIFY  | 
|
| 17412 | 553  | 
else Vartab.update_new (v, (S, T)) tye  | 
| 16946 | 554  | 
| (T, TVar (v, S)) =>  | 
555  | 
if occurs v tye T then raise TUNIFY  | 
|
| 17412 | 556  | 
else Vartab.update_new (v, (S, T)) tye  | 
| 16946 | 557  | 
| (Type (a, Ts), Type (b, Us)) =>  | 
558  | 
if a <> b then raise TUNIFY  | 
|
559  | 
else raw_unifys (Ts, Us) tye  | 
|
560  | 
| (T, U) => if T = U then tye else raise TUNIFY)  | 
|
561  | 
and raw_unifys (T :: Ts, U :: Us) tye = raw_unifys (Ts, Us) (raw_unify (T, U) tye)  | 
|
| 19696 | 562  | 
| raw_unifys ([], []) tye = tye  | 
563  | 
| raw_unifys _ _ = raise TUNIFY;  | 
|
| 16946 | 564  | 
|
| 17804 | 565  | 
(*fast unification filter*)  | 
566  | 
fun could_unify (Type (a, Ts), Type (b, Us)) = a = b andalso could_unifys (Ts, Us)  | 
|
567  | 
| could_unify (TFree (a, _), TFree (b, _)) = a = b  | 
|
568  | 
| could_unify (TVar _, _) = true  | 
|
569  | 
| could_unify (_, TVar _) = true  | 
|
570  | 
| could_unify _ = false  | 
|
571  | 
and could_unifys (T :: Ts, U :: Us) = could_unify (T, U) andalso could_unifys (Ts, Us)  | 
|
| 19696 | 572  | 
| could_unifys ([], []) = true  | 
573  | 
| could_unifys _ = false;  | 
|
| 450 | 574  | 
|
| 17804 | 575  | 
|
576  | 
(*equality with respect to a type environment*)  | 
|
| 32030 | 577  | 
fun equal_type tye (T, T') =  | 
| 16885 | 578  | 
(case (devar tye T, devar tye T') of  | 
| 16650 | 579  | 
(Type (s, Ts), Type (s', Ts')) =>  | 
| 32030 | 580  | 
s = s' andalso ListPair.all (equal_type tye) (Ts, Ts')  | 
| 16650 | 581  | 
| (U, U') => U = U');  | 
582  | 
||
| 32030 | 583  | 
fun eq_type tye =  | 
584  | 
if Vartab.is_empty tye then op = else equal_type tye;  | 
|
585  | 
||
| 450 | 586  | 
|
| 0 | 587  | 
|
| 
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
 | 
588  | 
(** 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
 | 
589  | 
|
| 
 
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
 | 
590  | 
(* 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
 | 
591  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
592  | 
fun add_class context (c, cs) tsig =  | 
| 19642 | 593  | 
tsig |> map_tsig (fn ((space, classes), default, types) =>  | 
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
594  | 
let  | 
| 
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
595  | 
val cs' = map (cert_class tsig) cs  | 
| 
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
596  | 
handle TYPE (msg, _, _) => error msg;  | 
| 
41254
 
78c3e472bb35
extra checking of name bindings for classes, types, consts;
 
wenzelm 
parents: 
39997 
diff
changeset
 | 
597  | 
val _ = Binding.check c;  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
598  | 
val (c', space') = space |> Name_Space.declare context true c;  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
599  | 
val classes' = classes |> Sorts.add_class (Context.pretty_generic context) (c', cs');  | 
| 19642 | 600  | 
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
 | 
601  | 
|
| 26669 | 602  | 
fun hide_class fully c = map_tsig (fn ((space, classes), default, types) =>  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33094 
diff
changeset
 | 
603  | 
((Name_Space.hide fully c space, classes), default, types));  | 
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
604  | 
|
| 
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  | 
|
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
606  | 
(* arities *)  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
607  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
608  | 
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
 | 
609  | 
let  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
610  | 
val _ =  | 
| 27302 | 611  | 
(case lookup_type tsig t of  | 
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33097 
diff
changeset
 | 
612  | 
SOME (LogicalType n) => if length Ss <> n then error (bad_nargs t) else ()  | 
| 
33094
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
613  | 
      | SOME _ => error ("Logical type constructor expected: " ^ quote t)
 | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
614  | 
| NONE => error (undecl_type t));  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
615  | 
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
 | 
616  | 
handle TYPE (msg, _, _) => error msg;  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
617  | 
val classes' = classes |> Sorts.add_arities pp ((t, map (fn c' => (c', Ss')) S'));  | 
| 19642 | 618  | 
in ((space, classes'), default, types) end);  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
619  | 
|
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
620  | 
|
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
621  | 
(* classrel *)  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
622  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
623  | 
fun add_classrel pp rel tsig =  | 
| 19642 | 624  | 
tsig |> map_tsig (fn ((space, classes), default, types) =>  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
625  | 
let  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
626  | 
val rel' = pairself (cert_class tsig) rel  | 
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
627  | 
handle TYPE (msg, _, _) => error msg;  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
628  | 
val classes' = classes |> Sorts.add_classrel pp rel';  | 
| 19642 | 629  | 
in ((space, classes'), default, types) end);  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
630  | 
|
| 
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
631  | 
|
| 
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
 | 
632  | 
(* 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
 | 
633  | 
|
| 19642 | 634  | 
fun set_defsort S tsig = tsig |> map_tsig (fn (classes, _, types) =>  | 
635  | 
(classes, cert_sort tsig S handle TYPE (msg, _, _) => error msg, types));  | 
|
| 0 | 636  | 
|
637  | 
||
| 
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
 | 
638  | 
(* 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
 | 
639  | 
|
| 
 
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
 | 
640  | 
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
 | 
641  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
642  | 
fun new_decl context (c, decl) types =  | 
| 
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
643  | 
(Binding.check c; #2 (Name_Space.define context true (c, decl) 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
 | 
644  | 
|
| 19642 | 645  | 
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
 | 
646  | 
let  | 
| 
 
b802d1804b77
replaced Sorts.DOMAIN by general Sorts.CLASS_ERROR;
 
wenzelm 
parents: 
19530 
diff
changeset
 | 
647  | 
val (space', tab') = f types;  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33094 
diff
changeset
 | 
648  | 
val _ = Name_Space.intern space' "dummy" = "dummy" orelse  | 
| 
21858
 
05f57309170c
avoid conflict with Alice keywords: renamed pack -> implode, unpack -> explode, any -> many, avoided assert;
 
wenzelm 
parents: 
21116 
diff
changeset
 | 
649  | 
error "Illegal declaration of dummy type";  | 
| 19642 | 650  | 
in (classes, default, (space', tab')) end);  | 
| 2964 | 651  | 
|
| 14989 | 652  | 
fun syntactic types (Type (c, Ts)) =  | 
| 
33173
 
b8ca12f6681a
eliminated obsolete tags for types/consts -- now handled via name space, in strongly typed fashion;
 
wenzelm 
parents: 
33097 
diff
changeset
 | 
653  | 
(case Symtab.lookup types c of SOME Nonterminal => true | _ => false)  | 
| 14989 | 654  | 
orelse exists (syntactic types) Ts  | 
655  | 
| syntactic _ _ = false;  | 
|
656  | 
||
| 27302 | 657  | 
in  | 
658  | 
||
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
659  | 
fun add_type context (c, n) =  | 
| 
42381
 
309ec68442c6
added Binding.print convenience, which includes quote already;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
660  | 
  if n < 0 then error ("Bad type constructor declaration " ^ Binding.print c)
 | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
661  | 
else map_types (new_decl context (c, LogicalType n));  | 
| 27302 | 662  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
663  | 
fun add_abbrev context (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
 | 
664  | 
let  | 
| 
30343
 
79f022df8527
replace old bstring by binding for logical primitives: class, type, const etc.;
 
wenzelm 
parents: 
29275 
diff
changeset
 | 
665  | 
fun err msg =  | 
| 
42381
 
309ec68442c6
added Binding.print convenience, which includes quote already;
 
wenzelm 
parents: 
42375 
diff
changeset
 | 
666  | 
      cat_error msg ("The error(s) above occurred in type abbreviation " ^ Binding.print a);
 | 
| 
24274
 
cb9236269af1
type mode: models certification mode (default, syntax, abbrev);
 
wenzelm 
parents: 
23655 
diff
changeset
 | 
667  | 
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
 | 
668  | 
handle TYPE (msg, _, _) => err msg;  | 
| 
33094
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
669  | 
val _ =  | 
| 
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
670  | 
(case duplicates (op =) vs of  | 
| 
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
671  | 
[] => []  | 
| 
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
672  | 
      | dups => err ("Duplicate variables on lhs: " ^ commas_quote dups));
 | 
| 
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
673  | 
val _ =  | 
| 
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
674  | 
(case subtract (op =) vs (map #1 (Term.add_tfreesT rhs' [])) of  | 
| 
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
675  | 
[] => []  | 
| 
 
ef0d77b1e687
eliminated separate stamp -- NameSpace.define/merge etc. ensure uniqueness already;
 
wenzelm 
parents: 
32784 
diff
changeset
 | 
676  | 
      | extras => err ("Extra variables on rhs: " ^ commas_quote extras));
 | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
677  | 
in types |> new_decl context (a, Abbreviation (vs, rhs', syntactic (#2 types) rhs')) end);  | 
| 0 | 678  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
679  | 
fun add_nonterminal context = map_types o new_decl context 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
 | 
680  | 
|
| 
 
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
 | 
681  | 
end;  | 
| 0 | 682  | 
|
| 26669 | 683  | 
fun hide_type fully c = map_tsig (fn (classes, default, (space, types)) =>  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33094 
diff
changeset
 | 
684  | 
(classes, default, (Name_Space.hide fully c space, types)));  | 
| 
16370
 
033d890fe91f
name space of classes and types maintained in tsig;
 
wenzelm 
parents: 
16340 
diff
changeset
 | 
685  | 
|
| 3790 | 686  | 
|
| 
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
 | 
687  | 
(* 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
 | 
688  | 
|
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
689  | 
fun merge_tsig 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
 | 
690  | 
let  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
691  | 
    val (TSig {classes = (space1, classes1), default = default1, types = types1,
 | 
| 26641 | 692  | 
log_types = _}) = tsig1;  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
693  | 
    val (TSig {classes = (space2, classes2), default = default2, types = types2,
 | 
| 26641 | 694  | 
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
 | 
695  | 
|
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33094 
diff
changeset
 | 
696  | 
val space' = Name_Space.merge (space1, space2);  | 
| 
47005
 
421760a1efe7
maintain generic context naming in structure Name_Space (NB: empty = default_naming, init = local_naming);
 
wenzelm 
parents: 
46649 
diff
changeset
 | 
697  | 
val classes' = Sorts.merge_algebra pp (classes1, classes2);  | 
| 
19515
 
9f650083da65
build classes/arities: refer to operations in sorts.ML;
 
wenzelm 
parents: 
19464 
diff
changeset
 | 
698  | 
val default' = Sorts.inter_sort classes' (default1, default2);  | 
| 
33095
 
bbd52d2f8696
renamed NameSpace to Name_Space -- also to emphasize its subtle change in semantics;
 
wenzelm 
parents: 
33094 
diff
changeset
 | 
699  | 
val types' = Name_Space.merge_tables (types1, types2);  | 
| 19642 | 700  | 
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
 | 
701  | 
|
| 0 | 702  | 
end;  |