Built with Alectryon, running Coq+SerAPI v8.15.0+0.15.0. Bubbles () indicate interactive fragments: hover for details, tap to reveal contents. Use Ctrl+↑ Ctrl+↓ to navigate, Ctrl+🖱️ to focus. On Mac, use instead of Ctrl.
(** Setoid Rewriting with boolean inequalities of ssreflect. Solution
    suggested by Georges Gonthier (ssreflect mailinglist @ 18.12.2016) *)

From Coq Require Import Basics Setoid Morphisms.
Notation "[ rel _ _ | _ ]" was already used in scope fun_scope. [notation-overridden,parsing]
Notation "[ rel _ _ : _ | _ ]" was already used in scope fun_scope. [notation-overridden,parsing]
Notation "[ rel _ _ in _ & _ | _ ]" was already used in scope fun_scope. [notation-overridden,parsing]
Notation "[ rel _ _ in _ & _ ]" was already used in scope fun_scope. [notation-overridden,parsing]
Notation "[ rel _ _ in _ | _ ]" was already used in scope fun_scope. [notation-overridden,parsing]
Notation "[ rel _ _ in _ ]" was already used in scope fun_scope. [notation-overridden,parsing]
Notation "_ + _" was already used in scope nat_scope. [notation-overridden,parsing]
Notation "_ - _" was already used in scope nat_scope. [notation-overridden,parsing]
Notation "_ <= _" was already used in scope nat_scope. [notation-overridden,parsing]
Notation "_ < _" was already used in scope nat_scope. [notation-overridden,parsing]
Notation "_ >= _" was already used in scope nat_scope. [notation-overridden,parsing]
Notation "_ > _" was already used in scope nat_scope. [notation-overridden,parsing]
Notation "_ <= _ <= _" was already used in scope nat_scope. [notation-overridden,parsing]
Notation "_ < _ <= _" was already used in scope nat_scope. [notation-overridden,parsing]
Notation "_ <= _ < _" was already used in scope nat_scope. [notation-overridden,parsing]
Notation "_ < _ < _" was already used in scope nat_scope. [notation-overridden,parsing]
Notation "_ * _" was already used in scope nat_scope. [notation-overridden,parsing]
(** This construction allows us to "rewrite" inequalities with ≤. *) Inductive leb a b := Leb of (a ==> b).
a, b: bool

leb a b <-> (a -> b)
a, b: bool

leb a b <-> (a -> b)
move: a b => [|] [|]; firstorder. Qed.

PreOrder leb

PreOrder leb
split => [[|]|[|][|][|][?][?]]; try exact: Leb. Qed.

Proper (leb ==> leb ==> leb) andb

Proper (leb ==> leb ==> leb) andb
move => [|] [|] [A] c d [B]; exact: Leb. Qed.

Proper (leb ==> leb ==> leb) orb

Proper (leb ==> leb ==> leb) orb
move => [|] [|] [A] [|] d [B]; exact: Leb. Qed.

Proper (leb ==> impl) is_true

Proper (leb ==> impl) is_true
a, b: bool

a ==> b -> impl a b
exact: implyP. Qed.

Proper (le --> le ==> leb) leq

Proper (le --> le ==> leb) leq
n, m: nat
__view_subject_1_: m <= n
n', m': nat
__view_subject_3_: n' <= m'

leb (n <= n') (m <= m')
n, m: nat
__view_subject_1_: m <= n
n', m': nat
__view_subject_3_: n' <= m'
_Hyp_: n <= n'

m <= m'
eauto using leq_trans. Qed.

Proper (le ==> le ==> le) addn

Proper (le ==> le ==> le) addn
n, m: nat
__view_subject_1_: n <= m
n', m': nat
__view_subject_3_: n' <= m'

(n + n' <= m + m')%coq_nat
n, m: nat
__view_subject_1_: n <= m
n', m': nat
__view_subject_3_: n' <= m'

n + n' <= m + m'
exact: leq_add. Qed.

Proper (le ==> le --> le) subn

Proper (le ==> le --> le) subn
n, m: nat
__view_subject_1_: n <= m
n', m': nat
__view_subject_3_: m' <= n'

(n - n' <= m - m')%coq_nat
n, m: nat
__view_subject_1_: n <= m
n', m': nat
__view_subject_3_: m' <= n'

n - n' <= m - m'
exact: leq_sub. Qed.

Proper (le ==> le) succn

Proper (le ==> le) succn
n, m: nat
__view_subject_1_: n <= m

(n.+1 <= m.+1)%coq_nat
n, m: nat
__view_subject_1_: n <= m

n < m.+1
by rewrite ltnS. Qed. #[global] Instance : RewriteRelation le. Defined. Definition leqRW {m n} : m <= n -> le m n := leP. (** To see the benefits, consider the following example. *) (* [Goal ] [ forall a b c d x y, ] [ x <= y -> ] [ a + (b + (x + c)) + d <= a + (b + (y + c)) + d. ] [Proof. ] [ move => a b c d x y LE. ] [ (* This could be an unpleasant proof, but we can just [rewrite LE] *) ] [ by rewrite (leqRW LE). (* magic here *) ] [Qed. ] *) (** Another benefit of [leqRW] is that it allows to avoid unnecessary [eapply leq_trans]. *) (* [Goal ] [ forall a b x y z, ] [ a <= x -> ] [ b <= y -> ] [ x + y <= z -> ] [ a + b <= z. ] [Proof. ] [ move => a b x y z LE1 LE2 LE3. ] [ rewrite -(leqRW LE3) leq_add //. ] [Qed. ] *)