| author | wenzelm | 
| Mon, 12 Dec 2022 13:28:18 +0100 | |
| changeset 76627 | 2542ea382215 | 
| parent 69575 | f77cc54f6d47 | 
| child 78652 | 3c57995c255c | 
| permissions | -rw-r--r-- | 
| 15797 | 1  | 
(* Title: Pure/unify.ML  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15797 
diff
changeset
 | 
2  | 
Author: Lawrence C Paulson, Cambridge University Computer Laboratory  | 
| 0 | 3  | 
Copyright Cambridge University 1992  | 
4  | 
||
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15797 
diff
changeset
 | 
5  | 
Higher-Order Unification.  | 
| 0 | 6  | 
|
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15797 
diff
changeset
 | 
7  | 
Types as well as terms are unified. The outermost functions assume  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15797 
diff
changeset
 | 
8  | 
the terms to be unified already have the same type. In resolution,  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15797 
diff
changeset
 | 
9  | 
this is assured because both have type "prop".  | 
| 0 | 10  | 
*)  | 
11  | 
||
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15797 
diff
changeset
 | 
12  | 
signature UNIFY =  | 
| 
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15797 
diff
changeset
 | 
13  | 
sig  | 
| 
24178
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
14  | 
val trace_bound: int Config.T  | 
| 
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
15  | 
val search_bound: int Config.T  | 
| 
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
16  | 
val trace_simp: bool Config.T  | 
| 
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
17  | 
val trace_types: bool Config.T  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
18  | 
val hounifiers: Context.generic * Envir.env * ((term * term) list) ->  | 
| 52126 | 19  | 
(Envir.env * (term * term) list) Seq.seq  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
20  | 
val unifiers: Context.generic * Envir.env * ((term * term) list) ->  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15797 
diff
changeset
 | 
21  | 
(Envir.env * (term * term) list) Seq.seq  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
22  | 
val smash_unifiers: Context.generic -> (term * term) list -> Envir.env -> Envir.env Seq.seq  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15797 
diff
changeset
 | 
23  | 
end  | 
| 0 | 24  | 
|
| 19864 | 25  | 
structure Unify : UNIFY =  | 
| 0 | 26  | 
struct  | 
27  | 
||
28  | 
(*Unification options*)  | 
|
29  | 
||
| 
24178
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
30  | 
(*tracing starts above this depth, 0 for full*)  | 
| 69575 | 31  | 
val trace_bound = Config.declare_int ("unify_trace_bound", \<^here>) (K 50);
 | 
| 
24178
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
32  | 
|
| 
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
33  | 
(*unification quits above this depth*)  | 
| 69575 | 34  | 
val search_bound = Config.declare_int ("unify_search_bound", \<^here>) (K 60);
 | 
| 
24178
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
35  | 
|
| 
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
36  | 
(*print dpairs before calling SIMPL*)  | 
| 69575 | 37  | 
val trace_simp = Config.declare_bool ("unify_trace_simp", \<^here>) (K false);
 | 
| 
24178
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
38  | 
|
| 
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
39  | 
(*announce potential incompleteness of type unification*)  | 
| 69575 | 40  | 
val trace_types = Config.declare_bool ("unify_trace_types", \<^here>) (K false);
 | 
| 
24178
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
41  | 
|
| 0 | 42  | 
|
| 52126 | 43  | 
type binderlist = (string * typ) list;  | 
| 0 | 44  | 
|
45  | 
type dpair = binderlist * term * term;  | 
|
46  | 
||
| 
12231
 
4a25f04bea61
Moved head_norm and fastype from unify.ML to envir.ML
 
berghofe 
parents: 
8406 
diff
changeset
 | 
47  | 
fun fastype env (Ts, t) = Envir.fastype env (map snd Ts) t;  | 
| 0 | 48  | 
|
49  | 
||
| 32032 | 50  | 
(* eta normal form *)  | 
51  | 
||
52  | 
fun eta_norm env =  | 
|
53  | 
let  | 
|
54  | 
val tyenv = Envir.type_env env;  | 
|
55  | 
    fun etif (Type ("fun", [T, U]), t) =
 | 
|
56  | 
          Abs ("", T, etif (U, incr_boundvars 1 t $ Bound 0))
 | 
