How to dependent match on a list with two elements? - coq

I'm trying to understand dependent types and dependent match in Coq, I have the code below, at the bottom I have the add_pair function, the idea is that I want to receive a (dependent) list with exactly two elements and return the sum of the two elements in the list. Since the size of the list is encoded in the type, I should be able to define it as a total function, but I got an error saying that the match is not exaustive
Here is the code
Module IList.
Section Lists.
(* o tipo generico dos elementos da lista *)
Variable A : Set.
(* data type da lista, recebe um nat que é
o tamanho da lista *)
Inductive t : nat -> Set :=
| nil : t 0 (* lista vazia, n = 0 *)
| cons : forall (n : nat), A -> t n -> t (S n). (* cons de uma lista
n = n + 1 *)
(* concatena duas listas, n = n1 + n2 *)
Fixpoint concat n1 (ls1 : t n1) n2 (ls2 : t n2) : t (n1 + n2) :=
match ls1 in (t n1) return (t (n1 + n2)) with
| nil => ls2
| cons x ls1' => cons x (concat ls1' ls2)
end.
(* computar o tamanho da lista é O(1) *)
Definition length n (l : t n) : nat := n.
End Lists.
Arguments nil {_}.
(* Isso aqui serve pra introduzir notações pra gente poder
falar [1;2;3] em vez de (cons 1 (cons 2 (cons 3 nil))) *)
Module Notations.
Notation "a :: b" := (cons a b) (right associativity, at level 60) : list_scope.
Notation "[ ]" := nil : list_scope.
Notation "[ x ]" := (cons x nil) : list_scope.
Notation "[ x ; y ; .. ; z ]" := (cons x (cons y .. (cons z nil) ..)) : list_scope.
Open Scope list_scope.
End Notations.
Import Notations.
(* Error: Non exhaustive pattern-matching: no clause found for pattern [_] *)
Definition add_pair (l : t nat 2) : nat :=
match l in (t _ 2) return nat with
| (cons x (cons y nil)) => x + y
end.
End IList.

