Skip to content

Commit 50654a3

Browse files
Merge PR rocq-prover#13907: Algorithmically faster algorithm for term replacing.
Reviewed-by: SkySkimmer
2 parents d332666 + 285419e commit 50654a3

File tree

4 files changed

+78
-58
lines changed

4 files changed

+78
-58
lines changed

engine/termops.ml

+31-48
Original file line numberDiff line numberDiff line change
@@ -979,69 +979,52 @@ let collapse_appl sigma c = match EConstr.kind sigma c with
979979

980980
(* First utilities for avoiding telescope computation for subst_term *)
981981

982-
let prefix_application sigma eq_fun (k,c) t =
982+
let prefix_application sigma eq_fun k l1 t =
983983
let open EConstr in
984-
let c' = collapse_appl sigma c and t' = collapse_appl sigma t in
985-
match EConstr.kind sigma c', EConstr.kind sigma t' with
986-
| App (f1,cl1), App (f2,cl2) ->
987-
let l1 = Array.length cl1
988-
and l2 = Array.length cl2 in
984+
let t' = collapse_appl sigma t in
985+
if 0 < l1 then match EConstr.kind sigma t' with
986+
| App (f2,cl2) ->
987+
let l2 = Array.length cl2 in
989988
if l1 <= l2
990-
&& eq_fun sigma c' (mkApp (f2, Array.sub cl2 0 l1)) then
991-
Some (mkApp (mkRel k, Array.sub cl2 l1 (l2 - l1)))
989+
&& eq_fun sigma k (mkApp (f2, Array.sub cl2 0 l1)) then
990+
Some (Array.sub cl2 l1 (l2 - l1))
992991
else
993992
None
994993
| _ -> None
994+
else None
995995

996-
let my_prefix_application sigma eq_fun (k,c) by_c t =
997-
let open EConstr in
998-
let c' = collapse_appl sigma c and t' = collapse_appl sigma t in
999-
match EConstr.kind sigma c', EConstr.kind sigma t' with
1000-
| App (f1,cl1), App (f2,cl2) ->
1001-
let l1 = Array.length cl1
1002-
and l2 = Array.length cl2 in
1003-
if l1 <= l2
1004-
&& eq_fun sigma c' (mkApp (f2, Array.sub cl2 0 l1)) then
1005-
Some (mkApp ((Vars.lift k by_c), Array.sub cl2 l1 (l2 - l1)))
1006-
else
1007-
None
1008-
| _ -> None
1009-
1010-
(* Recognizing occurrences of a given subterm in a term: [subst_term c t]
1011-
substitutes [(Rel 1)] for all occurrences of term [c] in a term [t];
1012-
works if [c] has rels *)
1013-
1014-
let subst_term_gen sigma eq_fun c t =
1015-
let open EConstr in
1016-
let open Vars in
1017-
let rec substrec (k,c as kc) t =
1018-
match prefix_application sigma eq_fun kc t with
1019-
| Some x -> x
1020-
| None ->
1021-
if eq_fun sigma c t then mkRel k
1022-
else
1023-
EConstr.map_with_binders sigma (fun (k,c) -> (k+1,lift 1 c)) substrec kc t
996+
let eq_upto_lift cache c sigma k t =
997+
let c =
998+
try Int.Map.find k !cache
999+
with Not_found ->
1000+
let c = EConstr.Vars.lift k c in
1001+
let () = cache := Int.Map.add k c !cache in
1002+
c
10241003
in
1025-
substrec (1,c) t
1026-
1027-
let subst_term sigma c t = subst_term_gen sigma EConstr.eq_constr c t
1004+
EConstr.eq_constr sigma c t
10281005

10291006
(* Recognizing occurrences of a given subterm in a term :
10301007
[replace_term c1 c2 t] substitutes [c2] for all occurrences of
10311008
term [c1] in a term [t]; works if [c1] and [c2] have rels *)
10321009

1033-
let replace_term_gen sigma eq_fun c by_c in_t =
1034-
let rec substrec (k,c as kc) t =
1035-
match my_prefix_application sigma eq_fun kc by_c t with
1036-
| Some x -> x
1010+
let replace_term_gen sigma eq_fun ar by_c in_t =
1011+
let rec substrec k t =
1012+
match prefix_application sigma eq_fun k ar t with
1013+
| Some args -> EConstr.mkApp (EConstr.Vars.lift k by_c, args)
10371014
| None ->
1038-
(if eq_fun sigma c t then (EConstr.Vars.lift k by_c) else
1039-
EConstr.map_with_binders sigma (fun (k,c) -> (k+1,EConstr.Vars.lift 1 c))
1040-
substrec kc t)
1015+
(if eq_fun sigma k t then (EConstr.Vars.lift k by_c) else
1016+
EConstr.map_with_binders sigma succ substrec k t)
10411017
in
1042-
substrec (0,c) in_t
1018+
substrec 0 in_t
1019+
1020+
let replace_term sigma c byc t =
1021+
let cache = ref Int.Map.empty in
1022+
let c = collapse_appl sigma c in
1023+
let ar = Array.length (snd (decompose_app_vect sigma c)) in
1024+
let eq sigma k t = eq_upto_lift cache c sigma k t in
1025+
replace_term_gen sigma eq ar byc t
10431026