|
57  | 
| etif (TVar v, t) =  | 
|
58  | 
(case Type.lookup tyenv v of  | 
|
59  | 
NONE => t  | 
|
60  | 
| SOME T => etif (T, t))  | 
|
61  | 
| etif (_, t) = t;  | 
|
62  | 
fun eta_nm (rbinder, Abs (a, T, body)) =  | 
|
63  | 
Abs (a, T, eta_nm ((a, T) :: rbinder, body))  | 
|
64  | 
| eta_nm (rbinder, t) = etif (fastype env (rbinder, t), t);  | 
|
| 0 | 65  | 
in eta_nm end;  | 
66  | 
||
67  | 
||
68  | 
(*OCCURS CHECK  | 
|
| 19864 | 69  | 
Does the uvar occur in the term t?  | 
| 0 | 70  | 
two forms of search, for whether there is a rigid path to the current term.  | 
71  | 
"seen" is list of variables passed thru, is a memo variable for sharing.  | 
|
| 15797 | 72  | 
This version searches for nonrigid occurrence, returns true if found.  | 
73  | 
Since terms may contain variables with same name and different types,  | 
|
74  | 
the occurs check must ignore the types of variables. This avoids  | 
|
75  | 
that ?x::?'a is unified with f(?x::T), which may lead to a cyclic  | 
|
76  | 
substitution when ?'a is instantiated with T later. *)  | 
|
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
77  | 
fun occurs_terms (seen: indexname list Unsynchronized.ref,  | 
| 19864 | 78  | 
env: Envir.env, v: indexname, ts: term list): bool =  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
79  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
80  | 
fun occurs [] = false  | 
| 37636 | 81  | 
| occurs (t :: ts) = occur t orelse occurs ts  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
82  | 
and occur (Const _) = false  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
83  | 
| occur (Bound _) = false  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
84  | 
| occur (Free _) = false  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
85  | 
| occur (Var (w, T)) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
86  | 
if member (op =) (!seen) w then false  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
87  | 
else if Term.eq_ix (v, w) then true  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
88  | 
(*no need to lookup: v has no assignment*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
89  | 
else  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
90  | 
(seen := w :: !seen;  | 
| 51700 | 91  | 
case Envir.lookup env (w, T) of  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
92  | 
NONE => false  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
93  | 
| SOME t => occur t)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
94  | 
| occur (Abs (_, _, body)) = occur body  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
95  | 
| occur (f $ t) = occur t orelse occur f;  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
96  | 
in occurs ts end;  | 
| 0 | 97  | 
|
98  | 
||
| 52126 | 99  | 
(* f a1 ... an ----> f using the assignments*)  | 
100  | 
fun head_of_in env t =  | 
|
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
101  | 
(case t of  | 
| 52126 | 102  | 
f $ _ => head_of_in env f  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
103  | 
| Var vT =>  | 
| 51700 | 104  | 
(case Envir.lookup env vT of  | 
| 52126 | 105  | 
SOME u => head_of_in env u  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
106  | 
| NONE => t)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
107  | 
| _ => t);  | 
| 0 | 108  | 
|
109  | 
||
110  | 
datatype occ = NoOcc | Nonrigid | Rigid;  | 
|
111  | 
||
112  | 
(* Rigid occur check  | 
|
113  | 
Returns Rigid if it finds a rigid occurrence of the variable,  | 
|
114  | 
Nonrigid if it finds a nonrigid path to the variable.  | 
|
115  | 
NoOcc otherwise.  | 
|
116  | 
Continues searching for a rigid occurrence even if it finds a nonrigid one.  | 
|
117  | 
||
118  | 
Condition for detecting non-unifable terms: [ section 5.3 of Huet (1975) ]  | 
|
119  | 
a rigid path to the variable, appearing with no arguments.  | 
|
120  | 
Here completeness is sacrificed in order to reduce danger of divergence:  | 
|
121  | 
reject ALL rigid paths to the variable.  | 
|
| 19864 | 122  | 
Could check for rigid paths to bound variables that are out of scope.  | 
| 0 | 123  | 
Not necessary because the assignment test looks at variable's ENTIRE rbinder.  | 
124  | 
||
125  | 
Treatment of head(arg1,...,argn):  | 
|
126  | 
If head is a variable then no rigid path, switch to nonrigid search  | 
|
| 19864 | 127  | 
for arg1,...,argn.  | 
128  | 
If head is an abstraction then possibly no rigid path (head could be a  | 
|
| 0 | 129  | 
constant function) so again use nonrigid search. Happens only if  | 
| 19864 | 130  | 
term is not in normal form.  | 
| 0 | 131  | 
|
132  | 
Warning: finds a rigid occurrence of ?f in ?f(t).  | 
|
133  | 
Should NOT be called in this case: there is a flex-flex unifier  | 
|
134  | 
*)  | 
|
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
135  | 
fun rigid_occurs_term (seen: indexname list Unsynchronized.ref, env, v: indexname, t) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
136  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
137  | 
fun nonrigid t =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
138  | 
if occurs_terms (seen, env, v, [t]) then Nonrigid  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
139  | 
else NoOcc  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
140  | 
fun occurs [] = NoOcc  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
141  | 
| occurs (t :: ts) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
142  | 
(case occur t of  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
143  | 
Rigid => Rigid  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
144  | 
| oc => (case occurs ts of NoOcc => oc | oc2 => oc2))  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
145  | 
and occomb (f $ t) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
146  | 
(case occur t of  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
147  | 
Rigid => Rigid  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
148  | 
| oc => (case occomb f of NoOcc => oc | oc2 => oc2))  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
149  | 
| occomb t = occur t  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
150  | 
and occur (Const _) = NoOcc  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
151  | 
| occur (Bound _) = NoOcc  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
152  | 
| occur (Free _) = NoOcc  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
153  | 
| occur (Var (w, T)) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
154  | 
if member (op =) (!seen) w then NoOcc  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
155  | 
else if Term.eq_ix (v, w) then Rigid  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
156  | 
else  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
157  | 
(seen := w :: !seen;  | 
| 51700 | 158  | 
case Envir.lookup env (w, T) of  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
159  | 
NONE => NoOcc  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
160  | 
| SOME t => occur t)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
161  | 
| occur (Abs (_, _, body)) = occur body  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
162  | 
| occur (t as f $ _) = (*switch to nonrigid search?*)  | 
| 52126 | 163  | 
(case head_of_in env f of  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
164  | 
Var (w,_) => (*w is not assigned*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
165  | 
if Term.eq_ix (v, w) then Rigid  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
166  | 
else nonrigid t  | 
| 37636 | 167  | 
| Abs _ => nonrigid t (*not in normal form*)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
168  | 
| _ => occomb t)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
169  | 
in occur t end;  | 
| 0 | 170  | 
|
171  | 
||
| 19864 | 172  | 
exception CANTUNIFY; (*Signals non-unifiability. Does not signal errors!*)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
173  | 
exception ASSIGN; (*Raised if not an assignment*)  | 
| 0 | 174  | 
|
175  | 
||
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
176  | 
fun unify_types context TU env =  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
177  | 
Pattern.unify_types context TU env handle Pattern.Unif => raise CANTUNIFY;  | 
| 0 | 178  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
179  | 
fun test_unify_types context (T, U) env =  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
180  | 
let  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
181  | 
fun trace () =  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
182  | 
if Context_Position.is_visible_generic context then  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
183  | 
let val str_of = Syntax.string_of_typ (Context.proof_of context)  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
184  | 
        in tracing ("Potential loss of completeness: " ^ str_of U ^ " = " ^ str_of T) end
 | 
| 
52701
 
51dfdcd88e84
guard unify tracing via visible status of global theory;
 
wenzelm 
parents: 
52698 
diff
changeset
 | 
185  | 
else ();  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
186  | 
val env' = unify_types context (T, U) env;  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
187  | 
in if is_TVar T orelse is_TVar U then trace () else (); env' end;  | 
| 0 | 188  | 
|
189  | 
(*Is the term eta-convertible to a single variable with the given rbinder?  | 
|
190  | 
Examples: ?a ?f(B.0) ?g(B.1,B.0)  | 
|
191  | 
Result is var a for use in SIMPL. *)  | 
|
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
192  | 
fun get_eta_var ([], _, Var vT) = vT  | 
| 
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
193  | 
| get_eta_var (_::rbinder, n, f $ Bound i) =  | 
| 
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
194  | 
if n = i then get_eta_var (rbinder, n + 1, f)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
195  | 
else raise ASSIGN  | 
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
196  | 
| get_eta_var _ = raise ASSIGN;  | 
| 0 | 197  | 
|
198  | 
||
199  | 
(*Solve v=u by assignment -- "fixedpoint" to Huet -- if v not in u.  | 
|
200  | 
If v occurs rigidly then nonunifiable.  | 
|
201  | 
If v occurs nonrigidly then must use full algorithm. *)  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
202  | 
fun assignment context (rbinder, t, u) env =  | 
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
203  | 
let val vT as (v,T) = get_eta_var (rbinder, 0, t) in  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
204  | 
(case rigid_occurs_term (Unsynchronized.ref [], env, v, u) of  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
205  | 
NoOcc =>  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
206  | 
let val env = unify_types context (Envir.body_type env T, fastype env (rbinder, u)) env  | 
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
207  | 
in Envir.update (vT, Logic.rlist_abs (rbinder, u)) env end  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
208  | 
| Nonrigid => raise ASSIGN  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
209  | 
| Rigid => raise CANTUNIFY)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
210  | 
end;  | 
| 0 | 211  | 
|
212  | 
||
213  | 
(*Extends an rbinder with a new disagreement pair, if both are abstractions.  | 
|
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
214  | 
Tries to unify types of the bound variables!  | 
| 0 | 215  | 
Checks that binders have same length, since terms should be eta-normal;  | 
216  | 
if not, raises TERM, probably indicating type mismatch.  | 
|
| 19864 | 217  | 
Uses variable a (unless the null string) to preserve user's naming.*)  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
218  | 
fun new_dpair context (rbinder, Abs (a, T, body1), Abs (b, U, body2)) env =  | 
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
219  | 
let  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
220  | 
val env' = unify_types context (T, U) env;  | 
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
221  | 
val c = if a = "" then b else a;  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
222  | 
in new_dpair context ((c,T) :: rbinder, body1, body2) env' end  | 
| 52126 | 223  | 
  | new_dpair _ (_, Abs _, _) _ = raise TERM ("new_dpair", [])
 | 
224  | 
  | new_dpair _ (_, _, Abs _) _ = raise TERM ("new_dpair", [])
 | 
|
225  | 
| new_dpair _ (rbinder, t1, t2) env = ((rbinder, t1, t2), env);  | 
|
| 0 | 226  | 
|
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
227  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
228  | 
fun head_norm_dpair context (env, (rbinder, t, u)) : dpair * Envir.env =  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
229  | 
new_dpair context (rbinder,  | 
| 19864 | 230  | 
eta_norm env (rbinder, Envir.head_norm env t),  | 
| 52126 | 231  | 
eta_norm env (rbinder, Envir.head_norm env u)) env;  | 
| 0 | 232  | 
|
233  | 
||
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
234  | 
|
| 0 | 235  | 
(*flexflex: the flex-flex pairs, flexrigid: the flex-rigid pairs  | 
236  | 
Does not perform assignments for flex-flex pairs:  | 
|
| 646 | 237  | 
may create nonrigid paths, which prevent other assignments.  | 
| 67725 | 238  | 
Does not even identify Vars in dpairs such as ?a \<equiv>\<^sup>? ?b; an attempt to  | 
| 646 | 239  | 
do so caused numerous problems with no compensating advantage.  | 
240  | 
*)  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
241  | 
fun SIMPL0 context dp0 (env,flexflex,flexrigid) : Envir.env * dpair list * dpair list =  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
242  | 
let  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
243  | 
val (dp as (rbinder, t, u), env) = head_norm_dpair context (env, dp0);  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
244  | 
fun SIMRANDS (f $ t, g $ u, env) =  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
245  | 
SIMPL0 context (rbinder, t, u) (SIMRANDS (f, g, env))  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
246  | 
| SIMRANDS (t as _$_, _, _) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
247  | 
          raise TERM ("SIMPL: operands mismatch", [t, u])
 | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
248  | 
| SIMRANDS (t, u as _ $ _, _) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
249  | 
          raise TERM ("SIMPL: operands mismatch", [t, u])
 | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
250  | 
| SIMRANDS (_, _, env) = (env, flexflex, flexrigid);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
251  | 
in  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
252  | 
(case (head_of t, head_of u) of  | 
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
253  | 
(Var (_, T), Var (_, U)) =>  | 
| 
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
254  | 
let  | 
| 52221 | 255  | 
val T' = Envir.body_type env T and U' = Envir.body_type env U;  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
256  | 
val env = unify_types context (T', U') env;  | 
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
257  | 
in (env, dp :: flexflex, flexrigid) end  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
258  | 
| (Var _, _) =>  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
259  | 
((assignment context (rbinder,t,u) env, flexflex, flexrigid)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
260  | 
handle ASSIGN => (env, flexflex, dp :: flexrigid))  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
261  | 
| (_, Var _) =>  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
262  | 
((assignment context (rbinder, u, t) env, flexflex, flexrigid)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
263  | 
handle ASSIGN => (env, flexflex, (rbinder, u, t) :: flexrigid))  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
264  | 
| (Const (a, T), Const (b, U)) =>  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
265  | 
if a = b then SIMRANDS (t, u, unify_types context (T, U) env)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
266  | 
else raise CANTUNIFY  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
267  | 
| (Bound i, Bound j) =>  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
268  | 
if i = j then SIMRANDS (t, u, env) else raise CANTUNIFY  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
269  | 
| (Free (a, T), Free (b, U)) =>  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
270  | 
if a = b then SIMRANDS (t, u, unify_types context (T, U) env)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
271  | 
else raise CANTUNIFY  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
272  | 
| _ => raise CANTUNIFY)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
273  | 
end;  | 
| 0 | 274  | 
|
275  | 
||
276  | 
(* changed(env,t) checks whether the head of t is a variable assigned in env*)  | 
|
| 52126 | 277  | 
fun changed env (f $ _) = changed env f  | 
278  | 
| changed env (Var v) = (case Envir.lookup env v of NONE => false | _ => true)  | 
|
279  | 
| changed _ _ = false;  | 
|
| 0 | 280  | 
|
281  | 
||
282  | 
(*Recursion needed if any of the 'head variables' have been updated  | 
|
283  | 
Clever would be to re-do just the affected dpairs*)  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
284  | 
fun SIMPL context (env,dpairs) : Envir.env * dpair list * dpair list =  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
285  | 
let  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
286  | 
val all as (env', flexflex, flexrigid) = fold_rev (SIMPL0 context) dpairs (env, [], []);  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
287  | 
val dps = flexrigid @ flexflex;  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
288  | 
in  | 
| 52126 | 289  | 
if exists (fn (_, t, u) => changed env' t orelse changed env' u) dps  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
290  | 
then SIMPL context (env', dps) else all  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
291  | 
end;  | 
| 0 | 292  | 
|
293  | 
||
| 19864 | 294  | 
(*Makes the terms E1,...,Em, where Ts = [T...Tm].  | 
| 0 | 295  | 
Each Ei is ?Gi(B.(n-1),...,B.0), and has type Ti  | 
296  | 
The B.j are bound vars of binder.  | 
|
| 19864 | 297  | 
The terms are not made in eta-normal-form, SIMPL does that later.  | 
| 0 | 298  | 
If done here, eta-expansion must be recursive in the arguments! *)  | 
| 37636 | 299  | 
fun make_args _ (_, env, []) = (env, []) (*frequent case*)  | 
| 0 | 300  | 
| make_args name (binder: typ list, env, Ts) : Envir.env * term list =  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
301  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
302  | 
fun funtype T = binder ---> T;  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
303  | 
val (env', vars) = Envir.genvars name (env, map funtype Ts);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
304  | 
in (env', map (fn var => Logic.combound (var, 0, length binder)) vars) end;  | 
| 0 | 305  | 
|
306  | 
||
| 
46219
 
426ed18eba43
discontinued old-style Term.list_abs in favour of plain Term.abs;
 
wenzelm 
parents: 
42279 
diff
changeset
 | 
307  | 
(*Abstraction over a list of types*)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
308  | 
fun types_abs ([], u) = u  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
309  | 
  | types_abs (T :: Ts, u) = Abs ("", T, types_abs (Ts, u));
 | 
| 0 | 310  | 
|
311  | 
(*Abstraction over the binder of a type*)  | 
|
| 52221 | 312  | 
fun type_abs (env, T, t) = types_abs (Envir.binder_types env T, t);  | 
| 0 | 313  | 
|
314  | 
||
315  | 
(*MATCH taking "big steps".  | 
|
316  | 
Copies u into the Var v, using projection on targs or imitation.  | 
|
317  | 
A projection is allowed unless SIMPL raises an exception.  | 
|
318  | 
Allocates new variables in projection on a higher-order argument,  | 
|
319  | 
or if u is a variable (flex-flex dpair).  | 
|
320  | 
Returns long sequence of every way of copying u, for backtracking  | 
|
321  | 
For example, projection in ?b'(?a) may be wrong if other dpairs constrain ?a.  | 
|
| 19864 | 322  | 
The order for trying projections is crucial in ?b'(?a)  | 
| 0 | 323  | 
NB "vname" is only used in the call to make_args!! *)  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
324  | 
fun matchcopy context vname =  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
325  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
326  | 
fun mc (rbinder, targs, u, ed as (env, dpairs)) : (term * (Envir.env * dpair list)) Seq.seq =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
327  | 
let  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
328  | 
val trace_types = Config.get_generic context trace_types;  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
329  | 
(*Produce copies of uarg and cons them in front of uargs*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
330  | 
fun copycons uarg (uargs, (env, dpairs)) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
331  | 
Seq.map (fn (uarg', ed') => (uarg' :: uargs, ed'))  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
332  | 
(mc (rbinder, targs,eta_norm env (rbinder, Envir.head_norm env uarg),  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
333  | 
(env, dpairs)));  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
334  | 
(*Produce sequence of all possible ways of copying the arg list*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
335  | 
fun copyargs [] = Seq.cons ([], ed) Seq.empty  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
336  | 
| copyargs (uarg :: uargs) = Seq.maps (copycons uarg) (copyargs uargs);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
337  | 
val (uhead, uargs) = strip_comb u;  | 
| 52221 | 338  | 
val base = Envir.body_type env (fastype env (rbinder, uhead));  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
339  | 
fun joinargs (uargs', ed') = (list_comb (uhead, uargs'), ed');  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
340  | 
(*attempt projection on argument with given typ*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
341  | 
val Ts = map (curry (fastype env) rbinder) targs;  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
342  | 
fun projenv (head, (Us, bary), targ, tail) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
343  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
344  | 
val env =  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
345  | 
if trace_types then test_unify_types context (base, bary) env  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
346  | 
else unify_types context (base, bary) env  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
347  | 
in  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
348  | 
Seq.make (fn () =>  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
349  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
350  | 
val (env', args) = make_args vname (Ts, env, Us);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
351  | 
(*higher-order projection: plug in targs for bound vars*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
352  | 
fun plugin arg = list_comb (head_of arg, targs);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
353  | 
val dp = (rbinder, list_comb (targ, map plugin args), u);  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
354  | 
val (env2, frigid, fflex) = SIMPL context (env', dp :: dpairs);  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
355  | 
(*may raise exception CANTUNIFY*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
356  | 
in  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
357  | 
SOME ((list_comb (head, args), (env2, frigid @ fflex)), tail)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
358  | 
end handle CANTUNIFY => Seq.pull tail)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
359  | 
end handle CANTUNIFY => tail;  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
360  | 
(*make a list of projections*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
361  | 
fun make_projs (T::Ts, targ::targs) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
362  | 
(Bound(length Ts), T, targ) :: make_projs (Ts,targs)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
363  | 
| make_projs ([],[]) = []  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
364  | 
          | make_projs _ = raise TERM ("make_projs", u::targs);
 | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
365  | 
(*try projections and imitation*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
366  | 
fun matchfun ((bvar,T,targ)::projs) =  | 
| 52221 | 367  | 
(projenv(bvar, Envir.strip_type env T, targ, matchfun projs))  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
368  | 
| matchfun [] = (*imitation last of all*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
369  | 
(case uhead of  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
370  | 
Const _ => Seq.map joinargs (copyargs uargs)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
371  | 
| Free _ => Seq.map joinargs (copyargs uargs)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
372  | 
| _ => Seq.empty) (*if Var, would be a loop!*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
373  | 
in  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
374  | 
(case uhead of  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
375  | 
Abs (a, T, body) =>  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
376  | 
Seq.map (fn (body', ed') => (Abs (a, T, body'), ed'))  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
377  | 
(mc ((a, T) :: rbinder, (map (incr_boundvars 1) targs) @ [Bound 0], body, ed))  | 
| 37636 | 378  | 
| Var (w, _) =>  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
379  | 
(*a flex-flex dpair: make variable for t*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
380  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
381  | 
val (env', newhd) = Envir.genvar (#1 w) (env, Ts ---> base);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
382  | 
val tabs = Logic.combound (newhd, 0, length Ts);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
383  | 
val tsub = list_comb (newhd, targs);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
384  | 
in Seq.single (tabs, (env', (rbinder, tsub, u) :: dpairs)) end  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
385  | 
| _ => matchfun (rev (make_projs (Ts, targs))))  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
386  | 
end;  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
387  | 
in mc end;  | 
| 0 | 388  | 
|
389  | 
||
390  | 
(*Call matchcopy to produce assignments to the variable in the dpair*)  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
391  | 
fun MATCH context (env, (rbinder, t, u), dpairs) : (Envir.env * dpair list) Seq.seq =  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
392  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
393  | 
val (Var (vT as (v, T)), targs) = strip_comb t;  | 
| 52221 | 394  | 
val Ts = Envir.binder_types env T;  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
395  | 
fun new_dset (u', (env', dpairs')) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
396  | 
(*if v was updated to s, must unify s with u' *)  | 
| 51700 | 397  | 
(case Envir.lookup env' vT of  | 
398  | 
NONE => (Envir.update (vT, types_abs (Ts, u')) env', dpairs')  | 
|
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
399  | 
| SOME s => (env', ([], s, types_abs (Ts, u')) :: dpairs'));  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
400  | 
in  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
401  | 
Seq.map new_dset (matchcopy context (#1 v) (rbinder, targs, u, (env, dpairs)))  | 
| 0 | 402  | 
end;  | 
403  | 
||
404  | 
||
405  | 
||
406  | 
(**** Flex-flex processing ****)  | 
|
407  | 
||
| 19864 | 408  | 
(*At end of unification, do flex-flex assignments like ?a -> ?f(?b)  | 
| 0 | 409  | 
Attempts to update t with u, raising ASSIGN if impossible*)  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
410  | 
fun ff_assign context (env, rbinder, t, u) : Envir.env =  | 
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
411  | 
let val vT as (v, T) = get_eta_var (rbinder, 0, t) in  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
412  | 
if occurs_terms (Unsynchronized.ref [], env, v, [u]) then raise ASSIGN  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
413  | 
else  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
414  | 
let val env = unify_types context (Envir.body_type env T, fastype env (rbinder, u)) env  | 
| 
52220
 
c4264f71dc3b
backout 3b9c31867707 -- too risky to "amend" modules from 25 years ago that don't handle Vars with different types;
 
wenzelm 
parents: 
52179 
diff
changeset
 | 
415  | 
in Envir.vupdate (vT, Logic.rlist_abs (rbinder, u)) env end  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
416  | 
end;  | 
| 0 | 417  | 
|
418  | 
||
| 
37720
 
50a9e2fa4f6b
Unification (flexflex) bug fix; making "auto" deterministic
 
paulson 
parents: 
37637 
diff
changeset
 | 
419  | 
(*If an argument contains a banned Bound, then it should be deleted.  | 
| 
 
50a9e2fa4f6b
Unification (flexflex) bug fix; making "auto" deterministic
 
paulson 
parents: 
37637 
diff
changeset
 | 
420  | 
But if the only path is flexible, this is difficult; the code gives up!  | 
| 67725 | 421  | 
In \<lambda>x y. ?a x \<equiv>\<^sup>? \<lambda>x y. ?b (?c y) should we instantiate ?b or ?c *)  | 
| 
37720
 
50a9e2fa4f6b
Unification (flexflex) bug fix; making "auto" deterministic
 
paulson 
parents: 
37637 
diff
changeset
 | 
422  | 
exception CHANGE_FAIL; (*flexible occurrence of banned variable, or other reason to quit*)  | 
| 
 
50a9e2fa4f6b
Unification (flexflex) bug fix; making "auto" deterministic
 
paulson 
parents: 
37637 
diff
changeset
 | 
423  | 
|
| 
 
50a9e2fa4f6b
Unification (flexflex) bug fix; making "auto" deterministic
 
paulson 
parents: 
37637 
diff
changeset
 | 
424  | 
|
| 0 | 425  | 
(*Flex argument: a term, its type, and the index that refers to it.*)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
426  | 
type flarg = {t: term, T: typ, j: int};
 | 
| 0 | 427  | 
|
428  | 
(*Form the arguments into records for deletion/sorting.*)  | 
|
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
429  | 
fun flexargs ([], [], []) = [] : flarg list  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
430  | 
  | flexargs (j :: js, t :: ts, T :: Ts) = {j = j, t = t, T = T} :: flexargs (js, ts, Ts)
 | 
| 
37720
 
50a9e2fa4f6b
Unification (flexflex) bug fix; making "auto" deterministic
 
paulson 
parents: 
37637 
diff
changeset
 | 
431  | 
| flexargs _ = raise CHANGE_FAIL;  | 
| 
41422
 
8a765db7e0f8
uniform treatment of type vs. term environment (cf. b654fa27fbc4);
 
wenzelm 
parents: 
39997 
diff
changeset
 | 
432  | 
(*We give up if we see a variable of function type not applied to a full list of  | 
| 
 
8a765db7e0f8
uniform treatment of type vs. term environment (cf. b654fa27fbc4);
 
wenzelm 
parents: 
39997 
diff
changeset
 | 
433  | 
arguments (remember, this code assumes that terms are fully eta-expanded). This situation  | 
| 
37720
 
50a9e2fa4f6b
Unification (flexflex) bug fix; making "auto" deterministic
 
paulson 
parents: 
37637 
diff
changeset
 | 
434  | 
can occur if a type variable is instantiated with a function type.  | 
| 
 
50a9e2fa4f6b
Unification (flexflex) bug fix; making "auto" deterministic
 
paulson 
parents: 
37637 
diff
changeset
 | 
435  | 
*)  | 
| 0 | 436  | 
|
| 
651
 
4b0455fbcc49
Pure/Unify/IMPROVING "CLEANING" OF FLEX-FLEX PAIRS: Old code would refuse
 
lcp 
parents: 
646 
diff
changeset
 | 
437  | 
(*Check whether the 'banned' bound var indices occur rigidly in t*)  | 
| 19864 | 438  | 
fun rigid_bound (lev, banned) t =  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
439  | 
let val (head,args) = strip_comb t in  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
440  | 
(case head of  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
441  | 
Bound i =>  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
442  | 
member (op =) banned (i - lev) orelse exists (rigid_bound (lev, banned)) args  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
443  | 
| Var _ => false (*no rigid occurrences here!*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
444  | 
| Abs (_, _, u) =>  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
445  | 
rigid_bound (lev + 1, banned) u orelse  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
446  | 
exists (rigid_bound (lev, banned)) args  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
447  | 
| _ => exists (rigid_bound (lev, banned)) args)  | 
| 0 | 448  | 
end;  | 
449  | 
||
| 
651
 
4b0455fbcc49
Pure/Unify/IMPROVING "CLEANING" OF FLEX-FLEX PAIRS: Old code would refuse
 
lcp 
parents: 
646 
diff
changeset
 | 
450  | 
(*Squash down indices at level >=lev to delete the banned from a term.*)  | 
| 
 
4b0455fbcc49
Pure/Unify/IMPROVING "CLEANING" OF FLEX-FLEX PAIRS: Old code would refuse
 
lcp 
parents: 
646 
diff
changeset
 | 
451  | 
fun change_bnos banned =  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
452  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
453  | 
fun change lev (Bound i) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
454  | 
if i < lev then Bound i  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
455  | 
else if member (op =) banned (i - lev) then  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
456  | 
raise CHANGE_FAIL (**flexible occurrence: give up**)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
457  | 
else Bound (i - length (filter (fn j => j < i - lev) banned))  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
458  | 
| change lev (Abs (a, T, t)) = Abs (a, T, change(lev + 1) t)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
459  | 
| change lev (t $ u) = change lev t $ change lev u  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
460  | 
| change lev t = t;  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
461  | 
in change 0 end;  | 
| 0 | 462  | 
|
463  | 
(*Change indices, delete the argument if it contains a banned Bound*)  | 
|
| 48263 | 464  | 
fun change_arg banned {j, t, T} args : flarg list =
 | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
465  | 
if rigid_bound (0, banned) t then args (*delete argument!*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
466  | 
  else {j = j, t = change_bnos banned t, T = T} :: args;
 | 
| 0 | 467  | 
|
468  | 
||
469  | 
(*Sort the arguments to create assignments if possible:  | 
|
| 
48262
 
a0d8abca8d7a
back to naive insertion sort before 1997 to accommodate peculiar less_arg relation -- NB: make_ord arg_less was not a quasi-order and thus inappropriate for generic sort (cf. de74b549f976, ecfeff48bf0c);
 
wenzelm 
parents: 
46219 
diff
changeset
 | 
470  | 
create eta-terms like ?g B.1 B.0*)  | 
| 
 
a0d8abca8d7a
back to naive insertion sort before 1997 to accommodate peculiar less_arg relation -- NB: make_ord arg_less was not a quasi-order and thus inappropriate for generic sort (cf. de74b549f976, ecfeff48bf0c);
 
wenzelm 
parents: 
46219 
diff
changeset
 | 
471  | 
local  | 
| 
 
a0d8abca8d7a
back to naive insertion sort before 1997 to accommodate peculiar less_arg relation -- NB: make_ord arg_less was not a quasi-order and thus inappropriate for generic sort (cf. de74b549f976, ecfeff48bf0c);
 
wenzelm 
parents: 
46219 
diff
changeset
 | 
472  | 
  fun less_arg ({t = Bound i1, ...}, {t = Bound i2, ...}) = (i2 < i1)
 | 
| 
 
a0d8abca8d7a
back to naive insertion sort before 1997 to accommodate peculiar less_arg relation -- NB: make_ord arg_less was not a quasi-order and thus inappropriate for generic sort (cf. de74b549f976, ecfeff48bf0c);
 
wenzelm 
parents: 
46219 
diff
changeset
 | 
473  | 
| less_arg (_: flarg, _: flarg) = false;  | 
| 
 
a0d8abca8d7a
back to naive insertion sort before 1997 to accommodate peculiar less_arg relation -- NB: make_ord arg_less was not a quasi-order and thus inappropriate for generic sort (cf. de74b549f976, ecfeff48bf0c);
 
wenzelm 
parents: 
46219 
diff
changeset
 | 
474  | 
|
| 
 
a0d8abca8d7a
back to naive insertion sort before 1997 to accommodate peculiar less_arg relation -- NB: make_ord arg_less was not a quasi-order and thus inappropriate for generic sort (cf. de74b549f976, ecfeff48bf0c);
 
wenzelm 
parents: 
46219 
diff
changeset
 | 
475  | 
fun ins_arg x [] = [x]  | 
| 
 
a0d8abca8d7a
back to naive insertion sort before 1997 to accommodate peculiar less_arg relation -- NB: make_ord arg_less was not a quasi-order and thus inappropriate for generic sort (cf. de74b549f976, ecfeff48bf0c);
 
wenzelm 
parents: 
46219 
diff
changeset
 | 
476  | 
| ins_arg x (y :: ys) =  | 
| 
 
a0d8abca8d7a
back to naive insertion sort before 1997 to accommodate peculiar less_arg relation -- NB: make_ord arg_less was not a quasi-order and thus inappropriate for generic sort (cf. de74b549f976, ecfeff48bf0c);
 
wenzelm 
parents: 
46219 
diff
changeset
 | 
477  | 
if less_arg (y, x) then y :: ins_arg x ys else x :: y :: ys;  | 
| 
 
a0d8abca8d7a
back to naive insertion sort before 1997 to accommodate peculiar less_arg relation -- NB: make_ord arg_less was not a quasi-order and thus inappropriate for generic sort (cf. de74b549f976, ecfeff48bf0c);
 
wenzelm 
parents: 
46219 
diff
changeset
 | 
478  | 
in  | 
| 
 
a0d8abca8d7a
back to naive insertion sort before 1997 to accommodate peculiar less_arg relation -- NB: make_ord arg_less was not a quasi-order and thus inappropriate for generic sort (cf. de74b549f976, ecfeff48bf0c);
 
wenzelm 
parents: 
46219 
diff
changeset
 | 
479  | 
fun sort_args [] = []  | 
| 
 
a0d8abca8d7a
back to naive insertion sort before 1997 to accommodate peculiar less_arg relation -- NB: make_ord arg_less was not a quasi-order and thus inappropriate for generic sort (cf. de74b549f976, ecfeff48bf0c);
 
wenzelm 
parents: 
46219 
diff
changeset
 | 
480  | 
| sort_args (x :: xs) = ins_arg x (sort_args xs);  | 
| 
 
a0d8abca8d7a
back to naive insertion sort before 1997 to accommodate peculiar less_arg relation -- NB: make_ord arg_less was not a quasi-order and thus inappropriate for generic sort (cf. de74b549f976, ecfeff48bf0c);
 
wenzelm 
parents: 
46219 
diff
changeset
 | 
481  | 
end;  | 
| 0 | 482  | 
|
483  | 
(*Test whether the new term would be eta-equivalent to a variable --  | 
|
484  | 
if so then there is no point in creating a new variable*)  | 
|
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
485  | 
fun decreasing n ([]: flarg list) = (n = 0)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
486  | 
  | decreasing n ({j, ...} :: args) = j = n - 1 andalso decreasing (n - 1) args;
 | 
| 0 | 487  | 
|
488  | 
(*Delete banned indices in the term, simplifying it.  | 
|
489  | 
Force an assignment, if possible, by sorting the arguments.  | 
|
490  | 
Update its head; squash indices in arguments. *)  | 
|
491  | 
fun clean_term banned (env,t) =  | 
|
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
492  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
493  | 
val (Var (v, T), ts) = strip_comb t;  | 
| 52221 | 494  | 
val (Ts, U) = Envir.strip_type env T  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
495  | 
and js = length ts - 1 downto 0;  | 
| 48263 | 496  | 
val args = sort_args (fold_rev (change_arg banned) (flexargs (js, ts, Ts)) [])  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
497  | 
val ts' = map #t args;  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
498  | 
in  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
499  | 
if decreasing (length Ts) args then (env, (list_comb (Var (v, T), ts')))  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
500  | 
else  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
501  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
502  | 
val (env', v') = Envir.genvar (#1 v) (env, map #T args ---> U);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
503  | 
val body = list_comb (v', map (Bound o #j) args);  | 
| 51700 | 504  | 
val env2 = Envir.vupdate ((v, T), types_abs (Ts, body)) env';  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
505  | 
(*the vupdate affects ts' if they contain v*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
506  | 
in (env2, Envir.norm_term env2 (list_comb (v', ts'))) end  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
507  | 
end;  | 
| 0 | 508  | 
|
509  | 
||
510  | 
(*Add tpair if not trivial or already there.  | 
|
511  | 
Should check for swapped pairs??*)  | 
|
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
512  | 
fun add_tpair (rbinder, (t0, u0), tpairs) : (term * term) list =  | 
| 19864 | 513  | 
if t0 aconv u0 then tpairs  | 
| 0 | 514  | 
else  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
515  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
516  | 
val t = Logic.rlist_abs (rbinder, t0)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
517  | 
and u = Logic.rlist_abs (rbinder, u0);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
518  | 
fun same (t', u') = (t aconv t') andalso (u aconv u')  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
519  | 
in if exists same tpairs then tpairs else (t, u) :: tpairs end;  | 
| 0 | 520  | 
|
521  | 
||
522  | 
(*Simplify both terms and check for assignments.  | 
|
523  | 
Bound vars in the binder are "banned" unless used in both t AND u *)  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
524  | 
fun clean_ffpair context ((rbinder, t, u), (env, tpairs)) =  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
525  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
526  | 
val loot = loose_bnos t and loou = loose_bnos u  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
527  | 
fun add_index (j, (a, T)) (bnos, newbinder) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
528  | 
if member (op =) loot j andalso member (op =) loou j  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
529  | 
then (bnos, (a, T) :: newbinder) (*needed by both: keep*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
530  | 
else (j :: bnos, newbinder); (*remove*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
531  | 
val (banned, rbin') = fold_rev add_index ((0 upto (length rbinder - 1)) ~~ rbinder) ([], []);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
532  | 
val (env', t') = clean_term banned (env, t);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
533  | 
val (env'',u') = clean_term banned (env',u);  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
534  | 
in  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
535  | 
(ff_assign context (env'', rbin', t', u'), tpairs)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
536  | 
handle ASSIGN =>  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
537  | 
(ff_assign context (env'', rbin', u', t'), tpairs)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
538  | 
handle ASSIGN => (env'', add_tpair (rbin', (t', u'), tpairs))  | 
| 0 | 539  | 
end  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
540  | 
handle CHANGE_FAIL => (env, add_tpair (rbinder, (t, u), tpairs));  | 
| 0 | 541  | 
|
542  | 
||
543  | 
(*IF the flex-flex dpair is an assignment THEN do it ELSE put in tpairs  | 
|
544  | 
eliminates trivial tpairs like t=t, as well as repeated ones  | 
|
| 19864 | 545  | 
trivial tpairs can easily escape SIMPL: ?A=t, ?A=?B, ?B=t gives t=t  | 
| 0 | 546  | 
Resulting tpairs MAY NOT be in normal form: assignments may occur here.*)  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
547  | 
fun add_ffpair context (rbinder,t0,u0) (env,tpairs) : Envir.env * (term * term) list =  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
548  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
549  | 
val t = Envir.norm_term env t0  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
550  | 
and u = Envir.norm_term env u0;  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
551  | 
in  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
552  | 
(case (head_of t, head_of u) of  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
553  | 
(Var (v, T), Var (w, U)) => (*Check for identical variables...*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
554  | 
if Term.eq_ix (v, w) then (*...occur check would falsely return true!*)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
555  | 
if T = U then (env, add_tpair (rbinder, (t, u), tpairs))  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
556  | 
          else raise TERM ("add_ffpair: Var name confusion", [t, u])
 | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
557  | 
else if Term_Ord.indexname_ord (v, w) = LESS then (*prefer to update the LARGER variable*)  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
558  | 
clean_ffpair context ((rbinder, u, t), (env, tpairs))  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
559  | 
else clean_ffpair context ((rbinder, t, u), (env, tpairs))  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
560  | 
    | _ => raise TERM ("add_ffpair: Vars expected", [t, u]))
 | 
| 0 | 561  | 
end;  | 
562  | 
||
563  | 
||
564  | 
(*Print a tracing message + list of dpairs.  | 
|
| 67721 | 565  | 
In t \<equiv> u print u first because it may be rigid or flexible --  | 
| 0 | 566  | 
t is always flexible.*)  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
567  | 
fun print_dpairs context msg (env, dpairs) =  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
568  | 
if Context_Position.is_visible_generic context then  | 
| 
52701
 
51dfdcd88e84
guard unify tracing via visible status of global theory;
 
wenzelm 
parents: 
52698 
diff
changeset
 | 
569  | 
let  | 
| 
 
51dfdcd88e84
guard unify tracing via visible status of global theory;
 
wenzelm 
parents: 
52698 
diff
changeset
 | 
570  | 
fun pdp (rbinder, t, u) =  | 
| 
 
51dfdcd88e84
guard unify tracing via visible status of global theory;
 
wenzelm 
parents: 
52698 
diff
changeset
 | 
571  | 
let  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
572  | 
val ctxt = Context.proof_of context;  | 
| 
52701
 
51dfdcd88e84
guard unify tracing via visible status of global theory;
 
wenzelm 
parents: 
52698 
diff
changeset
 | 
573  | 
fun termT t =  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
574  | 
Syntax.pretty_term ctxt (Envir.norm_term env (Logic.rlist_abs (rbinder, t)));  | 
| 67725 | 575  | 
val prt = Pretty.blk (0, [termT u, Pretty.str " \<equiv>\<^sup>?", Pretty.brk 1, termT t]);  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
576  | 
in tracing (Pretty.string_of prt) end;  | 
| 
52701
 
51dfdcd88e84
guard unify tracing via visible status of global theory;
 
wenzelm 
parents: 
52698 
diff
changeset
 | 
577  | 
in tracing msg; List.app pdp dpairs end  | 
| 
 
51dfdcd88e84
guard unify tracing via visible status of global theory;
 
wenzelm 
parents: 
52698 
diff
changeset
 | 
578  | 
else ();  | 
| 0 | 579  | 
|
580  | 
||
581  | 
(*Unify the dpairs in the environment.  | 
|
| 19864 | 582  | 
Returns flex-flex disagreement pairs NOT IN normal form.  | 
| 0 | 583  | 
SIMPL may raise exception CANTUNIFY. *)  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
584  | 
fun hounifiers (context, env, tus : (term * term) list) : (Envir.env * (term * term) list) Seq.seq =  | 
| 
24178
 
4ff1dc2aa18d
turned Unify flags into configuration options (global only);
 
wenzelm 
parents: 
23178 
diff
changeset
 | 
585  | 
let  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
586  | 
val trace_bound = Config.get_generic context trace_bound;  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
587  | 
val search_bound = Config.get_generic context search_bound;  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
588  | 
val trace_simp = Config.get_generic context trace_simp;  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
589  | 
fun add_unify tdepth ((env, dpairs), reseq) =  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
590  | 
Seq.make (fn () =>  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
591  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
592  | 
val (env', flexflex, flexrigid) =  | 
| 52698 | 593  | 
(if tdepth > trace_bound andalso trace_simp  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
594  | 
then print_dpairs context "Enter SIMPL" (env, dpairs) else ();  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
595  | 
SIMPL context (env, dpairs));  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
596  | 
in  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
597  | 
(case flexrigid of  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
598  | 
[] => SOME (fold_rev (add_ffpair context) flexflex (env', []), reseq)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
599  | 
| dp :: frigid' =>  | 
| 52698 | 600  | 
if tdepth > search_bound then  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
601  | 
(if Context_Position.is_visible_generic context  | 
| 
56294
 
85911b8a6868
prefer Context_Position where a context is available;
 
wenzelm 
parents: 
52701 
diff
changeset
 | 
602  | 
then warning "Unification bound exceeded" else (); Seq.pull reseq)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
603  | 
else  | 
| 52698 | 604  | 
(if tdepth > trace_bound then  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
605  | 
print_dpairs context "Enter MATCH" (env',flexrigid@flexflex)  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
606  | 
else ();  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
607  | 
Seq.pull (Seq.it_right  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
608  | 
(add_unify (tdepth + 1)) (MATCH context (env',dp, frigid'@flexflex), reseq))))  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
609  | 
end  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
610  | 
handle CANTUNIFY =>  | 
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
611  | 
(if tdepth > trace_bound andalso Context_Position.is_visible_generic context  | 
| 
56294
 
85911b8a6868
prefer Context_Position where a context is available;
 
wenzelm 
parents: 
52701 
diff
changeset
 | 
612  | 
then tracing "Failure node"  | 
| 
52701
 
51dfdcd88e84
guard unify tracing via visible status of global theory;
 
wenzelm 
parents: 
52698 
diff
changeset
 | 
613  | 
else (); Seq.pull reseq));  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
614  | 
val dps = map (fn (t, u) => ([], t, u)) tus;  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15797 
diff
changeset
 | 
615  | 
in add_unify 1 ((env, dps), Seq.empty) end;  | 
| 0 | 616  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
617  | 
fun unifiers (params as (context, env, tus)) =  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
618  | 
Seq.cons (fold (Pattern.unify context) tus env, []) Seq.empty  | 
| 
16425
 
2427be27cc60
accomodate identification of type Sign.sg and theory;
 
wenzelm 
parents: 
15797 
diff
changeset
 | 
619  | 
handle Pattern.Unif => Seq.empty  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
620  | 
| Pattern.Pattern => hounifiers params;  | 
| 0 | 621  | 
|
622  | 
||
623  | 
(*For smash_flexflex1*)  | 
|
624  | 
fun var_head_of (env,t) : indexname * typ =  | 
|
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
625  | 
(case head_of (strip_abs_body (Envir.norm_term env t)) of  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
626  | 
Var (v, T) => (v, T)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
627  | 
| _ => raise CANTUNIFY); (*not flexible, cannot use trivial substitution*)  | 
| 0 | 628  | 
|
629  | 
||
630  | 
(*Eliminate a flex-flex pair by the trivial substitution, see Huet (1975)  | 
|
| 67721 | 631  | 
Unifies ?f t1 ... rm with ?g u1 ... un by ?f -> \<lambda>x1...xm. ?a, ?g -> \<lambda>x1...xn. ?a  | 
632  | 
Unfortunately, unifies ?f t u with ?g t u by ?f, ?g -> \<lambda>x y. ?a,  | 
|
| 19864 | 633  | 
though just ?g->?f is a more general unifier.  | 
| 0 | 634  | 
Unlike Huet (1975), does not smash together all variables of same type --  | 
635  | 
requires more work yet gives a less general unifier (fewer variables).  | 
|
| 67721 | 636  | 
Handles ?f t1 ... rm with ?f u1 ... um to avoid multiple updates. *)  | 
| 48263 | 637  | 
fun smash_flexflex1 (t, u) env : Envir.env =  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
638  | 
let  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
639  | 
val vT as (v, T) = var_head_of (env, t)  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
640  | 
and wU as (w, U) = var_head_of (env, u);  | 
| 52221 | 641  | 
val (env', var) = Envir.genvar (#1 v) (env, Envir.body_type env T);  | 
| 51700 | 642  | 
val env'' = Envir.vupdate (wU, type_abs (env', U, var)) env';  | 
| 
37635
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
643  | 
in  | 
| 
 
484efa650907
recovered some indentation from the depths of time;
 
wenzelm 
parents: 
36787 
diff
changeset
 | 
644  | 
if vT = wU then env'' (*the other update would be identical*)  | 
| 51700 | 645  | 
else Envir.vupdate (vT, type_abs (env', T, var)) env''  | 
| 0 | 646  | 
end;  | 
647  | 
||
648  | 
||
649  | 
(*Smash all flex-flexpairs. Should allow selection of pairs by a predicate?*)  | 
|
| 37636 | 650  | 
fun smash_flexflex (env, tpairs) : Envir.env =  | 
| 48263 | 651  | 
fold_rev smash_flexflex1 tpairs env;  | 
| 0 | 652  | 
|
653  | 
(*Returns unifiers with no remaining disagreement pairs*)  | 
|
| 
58950
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
654  | 
fun smash_unifiers context tus env =  | 
| 
 
d07464875dd4
optional proof context for unify operations, for the sake of proper local options;
 
wenzelm 
parents: 
56438 
diff
changeset
 | 
655  | 
Seq.map smash_flexflex (unifiers (context, env, tus));  | 
| 0 | 656  | 
|
657  | 
end;  |