frap/Invariant.v

45 lines
1.2 KiB
Coq
Raw Normal View History

2015-12-31 20:44:34 +00:00
Require Import Relations.
2016-02-08 23:14:11 +00:00
Set Implicit Arguments.
2015-12-31 20:44:34 +00:00
2016-02-08 23:14:11 +00:00
Record trsys state := {
Initial : state -> Prop;
Step : state -> state -> Prop
}.
Definition invariantFor {state} (sys : trsys state) (invariant : state -> Prop) :=
forall s, sys.(Initial) s
-> forall s', sys.(Step)^* s s'
2016-02-08 23:04:14 +00:00
-> invariant s'.
2016-02-08 23:14:11 +00:00
Theorem use_invariant : forall {state} (sys : trsys state) (invariant : state -> Prop) s s',
sys.(Step)^* s s'
-> sys.(Initial) s
-> invariantFor sys invariant
2016-02-08 23:04:14 +00:00
-> invariant s'.
Proof.
firstorder.
Qed.
2016-02-08 23:14:11 +00:00
Theorem invariantFor_monotone : forall {state} (sys : trsys state)
(invariant1 invariant2 : state -> Prop),
2016-02-08 23:04:14 +00:00
(forall s, invariant1 s -> invariant2 s)
2016-02-08 23:14:11 +00:00
-> invariantFor sys invariant1
-> invariantFor sys invariant2.
2016-02-08 23:04:14 +00:00
Proof.
unfold invariantFor; intuition eauto.
Qed.
2016-02-08 23:14:11 +00:00
Theorem invariant_induction : forall {state} (sys : trsys state)
(invariant : state -> Prop),
(forall s, sys.(Initial) s -> invariant s)
-> (forall s, invariant s -> forall s', sys.(Step) s s' -> invariant s')
-> invariantFor sys invariant.
2016-02-08 23:04:14 +00:00
Proof.
unfold invariantFor; intros.
assert (invariant s) by eauto.
clear H1.
induction H2; eauto.
Qed.