From ad2e11471dbfc0894b4fdfedd895e7f0a75bd930 Mon Sep 17 00:00:00 2001 From: Matthieu Sozeau Date: Wed, 24 Sep 2014 11:16:11 +0200 Subject: First version of keyed subterm selection in unification. --- library/keys.ml | 46 +++++++++++++++++++++++++++++++++++++++ library/keys.mli | 15 +++++++++++++ library/library.mllib | 1 + pretyping/unification.ml | 36 ++++++++++++++++++++++++++++-- test-suite/success/keyedrewrite.v | 1 + 5 files changed, 97 insertions(+), 2 deletions(-) create mode 100644 library/keys.ml create mode 100644 library/keys.mli create mode 100644 test-suite/success/keyedrewrite.v diff --git a/library/keys.ml b/library/keys.ml new file mode 100644 index 000000000..5c295a385 --- /dev/null +++ b/library/keys.ml @@ -0,0 +1,46 @@ +open Globnames +open Term +open Libobject + +(* Union-find structure for references to be considered equivalent *) + +module RefUF = Unionfind.Make(Refset)(Refmap) + +let keys = (* Summary.ref ( *)RefUF.create ()(* ) ~name:"Keys_decl" *) + +let add_key k v = RefUF.union k v keys + +let equiv_keys k k' = RefUF.find k keys == RefUF.find k' keys + +(** Registration of keys as an object *) + +let load_key _ (_,(ref,ref')) = + add_key ref ref' + +let cache_key o = + load_key 1 o + +let subst_key (subst,(ref,ref')) = + (subst_global_reference subst ref, + subst_global_reference subst ref') + +let discharge_key (_,refs) = + match refs with + | VarRef _, _ | _, VarRef _ -> None + | ref, ref' -> Some (pop_global_reference ref, pop_global_reference ref') + +let rebuild_key (ref,ref') = (ref, ref') + +type key_obj = global_reference * global_reference + +let inKeys : key_obj -> obj = + declare_object {(default_object "KEYS") with + cache_function = cache_key; + load_function = load_key; + subst_function = subst_key; + classify_function = (fun x -> Substitute x); + discharge_function = discharge_key; + rebuild_function = rebuild_key } + +let declare_keys ref ref' = + Lib.add_anonymous_leaf (inKeys (ref,ref')) diff --git a/library/keys.mli b/library/keys.mli new file mode 100644 index 000000000..4e6a803e4 --- /dev/null +++ b/library/keys.mli @@ -0,0 +1,15 @@ +(************************************************************************) +(* v * The Coq Proof Assistant / The Coq Development Team *) +(* global_reference -> unit +(** Declare two keys as being equivalent. *) + +val equiv_keys : global_reference -> global_reference -> bool +(** Check equivalence of keys. *) diff --git a/library/library.mllib b/library/library.mllib index 6a58a1057..04244160b 100644 --- a/library/library.mllib +++ b/library/library.mllib @@ -17,3 +17,4 @@ Goptions Decls Heads Assumptions + diff --git a/pretyping/unification.ml b/pretyping/unification.ml index 050103542..1a1143fe7 100644 --- a/pretyping/unification.ml +++ b/pretyping/unification.ml @@ -28,6 +28,15 @@ open Locus open Locusops open Find_subterm +let keyed_unification = ref (false) +let _ = Goptions.declare_bool_option { + Goptions.optsync = true; Goptions.optdepr = false; + Goptions.optname = "Unification is keyed"; + Goptions.optkey = ["Keyed";"Unification"]; + Goptions.optread = (fun () -> !keyed_unification); + Goptions.optwrite = (fun a -> keyed_unification:=a); +} + let occur_meta_or_undefined_evar evd c = let rec occrec c = match kind_of_term c with | Meta _ -> raise Occur @@ -1448,6 +1457,29 @@ let make_abstraction env evd ccl abs = make_abstraction_core name (make_eq_test env evd c) (evd,c) ty occs check_occs env ccl +let rec constr_key c = + let open Globnames in + match kind_of_term c with + | Const (c, _) -> ConstRef c + | Ind (i, u) -> IndRef i + | Construct (c,u) -> ConstructRef c + | Var id -> VarRef id + | App (f, _) -> constr_key f + | Proj (p, _) -> ConstRef p + | Cast (p, _, _) -> constr_key p + | Lambda _ + | Prod _ + | Case _ + | Fix _ | CoFix _ + | Rel _ | Meta _ | Evar _ | Sort _ | LetIn _ -> raise Not_found + +let keyed_unify env evd op cl = + if not !keyed_unification then true + else + let k1 = constr_key op in + let k2 = constr_key cl in + Globnames.eq_gr k1 k2 || Keys.equiv_keys k1 k2 + (* Tries to find an instance of term [cl] in term [op]. Unifies [cl] to every subterm of [op] until it finds a match. Fails if no match is found *) @@ -1456,7 +1488,7 @@ let w_unify_to_subterm env evd ?(flags=default_unify_flags ()) (op,cl) = let rec matchrec cl = let cl = strip_outer_cast cl in (try - if closed0 cl && not (isEvar cl) then + if closed0 cl && not (isEvar cl) && keyed_unify env evd op cl then (try w_typed_unify env evd CONV flags op cl,cl with ex when Pretype_errors.unsatisfiable_exception ex -> bestexn := Some ex; error "Unsat") @@ -1588,7 +1620,7 @@ let w_unify_to_subterm_list env evd flags hdmeta oplist t = if isMeta op then if flags.allow_K_in_toplevel_higher_order_unification then (evd,op::l) else error_abstraction_over_meta env evd hdmeta (destMeta op) - else if occur_meta_or_existential op then + else if occur_meta_or_existential op || !keyed_unification then let (evd',cl) = try (* This is up to delta for subterms w/o metas ... *) diff --git a/test-suite/success/keyedrewrite.v b/test-suite/success/keyedrewrite.v new file mode 100644 index 000000000..438613b44 --- /dev/null +++ b/test-suite/success/keyedrewrite.v @@ -0,0 +1 @@ +Lemma foo : \ No newline at end of file -- cgit v1.2.3