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