Library Coq.Structures.Equalities
Structure with nothing inside.
Used to force a module type T into a module via Nop <+ T. (HACK!)
Module Type HasEq (Import T:Typ).
Parameter Inline(30) eq : t -> t -> Prop.
End HasEq.
Module Type Eq := Typ <+ HasEq.
Module Type EqNotation (Import E:Eq).
Infix "==" := eq (at level 70, no associativity).
Notation "x ~= y" := (~eq x y) (at level 70, no associativity).
End EqNotation.
Module Type Eq' := Eq <+ EqNotation.
Module Type IsEqOrig (Import E:Eq').
Axiom eq_refl : forall x : t, x==x.
Axiom eq_sym : forall x y : t, x==y -> y==x.
Axiom eq_trans : forall x y z : t, x==y -> y==z -> x==z.
#[global]
Hint Immediate eq_sym : core.
#[global]
Hint Resolve eq_refl eq_trans : core.
End IsEqOrig.
Module Type HasEqDec (Import E:Eq').
Parameter eq_dec : forall x y : t, { x==y } + { ~ x==y }.
End HasEqDec.
Module Type HasEqb (Import T:Typ).
Parameter Inline eqb : t -> t -> bool.
End HasEqb.
Module Type EqbSpec (T:Typ)(X:HasEq T)(Y:HasEqb T).
Parameter eqb_eq : forall x y, Y.eqb x y = true <-> X.eq x y.
End EqbSpec.
Module Type EqbNotation (T:Typ)(E:HasEqb T).
Infix "=?" := E.eqb (at level 70, no associativity).
End EqbNotation.
Module Type HasEqBool (E:Eq) := HasEqb E <+ EqbSpec E E.
From these basic blocks, we can build many combinations
of static standalone module types.
Module Type EqualityType := Eq <+ IsEq.
Module Type EqualityTypeOrig := Eq <+ IsEqOrig.
Module Type EqualityTypeBoth <: EqualityType <: EqualityTypeOrig
:= Eq <+ IsEq <+ IsEqOrig.
Module Type DecidableType <: EqualityType
:= Eq <+ IsEq <+ HasEqDec.
Module Type DecidableTypeOrig <: EqualityTypeOrig
:= Eq <+ IsEqOrig <+ HasEqDec.
Module Type DecidableTypeBoth <: DecidableType <: DecidableTypeOrig
:= EqualityTypeBoth <+ HasEqDec.
Module Type BooleanEqualityType <: EqualityType
:= Eq <+ IsEq <+ HasEqBool.
Module Type BooleanDecidableType <: DecidableType <: BooleanEqualityType
:= Eq <+ IsEq <+ HasEqDec <+ HasEqBool.
Module Type DecidableTypeFull <: DecidableTypeBoth <: BooleanDecidableType
:= Eq <+ IsEq <+ IsEqOrig <+ HasEqDec <+ HasEqBool.
Same, with notation for eq
Module Type EqualityType' := EqualityType <+ EqNotation.
Module Type EqualityTypeOrig' := EqualityTypeOrig <+ EqNotation.
Module Type EqualityTypeBoth' := EqualityTypeBoth <+ EqNotation.
Module Type DecidableType' := DecidableType <+ EqNotation.
Module Type DecidableTypeOrig' := DecidableTypeOrig <+ EqNotation.
Module Type DecidableTypeBoth' := DecidableTypeBoth <+ EqNotation.
Module Type BooleanEqualityType' :=
BooleanEqualityType <+ EqNotation <+ EqbNotation.
Module Type BooleanDecidableType' :=
BooleanDecidableType <+ EqNotation <+ EqbNotation.
Module Type DecidableTypeFull' := DecidableTypeFull <+ EqNotation.
Module BackportEq (E:Eq)(F:IsEq E) <: IsEqOrig E.
Definition eq_refl := @Equivalence_Reflexive _ _ F.eq_equiv.
Definition eq_sym := @Equivalence_Symmetric _ _ F.eq_equiv.
Definition eq_trans := @Equivalence_Transitive _ _ F.eq_equiv.
End BackportEq.
Module UpdateEq (E:Eq)(F:IsEqOrig E) <: IsEq E.
#[global]
Instance eq_equiv : Equivalence E.eq.
End UpdateEq.
Module Backport_ET (E:EqualityType) <: EqualityTypeBoth
:= E <+ BackportEq.
Module Update_ET (E:EqualityTypeOrig) <: EqualityTypeBoth
:= E <+ UpdateEq.
Module Backport_DT (E:DecidableType) <: DecidableTypeBoth
:= E <+ BackportEq.
Module Update_DT (E:DecidableTypeOrig) <: DecidableTypeBoth
:= E <+ UpdateEq.
Module HasEqDec2Bool (E:Eq)(F:HasEqDec E) <: HasEqBool E.
Definition eqb x y := if F.eq_dec x y then true else false.
Lemma eqb_eq : forall x y, eqb x y = true <-> E.eq x y.
End HasEqDec2Bool.
Module HasEqBool2Dec (E:Eq)(F:HasEqBool E) <: HasEqDec E.
Lemma eq_dec : forall x y, {E.eq x y}+{~E.eq x y}.
End HasEqBool2Dec.
Module Dec2Bool (E:DecidableType) <: BooleanDecidableType
:= E <+ HasEqDec2Bool.
Module Bool2Dec (E:BooleanEqualityType) <: BooleanDecidableType
:= E <+ HasEqBool2Dec.
Some properties of boolean equality
eqb is compatible with eq
Alternative specification of eqb based on reflect.
Negated form of eqb_eq
Basic equality laws for eqb
Transitivity is a particular case of eqb_compat
Module Type HasUsualEq (Import T:Typ) <: HasEq T.
Definition eq := @Logic.eq t.
End HasUsualEq.
Module Type UsualEq <: Eq := Typ <+ HasUsualEq.
Module Type UsualIsEq (E:UsualEq) <: IsEq E.
Definition eq_equiv : Equivalence E.eq := eq_equivalence.
End UsualIsEq.
Module Type UsualIsEqOrig (E:UsualEq) <: IsEqOrig E.
Definition eq_refl := @Logic.eq_refl E.t.
Definition eq_sym := @Logic.eq_sym E.t.
Definition eq_trans := @Logic.eq_trans E.t.
End UsualIsEqOrig.
Module Type UsualEqualityType <: EqualityType
:= UsualEq <+ UsualIsEq.
Module Type UsualDecidableType <: DecidableType
:= UsualEq <+ UsualIsEq <+ HasEqDec.
Module Type UsualDecidableTypeOrig <: DecidableTypeOrig
:= UsualEq <+ UsualIsEqOrig <+ HasEqDec.
Module Type UsualDecidableTypeBoth <: DecidableTypeBoth
:= UsualEq <+ UsualIsEq <+ UsualIsEqOrig <+ HasEqDec.
Module Type UsualBoolEq := UsualEq <+ HasEqBool.
Module Type UsualDecidableTypeFull <: DecidableTypeFull
:= UsualEq <+ UsualIsEq <+ UsualIsEqOrig <+ HasEqDec <+ HasEqBool.
Some shortcuts for easily building a UsualDecidableType
Module Type MiniDecidableType.
Include Typ.
Parameter eq_dec : forall x y : t, {x=y}+{~x=y}.
End MiniDecidableType.
Module Make_UDT (M:MiniDecidableType) <: UsualDecidableTypeBoth
:= M <+ HasUsualEq <+ UsualIsEq <+ UsualIsEqOrig.
Module Make_UDTF (M:UsualBoolEq) <: UsualDecidableTypeFull
:= M <+ UsualIsEq <+ UsualIsEqOrig <+ HasEqBool2Dec.