Library ListFromRight.foldl_issue
Author: Jean-François Monin, Verimag, Université Grenoble Alpes
List traversal from right to left: a first attempt.
List.fold_left
, which
follows the same structural pattern as List.fold_right
, but where
lists are traversed from right to left.
let rec foldl_ref f b l = fakematch l with | [] → b | u +: z → f (foldl_ref f b u) zThis can be actually programmed by reflecting this decomposition of lists from the right using l2r.
let rec foldl_ref f b l = match l2r l with | Nilr → b | Consr (u, z) → f (foldl_ref f b u) zThis definition corresponds to the usual informal drawings explaining the expected result. It is of course quite inefficient, however and it can be seen as a specification of OCaml
fold_left
,
named foldl here.
let rec fold_left f b l = match l with | [] → b | x :: l → fold_left f (f b x) lThis Coq script plays with a first direct attempt to define foldl_ref in Coq and to prove a simple property on it, in order to illustrate a simple situation where Coq standard inversion fails whereas the small inversion mechanism presented at TYPES'22 succeeds.
Require Import Utf8.
Require Import List.
Import ListNotations.
Require Import rl.
Section sec_context.
Context {A B: Type}.
Implicit Type l u v : list A.
Implicit Type x y z : A.
Implicit Type b : B.
A convenient helper
Lemma rew {u z} (d : 𝔻listz (l2r (u+:z))) : 𝔻listz (Consr u z).
Proof. case l2r_consr; exact d. Defined.
Proof. case l2r_consr; exact d. Defined.
A reference definition for
fold_left
Fixpoint foldl_ref l (d: 𝔻listz (l2r l)): B :=
match l2r l in rl l return 𝔻listz (l2r l) → B with
| Nilr ⇒ λ d, b
| Consr u z ⇒ λ d, f (foldl_ref u (π_𝔻listz_Consr (rew d))) z
end d.
match l2r l in rl l return 𝔻listz (l2r l) → B with
| Nilr ⇒ λ d, b
| Consr u z ⇒ λ d, f (foldl_ref u (π_𝔻listz_Consr (rew d))) z
end d.
A lemma where small inversion suceeds, in contrast with Coq inversion.
Lemma foldl_ref_nil d : foldl_ref [] d = b.
Proof.
cbn in d.
inversion d .
Fail dependent inversion d.
destruct (𝔻listz_inv d) .
cbn.
reflexivity.
Qed.
Proof.
cbn in d.
inversion d .
Fail dependent inversion d.
destruct (𝔻listz_inv d) .
cbn.
reflexivity.
Qed.
Unfortunately, a similar statement for the Consr case would need
a preliminary step which is not accepted.
Lemma foldl_ref_cons u z d : foldl_ref (u +: z) d = b.
Proof.
Fail rewrite l2r_consr in d.
revert d.
Fail rewrite l2r_consr.
Abort.
Proof.
Fail rewrite l2r_consr in d.
revert d.
Fail rewrite l2r_consr.
Abort.
A solution is presented in foldl.v