Instance packages and factories for them #
Instance packages are classes with the sole purpose to bundle together multiple smaller classes. They should not be used as hypotheses, but they make it more convenient to define multiple instances at once.
This class entails LE α
, LT α
and BEq α
instances as well as proofs that these operations
represent the same preorder structure on α
.
- decidableLE : DecidableLE α
- decidableLT : DecidableLT α
Instances
Instances For
If LT
can be characterized in terms of a decidable LE
, then LT
is decidable either.
Equations
Instances For
This structure contains all the data needed to create a PreorderPackage α
instance. Its fields
are automatically provided if possible. For the detailed rules how the fields are inferred, see
PreorderPackage.ofLE
.
- le : LE α
- decidableLE : DecidableLE α
- beq : let this := self.le; let this := self.decidableLE; BEq α
- decidableLT : let this := self.le; let this := self.decidableLE; let this := self.lt; have this_1 := ⋯; DecidableLT α
Instances For
Use this factory to conveniently define a preorder on a type α
and all the associated operations
and instances given an LE α
instance.
Creates a PreorderPackage α
instance. Such an instance entails LE α
, LT α
and
BEq α
instances as well as an IsPreorder α
instance and LawfulOrder*
instances proving the
compatibility of the operations with the preorder.
In the presence of LE α
, DecidableLE α
, Refl (· ≤ ·)
and Trans (· ≤ ·) (· ≤ ·) (· ≤ ·)
instances, no arguments are required and the factory can be used as in this example:
public instance : PreorderPackage X := .ofLE X
If not all of these instances are available via typeclass synthesis, it is necessary to explicitly provide some arguments:
public instance : PreorderPackage X := .ofLE X {
le_refl := sorry
le_trans := sorry }
It is also possible to do all of this by hand, without resorting to PreorderPackage
. This can
be useful if, say, one wants to avoid specifying an LT α
instance, which is not possible with
PreorderPackage
.
How the arguments are filled
Lean tries to fill all of the fields of the args : Packages.PreorderOfLEArgs α
parameter
automatically. If it fails, it is necessary to provide some of the fields manually.
- For the data-carrying typeclasses
LE
,LT
andBEq
, existing instances are always preferred. If no existing instances can be synthesized, it is attempted to derive an instance from theLE
instance. - Some proof obligations can be filled automatically if the data-carrying typeclasses have been
derived from the
LE
instance. For example: If thebeq
field is omitted and noBEq α
instance can be synthesized, it is derived from theLE α
instance. In this case,beq_iff_le_and_ge
can be omitted because Lean can infer thatBEq α
andLE α
are compatible. - Other proof obligations, namely
le_refl
andle_trans
, can be omitted ifRefl
andTrans
instances can be synthesized.
Equations
- One or more equations did not get rendered due to their size.
Instances For
This class entails LE α
, LT α
and BEq α
instances as well as proofs that these operations
represent the same partial order structure on α
.
Instances
This structure contains all the data needed to create a PartialOrderPakckage α
instance. Its
fields are automatically provided if possible. For the detailed rules how the fields are inferred,
see PartialOrderPackage.ofLE
.
- beq : let this := self.le; let this := self.decidableLE; BEq α
- decidableLT : let this := self.le; let this := self.decidableLE; let this := self.lt; have this_1 := ⋯; DecidableLT α
Instances For
Equations
- Std.PartialOrderPackage.ofLE α args = { toPreorderPackage := Std.PreorderPackage.ofLE α args.toPreorderOfLEArgs, le_antisymm := ⋯ }
Instances For
Equations
- Std.FactoryInstances.instOrdOfDecidableLE = { compare := fun (a b : α) => if a ≤ b then if b ≤ a then Ordering.eq else Ordering.lt else Ordering.gt }
Instances For
This class entails LE α
, LT α
, BEq α
and Ord α
instances as well as proofs that these
operations represent the same linear preorder structure on α
.
Instances
This structure contains all the data needed to create a LinearPreorderPackage α
instance. Its fields
are automatically provided if possible. For the detailed rules how the fields are inferred, see
LinearPreorderPackage.ofLE
.
- beq : let this := self.le; let this := self.decidableLE; BEq α
- decidableLT : let this := self.le; let this := self.decidableLE; let this := self.lt; have this_1 := ⋯; DecidableLT α
- ord : let this := self.le; let this := self.decidableLE; Ord α
Instances For
Use this factory to conveniently define a linear preorder on a type α
and all the associated
operations and instances given an LE α
instance.
Creates a LinearPreorderPackage α
instance. Such an instance entails LE α
, LT α
, BEq α
and
Ord α
instances as well as an IsLinearPreorder α
instance and LawfulOrder*
instances proving
the compatibility of the operations with the linear preorder.
In the presence of LE α
, DecidableLE α
, Total (· ≤ ·)
and Trans (· ≤ ·) (· ≤ ·) (· ≤ ·)
instances, no arguments are required and the factory can be used as in this example:
public instance : LinearPreorderPackage X := .ofLE X
If not all of these instances are available via typeclass synthesis, it is necessary to explicitly provide some arguments:
public instance : LinearPreorderPackage X := .ofLE X {
le_total := sorry
le_trans := sorry }
It is also possible to do all of this by hand, without resorting to LinearPreorderPackage
. This
can be useful if, say, one wants to avoid specifying an LT α
instance, which is not possible with
LinearPreorderPackage
.
How the arguments are filled
Lean tries to fill all of the fields of the args : Packages.LinearPreorderOfLEArgs α
parameter
automatically. If it fails, it is necessary to provide some of the fields manually.
- For the data-carrying typeclasses
LE
,LT
,BEq
andOrd
, existing instances are always preferred. If no existing instances can be synthesized, it is attempted to derive an instance from theLE
instance. - Some proof obligations can be filled automatically if the data-carrying typeclasses have been
derived from the
LE
instance. For example: If thebeq
field is omitted and noBEq α
instance can be synthesized, it is derived from theLE α
instance. In this case,beq_iff_le_and_ge
can be omitted because Lean can infer thatBEq α
andLE α
are compatible. - Other proof obligations, namely
le_total
andle_trans
, can be omitted ifTotal
andTrans
instances can be synthesized.
Equations
- Std.LinearPreorderPackage.ofLE α args = { toPreorderPackage := Std.PreorderPackage.ofLE α args.toPreorderOfLEArgs, toOrd := args.ord, toLawfulOrderOrd := ⋯, le_total := ⋯ }
Instances For
This class entails LE α
, LT α
, BEq α
, Ord α
, Min α
and Max α
instances as well as proofs
that these operations represent the same linear order structure on α
.
Instances
This structure contains all the data needed to create a LinearOrderPackage α
instance. Its fields
are automatically provided if possible. For the detailed rules how the fields are inferred, see
LinearOrderPackage.ofLE
.
- beq : let this := self.le; let this := self.decidableLE; BEq α
- decidableLT : let this := self.le; let this := self.decidableLE; let this := self.lt; have this_1 := ⋯; DecidableLT α
- ord : let this := self.le; let this := self.decidableLE; Ord α
- isLE_compare : let this := self.le; let this_1 := self.decidableLE; let this_2 := self.ord; ∀ (a b : α), (compare a b).isLE = true ↔ a ≤ b
- isGE_compare : let this := self.le; let this_1 := self.decidableLE; have this_2 := ⋯; let this_3 := self.ord; ∀ (a b : α), (compare a b).isGE = true ↔ b ≤ a
- le_antisymm : let this := self.le; ∀ (a b : α), a ≤ b → b ≤ a → a = b
- min : let this := self.le; let this := self.decidableLE; Min α
- max : let this := self.le; let this := self.decidableLE; Max α
Instances For
Use this factory to conveniently define a linear order on a type α
and all the associated
operations and instances given an LE α
instance.
Creates a LinearOrderPackage α
instance. Such an instance entails LE α
, LT α
, BEq α
,
Ord α
, Min α
and Max α
instances as well as an IsLinearOrder α
instance and LawfulOrder*
instances proving the compatibility of the operations with the linear order.
In the presence of LE α
, DecidableLE α
, Total (· ≤ ·)
, Trans (· ≤ ·) (· ≤ ·) (· ≤ ·)
and
Antisymm (· ≤ ·)
instances, no arguments are required and the factory can be used as in this
example:
public instance : LinearOrderPackage X := .ofLE X
If not all of these instances are available via typeclass synthesis, it is necessary to explicitly provide some arguments:
public instance : LinearOrderPackage X := .ofLE X {
le_total := sorry
le_trans := sorry
le_antisymm := sorry }
It is also possible to do all of this by hand, without resorting to LinearOrderPackage
. This
can be useful if, say, one wants to avoid specifying an LT α
instance, which is not possible with
LinearOrderPackage
.
How the arguments are filled
Lean tries to fill all of the fields of the args : Packages.LinearOrderOfLEArgs α
parameter
automatically. If it fails, it is necessary to provide some of the fields manually.
- For the data-carrying typeclasses
LE
,LT
,BEq
,Ord
,Min
andMax
, existing instances are always preferred. If no existing instances can be synthesized, it is attempted to derive an instance from theLE
instance. - Some proof obligations can be filled automatically if the data-carrying typeclasses have been
derived from the
LE
instance. For example: If thebeq
field is omitted and noBEq α
instance can be synthesized, it is derived from theLE α
instance. In this case,beq_iff_le_and_ge
can be omitted because Lean can infer thatBEq α
andLE α
are compatible. - Other proof obligations, namely
le_total
,le_trans
andle_antisymm
, can be omitted ifTotal
,Trans
andAntisymm
instances can be synthesized.
Equations
- One or more equations did not get rendered due to their size.
Instances For
This structure contains all the data needed to create a LinearPreorderPackage α
instance.
Its fields are automatically provided if possible. For the detailed rules how the fields are
inferred, see LinearPreorderPackage.ofOrd
.
- ord : Ord α
- transOrd : TransOrd α
- lawfulOrderOrd : let this := self.ord; let this_1 := ⋯; let this_2 := self.le; LawfulOrderOrd α
- decidableLE : let this := self.ord; let this := self.le; have this_1 := ⋯; DecidableLE α
- decidableLT : let this := self.ord; let this := self.lt; let this_1 := self.le; have this_2 := ⋯; have this_3 := ⋯; DecidableLT α
Instances For
Use this factory to conveniently define a linear preorder on a type α
and all the associated
operations and instances given an Ord α
instance.
Creates a LinearPreorderPackage α
instance. Such an instance entails LE α
, LT α
, BEq α
and
Ord α
instances as well as an IsLinearPreorder α
instance and LawfulOrder*
instances proving
the compatibility of the operations with the linear preorder.
In the presence of Ord α
and TransOrd α
instances, no arguments are required and the factory can
be used as in this example:
public instance : LinearPreorderPackage X := .ofOrd X
If not all of these instances are available via typeclass synthesis, it is necessary to explicitly provide some arguments:
public instance : LinearPreorderPackage X := .ofOrd X {
ord := sorry
transOrd := sorry }
It is also possible to do all of this by hand, without resorting to LinearPreorderPackage
. This
can be useful if, say, one wants to avoid specifying an LT α
instance, which is not possible with
LinearPreorderPackage
.
How the arguments are filled
Lean tries to fill all of the fields of the args : Packages.LinearPreorderOfOrdArgs α
parameter
automatically. If it fails, it is necessary to provide some of the fields manually.
- For the data-carrying typeclasses
LE
,LT
,BEq
andOrd
, existing instances are always preferred. If no existing instances can be synthesized, it is attempted to derive an instance from theOrd
instance. - Some proof obligations can be filled automatically if the data-carrying typeclasses have been
derived from the
Ord
instance. For example: If thebeq
field is omitted and noBEq α
instance can be synthesized, it is derived from theOrd α
instance. In this case,beq_iff
can be omitted because Lean can infer thatBEq α
andOrd α
are compatible. - Other proof obligations, for example
transOrd
, can be omitted if a matching instance can be synthesized.
Equations
- One or more equations did not get rendered due to their size.
Instances For
This structure contains all the data needed to create a LinearOrderPackage α
instance.
Its fields are automatically provided if possible. For the detailed rules how the fields are
inferred, see LinearOrderPackage.ofOrd
.
- lawfulOrderOrd : let this := self.ord; let this_1 := ⋯; let this_2 := self.le; LawfulOrderOrd α
- decidableLE : let this := self.ord; let this := self.le; have this_1 := ⋯; DecidableLE α
- decidableLT : let this := self.ord; let this := self.lt; let this_1 := self.le; have this_2 := ⋯; have this_3 := ⋯; DecidableLT α
Instances For
Use this factory to conveniently define a linear order on a type α
and all the associated
operations and instances given an Ord α
instance.
Creates a LinearOrderPackage α
instance. Such an instance entails LE α
, LT α
, BEq α
,
Ord α
, Min α
and Max α
instances as well as an IsLinearOrder α
instance and LawfulOrder*
instances proving the compatibility of the operations with the linear order.
In the presence of Ord α
, TransOrd α
and LawfulEqOrd α
instances, no arguments are required
and the factory can be used as in this
example:
public instance : LinearOrderPackage X := .ofLE X
If not all of these instances are available via typeclass synthesis, it is necessary to explicitly provide some arguments:
public instance : LinearOrderPackage X := .ofLE X {
transOrd := sorry
eq_of_compare := sorry }
It is also possible to do all of this by hand, without resorting to LinearOrderPackage
. This
can be useful if, say, one wants to avoid specifying an LT α
instance, which is not possible with
LinearOrderPackage
.
How the arguments are filled
Lean tries to fill all of the fields of the args : Packages.LinearOrderOfLEArgs α
parameter
automatically. If it fails, it is necessary to provide some of the fields manually.
- For the data-carrying typeclasses
LE
,LT
,BEq
,Ord
,Min
andMax
, existing instances are always preferred. If no existing instances can be synthesized, it is attempted to derive an instance from theOrd
instance. - Some proof obligations can be filled automatically if the data-carrying typeclasses have been
derived from the
Ord
instance. For example: If thebeq
field is omitted and noBEq α
instance can be synthesized, it is derived from theLE α
instance. In this case,beq_iff
can be omitted because Lean can infer thatBEq α
andOrd α
are compatible. - Other proof obligations, such as
transOrd
, can be omitted if matching instances can be synthesized.
Equations
- One or more equations did not get rendered due to their size.