-
Notifications
You must be signed in to change notification settings - Fork 0
/
Copy pathSequences.v
185 lines (146 loc) · 4.48 KB
/
Sequences.v
1
2
3
4
5
6
7
8
9
10
11
12
13
14
15
16
17
18
19
20
21
22
23
24
25
26
27
28
29
30
31
32
33
34
35
36
37
38
39
40
41
42
43
44
45
46
47
48
49
50
51
52
53
54
55
56
57
58
59
60
61
62
63
64
65
66
67
68
69
70
71
72
73
74
75
76
77
78
79
80
81
82
83
84
85
86
87
88
89
90
91
92
93
94
95
96
97
98
99
100
101
102
103
104
105
106
107
108
109
110
111
112
113
114
115
116
117
118
119
120
121
122
123
124
125
126
127
128
129
130
131
132
133
134
135
136
137
138
139
140
141
142
143
144
145
146
147
148
149
150
151
152
153
154
155
156
157
158
159
160
161
162
163
164
165
166
167
168
169
170
171
172
173
174
175
176
177
178
179
180
181
182
183
(** A library of relation operators defining sequences of transitions
and useful properties about them. *)
Require Import Classical.
Set Implicit Arguments.
Section SEQUENCES.
Variable A: Type.
Variable R: A -> A -> Prop.
(** Zero, one or several transitions: reflexive, transitive closure of [R]. *)
Inductive star: A -> A -> Prop :=
| star_refl: forall a,
star a a
| star_step: forall a b c,
R a b -> star b c -> star a c.
Lemma star_one:
forall (a b: A), R a b -> star a b.
Proof.
intros. econstructor; eauto. constructor.
Qed.
Lemma star_trans:
forall (a b: A), star a b -> forall c, star b c -> star a c.
Proof.
induction 1; intros. auto. econstructor; eauto.
Qed.
(** One or several transitions: transitive closure of [R]. *)
Inductive plus: A -> A -> Prop :=
| plus_left: forall a b c,
R a b -> star b c -> plus a c.
Lemma plus_one:
forall a b, R a b -> plus a b.
Proof.
intros. apply plus_left with b. auto. apply star_refl.
Qed.
Lemma plus_star:
forall a b,
plus a b -> star a b.
Proof.
intros. inversion H. eapply star_step; eauto.
Qed.
Lemma plus_star_trans:
forall a b c, plus a b -> star b c -> plus a c.
Proof.
intros. inversion H. eapply plus_left; eauto. eapply star_trans; eauto.
Qed.
Lemma star_plus_trans:
forall a b c, star a b -> plus b c -> plus a c.
Proof.
intros. inversion H0. inversion H. econstructor; eauto.
econstructor; eauto. eapply star_trans; eauto. econstructor; eauto.
Qed.
Lemma plus_right:
forall a b c, star a b -> R b c -> plus a c.
Proof.
intros. eapply star_plus_trans; eauto. apply plus_one; auto.
Qed.
(** Infinitely many transitions. *)
CoInductive infseq: A -> Prop :=
| infseq_step: forall a b,
R a b -> infseq b -> infseq a.
(** Coinduction principles to show the existence of infinite sequences. *)
Lemma infseq_coinduction_principle:
forall (X: A -> Prop),
(forall a, X a -> exists b, R a b /\ X b) ->
forall a, X a -> infseq a.
Proof.
intros X P. cofix COINDHYP; intros.
destruct (P a H) as [b [U V]]. apply infseq_step with b; auto.
Qed.
Lemma infseq_coinduction_principle_2:
forall (X: A -> Prop),
(forall a, X a -> exists b, plus a b /\ X b) ->
forall a, X a -> infseq a.
Proof.
intros.
apply infseq_coinduction_principle with
(X := fun a => exists b, star a b /\ X b).
intros.
destruct H1 as [b [STAR Xb]]. inversion STAR; subst.
destruct (H b Xb) as [c [PLUS Xc]]. inversion PLUS; subst.
exists b0; split. auto. exists c; auto.
exists b0; split. auto. exists b; auto.
exists a; split. apply star_refl. auto.
Qed.
(** An example of use of [infseq_coinduction_principle]. *)
Lemma infseq_alternate_characterization:
forall a,
(forall b, star a b -> exists c, R b c) ->
infseq a.
Proof.
apply infseq_coinduction_principle.
intros. destruct (H a) as [b Rb]. constructor.
exists b; split; auto.
intros. apply H. econstructor; eauto.
Qed.
(** A sequence is either infinite or stops on an irreducible term. *)
Definition irred (a: A) : Prop := forall b, ~(R a b).
Lemma infseq_or_finseq:
forall a, infseq a \/ exists b, star a b /\ irred b.
Proof.
intros.
destruct (classic (forall b, star a b -> exists c, R b c)).
left. apply infseq_alternate_characterization; auto.
right.
generalize (not_all_ex_not _ _ H). intros [b P].
generalize (imply_to_and _ _ P). intros [U V].
exists b; split. auto.
red; intros; red; intros. elim V. exists b0; auto.
Qed.
(** Additional properties for deterministic transition relations. *)
Hypothesis R_determ:
forall a b c, R a b -> R a c -> b = c.
(** Uniqueness of transition sequences. *)
Lemma star_star_inv:
forall a b, star a b -> forall c, star a c -> star b c \/ star c b.
Proof.
induction 1; intros.
auto.
inversion H1; subst. right. eapply star_step; eauto.
assert (b = b0). eapply R_determ; eauto. subst b0.
apply IHstar; auto.
Qed.
Lemma finseq_unique:
forall a b b',
star a b -> irred b ->
star a b' -> irred b' ->
b = b'.
Proof.
intros. destruct (star_star_inv H H1).
inversion H3; subst. auto. elim (H0 _ H4).
inversion H3; subst. auto. elim (H2 _ H4).
Qed.
Lemma infseq_star_inv:
forall a b, star a b -> infseq a -> infseq b.
Proof.
induction 1; intros. auto.
inversion H1; subst. assert (b = b0). eapply R_determ; eauto. subst b0.
apply IHstar. auto.
Qed.
Lemma infseq_finseq_excl:
forall a b,
star a b -> irred b -> infseq a -> False.
Proof.
intros.
assert (infseq b). eapply infseq_star_inv; eauto.
inversion H2. elim (H0 b0); auto.
Qed.
End SEQUENCES.