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