1044-
let replace_term sigma c byc t = replace_term_gen sigma EConstr.eq_constr c byc t
1027+
let subst_term sigma c t = replace_term sigma c (EConstr.mkRel 1) t
10451028

10461029
let vars_of_env env =
10471030
let s = Environ.ids_of_named_context_val (Environ.named_context_val env) in

engine/termops.mli

+5-9
Original file line numberDiff line numberDiff line change
@@ -122,16 +122,12 @@ val pop : constr -> constr
122122
(** Substitution of an arbitrary large term. Uses equality modulo
123123
reduction of let *)
124124

125-
(** [subst_term_gen eq d c] replaces [d] by [Rel 1] in [c] using [eq]
126-
as equality *)
127-
val subst_term_gen : Evd.evar_map ->
128-
(Evd.evar_map -> constr -> constr -> bool) -> constr -> constr -> constr
129-
130-
(** [replace_term_gen eq d e c] replaces [d] by [e] in [c] using [eq]
131-
as equality *)
125+
(** [replace_term_gen eq arity e c] replaces matching subterms according to
126+
[eq] by [e] in [c]. If [arity] is non-zero applications of larger length
127+
are handled atomically. *)
132128
val replace_term_gen :
133-
Evd.evar_map -> (Evd.evar_map -> constr -> constr -> bool) ->
134-
constr -> constr -> constr -> constr
129+
Evd.evar_map -> (Evd.evar_map -> int -> constr -> bool) ->
130+
int -> constr -> constr -> constr
135131

136132
(** [subst_term d c] replaces [d] by [Rel 1] in [c] *)
137133
val subst_term : Evd.evar_map -> constr -> constr -> constr

tactics/tactics.ml

+18-1
Original file line numberDiff line numberDiff line change
@@ -2796,7 +2796,24 @@ let generalize_goal_gen env sigma ids i ((occs,c,b),na) t cl =
27962796
let open Context.Rel.Declaration in
27972797
let decls,cl = decompose_prod_n_assum sigma i cl in
27982798
let dummy_prod = it_mkProd_or_LetIn mkProp decls in
2799-
let newdecls,_ = decompose_prod_n_assum sigma i (subst_term_gen sigma EConstr.eq_constr_nounivs c dummy_prod) in
2799+
let newdecls,_ =
2800+
let c = Termops.collapse_appl sigma c in
2801+
let arity = Array.length (snd (Termops.decompose_app_vect sigma c)) in
2802+
let cache = ref Int.Map.empty in
2803+
let eq sigma k t =
2804+
let c =
2805+
try Int.Map.find k !cache
2806+
with Not_found ->
2807+
let c = EConstr.Vars.lift k c in
2808+
let () = cache := Int.Map.add k c !cache in
2809+
c
2810+
in
2811+
(* We use a nounivs equality because generalize morally takes a pattern as
2812+
argument, so we have to ignore freshly generated sorts. *)
2813+
EConstr.eq_constr_nounivs sigma c t
2814+
in
2815+
decompose_prod_n_assum sigma i (replace_term_gen sigma eq arity (mkRel 1) dummy_prod)
2816+
in
28002817
let cl',sigma' = subst_closed_term_occ env sigma (AtOccs occs) c (it_mkProd_or_LetIn cl newdecls) in
28012818
let na = generalized_name env sigma c t ids cl' na in
28022819
let r = Retyping.relevance_of_type env sigma t in

test-suite/bugs/closed/bug_13896.v

+24
Original file line numberDiff line numberDiff line change
@@ -0,0 +1,24 @@
1+
Inductive type : Set :=
2+
Tptr : type -> type
3+
| Tref : type -> type
4+
| Trv_ref : type -> type
5+
| Tint : type -> type -> type
6+
| Tvoid : type
7+
| Tarray : type -> type -> type
8+
| Tnamed : type -> type
9+
| Tfunction : type -> type -> type -> type
10+
| Tbool : type
11+
| Tmember_pointer : type -> type -> type
12+
| Tfloat : type -> type
13+
| Tqualified : type -> type -> type
14+
| Tnullptr : type
15+
| Tarch : type -> type -> type
16+
.
17+
Definition type_eq_dec : forall (ty1 ty2 : type), { ty1 = ty2 } + { ty1 <> ty2 }.
18+
Proof. fix IHty1 1. decide equality. Defined.
19+
20+
Goal (if type_eq_dec (Tptr Tvoid) (Tptr Tvoid) then True else False).
21+
Proof.
22+
timeout 1 cbn.
23+
constructor.
24+
Qed.

0 commit comments

Comments
 (0)