Skip to content

Commit

Permalink
Fixes refolding of section-defined mutual fix/cofix in "cbn".
Browse files Browse the repository at this point in the history
This was mentioned e.g. in coq#4056.
  • Loading branch information
herbelin committed Feb 1, 2024
1 parent 511511c commit 76d44a2
Show file tree
Hide file tree
Showing 2 changed files with 36 additions and 4 deletions.
8 changes: 4 additions & 4 deletions tactics/cbn.ml
Expand Up @@ -70,7 +70,7 @@ module Cst_stack = struct
| _ -> None

let reference sigma t = match best_cst t with
| Some (c, _) when isConst sigma c -> Some (fst (destConst sigma c))
| Some (c, params) when isConst sigma c -> Some (fst (destConst sigma c), params)
| _ -> None

(** [best_replace d cst_l c] makes the best replacement for [d]
Expand Down Expand Up @@ -435,7 +435,7 @@ let apply_subst env sigma cst_l t stack =
f x := t. End M. Definition f := u. and say goodbye to any hope
of refolding M.f this way ...
*)
let magically_constant_of_fixbody env sigma reference bd = function
let magically_constant_of_fixbody env sigma (reference,args) bd = function
| Name.Anonymous -> bd
| Name.Name id ->
let open UnivProblem in
Expand All @@ -447,7 +447,7 @@ let magically_constant_of_fixbody env sigma reference bd = function
match constant_opt_value_in env (cst,u) with
| None -> bd
| Some t ->
let csts = EConstr.eq_constr_universes env sigma (EConstr.of_constr t) bd in
let csts = EConstr.eq_constr_universes env sigma (Reductionops.beta_applist sigma (EConstr.of_constr t,args)) bd in
begin match csts with
| Some csts ->
let addqs l r (qs,us) = Sorts.QVar.Map.add l r qs, us in
Expand All @@ -472,7 +472,7 @@ let magically_constant_of_fixbody env sigma reference bd = function
csts UVars.empty_sort_subst
in
let inst = UVars.subst_sort_level_instance subst u in
mkConstU (cst, EInstance.make inst)
applist (mkConstU (cst, EInstance.make inst), args)
| None -> bd
end

Expand Down
32 changes: 32 additions & 0 deletions test-suite/success/cbn.v
Expand Up @@ -39,3 +39,35 @@ Proof.
cbn.
match goal with |- _ + _ = _ => idtac end.
Abort.

Module MutualFixCoFixInSection.

Section S.
Variable p:nat.
Fixpoint f n := match n with 0 => p | S n => f n + g n end
with g n := match n with 0 => p | S n => f n + g n end.
End S.

Goal forall n, f n (S n) = g 0 (S n).
intros. cbn.
match goal with [ |- f n n + g n n = f 0 n + g 0 n ] => idtac end.
Abort.

CoInductive stream {A:Type} : Type :=
| scons: A->stream->stream.
Definition stream_unfold {A} (s: @ stream A) := match s with
| scons a s' => (a, scons a s')
end.

Section C.
Variable (x:nat).
CoFixpoint mut_stream1 (n:nat) := scons n (mut_stream2 (n+x))
with mut_stream2 (n:nat) := scons n (mut_stream1 (n+x)).
End C.

Goal (forall x n, stream_unfold (mut_stream1 x n) = stream_unfold (mut_stream2 x n)).
intros. cbn.
match goal with [ |- (n, scons n (mut_stream2 x (n + x))) = (n, scons n (mut_stream1 x (n + x))) ] => idtac end.
Abort.

End MutualFixCoFixInSection.

0 comments on commit 76d44a2

Please sign in to comment.