playground/coq/unfinished/stream-eq-casteran.v

65 lines
1.6 KiB
Coq

(* https://www.labri.fr/perso/casteran/RecTutorial.pdf *)
CoInductive stream (A:Type) : Type :=
| Cons : A -> stream A -> stream A.
Arguments Cons {A}.
Definition hd {A:Type} (s:stream A) : A :=
match s with
| Cons val _ => val
end.
Definition tl {A:Type} (s:stream A) : stream A :=
match s with
| Cons _ ss => ss
end.
CoFixpoint map {A B:Type} (f:A->B) (a:stream A) : stream B :=
match a with
| Cons val aa => Cons (f val) (map f aa)
end.
CoFixpoint iterate {A:Type} (f:A->A) (start:A) : stream A :=
Cons start (iterate f (f start)).
CoInductive EqSt {A:Type} : stream A -> stream A -> Prop :=
| eqst : forall a b:stream A,
(hd a = hd b) -> EqSt (tl a) (tl b) -> EqSt a b.
Section sec_park.
Context {A:Type}.
Variable R:stream A -> stream A -> Prop.
Hypothesis bisim_hd : forall a b:stream A,
R a b -> (hd a = hd b).
Hypothesis bisim_tl : forall a b:stream A,
R a b -> R (tl a) (tl b).
CoFixpoint park_ppl : forall (a b:stream A),
R a b -> EqSt a b := fun a b (prp:R a b) =>
eqst a b (bisim_hd a b prp)
(park_ppl (tl a) (tl b) (bisim_tl a b prp)).
End sec_park.
Theorem map_iterate : forall {A:Type} (f:A->A) (x:A),
EqSt (iterate f (f x)) (map f (iterate f x)).
Proof.
intros.
apply park_ppl with
(R := fun a b =>
exists x:A, a = iterate f (f x) /\
b = map f (iterate f x)).
- intros.
destruct H as (ex_x, (eq_a, eq_b)).
rewrite eq_a.
rewrite eq_b.
reflexivity.
- intros a b (ex_x, (eq_a, eq_b)).
exists (f ex_x).
split.
* rewrite eq_a.
reflexivity.
* rewrite eq_b.
reflexivity.
- exists x.
split; reflexivity.
Qed.