Indeed, it is true that the match you provided is exhaustive, but the pattern-matching algorithm of Coq is limited, and not able to detect it. The issue, I think, is that it compiles a nested pattern-matching such as yours (you have two imbricated cons) down to a successions of elementary pattern-matching (which have patterns of depth at most one). But in the cons branch of the outer match, the information that the index should be 1 is lost if you do not record it explicitly with an equality – something the current algorithm is not smart enough to do.
As a possible solution that avoids fiddling with impossible branches, equalities, and the like, I propose the following:
Definition head {A n} (l : t A (S n)) : A :=
match l with
| cons x _ => x
end.
Definition tail {A n} (l : t A (S n)) : t A n :=
match l with
| cons _ l' => l'
end.
Definition add_pair (l : t nat 2) : nat :=
head l + (head (tail l)).
For the record, a solution that does fiddle with the impossible branches and records the information of the index using equalities
(there’s probably a nicer version):
Definition add_pair (l : t nat 2) : nat :=
match l in (t _ m) return (m = 2) -> nat with
| [] => fun e => ltac:(lia)
| x :: l' => fun e =>
match l' in (t _ m') return (m' = 1) -> nat with
| [] => fun e' => ltac:(lia)
| x' :: l'' => fun e' =>
match l'' in (t _ m'') return (m'' = 0) -> nat with
| [] => fun _ => x + x'
| _ => fun e'' => ltac:(lia)
end ltac:(lia)
end ltac:(lia)
end eq_refl.
The interesting part is the use of explicit equalities to record the value of the index (these are used by the lia tactic to discard impossible branches).

Related

Coq: Comparing an Int to a Nat in Separation Logic Foundations

Going through Separation Logic Foundations and I'm stuck on the exercise triple_mlength in Repr.v. I think my current problem is that I don't know how to handle ints and nats in Coq.
Lemma triple_mlength: forall (L: list val) (p:loc),
triple (mlength p)
(MList L p)
(fun r => \[r = val_int (length L)] \* (MList L p))
Check (fun L => val_int (length L)) doesn't throw an error, so that means length is capable of being an int. However, length is opaque and I can't unfold it.
My current context and goal:
x : val
p : loc
C : p <> null
x0 : loc
H : p <> null
xs : list val
IH : forall y : list val,
list_sub y (x :: xs) ->
forall p, triple (mlength p)
(MList y p)
(fun r:val => \[r = length y] \* MList y p)
______________________________________________________________
length xs + 1 = length (x :: xs)
Unsetting print notation the goal transforms into:
eq (Z.add (length xs) (Zpos xH)) (length (cons x xs))
which I think is trying to add (1:Z) to (length xs: nat), then compare it to (length (cons x xs) : nat)
Types:
Inductive nat : Set := O : nat
| S : nat -> nat
Inductive Z : Set := Z0 : int
| Zpos : positive -> int
| Zneg : positive -> int
list: forall A, list A -> nat
length: forall A, list A -> nat
val_int: int -> val
Coq version is 8.12.2
There is a coercion nat_to_Z : nat -> int in scope that is converting length xs : nat and length (x :: xs) : nat to ints. This is separate from the notation mechanism and thus you don't see it when you only ask Coq to show notations. However, it is there and you need to handle it in your proofs. There are a bunch of lemmas floating around that prove equivalence between nat operations and Z/int operations.
Having loaded your file and looked around a bit (Search is your friend!), it appears the reason you cannot simplify length (x :: xs) = S (length xs) is because there is a lemma length_cons which gives length (x :: xs) = (1 + length xs)%nat, instead. I suppose the authors of this book thought that would be a good idea for some reason, so they disabled the usual simplification. Do note that "normally" length is transparent and simpl would work on this goal.
After using length_cons, you can use plus_nat_eq_plus_int to push the coercion down under the +, and then Z.add_comm finishes. This line should satisfy the goal.
now rewrite length_cons, plus_nat_eq_plus_int, Z.add_comm.

Building up tree and decreasing argument of fix

I'm attempting to implement a function to build up a Braun tree with n elements using the following function in Coq, but Coq gives me the error that it cannot guess decreasing argument of fix:
Fixpoint copy (x : V) (n : nat) : BraunTree :=
let
fix copy2 (a : V) (i : nat) : (BraunTree * BraunTree) :=
match i with
| 0 => (T a E E,E)
| _ => match Nat.odd i with
| true => let m := ((i - 1) / 2) in
let (s,t) := copy2 a m in
((T a s t),(T a t t))
| false => let m := ((i - 2) / 2) in
let (s,t) := copy2 a m in
((T a s s),(T a s t))
end
end
in
match copy2 x n with
|(_,snd) => snd
end.
I know that it is not the separate even and odd cases that is the problem because it gave the same error when I removed the even/odd cases:
Fixpoint copy (x : V) (n : nat) : BraunTree :=
let
fix copy2 (a : V) (i : nat) : (BraunTree * BraunTree) :=
match i with
| 0 => (T a E E,E)
| _ => let m := ((i - 1) / 2) in
let (s,t) := copy2 a m in
((T a s t),(T a t t))
end
in
match copy2 x n with
|(_,snd) => snd
end.
How can I convince Coq that i is in fact a decreasing argument?
EDIT Type of BraunTree:
Inductive BraunTree : Type :=
| E : BraunTree
| T: V -> BraunTree -> BraunTree -> BraunTree.
Fixpoint/fix only allows recursive calls on a syntactically smaller argument.
Fixpoint example (n : nat) :=
... (* There must be a match on [n] somewhere *)
... match n with
| O => base_case (* no recursive call allowed *)
| S m =>
... (example m)
(* We can only call [example] on [m], or some even smaller value obtained by matching on [m] *)
end ...
In particular, it's not allowed to make a recursive call on a value obtained via some arbitrary function (in this case, div and sub in copy2 a ((i-1) / 2)).
Here are three options:
Pick another representation of natural numbers so that pattern-matching on it naturally decomposes into the different branches of the desired definition (base case (zero), even, odd).
Use the fact that the recursion depth is actually bounded by n, so we can use n as "fuel", which we know will not actually deplete before we are done.
Cunningly extract a subterm of the decreasing argument to make the recursive call. This solution is less general and robust than the previous ones; it's a much harder fight against the termination checker.
Alternative representation
We have three cases: zero, even, and odd. Luckily the standard library has a type with almost the same structure, positive:
Inductive positive := (* p > 0 *)
| xH (* 1 *)
| xI (p : positive) (* 2p + 1 *)
| xO (p : positive) (* 2p *)
.
Pointing the type positive with an additional zero, we get N:
Inductive N :=
| N0 (* 0 *)
| Npos (p : positive) (* p > 0 *)
.
There is also a conversion function N.of_nat : nat -> N, although it might also be a good idea to use N everywhere instead of nat, if the conversions become too annoying.
The final definition starts by case analysis on N, and the case revealing a positive number is handled with a fix-point, where the base case is 1 instead of 0. We have to shift some details, because the even case is 2p instead of 2p+2, so instead of a pair of trees of size (i+1,i) we have to do (i-1,i). But overall the recursive cases still naturally match an informal specification:
Require Import NArith PArith.
Parameter V : Type.
Inductive BraunTree : Type :=
| E : BraunTree
| T: V -> BraunTree -> BraunTree -> BraunTree.
Definition copy (x : V) (n : N) : BraunTree :=
match n with
| N0 => E
| Npos p =>
let
(* copy2 a i : a tree of (i-1) copies of a, and another of i copies of a *)
fix copy2 (a : V) (i : positive) : (BraunTree * BraunTree) :=
match i with
| xH => (* i = 1 *)
(E, T a E E)
| xI p => (* i = 2p + 1 *)
let (s,t) := copy2 a p in
((T a t s),(T a t t))
| xO p => (* i = 2p *)
let (s,t) := copy2 a p in
((T a s s),(T a t s))
end
in
match copy2 x p with
|(_,snd) => snd
end
end.
Just enough fuel
We add fuel to the fix as the decreasing argument. We can only run out if n = i = 0, so we know what the result should be then.
(* note: This doesn't need to be a Fixpoint *)
Definition copy (x : V) (n : nat) : BraunTree :=
let
fix copy2 (a : V) (n : nat) (i : nat) : (BraunTree * BraunTree) :=
match n with
| O => (T a E E,E)
| S n' =>
match i with
| O => (T a E E,E)
| _ =>
if Nat.odd i then
let m := div2 ((i - 1) / 2) in
let (s,t) := copy2 a n' m in
((T a s t),(T a t t))
else
let m := div2 ((i - 2) / 2) in
let (s,t) := copy2 a n' m in
((T a s s),(T a s t))
end
end
in
match copy2 x n n with
|(_,snd) => snd
end.
This works nicely when:
we can compute the amount of fuel needed;
and there is a predictable answer to give when we run out of fuel.
If either of those assumption does not hold, we need to litter our code with option.
Nested recursion
As mentioned earlier, Coq has strict rules about decreasing arguments. The usual explanation is that we can only make a recursive call on a subterm obtained through pattern-matching on the decreasing argument (or transitively, one of its subterms).
One apparent restriction is that, because the condition is syntactic (i.e., Coq looks at the definition to track the provenance of the decreasing argument), the argument n can only decrease by a constant at most (constant with respect to n), since there are only finitely many match in a definition. In particular, there is no way to make a recursive call on the result of a division by two, as that represents a decrease by n/2, a value linear in n.
For better or for worse, Coq's termination criterion is actually a bit smarter than that: one can pass the decreasing argument to a nested fixpoint, and the "subterm" relation will be tracked through it.
Cons-free division
And indeed, the division of a Peano nat can be defined in such a way that Coq can tell that the result is a subterm of the dividend:
Definition div2 (n : nat) :=
let fix d2 (n1 : nat) (n2 : nat) {struct n1} :=
match n2 with
| S (S n2') =>
match n1 with
| O => n1
| S n1' => d2 n1' n2'
end
| _ => n1
end
in d2 n n.
The idea is to write a fix-point of two arguments (somewhat like the fuel solution), which start out equal (d2 n n), and we remove two S constructors from one (n2) of them for every one S we remove from the other (n1). Important details:
In all the non-recursing cases, we return n1 (and not 0 in any case), which is then guaranteed to be a subterm of the topmost n.
And the function must be decreasing in n1 (the term we return), rather than n2 (Coq only keeps track of subterms of decreasing arguments).
All that ensures that div2 n is a subterm of n (not a strict subterm (or proper subterm), because n could be O).
This has similarities to the previous fuel-based solution, but here the decreasing argument is a lot more relevant than just a device to trick the typechecker.
This technique is a variant of cons-free programming. (Note though that the constraints are not quite the same as what is discussed in the literature, for example when the focus is on avoiding memory allocations rather than ensuring termination by structural well-foundedness.)
Conclusion: definition of copy
Once we have div2, we can define copy with a few tweaks to obtain i-1 and i-2 as proper subterms of i, again by pattern-matching. Below, i' and i'' are proper subterms of i (by visual inspection), and div2 i' and div2 i'' are subterms of i' and i'' (by the definition of div2). By transitivity they are proper subterms of i, so the termination checker accepts.
Definition copy (x : V) (n : nat) : BraunTree :=
let
fix copy2 (a : V) (i : nat) : (BraunTree * BraunTree) :=
match i with
| 0 => (T a E E,E)
| S i' => (* i' = i-1 *)
if Nat.odd i then
let m := div2 i' in
let (s,t) := copy2 a m in
((T a s t),(T a t t))
else
match i' with
| O => (* Unreachable *) (E, E)
| S i'' => (* i'' = i-2 *)
let m := div2 i'' in
let (s,t) := copy2 a m in
((T a s s),(T a s t))
end
end
in
match copy2 x n with
|(_,snd) => snd
end.

Transform casual list into dependently typed list in Coq

I have following definition of list in Coq:
Variable A : Set.
Variable P : A -> Prop.
Hypothesis P_dec : forall x, {P x}+{~(P x)}.
Inductive plist : nat -> Set :=
pnil : plist O
| pcons : A -> forall n, plist n -> plist n
| pconsp : forall (a:A) n, plist n -> P a -> plist (S n)
.
It describes "list of elements of type A where at least n of them fulfill predicate P".
My task is to create function that will convert casual list into plist (with maximum possible n). My attempt was to first count all elements that match P and then set the output type according to the result:
Fixpoint pcount (l : list A) : nat :=
match l with
| nil => O
| h::t => if P_dec h then S(pcount t) else pcount t
end.
Fixpoint plistIn (l : list A) : (plist (pcount l)) :=
match l with
| nil => pnil
| h::t => match P_dec h with
| left proof => pconsp h _ (plistIn t) proof
| right _ => pcons h _ (plistIn t)
end
end.
However, I get an error in the line with left proof:
Error:
In environment
A : Set
P : A -> Prop
P_dec : forall x : A, {P x} + {~ P x}
plistIn : forall l : list A, plist (pcount l)
l : list A
h : A
t : list A
proof : P h
The term "pconsp h (pcount t) (plistIn t) proof" has type
"plist (S (pcount t))" while it is expected to have type
"plist (pcount (h :: t))".
The problem is that Coq cannot see that S (pcount t) equals pcount (h :: t) knowing that P h, which was already proven. I cannot let Coq know this truth.
How to define this function correctly? Is it even possible to do so?
You can use dependent pattern-matching, as the result type plist (pcount (h :: t)) depends on whether P_dec h is left or right.
Below, the keyword as introduces a new variable p, and return tells the type of the whole match expression, parameterized by p.
Fixpoint plistIn (l : list A) : (plist (pcount l)) :=
match l with
| nil => pnil
| h::t => match P_dec h as p return plist (if p then _ else _) with
| left proof => pconsp h (pcount t) (plistIn t) proof
| right _ => pcons h _ (plistIn t)
end
end.
The type plist (if p then _ else _) must be equal to plist (pcount (h :: t)) when substituting p := P_dec h. Then in each branch, say left proof, you need to produce plist (if left proof then _ else _) (which reduces to the left branch).
It's a bit magical that Coq can infer what goes in the underscores here, but to be safe you can always spell it out: if p then S (pcount t) else pcount t (which is meant to exactly match the definition of pcount).

Coq rewriting using lambda arguments

We have a function that inserts an element into a specific index of a list.
Fixpoint inject_into {A} (x : A) (l : list A) (n : nat) : option (list A) :=
match n, l with
| 0, _ => Some (x :: l)
| S k, [] => None
| S k, h :: t => let kwa := inject_into x t k
in match kwa with
| None => None
| Some l' => Some (h :: l')
end
end.
The following property of the aforementioned function is of relevance to the problem (proof omitted, straightforward induction on l with n not being fixed):
Theorem inject_correct_index : forall A x (l : list A) n,
n <= length l -> exists l', inject_into x l n = Some l'.
And we have a computational definition of permutations, with iota k being a list of nats [0...k]:
Fixpoint permute {A} (l : list A) : list (list A) :=
match l with
| [] => [[]]
| h :: t => flat_map (
fun x => map (
fun y => match inject_into h x y with
| None => []
| Some permutations => permutations
end
) (iota (length t))) (permute t)
end.
The theorem we're trying to prove:
Theorem num_permutations : forall A (l : list A) k,
length l = k -> length (permute l) = factorial k.
By induction on l we can (eventually) get to following goal: length (permute (a :: l)) = S (length l) * length (permute l). If we now simply cbn, the resulting goal is stated as follows:
length
(flat_map
(fun x : list A =>
map
(fun y : nat =>
match inject_into a x y with
| Some permutations => permutations
| None => []
end) (iota (length l))) (permute l)) =
length (permute l) + length l * length (permute l)
Here I would like to proceed by destruct (inject_into a x y), which is impossible considering x and y are lambda arguments. Please note that we will never get the None branch as a result of the lemma inject_correct_index.
How does one proceed from this proof state? (Please do note that I am not trying to simply complete the proof of the theorem, that's completely irrelevant.)
There is a way to rewrite under binders: the setoid_rewrite tactic (see §27.3.1 of the Coq Reference manual).
However, direct rewriting under lambdas is not possible without assuming an axiom as powerful as the axiom of functional extensionality (functional_extensionality).
Otherwise, we could have proved:
(* classical example *)
Goal (fun n => n + 0) = (fun n => n).
Fail setoid_rewrite <- plus_n_O.
Abort.
See here for more detail.
Nevertheless, if you are willing to accept such axiom, then you can use the approach described by Matthieu Sozeau in this Coq Club post to rewrite under lambdas like so:
Require Import Coq.Logic.FunctionalExtensionality.
Require Import Coq.Setoids.Setoid.
Require Import Coq.Classes.Morphisms.
Generalizable All Variables.
Instance pointwise_eq_ext {A B : Type} `(sb : subrelation B RB eq)
: subrelation (pointwise_relation A RB) eq.
Proof. intros f g Hfg. apply functional_extensionality. intro x; apply sb, (Hfg x). Qed.
Goal (fun n => n + 0) = (fun n => n).
setoid_rewrite <- plus_n_O.
reflexivity.
Qed.

How to make sublists in Coq?

I'm working in Coq and trying to figure out how to do the next thing: If I have a list of natural numbers and a given number n, I want to break my list in what goes before and after each of the n's. To make it clearer, if I have the list [1; 2; 0; 3; 4; 0; 9] and the number n = 0, then I want to have as output the three lists: [1;2], [3;4] and [9]. The main problem I have is that I don't know how to output several elements on a Fixpoint. I think I need to nest Fixpoints but I just don't see how. As a very raw idea with one too many issues I have:
Fixpoint SubLists (A : list nat)(m : nat) :=
match A with
|[] => []
|n::A0 => if n =? m then (SubLists L) else n :: (SubLists L)
end.
I would very much appreciate your input on how to do this, and how to navigate having an output of several elements.
You can do this by combining a few fixpoints:
Require Import Coq.Arith.Arith.
Require Import Coq.Lists.List.
Import ListNotations.
Fixpoint prefix n l :=
match l with
| [] => []
| m :: l' => if beq_nat n m then []
else m :: prefix n l'
end.
Fixpoint suffix n l :=
match l with
| [] => l
| m :: l' => if beq_nat n m then l'
else suffix n l'
end.
Fixpoint split_at n l :=
match l with
| [] => []
| m :: l' => prefix n (m :: l') :: split_at n (suffix n (m :: l'))
end.
Notice that Coq's termination checker accepts the recursive call to split_at, even though it is not done syntactically a subterm of l. The reason for that is that it is able to detect that suffix only outputs subterms of its argument. But in order for this to work, we must return l, and not [] on its first branch (try changing it to see what happens!).
In addition to Arthur's solution, you can use an accumulator, which is typical of Functional Programming style:
Require Import Coq.Arith.Arith.
Require Import Coq.Lists.List.
Import ListNotations.
Definition add_acc m (s : list (list nat)) :=
match s with
| [] => [[m]]
| s :: ss => (m :: s) :: ss
end.
Fixpoint split_seq n l acc :=
match l with
| [] => map (#rev _) (rev acc)
| m :: l' => if beq_nat n m then
split_seq n l' ([] :: acc)
else
split_seq n l' (add_acc m acc)
end.
Compute (split_seq 0 [1; 2; 0; 3; 4; 0; 9] []).
Note that the result is reversed so you need to use rev. A bonus exercise is to improve this.
EDIT: Provided second variant that doesn't add [] for repeated separators.
Definition reset_acc (s : list (list nat)) :=
match s with
| [] :: ss => [] :: ss
| ss => [] :: ss
end.
Fixpoint split_seq_nodup n l acc :=
match l with
| [] => map (#rev _) (rev acc)
| m :: l' => if beq_nat n m then
split_seq_nodup n l' (reset_acc acc)
else
split_seq_nodup n l' (add_acc m acc)
end.
Compute (split_seq_nodup 0 [1; 2; 0; 3; 4; 0; 9] []).
An alternative way to tackle this issue is to formally describe the problem you are trying to solve and then either write a dependently-typed function proving that this problem can indeed be solved or using tactics to slowly build up your proof.
This is, if I am not mistaken, a relation describing the relationship between the outputs n and ns you want to pass your function and the output mss you want to get back.
The (* ------- *) lines are simple comments used to suggest that these constructors should be seen as inference rules: whatever is under one such line is the conclusion one can make based on the assumptions above it.
Inductive SubListsRel (n : nat) : forall (ns : list nat)
(mss : list (list nat)), Prop :=
| base : SubListsRel n nil (nil :: nil)
| consEq : forall ns m mss,
n = m -> SubListsRel n ns mss ->
(* ----------------------------- *)
SubListsRel n (m :: ns) (nil :: mss)
| consNotEq : forall ns m ms mss,
(n <> m) -> SubListsRel n ns (ms :: mss) ->
(* ------------------------------------------------- *)
SubListsRel n (m :: ns) ((m :: ms) :: mss)
.
We can then express your Sublists problem as being, given inputs n and ns, the existence of an output mss such that SubListsRel n ns mss holds:
Definition SubLists (n : nat) (ns : list nat) : Set :=
{ mss | SubListsRel n ns mss }.
Using tactics we can readily generate such Sublists for concrete examples in order to sanity-check our specification. We can for instance take the example you had in your original post:
Example example1 : SubLists 0 (1 :: 2 :: 0 :: 3 :: 4 :: 0 :: 9 :: nil).
Proof.
eexists ; repeat econstructor ; intro Hf; inversion Hf.
Defined.
And check that the output is indeed the list you were expecting:
Check (eq_refl : proj1_sig example1
= ((1 :: 2 :: nil) :: (3 :: 4 :: nil) :: (9 :: nil) :: nil)).
Now comes the main part of this post: the proof that forall n ns, SubLists n ns. Given that the premise of consNotEq assumes that mss is non-empty, we will actually prove a strengthened statement in order to make our life easier:
Definition Strenghtened_SubLists (n : nat) (ns : list nat) : Set :=
{ mss | SubListsRel n ns mss /\ mss <> nil }.
And given that oftentimes we will have goals of the shape something_absurd -> False, I define a simple tactic to handle these things. It introduces the absurd assumption and inverts it immediately to make the goal disappear:
Ltac dismiss := intro Hf; inversion Hf.
We can now prove the main statement by proving the strengthened version by induction and deducing it. I guess that here it's better for you to step through it in Coq rather than me trying to explain what happens. The key steps are the cut (proving a stronger statement), induction and the case analysis on eq_nat_dec.
Lemma subLists : forall n ns, SubLists n ns.
Proof
intros n ns; cut (Strenghtened_SubLists n ns).
- intros [mss [Hmss _]]; eexists; eassumption.
- induction ns.
+ eexists; split; [econstructor | dismiss].
+ destruct IHns as [mss [Hmss mssNotNil]];
destruct (eq_nat_dec n a).
* eexists; split; [eapply consEq ; eassumption| dismiss].
* destruct mss; [apply False_rect, mssNotNil; reflexivity |].
eexists; split; [eapply consNotEq; eassumption| dismiss].
Defined.
Once we have this function, we can come back to our example and generate the appropriate Sublists this time not by calling tactics but by running the function subLists we just defined.
Example example2 : SubLists 0 (1 :: 2 :: 0 :: 3 :: 4 :: 0 :: 9 :: nil) :=
subLists _ _.
And we can Check that the computed list is indeed the same as the one obtained in example1:
Check (eq_refl : proj1_sig example1 = proj1_sig example2).
Nota Bene: It is paramount here that our proofs are ended with Defined rather than Qed in order for them to be unfolded when computing with them (which is what we want to do here: they give us the list (list nat) we are looking for!).
A gist with all the code and the right imports.
Here is another take, based on the standard library function List.fold_left.
It works by maintaining an accumulator, which is a pair of the overall reversed result (a list of lists) and a current sublist (also reversed while accumulating). Once we reach a delimiter, we reverse the current sublist and put it into the resulting list of sublists. After executing fold_left, we reverse the result in the outermost match expression.
Require Import Coq.Arith.Arith.
Require Import Coq.Lists.List.
Import ListNotations.
Definition split_skip_dup_delims (m : nat) (xs : list nat) :=
match fold_left
(fun (acctup: _ * _) x => let (acc, rev_subl) := acctup in
if beq_nat x m
then match rev_subl with (* a delimiter found *)
| [] => (acc, []) (* do not insert empty sublist *)
| _ => (rev rev_subl :: acc, []) end
else (acc, x :: rev_subl)) (* keep adding to the current sublist *)
xs
([],[]) with
| (acc, []) => rev acc (* list ends with a delimiter *)
| (acc, rev_subl) => rev (rev rev_subl :: acc) (* no delimiter at the end *)
end.
Eval compute in split_skip_dup_delims 0 [1; 2; 0; 0; 0; 3; 4; 0; 9].
(* = [[1; 2]; [3; 4]; [9]]
: list (list nat) *)