coq-club AT inria.fr
Subject: The Coq mailing list
List archive
- From: mukesh tiwari <mukeshtiwari.iiitm AT gmail.com>
- To: coq-club AT inria.fr
- Subject: [Coq-Club] Proof in functional style
- Date: Fri, 30 Nov 2018 15:24:23 +1100
- Authentication-results: mail2-smtp-roc.national.inria.fr; spf=None smtp.pra=mukeshtiwari.iiitm AT gmail.com; spf=Pass smtp.mailfrom=mukeshtiwari.iiitm AT gmail.com; spf=None smtp.helo=postmaster AT mail-wr1-f42.google.com
- Ironport-phdr: 9a23:I2o8Fhwj6E9RxnzXCy+O+j09IxM/srCxBDY+r6Qd1O0WIJqq85mqBkHD//Il1AaPAd2Lraocw8Pt8InYEVQa5piAtH1QOLdtbDQizfssogo7HcSeAlf6JvO5JwYzHcBFSUM3tyrjaRsdF8nxfUDdrWOv5jAOBBr/KRB1JuPoEYLOksi7ze+/94HQbglSmDaxfa55IQmrownWqsQYm5ZpJLwryhvOrHtIeuBWyn1tKFmOgRvy5dq+8YB6/ShItP0v68BPUaPhf6QlVrNYFygpM3o05MLwqxbOSxaE62YGXWUXlhpIBBXF7A3/U5zsvCb2qvZx1S+HNsDwULs6Wymt771zRRDniCkJOT03/nzJhMNsl69Uug6tqgZlzoLIfI2YNvxzdb7dc9MAQmpBW95cVylcAoO/cYQPFOoBNvtEr4n8qVoBtwG+BQixD+7ozz9Im3z20rMh0+QhDQHG3QIhEMgUsHTVttn1NaYSUeGpzKnN1jjDYPZW1i386IjMaBwuvfaMXbdpfMfX1EIhGQTFjlCKpozkOTOYzusNs2mH7+pgSOKgkHQrpB12ojiqwMonl4rHhpoNx1za6Sl0xJw5KN64RUJhf9KoDoZcuzuVOodoRM4pXntmtzwgyrIcvJ62ZCgKx4ojxx7Yc/GHdpKH4hPnVOqILzZ4nm9pdKuxhxu970Ss0OL8Vs6z0FZFqipKjMPAuWwK1xzW8sSHS/198Vm92TuXyQzf9uVJLVo3mKfbMZIt3KA8mocJvUnMACP6gED2g7WXdkUg9Oio8ePnYrD+q5+ZKo90iR/xMr8umsyjAOQ5PBIBX3Ka+eim1b3j/Ez5QKlPjvAtnanZtYrVJcUfpqKjHwBV1YMj5w6lDzi6yNQYgWUHLFVddR2biIjpIkjCL+z8DfeimFuhiyxrxvDDPr35GJrBNHnDkLH7fbZ88UFQ0gQzzcoMr65TX7oGObf4XlL7nN3eFB4wdQKukMj9D9Ao054dVHmPSrOYL6rIsBfc4/8sLvKMeI4KsSz8bfkk5uLrpXA8kF4ZO6Ku2M1EOziDAv16LhDBMjLXidAbHDJS51tsfKnRkFSHFAVrSTO3VqM46Cs8Ddv/X4jGT4GpxreG2XXiR8EEViV9ElmJVEzQWcCcQf5VMXCdJ8ZglnoPUr3zE9Z8hyHrjxfzzv9cFsSR+iAcssi+ht185umWigtrsDIoVYKS1GaCS2wylWQNFWc7
Hi Everyone, I have written a function (basically Haskell intersperse function)
Fixpoint intersperse (A : Type) (c : A) (l : list A) : list A :=
match l with
| [] => []
| [h] => [h]
| h :: t => h :: c :: intersperse _ c t
end.
Compute (intersperse _ 0 []).
Compute (intersperse _ 0 [1]).
Compute (intersperse _ 0 [1; 2; 3]).
match l with
| [] => []
| [h] => [h]
| h :: t => h :: c :: intersperse _ c t
end.
Compute (intersperse _ 0 []).
Compute (intersperse _ 0 [1]).
Compute (intersperse _ 0 [1; 2; 3]).
and proof about it's correctness
Lemma map_intersperse : forall (A B : Type) (f : A -> B) (c : A) (l : list A),
map f (intersperse _ c l) = intersperse _ (f c) (map f l).
Proof.
intros A B f c.
induction l.
+ cbn; auto.
+ destruct l.
++ cbn; auto.
++ cbn in *. repeat f_equal.
(* At this point, the induction hypothesis is [1], and assumption is able to discharge
the proof *)
assumption.
Qed.
Now I rewrote the same proof in function style using refine tactic.
Lemma map_intersperse_not_working : forall (A B : Type) (f : A -> B) (c : A) (l : list A),
map f (intersperse _ c l) = intersperse _ (f c) (map f l).
refine (fun A B f c =>
fix F l :=
match l as l'
return l = l' ->
map f (intersperse A c l') =
intersperse B (f c) (map f l') with
| [] => fun H => eq_refl
| [h] => fun H => eq_refl
| h1 :: h2 :: t => fun H => _
end eq_refl).
pose proof (F (h2 :: t)).
cbn in *. repeat f_equal. Guarded.
(* At this point assumption H0 is same as goal [2] so apply H0 should discharge the goal*)
map f (intersperse _ c l) = intersperse _ (f c) (map f l).
Proof.
intros A B f c.
induction l.
+ cbn; auto.
+ destruct l.
++ cbn; auto.
++ cbn in *. repeat f_equal.
(* At this point, the induction hypothesis is [1], and assumption is able to discharge
the proof *)
assumption.
Qed.
Now I rewrote the same proof in function style using refine tactic.
Lemma map_intersperse_not_working : forall (A B : Type) (f : A -> B) (c : A) (l : list A),
map f (intersperse _ c l) = intersperse _ (f c) (map f l).
refine (fun A B f c =>
fix F l :=
match l as l'
return l = l' ->
map f (intersperse A c l') =
intersperse B (f c) (map f l') with
| [] => fun H => eq_refl
| [h] => fun H => eq_refl
| h1 :: h2 :: t => fun H => _
end eq_refl).
pose proof (F (h2 :: t)).
cbn in *. repeat f_equal. Guarded.
(* At this point assumption H0 is same as goal [2] so apply H0 should discharge the goal*)
apply H0. Guarded.
(* I am getting Recursive definition of F is ill-formed. *)Could some one point me why it's saying F is ill-formed ? I am calling it smaller argument (At least that is my understanding).
I got working solution with help of Li-yao [3], but I am wondering why destruct t is making a difference in this solution which I am doing explicitly in previous one (matching on one element list [h])
Lemma map_intersperse_not_working : forall (A B : Type) (f : A -> B) (c : A) (l : list A),
map f (intersperse _ c l) = intersperse _ (f c) (map f l).
refine (fun A B f c =>
fix F l :=
match l as l'
return l = l' ->
map f (intersperse A c l') =
intersperse B (f c) (map f l') with
| [] => fun H => eq_refl
| h :: t => fun H => _
end eq_refl).
pose proof (F t).
destruct t.
+ cbn. exact eq_refl.
+ cbn in *. repeat f_equal. apply H0. Guarded.
Qed.
Best regards,
Mukesh Tiwari
- [Coq-Club] Proof in functional style, mukesh tiwari, 11/30/2018
- Re: [Coq-Club] Proof in functional style, mukesh tiwari, 11/30/2018
- Re: [Coq-Club] Proof in functional style, Larry Lee, 11/30/2018
- Re: [Coq-Club] Proof in functional style, Pierre Courtieu, 11/30/2018
Archive powered by MHonArc 2.6.18.