2015-12-31 20:44:34 +00:00
|
|
|
Set Implicit Arguments.
|
|
|
|
|
|
|
|
|
|
|
|
Section trc.
|
|
|
|
Variable A : Type.
|
|
|
|
Variable R : A -> A -> Prop.
|
|
|
|
|
|
|
|
Inductive trc : A -> A -> Prop :=
|
|
|
|
| TrcRefl : forall x, trc x x
|
|
|
|
| TrcFront : forall x y z,
|
|
|
|
R x y
|
|
|
|
-> trc y z
|
|
|
|
-> trc x z.
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Constructors trc : core.
|
2015-12-31 20:44:34 +00:00
|
|
|
|
2016-04-21 23:12:02 +00:00
|
|
|
Theorem trc_one : forall x y, R x y
|
|
|
|
-> trc x y.
|
|
|
|
Proof.
|
|
|
|
eauto.
|
|
|
|
Qed.
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Resolve trc_one : core.
|
2016-04-21 23:12:02 +00:00
|
|
|
|
2015-12-31 20:44:34 +00:00
|
|
|
Theorem trc_trans : forall x y, trc x y
|
|
|
|
-> forall z, trc y z
|
|
|
|
-> trc x z.
|
|
|
|
Proof.
|
|
|
|
induction 1; eauto.
|
|
|
|
Qed.
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Resolve trc_trans : core.
|
2015-12-31 20:44:34 +00:00
|
|
|
|
|
|
|
Inductive trcEnd : A -> A -> Prop :=
|
|
|
|
| TrcEndRefl : forall x, trcEnd x x
|
|
|
|
| TrcBack : forall x y z,
|
|
|
|
trcEnd x y
|
|
|
|
-> R y z
|
|
|
|
-> trcEnd x z.
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Constructors trcEnd : core.
|
2015-12-31 20:44:34 +00:00
|
|
|
|
|
|
|
Lemma TrcFront' : forall x y z,
|
|
|
|
R x y
|
|
|
|
-> trcEnd y z
|
|
|
|
-> trcEnd x z.
|
|
|
|
Proof.
|
|
|
|
induction 2; eauto.
|
|
|
|
Qed.
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Resolve TrcFront' : core.
|
2015-12-31 20:44:34 +00:00
|
|
|
|
|
|
|
Theorem trc_trcEnd : forall x y, trc x y
|
|
|
|
-> trcEnd x y.
|
|
|
|
Proof.
|
|
|
|
induction 1; eauto.
|
|
|
|
Qed.
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Resolve trc_trcEnd : core.
|
2015-12-31 20:44:34 +00:00
|
|
|
|
|
|
|
Lemma TrcBack' : forall x y z,
|
|
|
|
trc x y
|
|
|
|
-> R y z
|
|
|
|
-> trc x z.
|
|
|
|
Proof.
|
|
|
|
induction 1; eauto.
|
|
|
|
Qed.
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Resolve TrcBack' : core.
|
2015-12-31 20:44:34 +00:00
|
|
|
|
|
|
|
Theorem trcEnd_trans : forall x y, trcEnd x y
|
|
|
|
-> forall z, trcEnd y z
|
|
|
|
-> trcEnd x z.
|
|
|
|
Proof.
|
|
|
|
induction 1; eauto.
|
|
|
|
Qed.
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Resolve trcEnd_trans : core.
|
2015-12-31 20:44:34 +00:00
|
|
|
|
|
|
|
Theorem trcEnd_trc : forall x y, trcEnd x y
|
|
|
|
-> trc x y.
|
|
|
|
Proof.
|
|
|
|
induction 1; eauto.
|
|
|
|
Qed.
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Resolve trcEnd_trc : core.
|
2015-12-31 20:44:34 +00:00
|
|
|
|
|
|
|
Inductive trcLiteral : A -> A -> Prop :=
|
|
|
|
| TrcLiteralRefl : forall x, trcLiteral x x
|
|
|
|
| TrcTrans : forall x y z, trcLiteral x y
|
|
|
|
-> trcLiteral y z
|
|
|
|
-> trcLiteral x z
|
|
|
|
| TrcInclude : forall x y, R x y
|
|
|
|
-> trcLiteral x y.
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Constructors trcLiteral : core.
|
2015-12-31 20:44:34 +00:00
|
|
|
|
|
|
|
Theorem trc_trcLiteral : forall x y, trc x y
|
|
|
|
-> trcLiteral x y.
|
|
|
|
Proof.
|
|
|
|
induction 1; eauto.
|
|
|
|
Qed.
|
|
|
|
|
|
|
|
Theorem trcLiteral_trc : forall x y, trcLiteral x y
|
|
|
|
-> trc x y.
|
|
|
|
Proof.
|
|
|
|
induction 1; eauto.
|
|
|
|
Qed.
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Resolve trc_trcLiteral trcLiteral_trc : core.
|
2015-12-31 20:44:34 +00:00
|
|
|
|
|
|
|
Theorem trcEnd_trcLiteral : forall x y, trcEnd x y
|
|
|
|
-> trcLiteral x y.
|
|
|
|
Proof.
|
|
|
|
induction 1; eauto.
|
|
|
|
Qed.
|
|
|
|
|
|
|
|
Theorem trcLiteral_trcEnd : forall x y, trcLiteral x y
|
|
|
|
-> trcEnd x y.
|
|
|
|
Proof.
|
|
|
|
induction 1; eauto.
|
|
|
|
Qed.
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Resolve trcEnd_trcLiteral trcLiteral_trcEnd : core.
|
2015-12-31 20:44:34 +00:00
|
|
|
End trc.
|
|
|
|
|
|
|
|
Notation "R ^*" := (trc R) (at level 0).
|
|
|
|
Notation "*^ R" := (trcEnd R) (at level 0).
|
|
|
|
|
2020-02-08 20:15:38 +00:00
|
|
|
Hint Constructors trc : core.
|