Library mathcomp.order.order

(* (c) Copyright 2006-2019 Microsoft Corporation and Inria.                  
 Distributed under the terms of CeCILL-B.                                  *)

From HB Require Import structures.
From mathcomp Require Import ssreflect ssrfun ssrbool eqtype ssrnat choice seq.
From mathcomp Require Import path fintype tuple bigop finset div prime finfun.
From mathcomp Require Import finset.
From mathcomp Require Export preorder.

Types equipped with order relations NB: See CONTRIBUTING.md for an introduction to HB concepts and commands. This files defines types equipped with order relations.

How to use orders in MathComp?

Use one of the following modules implementing different theories (all located in the module Order): Order.LTheory: partially ordered types and lattices excluding complement and totality related theorems Order.CTheory: complemented lattices including Order.LTheory Order.TTheory: totally ordered types including Order.LTheory Order.Theory: ordered types including all of the above theory modules To access the definitions, notations, and the theory from, say, "Order.Xyz", insert "Import Order.Xyz." at the top of your scripts. You can also "Import Order.Def." to enjoy shorter notations (e.g., min instead of Order.min, nondecreasing instead of Order.nondecreasing, etc.). In order to reason about abstract orders, notations are accessible by opening the scope "order_scope" bound to the delimiting key "O"; however, when dealing with another notation scope providing order notations for a concrete instance (e.g., "ring_scope"), it is not recommended to open "order_scope" at the same time.

Control of inference (parsing) and printing

One characteristic of ordered types is that one carrier type may have several orders. For example, natural numbers can be totally or partially ordered by the less than or equal relation, the divisibility relation, and their dual relations. Therefore, we need a way to control inference of ordered type instances and printing of generic relations and operations on ordered types. As a rule of thumb, we use the carrier type or its "alias" (named copy) to control inference (using canonical structures), and use a "display" to control the printing of notations. Each generic interface and operation for ordered types has, as its first argument, a "display" of type Order.disp_t. For example, the less than or equal relation has type: Order.le : forall {d : Order.disp_t} {T : porderType d}, rel T, where porderType d is the structure of partially ordered types with display d. (@Order.le dvd_display _ m n) is printed as m %| n because ordered type instances associated to the display dvd_display is intended to represent natural numbers partially ordered by the divisibility relation. We stress that order structure inference can be triggered only from the carrier type (or its alias), but not the display. For example, writing m %| n for m and n of type nat does not trigger an inference of the divisibility relation on natural numbers, which is associated to an alias natdvd for nat; such an inference should be triggered through the use of the corresponding alias, i.e., (m : natdvd) %| n. In other words, displays are merely used to inform the user and the notation mechanism of what the inference did; they are not additional input for the inference. See below for various aliases and their associated displays. NB: algebra/ssrnum.v provides the display ring_display to change the scope of the usual notations to ring_scope. Instantiating d with Disp tt tt or an unknown display will lead to a default display for notations. Alternative notation displays can be defined by : 1. declaring a new opaque definition of type unit. Using the idiom `Fact my_display : Order.disp_t. Proof. exact: Disp tt tt. Qed.` 2. using this symbol to tag canonical porderType structures using `HB.instance Definition _ := isPOrder.Build my_display my_type ...`, 3. declaring notations for the main operations of this library, by setting the first argument of the definition to the display, e.g. `Notation my_syndef_le x y := @Order.le my_display _ x y.` or `Notation "x <=< y" := @Order.lt my_display _ x y (at level ...).` Non overloaded notations will default to the default display. We suggest the user to refer to the example of natdvd below as a guideline example to add their own displays.

Interfaces

We provide the following interfaces for types equipped with an order: porderType d == the type of partially ordered types The HB class is called POrder. bPOrderType d == porderType with a bottom element (\bot) The HB class is called BPOrder. tPOrderType d == porderType with a top element (\top) The HB class is called TPOrder. tbPOrderType d == porderType with both a top and a bottom The HB class is called TBPOrder. meetSemilatticeType d == the type of meet semilattices The HB class is called MeetSemilattice. bMeetSemilatticeType d == meetSemilatticeType with a bottom element The HB class is called BMeetSemilattice. tMeetSemilatticeType d == meetSemilatticeType with a top element The HB class is called TMeetSemilattice. tbMeetSemilatticeType d == meetSemilatticeType with both a top and a bottom The HB class is called TBMeetSemilattice. joinSemilatticeType d == the type of join semilattices The HB class is called JoinSemilattice. bJoinSemilatticeType d == joinSemilatticeType with a bottom element The HB class is called BJoinSemilattice. tJoinSemilatticeType d == joinSemilatticeType with a top element The HB class is called TJoinSemilattice. tbJoinSemilatticeType d == joinSemilatticeType with both a top and a bottom The HB class is called TBJoinSemilattice. latticeType d == the type of lattices The HB class is called Lattice. bLatticeType d == latticeType with a bottom element The HB class is called BLattice. tLatticeType d == latticeType with a top element The HB class is called TLattice. tbLatticeType d == latticeType with both a top and a bottom The HB class is called TBLattice. distrLatticeType d == the type of distributive lattices The HB class is called DistrLattice. bDistrLatticeType d == distrLatticeType with a bottom element The HB class is called BDistrLattice. tDistrLatticeType d == distrLatticeType with a top element The HB class is called TDistrLattice. tbDistrLatticeType d == distrLatticeType with both a top and a bottom The HB class is called TBDistrLattice. orderType d == the type of totally ordered types The HB class is called Total. bOrderType d == orderType with a bottom element The HB class is called BTotal. tOrderType d == orderType with a top element The HB class is called TTotal. tbOrderType d == orderType with both a top and a bottom The HB class is called TBTotal. cDistrLatticeType d == the type of relatively complemented distributive lattices, where each interval [a, b] is equipped with a complement operation The HB class is called CDistrLattice. cbDistrLatticeType d == the type of sectionally complemented distributive lattices, equipped with a bottom, a relative complement operation, and a difference operation, i.e., a complement operation for each interval of the form [\bot, b] The HB class is called CBDistrLattice. ctDistrLatticeType d == the type of dually sectionally complemented distributive lattices, equipped with a top, a relative complement operation, and a dual difference operation, i.e. a complement operation for each interval of the form [a, \top] The HB class is called CTDistrLattice. ctbDistrLatticeType d == the type of complemented distributive lattices, equipped with top, bottom, difference, dual difference, and complement The HB class is called CTBDistrLattice. finPOrderType d == the type of partially ordered finite types The HB class is called FinPOrder. finBPOrderType d == finPOrderType with a bottom element The HB class is called FinBPOrder. finTPOrderType d == finPOrderType with a top element The HB class is called FinTPOrder. finTBPOrderType d == finPOrderType with both a top and a bottom The HB class is called FinTBPOrder. finMeetSemilatticeType d == the type of finite meet semilattice types The HB class is called FinMeetSemilattice. finBMeetSemilatticeType d == finMeetSemilatticeType with a bottom element Note that finTMeetSemilatticeType is just finTBLatticeType. The HB class is called FinBMeetSemilattice. finJoinSemilatticeType d == the type of finite join semilattice types The HB class is called FinJoinSemilattice. finTJoinSemilatticeType d == finJoinSemilatticeType with a top element Note that finBJoinSemilatticeType is just finTBLatticeType. The HB class is called FinTJoinSemilattice. finLatticeType d == the type of finite lattices The HB class is called FinLattice. finTBLatticeType d == the type of nonempty finite lattices The HB class is called FinTBLattice. finDistrLatticeType d == the type of finite distributive lattices The HB class is called FinDistrLattice. finTBDistrLatticeType d == the type of nonempty finite distributive lattices The HB class is called FinTBDistrLattice. finOrderType d == the type of totally ordered finite types The HB class is called FinTotal. finTBOrderType d == the type of nonempty totally ordered finite types The HB class is called FinTBTotal. finCDistrLatticeType d == the type of finite relatively complemented distributive lattices The HB class is called FinCDistrLattice. finCTBDistrLatticeType d == the type of finite complemented distributive lattices The HB class is called FinCTBDistrLattice. and their joins with subType: subPOrder d T P d' == join of porderType d' and subType (P : pred T) such that val is monotonic The HB class is called SubPOrder. meetSubLattice d T P d' == join of latticeType d' and subType (P : pred T) such that val is monotonic and a morphism for meet The HB class is called MeetSubLattice. joinSubLattice d T P d' == join of latticeType d' and subType (P : pred T) such that val is monotonic and a morphism for join The HB class is called JoinSubLattice. subLattice d T P d' == join of JoinSubLattice and MeetSubLattice The HB class is called SubLattice. bJoinSubLattice d T P d' == join of JoinSubLattice and BLattice such that val is a morphism for \bot The HB class is called BJoinSubLattice. tMeetSubLattice d T P d' == join of MeetSubLattice and TLattice such that val is a morphism for \top The HB class is called TMeetSubLattice. bSubLattice d T P d' == join of SubLattice and BLattice such that val is a morphism for \bot The HB class is called BSubLattice. tSubLattice d T P d' == join of SubLattice and TLattice such that val is a morphism for \top The HB class is called BSubLattice. subOrder d T P d' == join of orderType d' and subLatticeType d T P d' The HB class is called SubOrder. subPOrderLattice d T P d' == join of SubPOrder and Lattice The HB class is called SubPOrderLattice. subPOrderBLattice d T P d' == join of SubPOrder and BLattice The HB class is called SubPOrderBLattice. subPOrderTLattice d T P d' == join of SubPOrder and TLattice The HB class is called SubPOrderTLattice. subPOrderTBLattice d T P d' == join of SubPOrder and TBLattice The HB class is called SubPOrderTBLattice. meetSubBLattice d T P d' == join of MeetSubLattice and BLattice The HB class is called MeetSubBLattice. meetSubTLattice d T P d' == join of MeetSubLattice and TLattice The HB class is called MeetSubTLattice. meetSubTBLattice d T P d' == join of MeetSubLattice and TBLattice The HB class is called MeetSubTBLattice. joinSubBLattice d T P d' == join of JoinSubLattice and BLattice The HB class is called JoinSubBLattice. joinSubTLattice d T P d' == join of JoinSubLattice and TLattice The HB class is called JoinSubTLattice. joinSubTBLattice d T P d' == join of JoinSubLattice and TBLattice The HB class is called JoinSubTBLattice. subBLattice d T P d' == join of SubLattice and BLattice The HB class is called SubBLattice. subTLattice d T P d' == join of SubLattice and TLattice The HB class is called SubTLattice. subTBLattice d T P d' == join of SubLattice and TBLattice The HB class is called SubTBLattice. bJoinSubTLattice d T P d' == join of BJoinSubLattice and TBLattice The HB class is called BJoinSubTLattice. tMeetSubBLattice d T P d' == join of TMeetSubLattice and TBLattice The HB class is called TMeetSubBLattice. bSubTLattice d T P d' == join of BSubLattice and TBLattice The HB class is called BSubTLattice. tSubBLattice d T P d' == join of TSubLattice and TBLattice The HB class is called TSubBLattice. tbSubBLattice d T P d' == join of BSubLattice and TSubLattice The HB class is called TBSubLattice. Morphisms between the above structures: MeetLatticeMorphism.type d T d' T', JoinLatticeMorphism.type d T d' T', LatticeMorphism.type d T d' T' == nondecreasing function between two lattices which are morphism for meet, join, and meet/join respectively BLatticeMorphism.type d T d' T' := {blmorphism T -> T'}, TLatticeMorphism.type d T d' T' := {tlmorphism T -> T'}, TBLatticeMorphism.type d T d' T' := {tblmorphism T -> T'} == nondecreasing function between two lattices with bottom/top which are morphism for bottom/top Closedness predicates for the algebraic structures: meetLatticeClosed d T == predicate closed under meet on T : latticeType d The HB class is MeetLatticeClosed. joinLatticeClosed d T == predicate closed under join on T : latticeType d The HB class is JoinLatticeClosed. latticeClosed d T == predicate closed under meet and join The HB class is JoinLatticeClosed. bLatticeClosed d T == predicate that contains bottom The HB class is BLatticeClosed. tLatticeClosed d T == predicate that contains top The HB class is TLatticeClosed. tbLatticeClosed d T == predicate that contains top and bottom the HB class ie TBLatticeClosed. bJoinLatticeClosed d T == predicate that contains bottom and is closed under join The HB class is BJoinLatticeClosed. tMeetLatticeClosed d T == predicate that contains top and is closed under meet The HB class is TMeetLatticeClosed.

Useful lemmas:

On orderType, leP, ltP, and ltgtP are the three main lemmas for case analysis. On porderType, one may use comparableP, comparable_leP, comparable_ltP, and comparable_ltgtP, which are the four main lemmas for case analysis.

Order relations and operations:

In general, an overloaded relation or operation on ordered types takes the following arguments: 1. a display d of type Order.disp_t, 2. an instance T of the minimal structure it operates on, and 3. operands. Here is the exhaustive list of all such operations together with their default notation (defined in order_scope unless specified otherwise). For T of type porderType d, x and y of type T, and C of type bool: x <= y := @Order.le d T x y <-> x is less than or equal to y. x < y := @Order.lt d T x y <-> x is less than y, i.e., (y != x) && (x <= y). x >= y := y <= x <-> x is greater than or equal to y. x > y := y < x <-> x is greater than y. x >=< y := @Order.comparable d T x y (:= (x <= y) || (y <= x)) <-> x and y are comparable. x >< y := ~~ x >=< y <-> x and y are incomparable. x <= y ?= iff C := @Order.leif d T x y C (:= (x <= y) * ((x == y) = C)) <-> x is less than y, or equal iff C is true. x < y ?<= if C := @Order.lteif d T x y C (:= if C then x <= y else x < y) <-> x is smaller than y, and strictly if C is false. Order.min x y := if x < y then x else y Order.max x y := if x < y then y else x f \min g == the function x |-> Order.min (f x) (g x); f \min g simplifies on application. f \max g == the function x |-> Order.max (f x) (g x); f \max g simplifies on application. nondecreasing f <-> the function f : T -> T' is nondecreasing, where T and T' are porderType := {homo f : x y / x <= y} Unary (partially applied) versions of order notations: >= y := @Order.le d T y == a predicate characterizing elements greater than or equal to y > y := @Order.lt d T y <= y := @Order.ge d T y < y := @Order.gt d T y >=< y := [pred x | @Order.comparable d T x y] >< y := [pred x | ~~ @Order.comparable d T x y] 0-ary versions of order notations (in function_scope): <=%O := @Order.le d T <%O := @Order.lt d T >=%O := @Order.ge d T >%O := @Order.gt d T >=<%O := @Order.comparable d T <?=%O := @Order.leif d T <?<=%O := @Order.lteif d T
  • > These conventions are compatible with Haskell's, where ((< y) x) = (x < y) = ((<) x y), except that we write <%O instead of (<).
For T of type bPOrderType d: \bot := @Order.bottom d T == the bottom element of type T \max_<range> e := \big[Order.max / Order.bottom]_<range> e == iterated max of a preorder with a bottom the possible <range>s are documented in bigop.v For T of type tPOrderType d: \top := @Order.top d T == the top element of type T \min_<range> e := \big[Order.max / Order.top]_<range> e == iterated min of a preorder with a top the possible <range>s are documented in bigop.v For T of type meetSemilatticeType d, and x, y of type T: x `&` y := @Order.meet d T x y == the meet of x and y For T of type joinSemilatticeType d, and x, y of type T: x `|` y := @Order.join d T x y == the join of x and y For T of type tMeetSemilatticeType d: \meet_<range> e := \big[Order.meet / Order.top]_<range> e == iterated meet of a meet-semilattice with a top the possible <range>s are documented in bigop.v For T of type bJoinSemilatticeType d: \join_<range> e := \big[Order.join / Order.bottom]_<range> e == iterated join of a join-semilattice with a bottom the possible <range>s are documented in bigop.v For T of type cDistrLatticeType d, and x, y, z of type T: rcompl x y z == the (relative) complement of z in [x, y] For T of type cbDistrLatticeType d, and x, y of type T: x `\` y := @Order.diff d T x y == the (sectional) complement of y in [\bot, x], i.e., rcompl \bot x y For T of type ctDistrLatticeType d, and x, y of type T: codiff x y == the (dual sectional) complement of y in [x, \top], i.e., rcompl x \top y For T of type ctbDistrLatticeType d, and x of type T: ~` x := @Order.compl d T x == the complement of x in [\bot, \top], i.e., rcompl \bot \top x For preorderType we provide the following operations: [arg min_(i < i0 | P) M] == a value i : T minimizing M : R, subject to the condition P (i may appear in P and M), and provided P holds for i0. [arg max_(i > i0 | P) M] == a value i maximizing M subject to P and provided P holds for i0. [arg min_(i < i0 in A) M] == an i \in A minimizing M if i0 \in A. [arg max_(i > i0 in A) M] == an i \in A maximizing M if i0 \in A. [arg min_(i < i0) M] == an i : T minimizing M, given i0 : T. [arg max_(i > i0) M] == an i : T maximizing M, given i0 : T. with head symbols Order.arg_min and Order.arg_max The user may use extremumP or extremum_inP to eliminate them.
  • > patterns for contextual rewriting: leLHS := (X in (X <= _)%O)%pattern leRHS := (X in (_ <= X)%O)%pattern ltLHS := (X in (X < _)%O)%pattern ltRHS := (X in (_ < X)%O)%pattern
We provide aliases for various types and their displays: natdvd := nat (associated with display dvd_display) == an alias for nat which is canonically ordered using divisibility predicate dvdn Notation %|, %<|, gcd, lcm are used instead of <=, <, meet and join. T^d := dual T, where dual is a new definition for (fun T => T) (associated with dual_display d where d is a display) == an alias for T, such that if T is canonically ordered, then T^d is canonically ordered with the dual order, and displayed with an extra ^d in the notation, i.e., <=^d, <^d, >=<^d, ><^d, `&`^d, `|`^d are used and displayed instead of <=, <, >=<, ><, `&`, `|` T *prod[d] T' := T * T' == an alias for the cartesian product such that, if T and T' are canonically ordered, then T *prod[d] T' is canonically ordered in product order, i.e., (x1, x2) <= (y1, y2) = (x1 <= y1) && (x2 <= y2), and displayed in display d T *p T' := T *prod[prod_display d d' ] T' where d and d' are the displays of T and T', respectively, and prod_display adds an extra ^p to all notations T *lexi[d] T' := T * T' == an alias for the cartesian product such that, if T and T' are canonically ordered, then T *lexi[d] T' is canonically ordered in lexicographic order, i.e., (x1, x2) <= (y1, y2) = (x1 <= y1) && ((x1 >= y1) ==> (x2 <= y2)) and (x1, x2) < (y1, y2) = (x1 <= y1) && ((x1 >= y1) ==> (x2 < y2)) and displayed in display d T *l T' := T *lexi[lexi_display d d' ] T' where d and d' are the displays of T and T', respectively, and lexi_display adds an extra ^l to all notations seqprod_with d T := seq T == an alias for seq, such that if T is canonically ordered, then seqprod_with d T is canonically ordered in product order, i.e., [:: x1, .., xn] <= [y1, .., yn] = (x1 <= y1) && ... && (xn <= yn) and displayed in display d n.-tupleprod[d] T == same with n.tuple T seqprod T := seqprod_with (seqprod_display d) T where d is the display of T, and seqprod_display adds an extra ^sp to all notations n.-tupleprod T := n.-tuple[seqprod_display d] T where d is the display of T seqlexi_with d T := seq T == an alias for seq, such that if T is canonically ordered, then seqprod_with d T is canonically ordered in lexicographic order, i.e., [:: x1, .., xn] <= [y1, .., yn] = (x1 <= x2) && ((x1 >= y1) ==> ((x2 <= y2) && ...)) and displayed in display d n.-tuplelexi[d] T == same with n.tuple T seqlexi T := lexiprod_with (seqlexi_display d) T where d is the display of T, and seqlexi_display adds an extra ^sl to all notations n.-tuplelexi T := n.-tuple[seqlexi_display d] T where d is the display of T {subset[d] T} := {set T} == an alias for set which is canonically ordered by the subset order and displayed in display d {subset T} := {subset[subset_display] T} The following notations are provided to build substructures: [SubChoice_isSubPOrder of U by <: with disp] == [SubChoice_isSubPOrder of U by <: ] == porderType mixin for a subType whose base type is a porderType [SubPOrder_isSubLattice of U by <: with disp] == [SubPOrder_isSubLattice of U by <: ] == [SubChoice_isSubLattice of U by <: with disp] == [SubChoice_isSubLattice of U by <: ] == latticeType mixin for a subType whose base type is a latticeType and whose predicate is a latticeClosed [SubPOrder_isBSubLattice of U by <: with disp] == [SubPOrder_isBSubLattice of U by <: ] == [SubChoice_isBSubLattice of U by <: with disp] == [SubChoice_isBSubLattice of U by <: ] == blatticeType mixin for a subType whose base type is a blatticeType and whose predicate is both a latticeClosed and a bLatticeClosed [SubPOrder_isTSubLattice of U by <: with disp] == [SubPOrder_isTSubLattice of U by <: ] == [SubChoice_isTSubLattice of U by <: with disp] == [SubChoice_isTSubLattice of U by <: ] == tlatticeType mixin for a subType whose base type is a tlatticeType and whose predicate is both a latticeClosed and a tLatticeClosed [SubPOrder_isTBSubLattice of U by <: with disp] == [SubPOrder_isTBSubLattice of U by <: ] == [SubChoice_isTBSubLattice of U by <: with disp] == [SubChoice_isTBSubLattice of U by <: ] == tblatticeType mixin for a subType whose base type is a tblatticeType and whose predicate is both a latticeClosed and a tbLatticeClosed [SubLattice_isSubOrder of U by <: with disp] == [SubLattice_isSubOrder of U by <: ] == [SubPOrder_isSubOrder of U by <: with disp] == [SubPOrder_isSubOrder of U by <: ] == [SubChoice_isSubOrder of U by <: with disp] == [SubChoice_isSubOrder of U by <: ] == orderType mixin for a subType whose base type is an orderType [POrder of U by <: ] == porderType mixin for a subType whose base type is a porderType [Order of U by <: ] == orderType mixin for a subType whose base type is an orderType We provide expected instances of ordered types for bool, nat (for leq and and dvdn), 'I_n, 'I_n.+1 (with a top and bottom), nat for dvdn, T *prod[disp] T', T *lexi[disp] T', {t : T & T' x} (with lexicographic ordering), seqprod_with d T (using product order), seqlexi_with d T (with lexicographic ordering), n.-tupleprod[disp] (using product order), n.-tuplelexi[d] T (with lexicographic ordering), on {subset[disp] T} (using subset order) and all possible finite type instances. (Use `HB.about type` to discover the instances on type.) In order to get a canonical order on prod, seq, tuple or set, one may import modules DefaultProdOrder or DefaultProdLexiOrder, DefaultSeqProdOrder or DefaultSeqLexiOrder, DefaultTupleProdOrder or DefaultTupleLexiOrder, and DefaultSetSubsetOrder. We also provide specialized versions of some theorems from path.v. We provide Order.enum_val, Order.enum_rank, and Order.enum_rank_in, which are monotonic variations of enum_val, enum_rank, and enum_rank_in whenever the type is porderType, and their monotonicity is provided if this order is total. The theory is in the module Order (Order.enum_valK, Order.enum_rank_inK, etc) but Order.Enum can be imported to shorten these. We provide an opaque monotonous bijection tagnat.sig / tagnat.rank between the finite types {i : 'I_n & 'I_(p_ i)} and 'I_(\sum_i p_ i): tagnat.sig : 'I_(\sum_i p_ i) -> {i : 'I_n & 'I_(p_ i)} tagnat.rank : {i : 'I_n & 'I_(p_ i)} -> 'I_(\sum_i p_ i) tagnat.sig1 : 'I_(\sum_i p_ i) -> 'I_n tagnat.sig2 : forall p : 'I_(\sum_i p_ i), 'I_(p_ (tagnat.sig1 p)) tagnat.Rank : forall i, 'I_(p_ i) -> 'I_(\sum_i p_ i) Acknowledgments: This file is based on prior work by D. Dreyer, G. Gonthier, A. Nanevski, P-Y Strub, B. Ziliani

Set Implicit Arguments.

Declare Scope order_scope.

Delimit Scope order_scope with O.
Local Open Scope order_scope.

Reserved notations for lattice operations
Reserved Notation "A `&` B" (at level 48, left associativity).
Reserved Notation "A `|` B" (at level 52, left associativity).
Reserved Notation "A `\` B" (at level 50, left associativity).
Reserved Notation "~` A" (at level 35, right associativity).

Reserved notations for dual order
Reserved Notation "A `&^d` B" (at level 48, left associativity).
Reserved Notation "A `|^d` B" (at level 52, left associativity).
Reserved Notation "A `\^d` B" (at level 50, left associativity).
Reserved Notation "~^d` A" (at level 35, right associativity).

Reserved notations for product ordering of prod
Reserved Notation "A `&^p` B" (at level 48, left associativity).
Reserved Notation "A `|^p` B" (at level 52, left associativity).
Reserved Notation "A `\^p` B" (at level 50, left associativity).
Reserved Notation "~^p` A" (at level 35, right associativity).

Reserved notations for product ordering of seq
Reserved Notation "A `&^sp` B" (at level 48, left associativity).
Reserved Notation "A `|^sp` B" (at level 52, left associativity).
Reserved Notation "A `\^sp` B" (at level 50, left associativity).
Reserved Notation "~^sp` A" (at level 35, right associativity).

Reserved notations for lexicographic ordering of prod
Reserved Notation "A `&^l` B" (at level 48, left associativity).
Reserved Notation "A `|^l` B" (at level 52, left associativity).
Reserved Notation "A `\^l` B" (at level 50, left associativity).
Reserved Notation "~^l` A" (at level 35, right associativity).

Reserved notations for lexicographic ordering of seq
Reserved Notation "A `&^sl` B" (at level 48, left associativity).
Reserved Notation "A `|^sl` B" (at level 52, left associativity).
Reserved Notation "A `\^sl` B" (at level 50, left associativity).
Reserved Notation "~^sl` A" (at level 35, right associativity).

Reserved notations for divisibility

Reserved Notation "\gcd_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \gcd_ i '/ ' F ']'").
Reserved Notation "\gcd_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \gcd_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\gcd_ ( i <- r ) F"
  (F at level 41,
           format "'[' \gcd_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\gcd_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \gcd_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\gcd_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \gcd_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\gcd_ ( i | P ) F"
  (F at level 41,
           format "'[' \gcd_ ( i | P ) '/ ' F ']'").
Reserved Notation "\gcd_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\gcd_ ( i : t ) F" (F at level 41).
Reserved Notation "\gcd_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \gcd_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\gcd_ ( i < n ) F"
  (F at level 41,
           format "'[' \gcd_ ( i < n ) F ']'").
Reserved Notation "\gcd_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \gcd_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\gcd_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \gcd_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\lcm_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \lcm_ i '/ ' F ']'").
Reserved Notation "\lcm_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \lcm_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\lcm_ ( i <- r ) F"
  (F at level 41,
           format "'[' \lcm_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\lcm_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \lcm_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\lcm_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \lcm_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\lcm_ ( i | P ) F"
  (F at level 41,
           format "'[' \lcm_ ( i | P ) '/ ' F ']'").
Reserved Notation "\lcm_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\lcm_ ( i : t ) F" (F at level 41).
Reserved Notation "\lcm_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \lcm_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\lcm_ ( i < n ) F"
  (F at level 41,
           format "'[' \lcm_ ( i < n ) F ']'").
Reserved Notation "\lcm_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \lcm_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\lcm_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \lcm_ ( i 'in' A ) '/ ' F ']'").

Reserved notations for iterative meet and join
Reserved Notation "\meet_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \meet_ i '/ ' F ']'").
Reserved Notation "\meet_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \meet_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\meet_ ( i <- r ) F"
  (F at level 41,
           format "'[' \meet_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\meet_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \meet_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\meet_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \meet_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\meet_ ( i | P ) F"
  (F at level 41,
           format "'[' \meet_ ( i | P ) '/ ' F ']'").
Reserved Notation "\meet_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\meet_ ( i : t ) F" (F at level 41).
Reserved Notation "\meet_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \meet_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\meet_ ( i < n ) F"
  (F at level 41,
           format "'[' \meet_ ( i < n ) F ']'").
Reserved Notation "\meet_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \meet_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\meet_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \meet_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\join_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \join_ i '/ ' F ']'").
Reserved Notation "\join_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \join_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\join_ ( i <- r ) F"
  (F at level 41,
           format "'[' \join_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\join_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \join_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\join_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \join_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\join_ ( i | P ) F"
  (F at level 41,
           format "'[' \join_ ( i | P ) '/ ' F ']'").
Reserved Notation "\join_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\join_ ( i : t ) F" (F at level 41).
Reserved Notation "\join_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \join_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\join_ ( i < n ) F"
  (F at level 41,
           format "'[' \join_ ( i < n ) F ']'").
Reserved Notation "\join_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \join_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\join_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \join_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\meet^d_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \meet^d_ i '/ ' F ']'").
Reserved Notation "\meet^d_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \meet^d_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\meet^d_ ( i <- r ) F"
  (F at level 41,
           format "'[' \meet^d_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\meet^d_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \meet^d_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\meet^d_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \meet^d_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\meet^d_ ( i | P ) F"
  (F at level 41,
           format "'[' \meet^d_ ( i | P ) '/ ' F ']'").
Reserved Notation "\meet^d_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\meet^d_ ( i : t ) F" (F at level 41).
Reserved Notation "\meet^d_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \meet^d_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\meet^d_ ( i < n ) F"
  (F at level 41,
           format "'[' \meet^d_ ( i < n ) F ']'").
Reserved Notation "\meet^d_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \meet^d_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\meet^d_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \meet^d_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\join^d_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \join^d_ i '/ ' F ']'").
Reserved Notation "\join^d_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \join^d_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\join^d_ ( i <- r ) F"
  (F at level 41,
           format "'[' \join^d_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\join^d_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \join^d_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\join^d_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \join^d_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\join^d_ ( i | P ) F"
  (F at level 41,
           format "'[' \join^d_ ( i | P ) '/ ' F ']'").
Reserved Notation "\join^d_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\join^d_ ( i : t ) F" (F at level 41).
Reserved Notation "\join^d_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \join^d_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\join^d_ ( i < n ) F"
  (F at level 41,
           format "'[' \join^d_ ( i < n ) F ']'").
Reserved Notation "\join^d_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \join^d_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\join^d_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \join^d_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\meet^p_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \meet^p_ i '/ ' F ']'").
Reserved Notation "\meet^p_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \meet^p_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\meet^p_ ( i <- r ) F"
  (F at level 41,
           format "'[' \meet^p_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\meet^p_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \meet^p_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\meet^p_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \meet^p_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\meet^p_ ( i | P ) F"
  (F at level 41,
           format "'[' \meet^p_ ( i | P ) '/ ' F ']'").
Reserved Notation "\meet^p_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\meet^p_ ( i : t ) F" (F at level 41).
Reserved Notation "\meet^p_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \meet^p_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\meet^p_ ( i < n ) F"
  (F at level 41,
           format "'[' \meet^p_ ( i < n ) F ']'").
Reserved Notation "\meet^p_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \meet^p_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\meet^p_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \meet^p_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\join^p_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \join^p_ i '/ ' F ']'").
Reserved Notation "\join^p_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \join^p_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\join^p_ ( i <- r ) F"
  (F at level 41,
           format "'[' \join^p_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\join^p_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \join^p_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\join^p_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \join^p_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\join^p_ ( i | P ) F"
  (F at level 41,
           format "'[' \join^p_ ( i | P ) '/ ' F ']'").
Reserved Notation "\join^p_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\join^p_ ( i : t ) F" (F at level 41).
Reserved Notation "\join^p_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \join^p_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\join^p_ ( i < n ) F"
  (F at level 41,
           format "'[' \join^p_ ( i < n ) F ']'").
Reserved Notation "\join^p_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \join^p_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\join^p_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \join^p_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\min^p_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \min^p_ i '/ ' F ']'").
Reserved Notation "\min^p_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \min^p_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\min^p_ ( i <- r ) F"
  (F at level 41,
           format "'[' \min^p_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\min^p_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \min^p_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\min^p_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \min^p_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\min^p_ ( i | P ) F"
  (F at level 41,
           format "'[' \min^p_ ( i | P ) '/ ' F ']'").
Reserved Notation "\min^p_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\min^p_ ( i : t ) F" (F at level 41).
Reserved Notation "\min^p_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \min^p_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\min^p_ ( i < n ) F"
  (F at level 41,
           format "'[' \min^p_ ( i < n ) F ']'").
Reserved Notation "\min^p_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \min^p_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\min^p_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \min^p_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\max^p_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \max^p_ i '/ ' F ']'").
Reserved Notation "\max^p_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \max^p_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\max^p_ ( i <- r ) F"
  (F at level 41,
           format "'[' \max^p_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\max^p_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \max^p_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\max^p_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \max^p_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\max^p_ ( i | P ) F"
  (F at level 41,
           format "'[' \max^p_ ( i | P ) '/ ' F ']'").
Reserved Notation "\max^p_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\max^p_ ( i : t ) F" (F at level 41).
Reserved Notation "\max^p_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \max^p_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\max^p_ ( i < n ) F"
  (F at level 41,
           format "'[' \max^p_ ( i < n ) F ']'").
Reserved Notation "\max^p_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \max^p_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\max^p_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \max^p_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\meet^sp_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \meet^sp_ i '/ ' F ']'").
Reserved Notation "\meet^sp_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \meet^sp_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\meet^sp_ ( i <- r ) F"
  (F at level 41,
           format "'[' \meet^sp_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\meet^sp_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \meet^sp_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\meet^sp_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \meet^sp_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\meet^sp_ ( i | P ) F"
  (F at level 41,
           format "'[' \meet^sp_ ( i | P ) '/ ' F ']'").
Reserved Notation "\meet^sp_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\meet^sp_ ( i : t ) F" (F at level 41).
Reserved Notation "\meet^sp_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \meet^sp_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\meet^sp_ ( i < n ) F"
  (F at level 41,
           format "'[' \meet^sp_ ( i < n ) F ']'").
Reserved Notation "\meet^sp_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \meet^sp_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\meet^sp_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \meet^sp_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\join^sp_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \join^sp_ i '/ ' F ']'").
Reserved Notation "\join^sp_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \join^sp_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\join^sp_ ( i <- r ) F"
  (F at level 41,
           format "'[' \join^sp_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\join^sp_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \join^sp_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\join^sp_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \join^sp_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\join^sp_ ( i | P ) F"
  (F at level 41,
           format "'[' \join^sp_ ( i | P ) '/ ' F ']'").
Reserved Notation "\join^sp_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\join^sp_ ( i : t ) F" (F at level 41).
Reserved Notation "\join^sp_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \join^sp_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\join^sp_ ( i < n ) F"
  (F at level 41,
           format "'[' \join^sp_ ( i < n ) F ']'").
Reserved Notation "\join^sp_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \join^sp_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\join^sp_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \join^sp_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\min^sp_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \min^sp_ i '/ ' F ']'").
Reserved Notation "\min^sp_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \min^sp_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\min^sp_ ( i <- r ) F"
  (F at level 41,
           format "'[' \min^sp_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\min^sp_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \min^sp_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\min^sp_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \min^sp_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\min^sp_ ( i | P ) F"
  (F at level 41,
           format "'[' \min^sp_ ( i | P ) '/ ' F ']'").
Reserved Notation "\min^sp_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\min^sp_ ( i : t ) F" (F at level 41).
Reserved Notation "\min^sp_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \min^sp_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\min^sp_ ( i < n ) F"
  (F at level 41,
           format "'[' \min^sp_ ( i < n ) F ']'").
Reserved Notation "\min^sp_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \min^sp_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\min^sp_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \min^sp_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\max^sp_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \max^sp_ i '/ ' F ']'").
Reserved Notation "\max^sp_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \max^sp_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\max^sp_ ( i <- r ) F"
  (F at level 41,
           format "'[' \max^sp_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\max^sp_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \max^sp_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\max^sp_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \max^sp_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\max^sp_ ( i | P ) F"
  (F at level 41,
           format "'[' \max^sp_ ( i | P ) '/ ' F ']'").
Reserved Notation "\max^sp_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\max^sp_ ( i : t ) F" (F at level 41).
Reserved Notation "\max^sp_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \max^sp_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\max^sp_ ( i < n ) F"
  (F at level 41,
           format "'[' \max^sp_ ( i < n ) F ']'").
Reserved Notation "\max^sp_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \max^sp_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\max^sp_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \max^sp_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\meet^l_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \meet^l_ i '/ ' F ']'").
Reserved Notation "\meet^l_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \meet^l_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\meet^l_ ( i <- r ) F"
  (F at level 41,
           format "'[' \meet^l_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\meet^l_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \meet^l_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\meet^l_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \meet^l_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\meet^l_ ( i | P ) F"
  (F at level 41,
           format "'[' \meet^l_ ( i | P ) '/ ' F ']'").
Reserved Notation "\meet^l_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\meet^l_ ( i : t ) F" (F at level 41).
Reserved Notation "\meet^l_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \meet^l_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\meet^l_ ( i < n ) F"
  (F at level 41,
           format "'[' \meet^l_ ( i < n ) F ']'").
Reserved Notation "\meet^l_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \meet^l_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\meet^l_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \meet^l_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\join^l_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \join^l_ i '/ ' F ']'").
Reserved Notation "\join^l_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \join^l_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\join^l_ ( i <- r ) F"
  (F at level 41,
           format "'[' \join^l_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\join^l_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \join^l_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\join^l_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \join^l_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\join^l_ ( i | P ) F"
  (F at level 41,
           format "'[' \join^l_ ( i | P ) '/ ' F ']'").
Reserved Notation "\join^l_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\join^l_ ( i : t ) F" (F at level 41).
Reserved Notation "\join^l_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \join^l_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\join^l_ ( i < n ) F"
  (F at level 41,
           format "'[' \join^l_ ( i < n ) F ']'").
Reserved Notation "\join^l_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \join^l_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\join^l_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \join^l_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\min^l_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \min^l_ i '/ ' F ']'").
Reserved Notation "\min^l_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \min^l_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\min^l_ ( i <- r ) F"
  (F at level 41,
           format "'[' \min^l_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\min^l_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \min^l_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\min^l_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \min^l_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\min^l_ ( i | P ) F"
  (F at level 41,
           format "'[' \min^l_ ( i | P ) '/ ' F ']'").
Reserved Notation "\min^l_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\min^l_ ( i : t ) F" (F at level 41).
Reserved Notation "\min^l_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \min^l_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\min^l_ ( i < n ) F"
  (F at level 41,
           format "'[' \min^l_ ( i < n ) F ']'").
Reserved Notation "\min^l_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \min^l_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\min^l_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \min^l_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "\max^l_ i F"
  (at level 34, F at level 41, i at level 0,
           format "'[' \max^l_ i '/ ' F ']'").
Reserved Notation "\max^l_ ( i <- r | P ) F"
  (F at level 41, i, r at level 60,
           format "'[' \max^l_ ( i <- r | P ) '/ ' F ']'").
Reserved Notation "\max^l_ ( i <- r ) F"
  (F at level 41,
           format "'[' \max^l_ ( i <- r ) '/ ' F ']'").
Reserved Notation "\max^l_ ( m <= i < n | P ) F"
  (F at level 41, i, n at level 60,
           format "'[' \max^l_ ( m <= i < n | P ) '/ ' F ']'").
Reserved Notation "\max^l_ ( m <= i < n ) F"
  (F at level 41,
           format "'[' \max^l_ ( m <= i < n ) '/ ' F ']'").
Reserved Notation "\max^l_ ( i | P ) F"
  (F at level 41,
           format "'[' \max^l_ ( i | P ) '/ ' F ']'").
Reserved Notation "\max^l_ ( i : t | P ) F" (F at level 41).
Reserved Notation "\max^l_ ( i : t ) F" (F at level 41).
Reserved Notation "\max^l_ ( i < n | P ) F"
  (F at level 41, n at level 60,
           format "'[' \max^l_ ( i < n | P ) '/ ' F ']'").
Reserved Notation "\max^l_ ( i < n ) F"
  (F at level 41,
           format "'[' \max^l_ ( i < n ) F ']'").
Reserved Notation "\max^l_ ( i 'in' A | P ) F"
  (F at level 41, A at level 60,
           format "'[' \max^l_ ( i 'in' A | P ) '/ ' F ']'").
Reserved Notation "\max^l_ ( i 'in' A ) F"
  (F at level 41,
           format "'[' \max^l_ ( i 'in' A ) '/ ' F ']'").

Reserved Notation "'{' 'omorphism' U '->' V '}'"
  (U at level 98, V at level 99, format "{ 'omorphism' U -> V }").
Reserved Notation "'{' 'mlmorphism' U '->' V '}'"
  (U at level 98, V at level 99, format "{ 'mlmorphism' U -> V }").
Reserved Notation "'{' 'jlmorphism' U '->' V '}'"
  (U at level 98, V at level 99, format "{ 'jlmorphism' U -> V }").
Reserved Notation "'{' 'lmorphism' U '->' V '}'"
  (U at level 98, V at level 99, format "{ 'lmorphism' U -> V }").
Reserved Notation "'{' 'blmorphism' U '->' V '}'"
  (U at level 98, V at level 99, format "{ 'blmorphism' U -> V }").
Reserved Notation "'{' 'tlmorphism' U '->' V '}'"
  (U at level 98, V at level 99, format "{ 'tlmorphism' U -> V }").
Reserved Notation "'{' 'tblmorphism' U '->' V '}'"
  (U at level 98, V at level 99, format "{ 'tblmorphism' U -> V }").

Module Order.

Export Order.

#[key="T", primitive]
HB.mixin Record Preorder_isDuallyPOrder (d : disp_t) T of Preorder d T := {
  le_anti : antisymmetric (@le d T);
  ge_anti : antisymmetric (fun x y @le d T y x);
}.

#[short(type="porderType")]
HB.structure Definition POrder (d : disp_t) :=
  { T of Preorder d T & Preorder_isDuallyPOrder d T }.

#[short(type="bPOrderType")]
HB.structure Definition BPOrder d := { T of hasBottom d T & POrder d T }.
#[short(type="tPOrderType")]
HB.structure Definition TPOrder d := { T of hasTop d T & POrder d T }.
#[short(type="tbPOrderType")]
HB.structure Definition TBPOrder d := { T of hasTop d T & BPOrder d T }.

Module POrderExports.
Arguments le_trans {d s} [_ _ _].
End POrderExports.

Bind Scope order_scope with POrder.sort.

#[key="T", primitive]
HB.mixin Record POrder_isMeetSemilattice d (T : Type) of POrder d T := {
  meet : T T T;
  lexI : x y z, (x meet y z) = (x y) && (x z);
}.

#[key="T", primitive]
HB.mixin Record POrder_isJoinSemilattice d T of POrder d T := {
  join : T T T;
  leUx : x y z, (join x y z) = (x z) && (y z);
}.

#[short(type="meetSemilatticeType")]
HB.structure Definition MeetSemilattice d :=
  { T of POrder d T & POrder_isMeetSemilattice d T }.

#[short(type="bMeetSemilatticeType")]
HB.structure Definition BMeetSemilattice d :=
  { T of MeetSemilattice d T & hasBottom d T }.

#[short(type="tMeetSemilatticeType")]
HB.structure Definition TMeetSemilattice d :=
  { T of MeetSemilattice d T & hasTop d T }.

#[short(type="tbMeetSemilatticeType")]
HB.structure Definition TBMeetSemilattice d :=
  { T of BMeetSemilattice d T & hasTop d T }.

#[short(type="joinSemilatticeType")]
HB.structure Definition JoinSemilattice d :=
  { T of POrder d T & POrder_isJoinSemilattice d T }.

#[short(type="bJoinSemilatticeType")]
HB.structure Definition BJoinSemilattice d :=
  { T of JoinSemilattice d T & hasBottom d T }.

#[short(type="tJoinSemilatticeType")]
HB.structure Definition TJoinSemilattice d :=
  { T of JoinSemilattice d T & hasTop d T }.

#[short(type="tbJoinSemilatticeType")]
HB.structure Definition TBJoinSemilattice d :=
  { T of BJoinSemilattice d T & hasTop d T }.

#[short(type="latticeType")]
HB.structure Definition Lattice d :=
  { T of JoinSemilattice d T & POrder_isMeetSemilattice d T }.

#[short(type="bLatticeType")]
HB.structure Definition BLattice d := { T of Lattice d T & hasBottom d T }.

#[short(type="tLatticeType")]
HB.structure Definition TLattice d := { T of Lattice d T & hasTop d T }.

#[short(type="tbLatticeType")]
HB.structure Definition TBLattice d := { T of BLattice d T & hasTop d T }.

Section LatticeDef.
Context {disp : disp_t} {T : latticeType disp}.

Variant lel_xor_gt (x y : T) :
  T T T T T T T T bool bool Set :=
  | LelNotGt of x y : lel_xor_gt x y x x y y x x y y true false
  | GtlNotLe of y < x : lel_xor_gt x y y y x x y y x x false true.

Variant ltl_xor_ge (x y : T) :
  T T T T T T T T bool bool Set :=
  | LtlNotGe of x < y : ltl_xor_ge x y x x y y x x y y false true
  | GelNotLt of y x : ltl_xor_ge x y y y x x y y x x true false.

Variant comparel (x y : T) :
   T T T T T T T T
   bool bool bool bool bool bool Set :=
  | ComparelLt of x < y : comparel x y
    x x y y x x y y false false false true false true
  | ComparelGt of y < x : comparel x y
    y y x x y y x x false false true false true false
  | ComparelEq of x = y : comparel x y
    x x x x x x x x true true true true false false.

Variant incomparel (x y : T) :
  T T T T T T T T
  bool bool bool bool bool bool bool bool Set :=
  | InComparelLt of x < y : incomparel x y
    x x y y x x y y false false false true false true true true
  | InComparelGt of y < x : incomparel x y
    y y x x y y x x false false true false true false true true
  | InComparel of x >< y : incomparel x y
    x y y x (meet y x) (meet x y) (join y x) (join x y)
    false false false false false false false false
  | InComparelEq of x = y : incomparel x y
    x x x x x x x x true true true true false false true true.

End LatticeDef.

Module LatticeSyntax.

Notation "x `&` y" := (meet x y) : order_scope.
Notation "x `|` y" := (join x y) : order_scope.

End LatticeSyntax.

Module BLatticeSyntax.

Notation "\join_ ( i <- r | P ) F" :=
  (\big[@join _ _ / \bot]_(i <- r | P%B) F%O) : order_scope.
Notation "\join_ ( i <- r ) F" :=
  (\big[@join _ _ / \bot]_(i <- r) F%O) : order_scope.
Notation "\join_ ( i | P ) F" :=
  (\big[@join _ _ / \bot]_(i | P%B) F%O) : order_scope.
Notation "\join_ i F" :=
  (\big[@join _ _ / \bot]_i F%O) : order_scope.
Notation "\join_ ( i : I | P ) F" :=
  (\big[@join _ _ / \bot]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\join_ ( i : I ) F" :=
  (\big[@join _ _ / \bot]_(i : I) F%O) (only parsing) : order_scope.
Notation "\join_ ( m <= i < n | P ) F" :=
  (\big[@join _ _ / \bot]_(m i < n | P%B) F%O) : order_scope.
Notation "\join_ ( m <= i < n ) F" :=
  (\big[@join _ _ / \bot]_(m i < n) F%O) : order_scope.
Notation "\join_ ( i < n | P ) F" :=
  (\big[@join _ _ / \bot]_(i < n | P%B) F%O) : order_scope.
Notation "\join_ ( i < n ) F" :=
  (\big[@join _ _ / \bot]_(i < n) F%O) : order_scope.
Notation "\join_ ( i 'in' A | P ) F" :=
  (\big[@join _ _ / \bot]_(i in A | P%B) F%O) : order_scope.
Notation "\join_ ( i 'in' A ) F" :=
  (\big[@join _ _ / \bot]_(i in A) F%O) : order_scope.

End BLatticeSyntax.

Module TLatticeSyntax.

Notation "\meet_ ( i <- r | P ) F" :=
  (\big[meet / \top]_(i <- r | P%B) F%O) : order_scope.
Notation "\meet_ ( i <- r ) F" :=
  (\big[meet / \top]_(i <- r) F%O) : order_scope.
Notation "\meet_ ( i | P ) F" :=
  (\big[meet / \top]_(i | P%B) F%O) : order_scope.
Notation "\meet_ i F" :=
  (\big[meet / \top]_i F%O) : order_scope.
Notation "\meet_ ( i : I | P ) F" :=
  (\big[meet / \top]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\meet_ ( i : I ) F" :=
  (\big[meet / \top]_(i : I) F%O) (only parsing) : order_scope.
Notation "\meet_ ( m <= i < n | P ) F" :=
 (\big[meet / \top]_(m i < n | P%B) F%O) : order_scope.
Notation "\meet_ ( m <= i < n ) F" :=
 (\big[meet / \top]_(m i < n) F%O) : order_scope.
Notation "\meet_ ( i < n | P ) F" :=
 (\big[meet / \top]_(i < n | P%B) F%O) : order_scope.
Notation "\meet_ ( i < n ) F" :=
 (\big[meet / \top]_(i < n) F%O) : order_scope.
Notation "\meet_ ( i 'in' A | P ) F" :=
 (\big[meet / \top]_(i in A | P%B) F%O) : order_scope.
Notation "\meet_ ( i 'in' A ) F" :=
 (\big[meet / \top]_(i in A) F%O) : order_scope.

End TLatticeSyntax.

#[key="T", primitive]
HB.mixin Record Lattice_isDistributive d (T : Type) of Lattice d T := {
  meetUl : @left_distributive T T meet join;
  joinIl : @left_distributive T T join meet; (* dual of meetUl *)
}.

#[short(type="distrLatticeType")]
HB.structure Definition DistrLattice d :=
  { T of Lattice_isDistributive d T & Lattice d T }.

#[short(type="bDistrLatticeType")]
HB.structure Definition BDistrLattice d :=
  { T of DistrLattice d T & hasBottom d T }.

#[short(type="tDistrLatticeType")]
HB.structure Definition TDistrLattice d :=
  { T of DistrLattice d T & hasTop d T }.

#[short(type="tbDistrLatticeType")]
HB.structure Definition TBDistrLattice d :=
  { T of BDistrLattice d T & hasTop d T }.

#[key="T", primitive]
HB.mixin Record DistrLattice_isTotal d T of DistrLattice d T :=
  { le_total : total (<=%O : rel T) }.

#[short(type="orderType")]
HB.structure Definition Total d :=
  { T of DistrLattice_isTotal d T & DistrLattice d T }.

#[short(type="bOrderType")]
HB.structure Definition BTotal d := { T of Total d T & hasBottom d T }.

#[short(type="tOrderType")]
HB.structure Definition TTotal d := { T of Total d T & hasTop d T }.

#[short(type="tbOrderType")]
HB.structure Definition TBTotal d := { T of BTotal d T & hasTop d T }.

#[key="T", primitive]
HB.mixin Record DistrLattice_hasRelativeComplement d T of DistrLattice d T := {
  (* rcompl x y z is the complement of z in the interval x, y. *)
  rcompl : T T T T;
  rcomplPmeet : x y z, ((x `&` y) `|` z) `&` rcompl x y z = x `&` y;
  rcomplPjoin : x y z, ((y `|` x) `&` z) `|` rcompl x y z = y `|` x;
}.

#[short(type="cDistrLatticeType")]
HB.structure Definition CDistrLattice d :=
  { T of DistrLattice d T & DistrLattice_hasRelativeComplement d T }.

#[key="T", primitive]
HB.mixin Record CDistrLattice_hasSectionalComplement d T
         of CDistrLattice d T & hasBottom d T := {
  diff : T T T;
  (* FIXME: a bug in HB prevents us writing "rcompl \bot x y" *)
  diffErcompl : x y, diff x y = rcompl (\bot : T) x y;
}.

#[short(type="cbDistrLatticeType")]
HB.structure Definition CBDistrLattice d :=
  { T of CDistrLattice d T & hasBottom d T &
         CDistrLattice_hasSectionalComplement d T }.

#[key="T", primitive]
HB.mixin Record CDistrLattice_hasDualSectionalComplement d T
         of CDistrLattice d T & hasTop d T := {
  codiff : T T T;
  codiffErcompl : x y, codiff x y = rcompl x \top y;
}.

#[short(type="ctDistrLatticeType")]
HB.structure Definition CTDistrLattice d :=
  { T of CDistrLattice d T & hasTop d T &
         CDistrLattice_hasDualSectionalComplement d T }.

Module Import CBDistrLatticeSyntax.
Notation "x `\` y" := (diff x y) : order_scope.
End CBDistrLatticeSyntax.

#[key="T", primitive]
HB.mixin Record CDistrLattice_hasComplement d T of
         CTDistrLattice d T & CBDistrLattice d T := {
  compl : T T;
  (* FIXME: a bug in HB prevents us writing "\top `\` x" and "codiff \bot x" *)
  complEdiff : x : T, compl x = (\top : T) `\` x;
  complEcodiff : x : T, compl x = codiff (\bot : T) x;
}.

#[short(type="ctbDistrLatticeType")]
HB.structure Definition CTBDistrLattice d :=
  { T of CBDistrLattice d T & CTDistrLattice d T &
         CDistrLattice_hasComplement d T }.

Module Import CTBDistrLatticeSyntax.
Notation "~` A" := (compl A) : order_scope.
End CTBDistrLatticeSyntax.

FINITE

#[short(type="finPOrderType")]
HB.structure Definition FinPOrder d := { T of Finite T & POrder d T }.

#[short(type="finBPOrderType")]
HB.structure Definition FinBPOrder d := { T of FinPOrder d T & hasBottom d T }.

#[short(type="finTPOrderType")]
HB.structure Definition FinTPOrder d := { T of FinPOrder d T & hasTop d T }.

#[short(type="finTBPOrderType")]
HB.structure Definition FinTBPOrder d := { T of FinBPOrder d T & hasTop d T }.

#[short(type="finMeetSemilatticeType")]
HB.structure Definition FinMeetSemilattice d :=
  { T of Finite T & MeetSemilattice d T }.

#[short(type="finBMeetSemilatticeType")]
HB.structure Definition FinBMeetSemilattice d :=
  { T of Finite T & BMeetSemilattice d T }.

#[short(type="finJoinSemilatticeType")]
HB.structure Definition FinJoinSemilattice d :=
  { T of Finite T & JoinSemilattice d T }.

#[short(type="finTJoinSemilatticeType")]
HB.structure Definition FinTJoinSemilattice d :=
  { T of Finite T & TJoinSemilattice d T }.

#[short(type="finLatticeType")]
HB.structure Definition FinLattice d := { T of Finite T & Lattice d T }.

#[short(type="finTBLatticeType")]
HB.structure Definition FinTBLattice d := { T of Finite T & TBLattice d T }.

#[short(type="finDistrLatticeType")]
HB.structure Definition FinDistrLattice d :=
  { T of Finite T & DistrLattice d T }.

#[short(type="finTBDistrLatticeType")]
HB.structure Definition FinTBDistrLattice d :=
  { T of Finite T & TBDistrLattice d T }.

#[short(type="finOrderType")]
HB.structure Definition FinTotal d := { T of Finite T & Total d T }.

#[short(type="finTBOrderType")]
HB.structure Definition FinTBTotal d := { T of Finite T & TBTotal d T }.

#[short(type="finCDistrLatticeType")]
HB.structure Definition FinCDistrLattice d :=
  { T of Finite T & CDistrLattice d T }.

#[short(type="finCTBDistrLatticeType")]
HB.structure Definition FinCTBDistrLattice d :=
  { T of Finite T & CTBDistrLattice d T }.

DUAL

Notation dual_meet := (@meet (dual_display _) _).
Notation dual_join := (@join (dual_display _) _).

Module Import DualSyntax.

Notation "x `&^d` y" := (dual_meet x y) : order_scope.
Notation "x `|^d` y" := (dual_join x y) : order_scope.

The following Local Notations are here to define the \join^d_ and \meet^d_ notations later. Do not remove them.
Local Notation "\bot" := dual_bottom.
Local Notation "\top" := dual_top.
Local Notation join := dual_join.
Local Notation meet := dual_meet.
Local Notation min := dual_min.
Local Notation max := dual_max.

Notation "\join^d_ ( i <- r | P ) F" :=
  (\big[join / \bot]_(i <- r | P%B) F%O) : order_scope.
Notation "\join^d_ ( i <- r ) F" :=
  (\big[join / \bot]_(i <- r) F%O) : order_scope.
Notation "\join^d_ ( i | P ) F" :=
  (\big[join / \bot]_(i | P%B) F%O) : order_scope.
Notation "\join^d_ i F" :=
  (\big[join / \bot]_i F%O) : order_scope.
Notation "\join^d_ ( i : I | P ) F" :=
  (\big[join / \bot]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\join^d_ ( i : I ) F" :=
  (\big[join / \bot]_(i : I) F%O) (only parsing) : order_scope.
Notation "\join^d_ ( m <= i < n | P ) F" :=
 (\big[join / \bot]_(m i < n | P%B) F%O) : order_scope.
Notation "\join^d_ ( m <= i < n ) F" :=
 (\big[join / \bot]_(m i < n) F%O) : order_scope.
Notation "\join^d_ ( i < n | P ) F" :=
 (\big[join / \bot]_(i < n | P%B) F%O) : order_scope.
Notation "\join^d_ ( i < n ) F" :=
 (\big[join / \bot]_(i < n) F%O) : order_scope.
Notation "\join^d_ ( i 'in' A | P ) F" :=
 (\big[join / \bot]_(i in A | P%B) F%O) : order_scope.
Notation "\join^d_ ( i 'in' A ) F" :=
 (\big[join / \bot]_(i in A) F%O) : order_scope.

Notation "\meet^d_ ( i <- r | P ) F" :=
  (\big[meet / \top]_(i <- r | P%B) F%O) : order_scope.
Notation "\meet^d_ ( i <- r ) F" :=
  (\big[meet / \top]_(i <- r) F%O) : order_scope.
Notation "\meet^d_ ( i | P ) F" :=
  (\big[meet / \top]_(i | P%B) F%O) : order_scope.
Notation "\meet^d_ i F" :=
  (\big[meet / \top]_i F%O) : order_scope.
Notation "\meet^d_ ( i : I | P ) F" :=
  (\big[meet / \top]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\meet^d_ ( i : I ) F" :=
  (\big[meet / \top]_(i : I) F%O) (only parsing) : order_scope.
Notation "\meet^d_ ( m <= i < n | P ) F" :=
 (\big[meet / \top]_(m i < n | P%B) F%O) : order_scope.
Notation "\meet^d_ ( m <= i < n ) F" :=
 (\big[meet / \top]_(m i < n) F%O) : order_scope.
Notation "\meet^d_ ( i < n | P ) F" :=
 (\big[meet / \top]_(i < n | P%B) F%O) : order_scope.
Notation "\meet^d_ ( i < n ) F" :=
 (\big[meet / \top]_(i < n) F%O) : order_scope.
Notation "\meet^d_ ( i 'in' A | P ) F" :=
 (\big[meet / \top]_(i in A | P%B) F%O) : order_scope.
Notation "\meet^d_ ( i 'in' A ) F" :=
 (\big[meet / \top]_(i in A) F%O) : order_scope.

Notation "\meet^d_ ( i <- r | P ) F" :=
  (\big[meet / \top]_(i <- r | P%B) F%O) : order_scope.
Notation "\meet^d_ ( i <- r ) F" :=
  (\big[meet / \top]_(i <- r) F%O) : order_scope.
Notation "\meet^d_ ( i | P ) F" :=
  (\big[meet / \top]_(i | P%B) F%O) : order_scope.
Notation "\meet^d_ i F" :=
  (\big[meet / \top]_i F%O) : order_scope.
Notation "\meet^d_ ( i : I | P ) F" :=
  (\big[meet / \top]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\meet^d_ ( i : I ) F" :=
  (\big[meet / \top]_(i : I) F%O) (only parsing) : order_scope.
Notation "\meet^d_ ( m <= i < n | P ) F" :=
 (\big[meet / \top]_(m i < n | P%B) F%O) : order_scope.
Notation "\meet^d_ ( m <= i < n ) F" :=
 (\big[meet / \top]_(m i < n) F%O) : order_scope.
Notation "\meet^d_ ( i < n | P ) F" :=
 (\big[meet / \top]_(i < n | P%B) F%O) : order_scope.
Notation "\meet^d_ ( i < n ) F" :=
 (\big[meet / \top]_(i < n) F%O) : order_scope.
Notation "\meet^d_ ( i 'in' A | P ) F" :=
 (\big[meet / \top]_(i in A | P%B) F%O) : order_scope.
Notation "\meet^d_ ( i 'in' A ) F" :=
 (\big[meet / \top]_(i in A) F%O) : order_scope.

Notation "\min^d_ i F" :=
  (\big[min/top]_i F) : order_scope.
Notation "\min^d_ ( i <- r | P ) F" :=
  (\big[min/top]_(i <- r | P%B) F%O) : order_scope.
Notation "\min^d_ ( i < r ) F" :=
  (\big[min/top]_(i <- r) F%O) : order_scope.
Notation "\min^d_ ( m <= i < n | P ) F" :=
  (\big[min/top]_(m i < n | P%B) F%O) : order_scope.
Notation "\min^d_ ( m <= i < n ) F" :=
  (\big[min/top]_(m i < n) F%O) : order_scope.
Notation "\min^d_ ( i | P ) F" :=
  (\big[min/top]_(i | P%B) F%O) : order_scope.
Notation "\min^d_ ( i : t | P ) F" :=
  (\big[min/top]_(i : t | P%B) F%O) (only parsing) : order_scope.
Notation "\min^d_ ( i : t ) F" :=
  (\big[min/top]_(i : t) F%O) (only parsing) : order_scope.
Notation "\min^d_ ( i < n | P ) F" :=
  (\big[min/top]_(i < n | P%B) F%O) : order_scope.
Notation "\min^d_ ( i < n ) F" :=
  (\big[min/top]_(i < n) F%O) : order_scope.
Notation "\min^d_ ( i 'in' A | P ) F" :=
  (\big[min/top]_(i in A | P%B) F%O) : order_scope.
Notation "\min^d_ ( i 'in' A ) F" :=
  (\big[min/top]_(i in A) F%O) : order_scope.

Notation "\max^d_ i F" :=
  (\big[max/bottom]_i F%O) : order_scope.
Notation "\max^d_ ( i <- r | P ) F" :=
  (\big[max/bottom]_(i <- r | P%B) F%O) : order_scope.
Notation "\max^d_ ( i < r ) F" :=
  (\big[max/bottom]_(i <- r) F%O) : order_scope.
Notation "\max^d_ ( m <= i < n | P ) F" :=
  (\big[max/bottom]_(m i < n | P%B) F%O) : order_scope.
Notation "\max^d_ ( m <= i < n ) F" :=
  (\big[max/bottom]_(m i < n) F%O) : order_scope.
Notation "\max^d_ ( i | P ) F" :=
  (\big[max/bottom]_(i | P%B) F%O) : order_scope.
Notation "\max^d_ ( i : t | P ) F" :=
  (\big[max/bottom]_(i : t | P%B) F%O) (only parsing) : order_scope.
Notation "\max^d_ ( i : t ) F" :=
  (\big[max/bottom]_(i : t) F%O) (only parsing) : order_scope.
Notation "\max^d_ ( i < n | P ) F" :=
  (\big[max/bottom]_(i < n | P%B) F%O) : order_scope.
Notation "\max^d_ ( i < n ) F" :=
  (\big[max/bottom]_(i < n) F%O) : order_scope.
Notation "\max^d_ ( i 'in' A | P ) F" :=
  (\big[max/bottom]_(i in A | P%B) F%O) : order_scope.
Notation "\max^d_ ( i 'in' A ) F" :=
  (\big[max/bottom]_(i in A) F%O) : order_scope.

End DualSyntax.

FIXME: we have two issues in the dual instance declarations in the DualOrder module below: 1. HB.saturate is slow, and 2. if we declare them by HB.instance, some declarations fail because it unfolds [dual] and the generated instance does not typecheck (math-comp/hierarchy-builder#257).
Module DualOrder.



Lemma meetEdual d (T : joinSemilatticeType d) (x y : T) :
  ((x : T^d) `&^d` y) = (x `|` y).


Lemma joinEdual d (T : meetSemilatticeType d) (x y : T) :
  ((x : T^d) `|^d` y) = (x `&` y).







End DualOrder.

THEORY

Module Import POrderTheory.

Include PreorderTheory.

Section POrderTheory.

Context {disp : disp_t} {T : porderType disp}.

Implicit Types (x y : T) (s : seq T).

Lemma le_anti: antisymmetric (<=%O : rel T).

Lemma ge_anti: antisymmetric (>=%O : rel T).

Lemma eq_le x y: (x == y) = (x y x).

Lemma lt_def x y : (x < y) = (y != x) && (x y).

Lemma lt_neqAle x y: (x < y) = (x != y) && (x y).

Lemma le_eqVlt x y: (x y) = (x == y) || (x < y).

Definition lte_anti := (=^~ eq_le, @lt_asym disp T, @lt_le_asym disp T, @le_lt_asym disp T).

Lemma eq_geP {x y} : reflect ( z, (z x) = (z y)) (x == y).

Lemma eq_leP {x y} : reflect ( z, (x z) = (y z)) (x == y).

Lemma lt_sorted_uniq_le s : sorted <%O s = uniq s && sorted <=%O s.

Lemma le_sorted_eq s1 s2 :
  sorted <=%O s1 sorted <=%O s2 perm_eq s1 s2 s1 = s2.

Lemma count_lt_le_mem x s : (count (< x) s < count ( x) s)%N = (x \in s).

Lemma comparable_ltgtP x y : x >=< y
  compare x y (min y x) (min x y) (max y x) (max x y)
  (y == x) (x == y) (x y) (x y) (x > y) (x < y).

Lemma comparable_leP x y : x >=< y
  le_xor_gt x y (min y x) (min x y) (max y x) (max x y) (x y) (y < x).

Lemma comparable_ltP x y : x >=< y
  lt_xor_ge x y (min y x) (min x y) (max y x) (max x y) (y x) (x < y).

Lemma comparableP x y : incompare x y
  (min y x) (min x y) (max y x) (max x y)
  (y == x) (x == y) (x y) (x y) (x > y) (x < y)
  (y >=< x) (x >=< y).

leif

Lemma leifP x y C : reflect (x y ?= iff C) (if C then x == y else x < y).

Lemma leif_trans x1 x2 x3 C12 C23 :
  x1 x2 ?= iff C12 x2 x3 ?= iff C23 x1 x3 ?= iff C12 && C23.

Lemma leif_le x y : x y x y ?= iff (x y).

Lemma leif_eq x y : x y x y ?= iff (x == y).

Lemma ge_leif x y C : x y ?= iff C (y x) = C.

Lemma lt_leif x y C : x y ?= iff C (x < y) = ~~ C.

Lemma ltNleif x y C : x y ?= iff ~~ C (x < y) = C.

lteif

Lemma lteif_anti C1 C2 x y :
  (x < y ?<= if C1) && (y < x ?<= if C2) = C1 && C2 && (x == y).

Lemma lteifN C x y : x < y ?<= if ~~ C ~~ (y < x ?<= if C).

min and max

Lemma minEle x y : min x y = if x y then x else y.

Lemma maxEle x y : max x y = if x y then y else x.

Lemma comparable_minEgt x y : x >=< y min x y = if x > y then y else x.
Lemma comparable_maxEgt x y : x >=< y max x y = if x > y then x else y.
Lemma comparable_minEge x y : x >=< y min x y = if x y then y else x.
Lemma comparable_maxEge x y : x >=< y max x y = if x y then x else y.

Lemma min_l x y : x y min x y = x.
Lemma min_r x y : y x min x y = y.
Lemma max_l x y : y x max x y = x.
Lemma max_r x y : x y max x y = y.

Lemma eq_minl x y : (min x y == x) = (x y).

Lemma eq_maxr x y : (max x y == y) = (x y).

Lemma min_idPl x y : reflect (min x y = x) (x y).

Lemma max_idPr x y : reflect (max x y = y) (x y).

Section Comparable2.
Context (z x y : T) (cmp_xy : x >=< y).

Lemma comparable_minC : min x y = min y x.

Lemma comparable_maxC : max x y = max y x.

Lemma comparable_eq_minr : (min x y == y) = (y x).

Lemma comparable_eq_maxl : (max x y == x) = (y x).

Lemma comparable_min_idPr : reflect (min x y = y) (y x).

Lemma comparable_max_idPl : reflect (max x y = x) (y x).

Lemma comparable_lteifNE C : x >=< y x < y ?<= if ~~ C = ~~ (y < x ?<= if C).

End Comparable2.

Section Comparable3.
Context (x y z : T) (cmp_xy : x >=< y) (cmp_xz : x >=< z) (cmp_yz : y >=< z).
Let P := comparableP.

Lemma comparable_max_minl : max (min x y) z = min (max x z) (max y z).

End Comparable3.

Section Comparable4.
Context (x y z w : T) (cmp_xy : x >=< y) (cmp_zw : z >=< w).

Lemma comparable_le_min2 : x z y w
  Order.min x y Order.min z w.

Lemma comparable_le_max2 : x z y w
  Order.max x y Order.max z w.

End Comparable4.

Lemma comparable_minAC x y z : x >=< y x >=< z y >=< z
  min (min x y) z = min (min x z) y.

Lemma comparable_maxAC x y z : x >=< y x >=< z y >=< z
  max (max x y) z = max (max x z) y.

Lemma comparable_minCA x y z : x >=< y x >=< z y >=< z
  min x (min y z) = min y (min x z).

Lemma comparable_maxCA x y z : x >=< y x >=< z y >=< z
  max x (max y z) = max y (max x z).

Lemma comparable_minACA x y z t :
    x >=< y x >=< z x >=< t y >=< z y >=< t z >=< t
  min (min x y) (min z t) = min (min x z) (min y t).

Lemma comparable_maxACA x y z t :
    x >=< y x >=< z x >=< t y >=< z y >=< t z >=< t
  max (max x y) (max z t) = max (max x z) (max y t).

Lemma comparable_min_maxr x y z : x >=< y x >=< z y >=< z
  min x (max y z) = max (min x y) (min x z).

monotonicity

Lemma mono_in_leif (A : {pred T}) (f : T T) C :
   {in A &, {mono f : x y / x y}}
  {in A &, x y, (f x f y ?= iff C) = (x y ?= iff C)}.

Lemma mono_leif (f : T T) C :
    {mono f : x y / x y}
   x y, (f x f y ?= iff C) = (x y ?= iff C).

Lemma nmono_in_leif (A : {pred T}) (f : T T) C :
    {in A &, {mono f : x y /~ x y}}
  {in A &, x y, (f x f y ?= iff C) = (y x ?= iff C)}.

Lemma nmono_leif (f : T T) C : {mono f : x y /~ x y}
   x y, (f x f y ?= iff C) = (y x ?= iff C).

Section bigminmax.
Context (I : Type) (r : seq I) (f : I T) (x0 x : T) (P : pred I).

Lemma bigmax_le : x0 x ( i, P i f i x)
  \big[max/x0]_(i <- r | P i) f i x.

Lemma le_bigmin : x x0 ( i, P i x f i)
  x \big[min/x0]_(i <- r | P i) f i.

End bigminmax.

End POrderTheory.

#[global] Hint Resolve comparable_minr comparable_minl : core.
#[global] Hint Resolve comparable_maxr comparable_maxl : core.

Section ContraTheory.
Context {disp1 disp2 : disp_t} {T1 : porderType disp1} {T2 : porderType disp2}.
Implicit Types (x y : T1) (z t : T2) (b : bool) (m n : nat) (P : Prop).

Lemma contra_leT b x y : (~~ b x < y) (y x b).

Lemma contra_ltT b x y : (~~ b x y) (y < x b).

Lemma contra_leN b x y : (b x < y) (y x ~~ b).

Lemma contra_ltN b x y : (b x y) (y < x ~~ b).

Lemma contra_le_not P x y : (P x < y) (y x ¬ P).

Lemma contra_lt_not P x y : (P x y) (y < x ¬ P).

Lemma contra_leF b x y : (b x < y) (y x b = false).

Lemma contra_ltF b x y : (b x y) (y < x b = false).

Lemma contra_le_leq x y m n : ((n < m)%N y < x) (x y (m n)%N).

Lemma contra_le_ltn x y m n : ((n m)%N y < x) (x y (m < n)%N).

Lemma contra_lt_leq x y m n : ((n < m)%N y x) (x < y (m n)%N).

Lemma contra_lt_ltn x y m n : ((n m)%N y x) (x < y (m < n)%N).

End ContraTheory.

Section POrderMonotonyTheory.
Context {disp disp' : disp_t} {T : porderType disp} {T' : porderType disp'}.
Context (D D' : {pred T}) (f : T T').

Let leT_anti := @le_anti _ T.
Hint Resolve lexx lt_neqAle : core.

Let ge_antiT : antisymmetric (>=%O : rel T).

Lemma ltW_homo : {homo f : x y / x < y} {homo f : x y / x y}.

Lemma ltW_nhomo : {homo f : x y /~ x < y} {homo f : x y /~ x y}.

Lemma inj_homo_lt :
  injective f {homo f : x y / x y} {homo f : x y / x < y}.

Lemma inj_nhomo_lt :
  injective f {homo f : x y /~ x y} {homo f : x y /~ x < y}.

Lemma inc_inj : {mono f : x y / x y} injective f.

Lemma dec_inj : {mono f : x y /~ x y} injective f.

Monotony in D D'
Lemma ltW_homo_in :
  {in D & D', {homo f : x y / x < y}} {in D & D', {homo f : x y / x y}}.

Lemma ltW_nhomo_in :
  {in D & D', {homo f : x y /~ x < y}} {in D & D', {homo f : x y /~ x y}}.

Lemma inj_homo_lt_in :
    {in D & D', injective f} {in D & D', {homo f : x y / x y}}
  {in D & D', {homo f : x y / x < y}}.

Lemma inj_nhomo_lt_in :
    {in D & D', injective f} {in D & D', {homo f : x y /~ x y}}
  {in D & D', {homo f : x y /~ x < y}}.

Lemma inc_inj_in : {in D &, {mono f : x y / x y}}
   {in D &, injective f}.

Lemma dec_inj_in :
  {in D &, {mono f : x y /~ x y}} {in D &, injective f}.

End POrderMonotonyTheory.

End POrderTheory.

Arguments leifP {disp T x y C}.
Arguments mono_in_leif [disp T A f C].
Arguments nmono_in_leif [disp T A f C].
Arguments mono_leif [disp T f C].
Arguments nmono_leif [disp T f C].
Arguments min_idPl {disp T x y}.
Arguments max_idPr {disp T x y}.
Arguments comparable_min_idPr {disp T x y _}.
Arguments comparable_max_idPl {disp T x y _}.

Module Import BPOrderTheory.

Export BPreorderTheory.

Section BPOrderTheory.
Context {disp : disp_t} {T : bPOrderType disp}.
Implicit Types (x y : T).

Lemma lex0 x : (x \bot) = (x == \bot).

Lemma lt0x x : (\bot < x) = (x != \bot).

Variant eq0_xor_gt0 x : bool bool Set :=
    Eq0NotPOs : x = \bot eq0_xor_gt0 x true false
  | POsNotEq0 : \bot < x eq0_xor_gt0 x false true.

Lemma posxP x : eq0_xor_gt0 x (x == \bot) (\bot < x).

End BPOrderTheory.
End BPOrderTheory.

Module Import TPOrderTheory.
Section TPOrderTheory.
Context {disp : disp_t} {T : tPOrderType disp}.
Implicit Types (x y : T).

Lemma le1x x : (\top x) = (x == \top).
Lemma ltx1 x : (x < \top) = (x != \top).

End TPOrderTheory.
End TPOrderTheory.

Module Import MeetTheory.
Section MeetTheory.
Context {disp : disp_t} {L : meetSemilatticeType disp}.
Implicit Types (x y : L).

interaction with order

Lemma lexI x y z : (x y `&` z) = (x y) && (x z).

Lemma leIr x y : y `&` x x.

Lemma leIl x y : x `&` y x.

Lemma leIxl x y z : y x y `&` z x.

Lemma leIxr x y z : z x y `&` z x.

Lemma leIx2 x y z : (y x) || (z x) y `&` z x.

Lemma leEmeet x y : (x y) = (x `&` y == x).

Lemma eq_meetl x y : (x `&` y == x) = (x y).

Lemma eq_meetr x y : (x `&` y == y) = (y x).

Lemma meet_idPl {x y} : reflect (x `&` y = x) (x y).
Lemma meet_idPr {x y} : reflect (y `&` x = x) (x y).

Lemma meet_l x y : x y x `&` y = x.
Lemma meet_r x y : y x x `&` y = y.

Lemma leIidl x y : (x x `&` y) = (x y).
Lemma leIidr x y : (x y `&` x) = (x y).

Lemma leI2 x y z t : x z y t x `&` y z `&` t.

algebraic properties

Lemma meetC : commutative (@meet _ L).

Lemma meetA : associative (@meet _ L).


Lemma meetxx : idempotent_op (@meet _ L).
Lemma meetAC : right_commutative (@meet _ L).
Lemma meetCA : left_commutative (@meet _ L).
Lemma meetACA : interchange (@meet _ L) (@meet _ L).

Lemma meetKI y x : x `&` (x `&` y) = x `&` y.
Lemma meetIK y x : (x `&` y) `&` y = x `&` y.
Lemma meetKIC y x : x `&` (y `&` x) = x `&` y.
Lemma meetIKC y x : y `&` x `&` y = x `&` y.

End MeetTheory.
End MeetTheory.

Arguments meet_idPl {disp L x y}.
Arguments meet_idPr {disp L x y}.

Module Import BMeetTheory.
Section BMeetTheory.
Context {disp : disp_t} {L : bMeetSemilatticeType disp}.

Lemma meet0x : left_zero \bot (@meet _ L).

Lemma meetx0 : right_zero \bot (@meet _ L).


End BMeetTheory.
End BMeetTheory.

Module Import TMeetTheory.
Section TMeetTheory.
Context {disp : disp_t} {L : tMeetSemilatticeType disp}.
Implicit Types (I : finType) (T : eqType) (x y : L).

Lemma meetx1 : right_id \top (@meet _ L).

Lemma meet1x : left_id \top (@meet _ L).

Lemma meet_eq1 x y : (x `&` y == \top) = (x == \top) && (y == \top).


Lemma meets_inf_seq T (r : seq T) (P : {pred T}) (F : T L) (x : T) :
  x \in r P x \meet_(i <- r | P i) F i F x.

Lemma meets_max_seq T (r : seq T) (P : {pred T}) (F : T L) (x : T) (u : L) :
  x \in r P x F x u \meet_(x <- r | P x) F x u.

Lemma meets_inf I (j : I) (P : {pred I}) (F : I L) :
   P j \meet_(i | P i) F i F j.

Lemma meets_max I (j : I) (u : L) (P : {pred I}) (F : I L) :
   P j F j u \meet_(i | P i) F i u.

Lemma meets_ge J (r : seq J) (P : {pred J}) (F : J L) (u : L) :
  ( x : J, P x u F x) u \meet_(x <- r | P x) F x.

Lemma meetsP_seq T (r : seq T) (P : {pred T}) (F : T L) (l : L) :
  reflect ( x : T, x \in r P x l F x)
          (l \meet_(x <- r | P x) F x).

Lemma meetsP I (l : L) (P : {pred I}) (F : I L) :
   reflect ( i : I, P i l F i) (l \meet_(i | P i) F i).

Lemma le_meets I (A B : {set I}) (F : I L) :
   A \subset B \meet_(i in B) F i \meet_(i in A) F i.

Lemma meets_setU I (A B : {set I}) (F : I L) :
   \meet_(i in (A :|: B)) F i = \meet_(i in A) F i `&` \meet_(i in B) F i.

Lemma meets_seq I (r : seq I) (F : I L) :
   \meet_(i <- r) F i = \meet_(i in r) F i.

End TMeetTheory.
End TMeetTheory.

Module Import JoinTheory.
Section JoinTheory.
Context {disp : disp_t} {L : joinSemilatticeType disp}.
Implicit Types (x y : L).

interaction with order

Lemma leUx x y z : (x `|` y z) = (x z) && (y z).

Lemma leUr x y : x y `|` x.
Lemma leUl x y : x x `|` y.

Lemma lexUl x y z : x y x y `|` z.
Lemma lexUr x y z : x z x y `|` z.
Lemma lexU2 x y z : (x y) || (x z) x y `|` z.

Lemma leEjoin x y : (x y) = (x `|` y == y).

Lemma eq_joinl x y : (x `|` y == x) = (y x).
Lemma eq_joinr x y : (x `|` y == y) = (x y).

Lemma join_idPl {x y} : reflect (y `|` x = y) (x y).
Lemma join_idPr {x y} : reflect (x `|` y = y) (x y).

Lemma join_l x y : y x x `|` y = x.
Lemma join_r x y : x y x `|` y = y.

Lemma leUidl x y : (x `|` y y) = (x y).
Lemma leUidr x y : (y `|` x y) = (x y).

Lemma leU2 x y z t : x z y t x `|` y z `|` t.

algebraic properties

Lemma joinC : commutative (@join _ L).
Lemma joinA : associative (@join _ L).


Lemma joinxx : idempotent_op (@join _ L).
Lemma joinAC : right_commutative (@join _ L).
Lemma joinCA : left_commutative (@join _ L).
Lemma joinACA : interchange (@join _ L) (@join _ L).

Lemma joinKU y x : x `|` (x `|` y) = x `|` y.
Lemma joinUK y x : (x `|` y) `|` y = x `|` y.
Lemma joinKUC y x : x `|` (y `|` x) = x `|` y.
Lemma joinUKC y x : y `|` x `|` y = x `|` y.

End JoinTheory.
End JoinTheory.

Arguments join_idPl {disp L x y}.
Arguments join_idPr {disp L x y}.

Module Import BJoinTheory.
Section BJoinTheory.
Context {disp : disp_t} {L : bJoinSemilatticeType disp}.
Implicit Types (I : finType) (T : eqType) (x y : L).

Lemma joinx0 : right_id \bot (@join _ L).
Lemma join0x : left_id \bot (@join _ L).

Lemma join_eq0 x y : (x `|` y == \bot) = (x == \bot) && (y == \bot).


Lemma joins_sup_seq T (r : seq T) (P : {pred T}) (F : T L) (x : T) :
  x \in r P x F x \join_(i <- r | P i) F i.

Lemma joins_min_seq T (r : seq T) (P : {pred T}) (F : T L) (x : T) (l : L) :
  x \in r P x l F x l \join_(x <- r | P x) F x.

Lemma joins_sup I (j : I) (P : {pred I}) (F : I L) :
  P j F j \join_(i | P i) F i.

Lemma joins_min I (j : I) (l : L) (P : {pred I}) (F : I L) :
  P j l F j l \join_(i | P i) F i.

Lemma joins_le J (r : seq J) (P : {pred J}) (F : J L) (u : L) :
  ( x : J, P x F x u) \join_(x <- r | P x) F x u.

Lemma joinsP_seq T (r : seq T) (P : {pred T}) (F : T L) (u : L) :
  reflect ( x : T, x \in r P x F x u)
          (\join_(x <- r | P x) F x u).

Lemma joinsP I (u : L) (P : {pred I}) (F : I L) :
  reflect ( i : I, P i F i u) (\join_(i | P i) F i u).

Lemma le_joins I (A B : {set I}) (F : I L) :
  A \subset B \join_(i in A) F i \join_(i in B) F i.

Lemma joins_setU I (A B : {set I}) (F : I L) :
  \join_(i in (A :|: B)) F i = \join_(i in A) F i `|` \join_(i in B) F i.

Lemma joins_seq I (r : seq I) (F : I L) :
  \join_(i <- r) F i = \join_(i in r) F i.

End BJoinTheory.
End BJoinTheory.

Module Import TJoinTheory.
Section TJoinTheory.
Context {disp : disp_t} {L : tJoinSemilatticeType disp}.

Lemma joinx1 : right_zero \top (@join _ L).
Lemma join1x : left_zero \top (@join _ L).


End TJoinTheory.
End TJoinTheory.

Module Import LatticeTheory.
Section LatticeTheory.
Context {disp : disp_t} {L : latticeType disp}.
Implicit Types (x y : L).

Lemma meetUK x y : (x `&` y) `|` y = y.
Lemma meetUKC x y : (y `&` x) `|` y = y.
Lemma meetKUC y x : x `|` (y `&` x) = x.
Lemma meetKU y x : x `|` (x `&` y) = x.

Lemma joinIK x y : (x `|` y) `&` y = y.
Lemma joinIKC x y : (y `|` x) `&` y = y.
Lemma joinKIC y x : x `&` (y `|` x) = x.
Lemma joinKI y x : x `&` (x `|` y) = x.

comparison predicates

Lemma lcomparableP x y : incomparel x y
  (min y x) (min x y) (max y x) (max x y)
  (y `&` x) (x `&` y) (y `|` x) (x `|` y)
  (y == x) (x == y) (x y) (x y) (x > y) (x < y) (y >=< x) (x >=< y).

Lemma lcomparable_ltgtP x y : x >=< y
  comparel x y (min y x) (min x y) (max y x) (max x y)
               (y `&` x) (x `&` y) (y `|` x) (x `|` y)
               (y == x) (x == y) (x y) (x y) (x > y) (x < y).

Lemma lcomparable_leP x y : x >=< y
  lel_xor_gt x y (min y x) (min x y) (max y x) (max x y)
                 (y `&` x) (x `&` y) (y `|` x) (x `|` y) (x y) (y < x).

Lemma lcomparable_ltP x y : x >=< y
  ltl_xor_ge x y (min y x) (min x y) (max y x) (max x y)
                 (y `&` x) (x `&` y) (y `|` x) (x `|` y) (y x) (x < y).

End LatticeTheory.
End LatticeTheory.

Module Import DistrLatticeTheory.
Section DistrLatticeTheory.
Context {disp : disp_t} {L : distrLatticeType disp}.

Lemma meetUl : left_distributive (@meet _ L) (@join _ L).

Lemma meetUr : right_distributive (@meet _ L) (@join _ L).

Lemma joinIl : left_distributive (@join _ L) (@meet _ L).

Lemma joinIr : right_distributive (@join _ L) (@meet _ L).

#[warning="-HB.no-new-instance"]
HB.instance Definition _ := Monoid.isAddLaw.Build L meet join meetUl meetUr.

End DistrLatticeTheory.
End DistrLatticeTheory.

Module Import BDistrLatticeTheory.
Section BDistrLatticeTheory.
Context {disp : disp_t} {L : bDistrLatticeType disp}.
Implicit Types (x y z : L).

Lemma leU2l_le y t x z : x `&` t = \bot x `|` y z `|` t x z.

Lemma leU2r_le y t x z : x `&` t = \bot y `|` x t `|` z x z.

Lemma disjoint_lexUl z x y : x `&` z = \bot (x y `|` z) = (x y).

Lemma disjoint_lexUr z x y : x `&` z = \bot (x z `|` y) = (x y).

Lemma leU2E x y z t : x `&` t = \bot y `&` z = \bot
  (x `|` y z `|` t) = (x z) && (y t).

Lemma joins_disjoint (I : finType) (d : L) (P : {pred I}) (F : I L) :
   ( i : I, P i d `&` F i = \bot) d `&` \join_(i | P i) F i = \bot.

End BDistrLatticeTheory.
End BDistrLatticeTheory.

Module Import TDistrLatticeTheory.
Section TDistrLatticeTheory.
Context {disp : disp_t} {L : tDistrLatticeType disp}.
Implicit Types (x y : L).

Lemma leI2l_le y t x z : y `|` z = \top x `&` y z `&` t x z.

Lemma leI2r_le y t x z : y `|` z = \top y `&` x t `&` z x z.

Lemma cover_leIxl z x y : z `|` y = \top (x `&` z y) = (x y).

Lemma cover_leIxr z x y : z `|` y = \top (z `&` x y) = (x y).

Lemma leI2E x y z t : x `|` t = \top y `|` z = \top
  (x `&` y z `&` t) = (x z) && (y t).

Lemma meets_total (I : finType) (d : L) (P : {pred I}) (F : I L) :
   ( i : I, P i d `|` F i = \top) d `|` \meet_(i | P i) F i = \top.

End TDistrLatticeTheory.
End TDistrLatticeTheory.

Module Import TotalTheory.
Section TotalTheory.
Context {disp : disp_t} {T : orderType disp}.
Implicit Types (x y z t : T) (s : seq T).

Definition le_total : total (<=%O : rel T) := le_total.
Hint Resolve le_total : core.

Lemma ge_total : total (>=%O : rel T).
Hint Resolve ge_total : core.

Lemma comparableT x y : x >=< y.
Hint Extern 0 (is_true (_ >=< _)%O) ⇒ solve [apply: comparableT] : core.

Lemma sort_le_sorted s : sorted <=%O (sort <=%O s).
Hint Resolve sort_le_sorted : core.

Lemma sort_lt_sorted s : sorted <%O (sort <=%O s) = uniq s.

Lemma perm_sort_leP s1 s2 : reflect (sort <=%O s1 = sort <=%O s2) (perm_eq s1 s2).

Lemma filter_sort_le p s : filter p (sort <=%O s) = sort <=%O (filter p s).

Lemma mask_sort_le s (m : bitseq) :
  {m_s : bitseq | mask m_s (sort <=%O s) = sort <=%O (mask m s)}.

Lemma sorted_mask_sort_le s (m : bitseq) :
  sorted <=%O (mask m s) {m_s : bitseq | mask m_s (sort <=%O s) = mask m s}.

Lemma subseq_sort_le : {homo sort <=%O : s1 s2 / @subseq T s1 s2}.

Lemma sorted_subseq_sort_le s1 s2 :
  subseq s1 s2 sorted <=%O s1 subseq s1 (sort <=%O s2).

Lemma mem2_sort_le s x y : x y mem2 s x y mem2 (sort <=%O s) x y.

Lemma leNgt x y : (x y) = ~~ (y < x).

Lemma ltNge x y : (x < y) = ~~ (y x).

Definition ltgtP x y := LatticeTheory.lcomparable_ltgtP (comparableT x y).
Definition leP x y := LatticeTheory.lcomparable_leP (comparableT x y).
Definition ltP x y := LatticeTheory.lcomparable_ltP (comparableT x y).

Lemma wlog_le P :
     ( x y, P y x P x y) ( x y, x y P x y)
    x y, P x y.

Lemma wlog_lt P :
    ( x, P x x)
    ( x y, (P y x P x y)) ( x y, x < y P x y)
   x y, P x y.

Lemma neq_lt x y : (x != y) = (x < y) || (y < x).

Lemma lt_total x y : x != y (x < y) || (y < x).

Lemma eq_leLR x y z t :
  (x y z t) (y < x t < z) (x y) = (z t).

Lemma eq_leRL x y z t :
  (x y z t) (y < x t < z) (z t) = (x y).

Lemma eq_ltLR x y z t :
  (x < y z < t) (y x t z) (x < y) = (z < t).

Lemma eq_ltRL x y z t :
  (x < y z < t) (y x t z) (z < t) = (x < y).

Lemma le_gtP {x y} : reflect ( z, z < x z < y) (x y).

Lemma le_ltP {x y} : reflect ( z, y < z x < z) (x y).

Lemma eq_gtP {x y} : reflect ( z, (z < x) = (z < y)) (x == y).

Lemma eq_ltP {x y} : reflect ( z, (x < z) = (y < z)) (x == y).

max and min is join and meet

Lemma meetEtotal x y : x `&` y = min x y.
Lemma joinEtotal x y : x `|` y = max x y.

max and min theory

Lemma minEgt x y : min x y = if x > y then y else x.
Lemma maxEgt x y : max x y = if x > y then x else y.
Lemma minEge x y : min x y = if x y then y else x.
Lemma maxEge x y : max x y = if x y then x else y.

Lemma minC : commutative (min : T T T).

Lemma maxC : commutative (max : T T T).

Lemma minA : associative (min : T T T).

Lemma maxA : associative (max : T T T).

Lemma minAC : right_commutative (min : T T T).

Lemma maxAC : right_commutative (max : T T T).

Lemma minCA : left_commutative (min : T T T).

Lemma maxCA : left_commutative (max : T T T).

Lemma minACA : interchange (min : T T T) min.

Lemma maxACA : interchange (max : T T T) max.

Lemma eq_minr x y : (min x y == y) = (y x).

Lemma eq_maxl x y : (max x y == x) = (y x).

Lemma min_idPr x y : reflect (min x y = y) (y x).

Lemma max_idPl x y : reflect (max x y = x) (y x).

Lemma le_min z x y : (z min x y) = (z x) && (z y).

Lemma ge_min z x y : (min x y z) = (x z) || (y z).

Lemma lt_min z x y : (z < min x y) = (z < x) && (z < y).

Lemma gt_min z x y : (min x y < z) = (x < z) || (y < z).

Lemma le_max z x y : (z max x y) = (z x) || (z y).

Lemma ge_max z x y : (max x y z) = (x z) && (y z).

Lemma lt_max z x y : (z < max x y) = (z < x) || (z < y).

Lemma gt_max z x y : (max x y < z) = (x < z) && (y < z).

Lemma minxK x y : max (min x y) y = y.
Lemma minKx x y : max x (min x y) = x.
Lemma maxxK x y : min (max x y) y = y.
Lemma maxKx x y : min x (max x y) = x.

Lemma max_minl : left_distributive (max : T T T) min.

Lemma min_maxl : left_distributive (min : T T T) max.

Lemma max_minr : right_distributive (max : T T T) min.

Lemma min_maxr : right_distributive (min : T T T) max.


Lemma leIx x y z : (meet y z x) = (y x) || (z x).

Lemma lexU x y z : (x join y z) = (x y) || (x z).

Lemma ltxI x y z : (x < meet y z) = (x < y) && (x < z).

Lemma ltIx x y z : (meet y z < x) = (y < x) || (z < x).

Lemma ltxU x y z : (x < join y z) = (x < y) || (x < z).

Lemma ltUx x y z : (join y z < x) = (y < x) && (z < x).

Definition ltexI := (@lexI _ T, ltxI).
Definition lteIx := (leIx, ltIx).
Definition ltexU := (lexU, ltxU).
Definition lteUx := (@leUx _ T, ltUx).

Lemma le_min2 x y z t : x z y t Order.min x y Order.min z t.

Lemma le_max2 x y z t : x z y t Order.max x y Order.max z t.

lteif

Lemma lteifNE x y C : x < y ?<= if ~~ C = ~~ (y < x ?<= if C).

Lemma lteif_minr z x y C :
  (z < min x y ?<= if C) = (z < x ?<= if C) && (z < y ?<= if C).

Lemma lteif_minl z x y C :
  (min x y < z ?<= if C) = (x < z ?<= if C) || (y < z ?<= if C).

Lemma lteif_maxr z x y C :
  (z < max x y ?<= if C) = (z < x ?<= if C) || (z < y ?<= if C).

Lemma lteif_maxl z x y C :
  (max x y < z ?<= if C) = (x < z ?<= if C) && (y < z ?<= if C).

Section ArgExtremum.
Context (I : finType) (i0 : I) (P : {pred I}) (F : I T) (Pi0 : P i0).

Lemma arg_minP: extremum_spec <=%O P F (arg_min i0 P F).

Lemma arg_maxP: extremum_spec >=%O P F (arg_max i0 P F).

End ArgExtremum.

Lemma count_le_gt x s : count ( x) s = size s - count (> x) s.

Lemma count_lt_ge x s : count (< x) s = size s - count ( x) s.

Section bigminmax_Type.
Context (I : Type) (r : seq I) (x : T).
Implicit Types (P : pred I) (F : I T).

Lemma bigmin_mkcond P F : \big[min/x]_(i <- r | P i) F i =
  \big[min/x]_(i <- r) (if P i then F i else x).

Lemma bigmax_mkcond P F :
  \big[max/x]_(i <- r | P i) F i = \big[max/x]_(i <- r) if P i then F i else x.

Lemma bigmin_mkcondl P Q F :
  \big[min/x]_(i <- r | P i && Q i) F i
  = \big[min/x]_(i <- r | Q i) if P i then F i else x.

Lemma bigmin_mkcondr P Q F :
  \big[min/x]_(i <- r | P i && Q i) F i
  = \big[min/x]_(i <- r | P i) if Q i then F i else x.

Lemma bigmax_mkcondl P Q F :
  \big[max/x]_(i <- r | P i && Q i) F i
  = \big[max/x]_(i <- r | Q i) if P i then F i else x.

Lemma bigmax_mkcondr P Q F :
  \big[max/x]_(i <- r | P i && Q i) F i
  = \big[max/x]_(i <- r | P i) if Q i then F i else x.

Lemma bigmin_split P F1 F2 :
  \big[min/x]_(i <- r | P i) (min (F1 i) (F2 i)) =
    min (\big[min/x]_(i <- r | P i) F1 i) (\big[min/x]_(i <- r | P i) F2 i).

Lemma bigmax_split P F1 F2 :
  \big[max/x]_(i <- r | P i) (max (F1 i) (F2 i)) =
    max (\big[max/x]_(i <- r | P i) F1 i) (\big[max/x]_(i <- r | P i) F2 i).

Lemma bigmin_idl P F :
  \big[min/x]_(i <- r | P i) F i = min x (\big[min/x]_(i <- r | P i) F i).

Lemma bigmax_idl P F :
  \big[max/x]_(i <- r | P i) F i = max x (\big[max/x]_(i <- r | P i) F i).

Lemma bigmin_idr P F :
  \big[min/x]_(i <- r | P i) F i = min (\big[min/x]_(i <- r | P i) F i) x.

Lemma bigmax_idr P F :
  \big[max/x]_(i <- r | P i) F i = max (\big[max/x]_(i <- r | P i) F i) x.

Lemma bigminID a P F : \big[min/x]_(i <- r | P i) F i =
  min (\big[min/x]_(i <- r | P i && a i) F i)
      (\big[min/x]_(i <- r | P i && ~~ a i) F i).

Lemma bigmaxID a P F : \big[max/x]_(i <- r | P i) F i =
  max (\big[max/x]_(i <- r | P i && a i) F i)
      (\big[max/x]_(i <- r | P i && ~~ a i) F i).

End bigminmax_Type.

Let ge_min_id (x y : T) : x min x y.
Let le_max_id (x y : T) : x max x y.

Lemma sub_bigmin [x0] I r (P P' : {pred I}) (F : I T) :
    ( i, P' i P i)
  \big[min/x0]_(i <- r | P i) F i \big[min/x0]_(i <- r | P' i) F i.

Lemma sub_bigmax [x0] I r (P P' : {pred I}) (F : I T) :
    ( i, P i P' i)
  \big[max/x0]_(i <- r | P i) F i \big[max/x0]_(i <- r | P' i) F i.

FIXME: Remove that.
Local Notation "'{subset' x '<=' y '}'" :=
  (sub_mem (mem x) (mem y)) (at level 0, x, y at level 1).

Lemma sub_bigmin_seq [x0] (I : eqType) r r' P (F : I T) : {subset r' r}
  \big[min/x0]_(i <- r | P i) F i \big[min/x0]_(i <- r' | P i) F i.

Lemma sub_bigmax_seq [x0] (I : eqType) r r' P (F : I T) : {subset r r'}
  \big[max/x0]_(i <- r | P i) F i \big[max/x0]_(i <- r' | P i) F i.

Lemma sub_bigmin_cond [x0] (I : eqType) r r' P P' (F : I T) :
    {subset ([seq i <- r | P i]) ([seq i <- r' | P' i])}
  \big[min/x0]_(i <- r' | P' i) F i \big[min/x0]_(i <- r | P i) F i.

Lemma sub_bigmax_cond [x0] (I : eqType) r r' P P' (F : I T) :
    {subset ([seq i <- r | P i]) ([seq i <- r' | P' i])}
  \big[max/x0]_(i <- r | P i) F i \big[max/x0]_(i <- r' | P' i) F i.

Lemma sub_in_bigmin [x0] [I : eqType] (r : seq I) (P P' : {pred I}) F :
    {in r, i, P' i P i}
  \big[min/x0]_(i <- r | P i) F i \big[min/x0]_(i <- r | P' i) F i.

Lemma sub_in_bigmax [x0] [I : eqType] (r : seq I) (P P' : {pred I}) F :
    {in r, i, P i P' i}
  \big[max/x0]_(i <- r | P i) F i \big[max/x0]_(i <- r | P' i) F i.

Lemma le_bigmin_nat [x0] n m n' m' P (F : nat T) :
    (n n')%N (m' m)%N
  \big[min/x0]_(n i < m | P i) F i \big[min/x0]_(n' i < m' | P i) F i.

Lemma le_bigmax_nat [x0] n m n' m' P (F : nat T) :
    (n' n)%N (m m')%N
  \big[max/x0]_(n i < m | P i) F i \big[max/x0]_(n' i < m' | P i) F i.

Lemma le_bigmin_nat_cond [x0] n m n' m' (P P' : pred nat) (F : nat T) :
    (n n')%N (m' m)%N ( i, (n' i < m')%N P' i P i)
  \big[min/x0]_(n i < m | P i) F i \big[min/x0]_(n' i < m' | P' i) F i.

Lemma le_bigmax_nat_cond [x0] n m n' m' (P P' : {pred nat}) (F : nat T) :
    (n' n)%N (m m')%N ( i, (n i < m)%N P i P' i)
  \big[max/x0]_(n i < m | P i) F i \big[max/x0]_(n' i < m' | P' i) F i.

Lemma le_bigmin_ord [x0] n m (P : pred nat) (F : nat T) : (m n)%N
  \big[min/x0]_(i < n | P i) F i \big[min/x0]_(i < m | P i) F i.

Lemma le_bigmax_ord [x0] n m (P : {pred nat}) (F : nat T) : (n m)%N
  \big[max/x0]_(i < n | P i) F i \big[max/x0]_(i < m | P i) F i.

Lemma le_bigmin_ord_cond [x0] n m (P P' : pred nat) (F : nat T) :
    (m n)%N ( i : 'I_m, P' i P i)
  \big[min/x0]_(i < n | P i) F i \big[min/x0]_(i < m | P' i) F i.

Lemma le_bigmax_ord_cond [x0] n m (P P' : {pred nat}) (F : nat T) :
    (n m)%N ( i : 'I_n, P i P' i)
  \big[max/x0]_(i < n | P i) F i \big[max/x0]_(i < m | P' i) F i.

Lemma subset_bigmin [x0] [I : finType] [A A' P : {pred I}] (F : I T) :
    A' \subset A
  \big[min/x0]_(i in A | P i) F i \big[min/x0]_(i in A' | P i) F i.

Lemma subset_bigmax [x0] [I : finType] (A A' P : {pred I}) (F : I T) :
    A \subset A'
  \big[max/x0]_(i in A | P i) F i \big[max/x0]_(i in A' | P i) F i.

Lemma subset_bigmin_cond [x0] (I : finType) (A A' P P' : {pred I}) (F : I T) :
    [set i in A' | P' i] \subset [set i in A | P i]
  \big[min/x0]_(i in A | P i) F i \big[min/x0]_(i in A' | P' i) F i.

Lemma subset_bigmax_cond [x0] (I : finType) (A A' P P' : {pred I}) (F : I T) :
    [set i in A | P i] \subset [set i in A' | P' i]
  \big[max/x0]_(i in A | P i) F i \big[max/x0]_(i in A' | P' i) F i.

Section bigminmax_Type.
Context (I : Type) (r : seq I) (x : T).
Implicit Types (P : pred I) (F : I T).

Lemma bigmin_le_id P F : \big[min/x]_(i <- r | P i) F i x.

Lemma bigmax_ge_id P F : \big[max/x]_(i <- r | P i) F i x.

Lemma bigmin_eq_id P F :
  ( i, P i x F i) \big[min/x]_(i <- r | P i) F i = x.

Lemma bigmax_eq_id P F :
  ( i, P i x F i) \big[max/x]_(i <- r | P i) F i = x.

End bigminmax_Type.

Section bigminmax_eqType.
Context (I : eqType) (r : seq I) (x : T).
Implicit Types (P : pred I) (F : I T).

Lemma ge_bigmin_seq i0 P F :
  i0 \in r P i0 \big[min/x]_(i <- r | P i) F i F i0.

Lemma le_bigmax_seq i0 P F :
  i0 \in r P i0 F i0 \big[max/x]_(i <- r | P i) F i.

Lemma bigmin_inf_seq i0 P t F :
  i0 \in r P i0 F i0 t \big[min/x]_(i <- r | P i) F i t.

Lemma bigmax_sup_seq i0 P t F :
  i0 \in r P i0 t F i0 t \big[max/x]_(i <- r | P i) F i.

End bigminmax_eqType.

Section bigminmax_finType.
Context (I : finType) (x : T).
Implicit Types (P : pred I) (F : I T).

Lemma bigminD1 j P F : P j
  \big[min/x]_(i | P i) F i = min (F j) (\big[min/x]_(i | P i && (i != j)) F i).

Lemma bigmaxD1 j P F : P j
  \big[max/x]_(i | P i) F i = max (F j) (\big[max/x]_(i | P i && (i != j)) F i).

Lemma bigmin_le_cond j P F : P j \big[min/x]_(i | P i) F i F j.

Lemma le_bigmax_cond j P F : P j F j \big[max/x]_(i | P i) F i.

Lemma bigmin_le j F : \big[min/x]_i F i F j.

Lemma le_bigmax F j : F j \big[max/x]_i F i.

Lemma bigmin_inf j P m F : P j F j m \big[min/x]_(i | P i) F i m.

NB: as of [2022-08-02], bigop.bigmax_sup already exists for nat
Lemma bigmax_sup j P m F : P j m F j m \big[max/x]_(i | P i) F i.

Lemma bigmin_geP m P F :
  reflect (m x i, P i m F i)
          (m \big[min/x]_(i | P i) F i).

Lemma bigmax_leP m P F :
  reflect (x m i, P i F i m)
          (\big[max/x]_(i | P i) F i m).

Lemma bigmin_gtP m P F :
  reflect (m < x i, P i m < F i) (m < \big[min/x]_(i | P i) F i).

Lemma bigmax_ltP m P F :
  reflect (x < m i, P i F i < m) (\big[max/x]_(i | P i) F i < m).

Lemma bigmin_eq_arg j P F : P j ( i, P i F i x)
  \big[min/x]_(i | P i) F i = F [arg min_(i < j | P i) F i].

Lemma bigmax_eq_arg j P F : P j ( i, P i x F i)
  \big[max/x]_(i | P i) F i = F [arg max_(i > j | P i) F i].

Lemma eq_bigmin j P F : P j ( i, P i F i x)
  {i0 | i0 \in P & \big[min/x]_(i | P i) F i = F i0}.

Lemma eq_bigmax j P F : P j ( i, P i x F i)
  {i0 | i0 \in P & \big[max/x]_(i | P i) F i = F i0}.

Lemma le_bigmin2 P F1 F2 : ( i, P i F1 i F2 i)
  \big[min/x]_(i | P i) F1 i \big[min/x]_(i | P i) F2 i.

Lemma le_bigmax2 P F1 F2 : ( i, P i F1 i F2 i)
  \big[max/x]_(i | P i) F1 i \big[max/x]_(i | P i) F2 i.

Lemma bigmaxUl (A B : {set I}) F :
  \big[max/x]_(i in A) F i \big[max/x]_(i in A :|: B) F i.

Lemma bigmaxUr (A B : {set I}) F :
  \big[max/x]_(i in B) F i \big[max/x]_(i in A :|: B) F i.

Lemma bigminUl (A B : {set I}) F :
  \big[min/x]_(i in A) F i \big[min/x]_(i in A :|: B) F i.

Lemma bigminUr (A B : {set I}) F :
  \big[min/x]_(i in B) F i \big[min/x]_(i in A :|: B) F i.

Lemma bigmaxIl (A B : {set I}) F :
  \big[max/x]_(i in A) F i \big[max/x]_(i in A :&: B) F i.

Lemma bigmaxIr (A B : {set I}) F :
  \big[max/x]_(i in B) F i \big[max/x]_(i in A :&: B) F i.

Lemma bigminIl (A B : {set I}) F :
  \big[min/x]_(i in A) F i \big[min/x]_(i in A :&: B) F i.

Lemma bigminIr (A B : {set I}) F :
  \big[min/x]_(i in B) F i \big[min/x]_(i in A :&: B) F i.

Lemma bigmaxD (A B : {set I}) F :
  \big[max/x]_(i in B) F i \big[max/x]_(i in B :\: A) F i.

Lemma bigminD (A B : {set I}) F :
  \big[min/x]_(i in B) F i \big[min/x]_(i in B :\: A) F i.

Lemma bigmaxU (A B : {set I}) F :
  \big[max/x]_(i in A :|: B) F i
  = max (\big[max/x]_(i in A) F i) (\big[max/x]_(i in B) F i).

Lemma bigminU (A B : {set I}) F :
  \big[min/x]_(i in A :|: B) F i
  = min (\big[min/x]_(i in A) F i) (\big[min/x]_(i in B) F i).

Lemma bigmin_set1 j F : \big[min/x]_(i in [set j]) F i = min (F j) x.

Lemma bigmax_set1 j F : \big[max/x]_(i in [set j]) F i = max (F j) x.

End bigminmax_finType.

Lemma bigmin_imset [I J : finType] x [h : I J] [A : {set I}] (F : J T) :
  \big[min/x]_(j in [set h x | x in A]) F j = \big[min/x]_(i in A) F (h i).

Lemma bigmax_imset [I J : finType] x [h : I J] [A : {set I}] (F : J T) :
  \big[max/x]_(j in [set h x | x in A]) F j = \big[max/x]_(i in A) F (h i).

End TotalTheory.

#[global] Hint Resolve le_total : core.
#[global] Hint Resolve ge_total : core.
#[global] Hint Extern 0 (is_true (_ >=< _)%O) ⇒ solve [apply: comparableT]
  : core.
#[global] Hint Resolve sort_le_sorted : core.

Arguments min_idPr {disp T x y}.
Arguments max_idPl {disp T x y}.
Arguments bigmin_mkcond {disp T I r}.
Arguments bigmax_mkcond {disp T I r}.
Arguments bigminID {disp T I r}.
Arguments bigmaxID {disp T I r}.
Arguments bigminD1 {disp T I x} j.
Arguments bigmaxD1 {disp T I x} j.
Arguments bigmin_inf {disp T I x} j.
Arguments bigmax_sup {disp T I x} j.
Arguments bigmin_eq_arg {disp T I} x j.
Arguments bigmax_eq_arg {disp T I} x j.
Arguments eq_bigmin {disp T I x} j.
Arguments eq_bigmax {disp T I x} j.
Arguments ge_bigmin_seq {disp T I r} x i0 P.
Arguments le_bigmax_seq {disp T I r} x i0 P.
Arguments bigmin_inf_seq {disp T I r} x i0 P.
Arguments bigmax_sup_seq {disp T I r} x i0 P.

FIXME: some lemmas in the following section should hold for any porderType
Module Import DualTotalTheory.
Section DualTotalTheory.
Context {disp : disp_t} {T : orderType disp}.
Implicit Type s : seq T.

Lemma sorted_filter_gt x s :
  sorted <=%O s [seq y <- s | x < y] = drop (count ( x) s) s.

Lemma sorted_filter_ge x s :
  sorted <=%O s [seq y <- s | x y] = drop (count (< x) s) s.

Lemma nth_count_ge x x0 s i : sorted <=%O s
  (count (< x) s i < size s)%N x nth x0 s i.

Lemma nth_count_gt x x0 s i : sorted <=%O s
  (count ( x) s i < size s)%N x < nth x0 s i.

Lemma nth_count_eq x x0 s i : sorted <=%O s
  (count (< x) s i < count ( x) s)%N nth x0 s i = x.

End DualTotalTheory.
End DualTotalTheory.

contra lemmas

Section ContraTheory.
Context {disp1 disp2 : disp_t} {T1 : porderType disp1} {T2 : orderType disp2}.
Implicit Types (x y : T1) (z t : T2) (b : bool) (m n : nat) (P : Prop).

Lemma contraTle b z t : (t < z ~~ b) (b z t).

Lemma contraTlt b z t : (t z ~~ b) (b z < t).

Lemma contraPle P z t : (t < z ¬ P) (P z t).

Lemma contraPlt P z t : (t z ¬ P) (P z < t).

Lemma contraNle b z t : (t < z b) (~~ b z t).

Lemma contraNlt b z t : (t z b) (~~ b z < t).

Lemma contra_not_le P z t : (t < z P) (¬ P z t).

Lemma contra_not_lt P z t : (t z P) (¬ P z < t).

Lemma contraFle b z t : (t < z b) (b = false z t).

Lemma contraFlt b z t : (t z b) (b = false z < t).

Lemma contra_leq_le m n z t : (t < z (n < m)%N) ((m n)%N z t).

Lemma contra_leq_lt m n z t : (t z (n < m)%N) ((m n)%N z < t).

Lemma contra_ltn_le m n z t : (t < z (n m)%N) ((m < n)%N z t).

Lemma contra_ltn_lt m n z t : (t z (n m)%N) ((m < n)%N z < t).

Lemma contra_le x y z t : (t < z y < x) (x y z t).

Lemma contra_le_lt x y z t : (t z y < x) (x y z < t).

Lemma contra_lt_le x y z t : (t < z y x) (x < y z t).

Lemma contra_lt x y z t : (t z y x) (x < y z < t).

End ContraTheory.

Section TotalMonotonyTheory.
Context {disp disp' : disp_t} {T : orderType disp} {T' : porderType disp'}.
Context (D : {pred T}) (f : T T').
Implicit Types (x y z : T) (u v w : T').

Let leT_anti := @le_anti _ T.
Let leT'_anti := @le_anti _ T'.
Let ltT_neqAle := @lt_neqAle _ T.
Let ltT'_neqAle := @lt_neqAle _ T'.
Let ltT_def := @lt_def _ T.
Let leT_total := @le_total _ T.

Lemma le_mono : {homo f : x y / x < y} {mono f : x y / x y}.

Lemma le_nmono : {homo f : x y /~ x < y} {mono f : x y /~ x y}.

Lemma le_mono_in :
  {in D &, {homo f : x y / x < y}} {in D &, {mono f : x y / x y}}.

Lemma le_nmono_in :
  {in D &, {homo f : x y /~ x < y}} {in D &, {mono f : x y /~ x y}}.

End TotalMonotonyTheory.

End TotalTheory.

Module Import CDistrLatticeTheory.
Section CDistrLatticeTheory.
Context {disp : disp_t} {L : cDistrLatticeType disp}.
Implicit Types (x y z : L).

Lemma rcomplPmeet x y z : ((x `&` y) `|` z) `&` rcompl x y z = x `&` y.

Lemma rcomplPjoin x y z : ((y `|` x) `&` z) `|` rcompl x y z = y `|` x.

Lemma rcomplKI x y z : x y (x `|` z) `&` rcompl x y z = x.

Lemma rcomplKU x y z : x y (y `&` z) `|` rcompl x y z = y.

End CDistrLatticeTheory.
End CDistrLatticeTheory.

Module Import CBDistrLatticeTheory.
Section CBDistrLatticeTheory.
Context {disp : disp_t} {L : cbDistrLatticeType disp}.
Implicit Types (x y z : L).

Lemma diffErcompl x y : x `\` y = rcompl \bot x y.

Lemma diffKI x y : y `&` (x `\` y) = \bot.

Lemma diffIK x y : (x `\` y) `&` y = \bot.

Lemma meetIB z x y : (z `&` y) `&` (x `\` y) = \bot.

Lemma meetBI z x y : (x `\` y) `&` (z `&` y) = \bot.

Lemma joinIB y x : (x `&` y) `|` (x `\` y) = x.

Lemma joinBI y x : (x `\` y) `|` (x `&` y) = x.

Lemma joinIBC y x : (y `&` x) `|` (x `\` y) = x.

Lemma joinBIC y x : (x `\` y) `|` (y `&` x) = x.

Lemma leBx x y : x `\` y x.
Hint Resolve leBx : core.

Lemma diffxx x : x `\` x = \bot.

Lemma leBl z x y : x y x `\` z y `\` z.

Lemma diffKU y x : y `|` (x `\` y) = y `|` x.

Lemma diffUK y x : (x `\` y) `|` y = x `|` y.

Lemma leBKU y x : y x y `|` (x `\` y) = x.

Lemma leBUK y x : y x (x `\` y) `|` y = x.

Lemma leBLR x y z : (x `\` y z) = (x y `|` z).

Lemma diffUx x y z : (x `|` y) `\` z = (x `\` z) `|` (y `\` z).

Lemma diff_eq0 x y : (x `\` y == \bot) = (x y).

Lemma joinxB x y z : x `|` (y `\` z) = ((x `|` y) `\` z) `|` (x `&` z).

Lemma joinBx x y z : (y `\` z) `|` x = ((y `|` x) `\` z) `|` (z `&` x).

Lemma leBr z x y : x y z `\` y z `\` x.

Lemma leB2 x y z t : x z t y x `\` y z `\` t.

Lemma meet_eq0E_diff z x y : x z (x `&` y == \bot) = (x z `\` y).

Lemma leBRL x y z : (x z `\` y) = (x z) && (x `&` y == \bot).

Lemma eq_diff x y z : (x `\` y == z) = (z x y `|` z) && (z `&` y == \bot).

Lemma diffxU x y z : z `\` (x `|` y) = (z `\` x) `&` (z `\` y).

Lemma diffx0 x : x `\` \bot = x.

Lemma diff0x x : \bot `\` x = \bot.

Lemma diffIx x y z : (x `&` y) `\` z = (x `\` z) `&` (y `\` z).

Lemma meetxB x y z : x `&` (y `\` z) = (x `&` y) `\` z.

Lemma meetBx x y z : (x `\` y) `&` z = (x `&` z) `\` y.

Lemma diffxI x y z : x `\` (y `&` z) = (x `\` y) `|` (x `\` z).

Lemma diffBx x y z : (x `\` y) `\` z = x `\` (y `|` z).

Lemma diffxB x y z : x `\` (y `\` z) = (x `\` y) `|` (x `&` z).

Lemma joinBK x y : (y `|` x) `\` x = (y `\` x).

Lemma joinBKC x y : (x `|` y) `\` x = (y `\` x).

Lemma disj_le x y : x `&` y == \bot x y = (x == \bot).

Lemma disj_leC x y : y `&` x == \bot x y = (x == \bot).

Lemma disj_diffl x y : x `&` y == \bot x `\` y = x.

Lemma disj_diffr x y : x `&` y == \bot y `\` x = y.

Lemma lt0B x y : x < y \bot < y `\` x.

End CBDistrLatticeTheory.
End CBDistrLatticeTheory.

Module Import CTDistrLatticeTheory.
Section CTDistrLatticeTheory.
Context {disp : disp_t} {L : ctDistrLatticeType disp}.
Implicit Types (x y z : L).

Lemma codiffErcompl x y : codiff x y = rcompl x \top y.

TODO: complete this theory module

End CTDistrLatticeTheory.
End CTDistrLatticeTheory.

Module Import CTBDistrLatticeTheory.
Section CTBDistrLatticeTheory.
Context {disp : disp_t} {L : ctbDistrLatticeType disp}.
Implicit Types (x y z : L).

Lemma complEdiff x : ~` x = \top `\` x.
#[deprecated(since="mathcomp 2.3.0", note="Use complEdiff instead.")]
Notation complE := complEdiff.

Lemma complEcodiff x : ~` x = codiff \bot x.

Lemma complErcompl x : ~` x = rcompl \bot \top x.

Lemma diff1x x : \top `\` x = ~` x.

Lemma diffE x y : x `\` y = x `&` ~` y.

Lemma complK : involutive (@compl _ L).

Lemma compl_inj : injective (@compl _ L).

Lemma disj_leC x y : (x `&` y == \bot) = (x ~` y).

Lemma leCx x y : (~` x y) = (~` y x).

Lemma lexC x y : (x ~` y) = (y ~` x).

Lemma leC x y : (~` x ~` y) = (y x).

Lemma complU x y : ~` (x `|` y) = ~` x `&` ~` y.

Lemma complI x y : ~` (x `&` y) = ~` x `|` ~` y.

Lemma joinxC x : x `|` ~` x = \top.

Lemma joinCx x : ~` x `|` x = \top.

Lemma meetxC x : x `&` ~` x = \bot.

Lemma meetCx x : ~` x `&` x = \bot.

Lemma compl1 : ~` \top = \bot :> L.

Lemma compl0 : ~` \bot = \top :> L.

Lemma complB x y : ~` (x `\` y) = ~` x `|` y.

Lemma leBC x y : x `\` y ~` y.

Lemma compl_joins (J : Type) (r : seq J) (P : {pred J}) (F : J L) :
   ~` (\join_(j <- r | P j) F j) = \meet_(j <- r | P j) ~` F j.

Lemma compl_meets (J : Type) (r : seq J) (P : {pred J}) (F : J L) :
   ~` (\meet_(j <- r | P j) F j) = \join_(j <- r | P j) ~` F j.

End CTBDistrLatticeTheory.
End CTBDistrLatticeTheory.

FACTORIES porderType



Let ge_anti : antisymmetric (fun x y ⇒ @le d T y x).





Let lt_le_def x y : lt x y = le x y && ~~ le y x.




TODO: print nice error message when keyed type is not provided
meetSemilatticeType and joinSemilatticeType
latticeType
distrLatticeType
complemented lattices


Module hasRelativeComplement.
#[deprecated(since="mathcomp 2.3.0",
             note="Use BDistrLattice_hasSectionalComplement.Build instead.")]
Notation Build d T :=
  (BDistrLattice_hasSectionalComplement.Build d T) (only parsing).
End hasRelativeComplement.

#[deprecated(since="mathcomp 2.3.0",
             note="Use BDistrLattice_hasSectionalComplement instead.")]
Notation hasRelativeComplement d T :=
  (BDistrLattice_hasSectionalComplement d T) (only parsing).


Definition rcompl x y z := (x `&` y) `|` diff (y `|` x) z.

Fact rcomplPmeet x y z : ((x `&` y) `|` z) `&` rcompl x y z = x `&` y.

Fact rcomplPjoin x y z : ((y `|` x) `&` z) `|` rcompl x y z = y `|` x.


Fact diffErcompl x y : diff x y = rcompl \bot x y.





Definition rcompl x y z := (y `|` x) `&` codiff (x `&` y) z.

Fact rcomplPmeet x y z : ((x `&` y) `|` z) `&` rcompl x y z = x `&` y.

Fact rcomplPjoin x y z : ((y `|` x) `&` z) `|` rcompl x y z = y `|` x.


Fact codiffErcompl x y : codiff x y = rcompl x \top y.




Module hasComplement.
#[deprecated(since="mathcomp 2.3.0",
             note="Use CBDistrLattice_hasComplement.Build instead.")]
Notation Build d T := (CBDistrLattice_hasComplement.Build d T) (only parsing).
End hasComplement.

#[deprecated(since="mathcomp 2.3.0",
             note="Use CBDistrLattice_hasComplement instead.")]
Notation hasComplement d T := (CBDistrLattice_hasComplement d T) (only parsing).



Fact complEcodiff (x : T) : compl x = codiff (\bot : T) x.






Fact complEdiff (x : T) : compl x = (\top : T) `\` x.





Definition diff x y := x `&` compl y.
Definition codiff x y := x `|` compl y.
Definition rcompl x y z := (x `&` y) `|` diff (y `|` x) z.

Fact diffKI x y : y `&` diff x y = \bot.

Fact joinIB x y : (x `&` y) `|` diff x y = x.


Fact codiffErcompl x y : codiff x y = rcompl x \top y.


Fact complEdiff x : compl x = diff \top x.
Fact complEcodiff x : compl x = codiff \bot x.



orderType



Fact meetUl : @left_distributive T T meet join.





Implicit Types (x y z : T).

Let comparableT x y : x >=< y := le_total x y.

Fact ltgtP x y :
  compare x y (min y x) (min x y) (max y x) (max x y)
              (y == x) (x == y) (x y) (x y) (x > y) (x < y).

Fact leP x y : le_xor_gt x y
  (min y x) (min x y) (max y x) (max x y) (x y) (y < x).

Definition meet := @min _ T.
Definition join := @max _ T.

Fact meetC : commutative meet.

Fact joinC : commutative join.

Fact meetA : associative meet.

Fact joinA : associative join.

Fact joinKI y x : meet x (join x y) = x.

Fact meetKU y x : join x (meet x y) = x.

Fact leEmeet x y : (x y) = (meet x y == x).




Fact le_refl : reflexive le.

Fact lt_le_def x y : lt x y = (le x y) && ~~ (le y x).


Section GeneratedOrder.

Local Definition T' := T.
Implicit Types (x y z : T').

Fact meetE x y : meet x y = x `&` y.
Fact joinE x y : join x y = x `|` y.
Fact meetC : commutative meet.
Fact joinC : commutative join.
Fact meetA : associative meet.
Fact joinA : associative join.
Fact joinKI y x : meet x (join x y) = x.
Fact meetKU y x : join x (meet x y) = x.
Fact meetUl : left_distributive meet join.
Fact meetxx : idempotent_op meet.
Fact le_def x y : x y = (meet x y == x).

End GeneratedOrder.





Fact lt_def x y : lt x y = (y != x) && le x y.

Fact meet_def_le x y : meet x y = if lt x y then x else y.

Fact join_def_le x y : join x y = if lt x y then y else x.

Fact le_anti : antisymmetric le.

Fact le_trans : transitive le.

Fact le_total : total le.


Fact totalT : total (<=%O : rel T).

Module CancelPartial.
Import PreCancelPartial.
Section CancelPartial.
Variables (disp : disp_t) (T : choiceType).
Variables (disp' : disp_t) (T' : porderType disp') (f : T T').

Section Pcan.
Variables (f' : T' option T) (f_can : pcancel f f').

Fact anti : antisymmetric (le f).

Fact lt_def x y :
  lt f x y = (y != x) && le f x y.

Definition Pcan := isPOrder.Build disp (Choice.Pack (Choice.class T))
  lt_def (@refl T disp' T' f) anti (@trans T disp' T' f).

End Pcan.

Definition Can f' (f_can : cancel f f') := Pcan (can_pcan f_can).

End CancelPartial.
End CancelPartial.

Notation PCanIsPartial := CancelPartial.Pcan.
Notation CanIsPartial := CancelPartial.Can.

#[export]
HB.instance Definition _ (disp : disp_t) (T : choiceType)
  (disp' : disp_t) (T' : porderType disp') (f : T T')
  (f' : T' option T) (f_can : pcancel f f') :=
  Preorder_isPOrder.Build disp (pcan_type f_can) (CancelPartial.anti f_can).

#[export]
HB.instance Definition _ (disp : disp_t) (T : choiceType)
  (disp' : disp_t) (T' : porderType disp') (f : T T') (f' : T' T)
  (f_can : cancel f f') :=
  Preorder_isPOrder.Build disp (can_type f_can)
    (CancelPartial.anti (can_pcan f_can)).

Section CancelTotal.
Variables (disp : disp_t) (T : choiceType).
Variables (disp' : disp_t) (T' : orderType disp') (f : T T').

Section PCan.

Variables (f' : T' option T) (f_can : pcancel f f').

#[local]
HB.instance Definition _ :=
   MonoTotal.Build disp (pcan_type f_can) (fun _ _ erefl).

Definition PCanIsTotal : DistrLattice_isTotal _ (pcan_type f_can) :=
  Total.on (pcan_type f_can).

End PCan.

Section Can.

Variables (f' : T' T) (f_can : cancel f f').

#[local]
HB.instance Definition _ :=
   MonoTotal.Build disp (can_type f_can) (fun _ _ erefl).

Definition CanIsTotal : DistrLattice_isTotal _ (can_type f_can) :=
  Total.on (can_type f_can).

End Can.
End CancelTotal.



Definition meet (x y : T) := f' (meet (f x) (f y)).
Definition join (x y : T) := f' (join (f x) (f y)).

Fact meetC : commutative meet.
Fact joinC : commutative join.
Fact meetA : associative meet.
Fact joinA : associative join.
Fact joinKI y x : meet x (join x y) = x.
Fact meetKI y x : join x (meet x y) = x.
Fact meet_eql x y : (x y) = (meet x y == x).






Fact meetUl : left_distributive (meet : T T T) join.



Morphism hierarchy.

Export OrderMorphismTheory.

Lemma omorph_lt (d : disp_t) (T : porderType d) (d' : disp_t) (T' : porderType d')
(f : {omorphism T T'}) : injective f {homo f : x y / x < y}.

Definition meet_morphism d (T : latticeType d) d' (T' : latticeType d')
  (f : T T') : Prop := {morph f : x y / x `&` y}.

Definition join_morphism d (T : latticeType d) d' (T' : latticeType d')
  (f : T T') : Prop := {morph f : x y / x `|` y}.








Module LatticeMorphismExports.
Notation "{ 'mlmorphism' T -> T' }" :=
  (@MeetLatticeMorphism.type _ T%type _ T'%type) : type_scope.
Notation "{ 'jlmorphism' T -> T' }" :=
  (@JoinLatticeMorphism.type _ T%type _ T'%type) : type_scope.
Notation "{ 'lmorphism' T -> T' }" :=
  (@LatticeMorphism.type _ T%type _ T'%type) : type_scope.
Notation "[ 'mlmorphism' 'of' f 'as' g ]" :=
  (MeetLatticeMorphism.clone _ _ _ _ f%function g)
  (format "[ 'mlmorphism' 'of' f 'as' g ]") : form_scope.
Notation "[ 'mlmorphism' 'of' f ]" :=
  (MeetLatticeMorphism.clone _ _ _ _ f%function _)
  (format "[ 'mlmorphism' 'of' f ]") : form_scope.
Notation "[ 'jlmorphism' 'of' f 'as' g ]" :=
  (JoinLatticeMorphism.clone _ _ _ _ f%function g)
  (format "[ 'jlmorphism' 'of' f 'as' g ]") : form_scope.
Notation "[ 'jlmorphism' 'of' f ]" :=
  (JoinLatticeMorphism.clone _ _ _ _ f%function _)
  (format "[ 'jlmorphism' 'of' f ]") : form_scope.
Notation "[ 'lmorphism' 'of' f 'as' g ]" :=
  (LatticeMorphism.clone _ _ _ _ f%function g)
  (format "[ 'lmorphism' 'of' f 'as' g ]") : form_scope.
Notation "[ 'lmorphism' 'of' f ]" :=
  (LatticeMorphism.clone _ _ _ _ f%function _)
  (format "[ 'lmorphism' 'of' f ]") : form_scope.
End LatticeMorphismExports.

Module Import LatticeMorphismTheory.
Section LatticeMorphismTheory.

Section Properties.

Variables (d : disp_t) (T : latticeType d) (d' : disp_t) (T' : latticeType d').

Lemma omorphI (f : {mlmorphism T T'}) : {morph f : x y / x `&` y}.

Lemma omorphU (f : {jlmorphism T T'}) : {morph f : x y / x `|` y}.

End Properties.

Section IdCompFun.

Variables (d : disp_t) (T : latticeType d) (d' : disp_t) (T' : latticeType d').
Variables (d'' : disp_t) (T'' : latticeType d'').

Section MeetCompFun.

Variables (f : {mlmorphism T' T''}) (g : {mlmorphism T T'}).

Fact idfun_is_meet_morphism : meet_morphism (@idfun T).
#[export]
HB.instance Definition _ := isMeetLatticeMorphism.Build d T d T idfun
  idfun_is_meet_morphism.

Fact comp_is_meet_morphism : meet_morphism (f \o g).
#[export]
HB.instance Definition _ := isMeetLatticeMorphism.Build d T d'' T'' (f \o g)
  comp_is_meet_morphism.

End MeetCompFun.

Section JoinCompFun.

Variables (f : {jlmorphism T' T''}) (g : {jlmorphism T T'}).

Fact idfun_is_join_morphism : join_morphism (@idfun T).
#[export]
HB.instance Definition _ := isJoinLatticeMorphism.Build d T d T idfun
  idfun_is_join_morphism.

Fact comp_is_join_morphism : join_morphism (f \o g).
#[export]
HB.instance Definition _ := isJoinLatticeMorphism.Build d T d'' T'' (f \o g)
  comp_is_join_morphism.

End JoinCompFun.

End IdCompFun.

End LatticeMorphismTheory.
End LatticeMorphismTheory.






Module TBLatticeMorphismExports.
Notation "{ 'blmorphism' T -> T' }" :=
  (@BLatticeMorphism.type _ T%type _ T'%type) : type_scope.
Notation "{ 'tlmorphism' T -> T' }" :=
  (@TLatticeMorphism.type _ T%type _ T'%type) : type_scope.
Notation "{ 'tblmorphism' T -> T' }" :=
  (@TBLatticeMorphism.type _ T%type _ T'%type) : type_scope.
End TBLatticeMorphismExports.

Module Import BLatticeMorphismTheory.
Section BLatticeMorphismTheory.

Section Properties.

Variables (d : disp_t) (T : bLatticeType d).
Variables (d' : disp_t) (T' : bLatticeType d').
Variables (f : {blmorphism T T'}).

Lemma omorph0 : f \bot = \bot.

End Properties.

Section IdCompFun.

Variables (d : disp_t) (T : bLatticeType d).
Variables (d' : disp_t) (T' : bLatticeType d').
Variables (d'' : disp_t) (T'' : bLatticeType d'').
Variables (f : {blmorphism T' T''}) (g : {blmorphism T T'}).

Fact idfun_is_bottom_morphism : (@idfun T) \bot = \bot.
#[export]
HB.instance Definition _ := isBLatticeMorphism.Build d T d T idfun
  idfun_is_bottom_morphism.

Fact comp_is_bottom_morphism : (f \o g) \bot = \bot.
#[export]
HB.instance Definition _ := isBLatticeMorphism.Build d T d'' T'' (f \o g)
  comp_is_bottom_morphism.

End IdCompFun.

End BLatticeMorphismTheory.
End BLatticeMorphismTheory.

Module Import TLatticeMorphismTheory.
Section TLatticeMorphismTheory.

Section Properties.

Variables (d : disp_t) (T : tLatticeType d).
Variables (d' : disp_t) (T' : tLatticeType d').
Variables (f : {tlmorphism T T'}).

Lemma omorph1 : f \top = \top.

End Properties.

Section IdCompFun.

Variables (d : disp_t) (T : tLatticeType d).
Variables (d' : disp_t) (T' : tLatticeType d').
Variables (d'' : disp_t) (T'' : tLatticeType d'').
Variables (f : {tlmorphism T' T''}) (g : {tlmorphism T T'}).

Fact idfun_is_top_morphism : (@idfun T) \top = \top.
#[export]
HB.instance Definition _ := isTLatticeMorphism.Build d T d T idfun
  idfun_is_top_morphism.

Fact comp_is_top_morphism : (f \o g) \top = \top.
#[export]
HB.instance Definition _ := isTLatticeMorphism.Build d T d'' T'' (f \o g)
  comp_is_top_morphism.

End IdCompFun.

End TLatticeMorphismTheory.
End TLatticeMorphismTheory.

Module Import ClosedPredicates.
Section ClosedPredicates.

Variable (d : disp_t) (T : latticeType d).
Variable S : {pred T}.

Definition meet_closed := {in S &, u v, u `&` v \in S}.

Definition join_closed := {in S &, u v, u `|` v \in S}.

End ClosedPredicates.
End ClosedPredicates.

Mixins for stability properties





Structures for stability properties

#[short(type="meetLatticeClosed")]
HB.structure Definition MeetLatticeClosed d T :=
  {S of isMeetLatticeClosed d T S}.

#[short(type="joinLatticeClosed")]
HB.structure Definition JoinLatticeClosed d T :=
  {S of isJoinLatticeClosed d T S}.

#[short(type="latticeClosed")]
HB.structure Definition LatticeClosed d T :=
  {S of @MeetLatticeClosed d T S & @JoinLatticeClosed d T S}.

#[short(type="bLatticeClosed")]
HB.structure Definition BLatticeClosed d T := {S of isBLatticeClosed d T S}.

#[short(type="bJoinLatticeClosed")]
HB.structure Definition BJoinLatticeClosed d T :=
  {S of isBLatticeClosed d T S & @JoinLatticeClosed d T S}.

#[short(type="tLatticeClosed")]
HB.structure Definition TLatticeClosed d T := {S of isTLatticeClosed d T S}.

#[short(type="tMeetLatticeClosed")]
HB.structure Definition TMeetLatticeClosed d T :=
  {S of isTLatticeClosed d T S & @MeetLatticeClosed d T S}.

#[short(type="tbLatticeClosed")]
HB.structure Definition TBLatticeClosed d (T : tbLatticeType d) :=
  {S of @BLatticeClosed d T S & @TLatticeClosed d T S}.





Module Import LatticePred.
Section LatticePred.

Variables (d : disp_t) (T : latticeType d).

Lemma opredI (S : meetLatticeClosed T) : {in S &, u v, u `&` v \in S}.

Lemma opredU (S : joinLatticeClosed T) : {in S &, u v, u `|` v \in S}.

End LatticePred.

Section BLatticePred.

Variables (d : disp_t) (T : bLatticeType d).

Lemma opred0 (S : bLatticeClosed T) : \bot \in S.

Lemma opred_joins (S : bJoinLatticeClosed T) I r (P : pred I) F :
  ( i, P i F i \in S) \join_(i <- r | P i) F i \in S.

End BLatticePred.

Section TLatticePred.

Variables (d : disp_t) (T : tLatticeType d).

Lemma opred1 (S : tLatticeClosed T) : \top \in S.

Lemma opred_meets (S : tMeetLatticeClosed T) I r (P : pred I) F :
  ( i, P i F i \in S) \meet_(i <- r | P i) F i \in S.

End TLatticePred.
End LatticePred.

#[short(type="subPOrder")]
HB.structure Definition SubPOrder d (T : porderType d) S d' :=
  { U of SubEquality T S U & POrder d' U & isSubPreorder d T S d' U }.



#[export]
HB.instance Definition _ d (T : porderType d) (S : pred T) (d' : disp_t)
  (U : subType S) := SubChoice_isSubPOrder.Build d T S d' (sub_type U).



#[short(type="subPOrderLattice")]
HB.structure Definition SubPOrderLattice d (T : latticeType d) S d' :=
  { U of @SubPOrder d T S d' U & Lattice d' U }.

#[short(type="subPOrderBLattice")]
HB.structure Definition SubPOrderBLattice d (T : latticeType d) S d' :=
  { U of @SubPOrderLattice d T S d' U & BLattice d' U }.

#[short(type="subPOrderTLattice")]
HB.structure Definition SubPOrderTLattice d (T : latticeType d) S d' :=
  { U of @SubPOrderLattice d T S d' U & TLattice d' U }.

#[short(type="subPOrderTBLattice")]
HB.structure Definition SubPOrderTBLattice d (T : latticeType d) S d' :=
  { U of @SubPOrderLattice d T S d' U & TBLattice d' U }.

#[short(type="meetSubLattice")]
HB.structure Definition MeetSubLattice d (T : latticeType d) S d' :=
  { U of @SubPOrderLattice d T S d' U & isMeetSubLattice d T S d' U }.

#[short(type="meetSubBLattice")]
HB.structure Definition MeetSubBLattice d (T : latticeType d) S d' :=
  { U of @MeetSubLattice d T S d' U & BLattice d' U }.

#[short(type="meetSubTLattice")]
HB.structure Definition MeetSubTLattice d (T : latticeType d) S d' :=
  { U of @MeetSubLattice d T S d' U & TLattice d' U }.

#[short(type="meetSubTBLattice")]
HB.structure Definition MeetSubTBLattice d (T : latticeType d) S d' :=
  { U of @MeetSubLattice d T S d' U & TBLattice d' U }.

#[short(type="joinSubLattice")]
HB.structure Definition JoinSubLattice d (T : latticeType d) S d' :=
  { U of @SubPOrderLattice d T S d' U & isJoinSubLattice d T S d' U }.

#[short(type="joinSubBLattice")]
HB.structure Definition JoinSubBLattice d (T : latticeType d) S d' :=
  { U of @JoinSubLattice d T S d' U & BLattice d' U }.

#[short(type="joinSubTLattice")]
HB.structure Definition JoinSubTLattice d (T : latticeType d) S d' :=
  { U of @JoinSubLattice d T S d' U & TLattice d' U }.

#[short(type="joinSubTBLattice")]
HB.structure Definition JoinSubTBLattice d (T : latticeType d) S d' :=
  { U of @JoinSubLattice d T S d' U & TBLattice d' U }.

#[short(type="subLattice")]
HB.structure Definition SubLattice d (T : latticeType d) S d' :=
  { U of @MeetSubLattice d T S d' U & @JoinSubLattice d T S d' U }.

#[short(type="subBLattice")]
HB.structure Definition SubBLattice d (T : latticeType d) S d' :=
  { U of @SubLattice d T S d' U & BLattice d' U }.

#[short(type="subTLattice")]
HB.structure Definition SubTLattice d (T : latticeType d) S d' :=
  { U of @SubLattice d T S d' U & TLattice d' U }.

#[short(type="subTBLattice")]
HB.structure Definition SubTBLattice d (T : latticeType d) S d' :=
  { U of @SubLattice d T S d' U & TBLattice d' U }.

#[export, warning="-HB.no-new-instance"]
HB.instance Definition _ (d : disp_t) (T : latticeType d) (S : pred T)
    d' (U : MeetSubLattice.type S d') :=
  isMeetLatticeMorphism.Build d' U d T val valI_subproof.

#[export, warning="-HB.no-new-instance"]
HB.instance Definition _ (d : disp_t) (T : latticeType d) (S : pred T)
    d' (U : JoinSubLattice.type S d') :=
  isJoinLatticeMorphism.Build d' U d T val valU_subproof.




Let inU v Sv : U := Sub v Sv.
Let meetU (u1 u2 : U) : U := inU (opredI (valP u1) (valP u2)).
Let joinU (u1 u2 : U) : U := inU (opredU (valP u1) (valP u2)).

Let meetUC : commutative meetU.
Let joinUC : commutative joinU.
Let meetUA : associative meetU.
Let joinUA : associative joinU.
Lemma joinUKI y x : meetU x (joinU x y) = x.
Let meetUKU y x : joinU x (meetU x y) = x.
Let le_meetU x y : (x y) = (meetU x y == x).

Fact valI : meet_morphism (val : U T).
Fact valU : join_morphism (val : U T).




#[short(type="bJoinSubLattice")]
HB.structure Definition BJoinSubLattice d (T : bLatticeType d) S d' :=
  { U of @JoinSubLattice d T S d' U & BLattice d' U & isBSubLattice d T S d' U }.

#[short(type="bJoinSubTLattice")]
HB.structure Definition BJoinSubTLattice d (T : bLatticeType d) S d' :=
  { U of @BJoinSubLattice d T S d' U & TBLattice d' U }.

#[short(type="bSubLattice")]
HB.structure Definition BSubLattice d (T : bLatticeType d) S d' :=
  { U of @SubLattice d T S d' U & @BJoinSubLattice d T S d' U }.

#[short(type="bSubTLattice")]
HB.structure Definition BSubTLattice d (T : bLatticeType d) S d' :=
  { U of @BSubLattice d T S d' U & TBLattice d' U }.

#[export]
HB.instance Definition _ (d : disp_t) (T : bLatticeType d) (S : pred T)
    d' (U : BJoinSubLattice.type S d') :=
  isBLatticeMorphism.Build d' U d T val val0_subproof.



Let inU v Sv : U := Sub v Sv.
Let zeroU : U := inU opred0_subproof.

Fact le0x x : zeroU x.

Fact val0 : (val : U T) \bot = \bot.
#[warning="-HB.no-new-instance"]
HB.instance Definition _ := isBSubLattice.Build d T S d' U val0.




#[short(type="tMeetSubLattice")]
HB.structure Definition TMeetSubLattice d (T : tLatticeType d) S d' :=
  { U of @MeetSubLattice d T S d' U & TLattice d' U & isTSubLattice d T S d' U }.

#[short(type="tMeetSubBLattice")]
HB.structure Definition TMeetSubBLattice d (T : tLatticeType d) S d' :=
  { U of @TMeetSubLattice d T S d' U & TBLattice d' U }.

#[short(type="tSubLattice")]
HB.structure Definition TSubLattice d (T : tLatticeType d) S d' :=
  { U of @SubLattice d T S d' U & @TMeetSubLattice d T S d' U }.

#[short(type="tSubBLattice")]
HB.structure Definition TSubBLattice d (T : tLatticeType d) S d' :=
  { U of @TSubLattice d T S d' U & TBLattice d' U }.

#[export]
HB.instance Definition _ (d : disp_t) (T : tLatticeType d) (S : pred T)
    d' (U : TMeetSubLattice.type S d') :=
  isTLatticeMorphism.Build d' U d T val val1_subproof.



Let inU v Sv : U := Sub v Sv.
Let oneU : U := inU opred1_subproof.

Fact lex1 x : x oneU.

Fact val1 : (val : U T) \top = \top.
#[warning="-HB.no-new-instance"]
HB.instance Definition _ := isTSubLattice.Build d T S d' U val1.



#[short(type="tbSubLattice")]
HB.structure Definition TBSubLattice d (T : tbLatticeType d) S d' :=
  { U of @BSubLattice d T S d' U & @TSubLattice d T S d' U}.

#[export]
HB.instance Definition _ (d : disp_t) (T : tbLatticeType d) (S : pred T) d'
    (U : TBSubLattice.type S d') := BLatticeMorphism.on (val : U T).





#[short(type="subOrder")]
HB.structure Definition SubOrder d (T : orderType d) S d' :=
  { U of @SubLattice d T S d' U & Total d' U }.


Lemma totalU : total (<=%O : rel U).


Fact opredI : meet_closed S.
Fact opredU : join_closed S.



Module SubOrderExports.

Notation "[ 'SubChoice_isSubPOrder' 'of' U 'by' <: ]" :=
  (SubChoice_isSubPOrder.Build _ _ _ _ U)
  (format "[ 'SubChoice_isSubPOrder' 'of' U 'by' <: ]")
  : form_scope.
Notation "[ 'SubChoice_isSubPOrder' 'of' U 'by' <: 'with' disp ]" :=
  (SubChoice_isSubPOrder.Build _ _ _ disp U)
  (format "[ 'SubChoice_isSubPOrder' 'of' U 'by' <: 'with' disp ]")
  : form_scope.
Notation "[ 'SubPOrder_isSubLattice' 'of' U 'by' <: ]" :=
  (SubPOrder_isSubLattice.Build _ _ _ _ U (@opredI _ _ _) (@opredU _ _ _))
  (format "[ 'SubPOrder_isSubLattice' 'of' U 'by' <: ]")
  : form_scope.
Notation "[ 'SubPOrder_isSubLattice' 'of' U 'by' <: 'with' disp ]" :=
  (SubPOrder_isSubLattice.Build _ _ _ disp U (@opredI _ _ _) (@opredU _ _ _))
  (format "[ 'SubPOrder_isSubLattice' 'of' U 'by' <: 'with' disp ]")
  : form_scope.
Notation "[ 'SubChoice_isSubLattice' 'of' U 'by' <: ]" :=
  (SubChoice_isSubLattice.Build _ _ _ _ U (@opredI _ _ _) (@opredU _ _ _))
  (format "[ 'SubChoice_isSubLattice' 'of' U 'by' <: ]")
    : form_scope.
Notation "[ 'SubChoice_isSubLattice' 'of' U 'by' <: 'with' disp ]" :=
  (SubChoice_isSubLattice.Build _ _ _ disp U (@opredI _ _ _) (@opredU _ _ _))
  (format "[ 'SubChoice_isSubLattice' 'of' U 'by' <: 'with' disp ]")
    : form_scope.
Notation "[ 'SubPOrder_isBSubLattice' 'of' U 'by' <: ]" :=
  (SubPOrder_isBSubLattice.Build _ _ _ _ U (opred0 _))
  (format "[ 'SubPOrder_isBSubLattice' 'of' U 'by' <: ]")
  : form_scope.
Notation "[ 'SubPOrder_isBSubLattice' 'of' U 'by' <: 'with' disp ]" :=
  (SubPOrder_isBSubLattice.Build _ _ _ disp U (opred0 _))
  (format "[ 'SubPOrder_isBSubLattice' 'of' U 'by' <: 'with' disp ]")
  : form_scope.
Notation "[ 'SubChoice_isBSubLattice' 'of' U 'by' <: ]" :=
  (SubChoice_isBSubLattice.Build _ _ _ _ U
     (@opredI _ _ _) (@opredU _ _ _) (opred0 _))
  (format "[ 'SubChoice_isBSubLattice' 'of' U 'by' <: ]")
  : form_scope.
Notation "[ 'SubChoice_isBSubLattice' 'of' U 'by' <: 'with' disp ]" :=
  (SubChoice_isBSubLattice.Build _ _ _ disp U
     (@opredI _ _ _) (@opredU _ _ _) (opred0 _))
  (format "[ 'SubChoice_isBSubLattice' 'of' U 'by' <: 'with' disp ]")
  : form_scope.
Notation "[ 'SubPOrder_isTSubLattice' 'of' U 'by' <: ]" :=
  (SubPOrder_isTSubLattice.Build _ _ _ _ U (opred1 _))
  (format "[ 'SubPOrder_isTSubLattice' 'of' U 'by' <: ]")
  : form_scope.
Notation "[ 'SubPOrder_isTSubLattice' 'of' U 'by' <: 'with' disp ]" :=
  (SubPOrder_isTSubLattice.Build _ _ _ disp U (opred1 _))
  (format "[ 'SubPOrder_isTSubLattice' 'of' U 'by' <: 'with' disp ]")
  : form_scope.
Notation "[ 'SubChoice_isTSubLattice' 'of' U 'by' <: ]" :=
  (SubChoice_isTSubLattice.Build _ _ _ _ U
     (@opredI _ _ _) (@opredU _ _ _) (opred1 _))
  (format "[ 'SubChoice_isTSubLattice' 'of' U 'by' <: ]")
  : form_scope.
Notation "[ 'SubChoice_isTSubLattice' 'of' U 'by' <: 'with' disp ]" :=
  (SubChoice_isTSubLattice.Build _ _ _ disp U
     (@opredI _ _ _) (@opredU _ _ _) (opred1 _))
  (format "[ 'SubChoice_isTSubLattice' 'of' U 'by' <: 'with' disp ]")
  : form_scope.
Notation "[ 'SubPOrder_isTBSubLattice' 'of' U 'by' <: ]" :=
  (SubPOrder_isTBSubLattice.Build _ _ _ _ U (opred0 _) (opred1 _))
  (format "[ 'SubPOrder_isTBSubLattice' 'of' U 'by' <: ]")
  : form_scope.
Notation "[ 'SubPOrder_isTBSubLattice' 'of' U 'by' <: 'with' disp ]" :=
  (SubPOrder_isTBSubLattice.Build _ _ _ disp U (opred0 _) (opred1 _))
  (format "[ 'SubPOrder_isTBSubLattice' 'of' U 'by' <: 'with' disp ]")
  : form_scope.
Notation "[ 'SubChoice_isTBSubLattice' 'of' U 'by' <: ]" :=
  (SubChoice_isTBSubLattice.Build _ _ _ _ U
     (@opredI _ _ _) (@opredU _ _ _) (opred0 _) (opred1 _))
  (format "[ 'SubChoice_isTBSubLattice' 'of' U 'by' <: ]")
  : form_scope.
Notation "[ 'SubChoice_isTBSubLattice' 'of' U 'by' <: 'with' disp ]" :=
  (SubChoice_isTBSubLattice.Build _ _ _ disp U
     (@opredI _ _ _) (@opredU _ _ _) (opred0 _) (opred1 _))
  (format "[ 'SubChoice_isTBSubLattice' 'of' U 'by' <: 'with' disp ]")
  : form_scope.
Notation "[ 'SubLattice_isSubOrder' 'of' U 'by' <: ]" :=
  (SubLattice_isSubOrder.Build _ _ _ _ U)
  (format "[ 'SubLattice_isSubOrder' 'of' U 'by' <: ]")
  : form_scope.
Notation "[ 'SubLattice_isSubOrder' 'of' U 'by' <: 'with' disp ]" :=
  (SubLattice_isSubOrder.Build _ _ _ disp U)
  (format "[ 'SubLattice_isSubOrder' 'of' U 'by' <: 'with' disp ]")
  : form_scope.
Notation "[ 'SubPOrder_isSubOrder' 'of' U 'by' <: ]" :=
  (SubPOrder_isSubOrder.Build _ _ _ _ U)
  (format "[ 'SubPOrder_isSubOrder' 'of' U 'by' <: ]")
  : form_scope.
Notation "[ 'SubPOrder_isSubOrder' 'of' U 'by' <: 'with' disp ]" :=
  (SubPOrder_isSubOrder.Build _ _ _ disp U)
  (format "[ 'SubPOrder_isSubOrder' 'of' U 'by' <: 'with' disp ]")
  : form_scope.
Notation "[ 'SubChoice_isSubOrder' 'of' U 'by' <: ]" :=
  (SubChoice_isSubOrder.Build _ _ _ _ U)
  (format "[ 'SubChoice_isSubOrder' 'of' U 'by' <: ]")
  : form_scope.
Notation "[ 'SubChoice_isSubOrder' 'of' U 'by' <: 'with' disp ]" :=
  (SubChoice_isSubOrder.Build _ _ _ disp U)
  (format "[ 'SubChoice_isSubOrder' 'of' U 'by' <: 'with' disp ]")
  : form_scope.

End SubOrderExports.

Module DeprecatedSubOrder.

Section Total.
Context {disp : disp_t} {T : orderType disp} (P : {pred T}) (sT : subType P).

#[export]
HB.instance Definition _ :=
  SubPOrder_isSubOrder.Build disp T P disp (sub_type sT).

End Total.

Module Exports.
Notation "[ 'POrder' 'of' T 'by' <: ]" :=
  (POrder.copy T%type (sub_type T%type))
  (format "[ 'POrder' 'of' T 'by' <: ]") : form_scope.
Notation "[ 'Order' 'of' T 'by' <: ]" :=
  (Total.copy T%type (sub_type T%type))
  (only parsing) : form_scope.
End Exports.
End DeprecatedSubOrder.

INSTANCES Instances on nat This is an example of creation of multiple instances on the same type, with distinct displays, using natural numbers. We declare two distinct canonical orders:
  • leq which is total, and where meet and join are minn and maxn, on nat
  • dvdn which is partial, and where meet and join are gcdn and lcmn, on natdvd
The Module NatOrder defines leq as the canonical order on the type nat, i.e., without creating an alias. We define and use nat_display and proceed like a standard canonical structure declaration, except that we use this display. We also use a single factory LeOrderMixin to instantiate three different canonical declarations porderType, distrLatticeType, orderType. We finish by providing theorems to convert the operations of ordered and lattice types to their definition without structure abstraction.

Module NatOrder.
Section NatOrder.

Export NatOrder.

#[export]
HB.instance Definition _ :=
  Preorder_isPOrder.Build nat_display nat anti_leq.
#[export]
HB.instance Definition _ :=
  POrder_isTotal.Build nat_display nat leq_total.

End NatOrder.

Module Exports.
End Exports.

End NatOrder.

Module NatMonotonyTheory.
Section NatMonotonyTheory.

Export NatMonotonyTheory.

Context {disp : disp_t} {T : porderType disp}.
Variables (D : {pred nat}) (f : nat T).
Hypothesis Dconvex : {in D &, i j k, i < k < j k \in D}.

Lemma incn_inP : {in D, i, i.+1 \in D f i < f i.+1}
  {in D &, {mono f : i j / i j}}.

Lemma decn_inP : {in D, i, i.+1 \in D f i > f i.+1}
  {in D &, {mono f : i j /~ i j}}.

Lemma incnP : ( i, f i < f i.+1) {mono f : i j / i j}.

Lemma decnP : ( i, f i > f i.+1) {mono f : i j /~ i j}.

End NatMonotonyTheory.
Arguments incn_inP {disp T} [D f].
Arguments decn_inP {disp T} [D f].
Arguments incnP {disp T} [f].
Arguments decnP {disp T} [f].
End NatMonotonyTheory.

The Module DvdSyntax introduces a new set of notations using the newly created display dvd_display. We first define the display as an opaque definition of type disp_t, and we use it as the first argument of the operator which display we want to change from the default one (here le, lt, dvd sdvd, meet, join, top and bottom, as well as big op notations on gcd and lcm). This notations will now be used for any ordered type which first parameter is set to dvd_display.

Module DvdSyntax.

Export DvdSyntax.

Notation gcd := (@meet dvd_display _).
Notation "@ 'gcd' T" := (@meet dvd_display T)
  (at level 10, T at level 8, only parsing) : function_scope.
Notation lcm := (@join dvd_display _).
Notation "@ 'lcm' T" := (@join dvd_display T)
  (at level 10, T at level 8, only parsing) : function_scope.

Notation "\gcd_ ( i <- r | P ) F" :=
  (\big[gcd/nat0]_(i <- r | P%B) F%O) : order_scope.
Notation "\gcd_ ( i <- r ) F" :=
  (\big[gcd/nat0]_(i <- r) F%O) : order_scope.
Notation "\gcd_ ( i | P ) F" :=
  (\big[gcd/nat0]_(i | P%B) F%O) : order_scope.
Notation "\gcd_ i F" :=
  (\big[gcd/nat0]_i F%O) : order_scope.
Notation "\gcd_ ( i : I | P ) F" :=
  (\big[gcd/nat0]_(i : I | P%B) F%O) (only parsing) :
  order_scope.
Notation "\gcd_ ( i : I ) F" :=
  (\big[gcd/nat0]_(i : I) F%O) (only parsing) : order_scope.
Notation "\gcd_ ( m <= i < n | P ) F" :=
  (\big[gcd/nat0]_(m i < n | P%B) F%O) : order_scope.
Notation "\gcd_ ( m <= i < n ) F" :=
  (\big[gcd/nat0]_(m i < n) F%O) : order_scope.
Notation "\gcd_ ( i < n | P ) F" :=
  (\big[gcd/nat0]_(i < n | P%B) F%O) : order_scope.
Notation "\gcd_ ( i < n ) F" :=
  (\big[gcd/nat0]_(i < n) F%O) : order_scope.
Notation "\gcd_ ( i 'in' A | P ) F" :=
  (\big[gcd/nat0]_(i in A | P%B) F%O) : order_scope.
Notation "\gcd_ ( i 'in' A ) F" :=
  (\big[gcd/nat0]_(i in A) F%O) : order_scope.

Notation "\lcm_ ( i <- r | P ) F" :=
  (\big[lcm/nat1]_(i <- r | P%B) F%O) : order_scope.
Notation "\lcm_ ( i <- r ) F" :=
  (\big[lcm/nat1]_(i <- r) F%O) : order_scope.
Notation "\lcm_ ( i | P ) F" :=
  (\big[lcm/nat1]_(i | P%B) F%O) : order_scope.
Notation "\lcm_ i F" :=
  (\big[lcm/nat1]_i F%O) : order_scope.
Notation "\lcm_ ( i : I | P ) F" :=
  (\big[lcm/nat1]_(i : I | P%B) F%O) (only parsing) :
  order_scope.
Notation "\lcm_ ( i : I ) F" :=
  (\big[lcm/nat1]_(i : I) F%O) (only parsing) : order_scope.
Notation "\lcm_ ( m <= i < n | P ) F" :=
  (\big[lcm/nat1]_(m i < n | P%B) F%O) : order_scope.
Notation "\lcm_ ( m <= i < n ) F" :=
  (\big[lcm/nat1]_(m i < n) F%O) : order_scope.
Notation "\lcm_ ( i < n | P ) F" :=
  (\big[lcm/nat1]_(i < n | P%B) F%O) : order_scope.
Notation "\lcm_ ( i < n ) F" :=
  (\big[lcm/nat1]_(i < n) F%O) : order_scope.
Notation "\lcm_ ( i 'in' A | P ) F" :=
  (\big[lcm/nat1]_(i in A | P%B) F%O) : order_scope.
Notation "\lcm_ ( i 'in' A ) F" :=
  (\big[lcm/nat1]_(i in A) F%O) : order_scope.

End DvdSyntax.

The Module NatDvd defines dvdn as the canonical order on NatDvd.t, which is abbreviated using the notation natdvd at the end of the module. We use the newly defined dvd_display, described above. We first recover structures that are common to both nat and natdvd (eqType, choiceType, countType) through the copy mechanism, then we use a single factory MeetJoinMixin to instantiate both porderType and distrLatticeType canonical structures, and end with top and bottom. We finish by providing theorems to convert the operations of ordered and lattice types to their definition without structure abstraction.

Module NatDvd.

Export NatDvd.

Section NatDvd.
Implicit Types (m n p : natdvd).

Lemma lcmnn n : lcmn n n = n.

Lemma le_def m n : m %| n = (gcdn m n == m)%N.

Lemma joinKI n m : gcdn m (lcmn m n) = m.

Lemma meetKU n m : lcmn m (gcdn m n) = m.

Lemma meetUl : left_distributive gcdn lcmn.

Fact dvdn_anti : antisymmetric dvdn.

#[export]
HB.instance Definition _ := Preorder_isPOrder.Build dvd_display t dvdn_anti.

#[export]
HB.instance Definition _ := @POrder_Meet_isDistrLattice.Build dvd_display t
  gcdn lcmn gcdnC lcmnC gcdnA lcmnA joinKI meetKU le_def meetUl.

Import DvdSyntax.
Lemma sdvdE (m n : t) : m %<| n = (n != m) && (m %| n).
Lemma gcdE : gcd = gcdn :> (t t t).
Lemma lcmE : lcm = lcmn :> (t t t).

End NatDvd.
Module Exports.
Definition sdvdEnat := sdvdE.
Definition gcdEnat := gcdE.
Definition lcmEnat := lcmE.
End Exports.
End NatDvd.

Instances on ordinal
Instances on bool
Definition of prod_display

Module Import ProdSyntax.

Export ProdSyntax.

The following Local Notations are here to define the \join^p_ and \meet^p_ notations later. Do not remove them.
Local Notation "\bot" := (@bottom (prod_display _ _) _).
Local Notation "\top" := (@top (prod_display _ _) _).
Local Notation meet := (@meet (prod_display _ _) _).
Local Notation join := (@join (prod_display _ _) _).
Local Notation min := (@min (prod_display _ _) _).
Local Notation max := (@max (prod_display _ _) _).

Notation "x `&^p` y" := (meet x y) : order_scope.
Notation "x `|^p` y" := (join x y) : order_scope.

Notation "\join^p_ ( i <- r | P ) F" :=
  (\big[join / \bot]_(i <- r | P%B) F%O) : order_scope.
Notation "\join^p_ ( i <- r ) F" :=
  (\big[join / \bot]_(i <- r) F%O) : order_scope.
Notation "\join^p_ ( i | P ) F" :=
  (\big[join / \bot]_(i | P%B) F%O) : order_scope.
Notation "\join^p_ i F" :=
  (\big[join / \bot]_i F%O) : order_scope.
Notation "\join^p_ ( i : I | P ) F" :=
  (\big[join / \bot]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\join^p_ ( i : I ) F" :=
  (\big[join / \bot]_(i : I) F%O) (only parsing) : order_scope.
Notation "\join^p_ ( m <= i < n | P ) F" :=
 (\big[join / \bot]_(m i < n | P%B) F%O) : order_scope.
Notation "\join^p_ ( m <= i < n ) F" :=
 (\big[join / \bot]_(m i < n) F%O) : order_scope.
Notation "\join^p_ ( i < n | P ) F" :=
 (\big[join / \bot]_(i < n | P%B) F%O) : order_scope.
Notation "\join^p_ ( i < n ) F" :=
 (\big[join / \bot]_(i < n) F%O) : order_scope.
Notation "\join^p_ ( i 'in' A | P ) F" :=
 (\big[join / \bot]_(i in A | P%B) F%O) : order_scope.
Notation "\join^p_ ( i 'in' A ) F" :=
 (\big[join / \bot]_(i in A) F%O) : order_scope.

Notation "\meet^p_ ( i <- r | P ) F" :=
  (\big[meet / \top]_(i <- r | P%B) F%O) : order_scope.
Notation "\meet^p_ ( i <- r ) F" :=
  (\big[meet / \top]_(i <- r) F%O) : order_scope.
Notation "\meet^p_ ( i | P ) F" :=
  (\big[meet / \top]_(i | P%B) F%O) : order_scope.
Notation "\meet^p_ i F" :=
  (\big[meet / \top]_i F%O) : order_scope.
Notation "\meet^p_ ( i : I | P ) F" :=
  (\big[meet / \top]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\meet^p_ ( i : I ) F" :=
  (\big[meet / \top]_(i : I) F%O) (only parsing) : order_scope.
Notation "\meet^p_ ( m <= i < n | P ) F" :=
 (\big[meet / \top]_(m i < n | P%B) F%O) : order_scope.
Notation "\meet^p_ ( m <= i < n ) F" :=
 (\big[meet / \top]_(m i < n) F%O) : order_scope.
Notation "\meet^p_ ( i < n | P ) F" :=
 (\big[meet / \top]_(i < n | P%B) F%O) : order_scope.
Notation "\meet^p_ ( i < n ) F" :=
 (\big[meet / \top]_(i < n) F%O) : order_scope.
Notation "\meet^p_ ( i 'in' A | P ) F" :=
 (\big[meet / \top]_(i in A | P%B) F%O) : order_scope.
Notation "\meet^p_ ( i 'in' A ) F" :=
 (\big[meet / \top]_(i in A) F%O) : order_scope.

Notation "\min^p_ i F" :=
  (\big[min/top]_i F) : order_scope.
Notation "\min^p_ ( i <- r | P ) F" :=
  (\big[min/top]_(i <- r | P%B) F%O) : order_scope.
Notation "\min^p_ ( i < r ) F" :=
  (\big[min/top]_(i <- r) F%O) : order_scope.
Notation "\min^p_ ( m <= i < n | P ) F" :=
  (\big[min/top]_(m i < n | P%B) F%O) : order_scope.
Notation "\min^p_ ( m <= i < n ) F" :=
  (\big[min/top]_(m i < n) F%O) : order_scope.
Notation "\min^p_ ( i | P ) F" :=
  (\big[min/top]_(i | P%B) F%O) : order_scope.
Notation "\min^p_ ( i : t | P ) F" :=
  (\big[min/top]_(i : t | P%B) F%O) (only parsing) : order_scope.
Notation "\min^p_ ( i : t ) F" :=
  (\big[min/top]_(i : t) F%O) (only parsing) : order_scope.
Notation "\min^p_ ( i < n | P ) F" :=
  (\big[min/top]_(i < n | P%B) F%O) : order_scope.
Notation "\min^p_ ( i < n ) F" :=
  (\big[min/top]_(i < n) F%O) : order_scope.
Notation "\min^p_ ( i 'in' A | P ) F" :=
  (\big[min/top]_(i in A | P%B) F%O) : order_scope.
Notation "\min^p_ ( i 'in' A ) F" :=
  (\big[min/top]_(i in A) F%O) : order_scope.

Notation "\max^p_ i F" :=
  (\big[max/bottom]_i F%O) : order_scope.
Notation "\max^p_ ( i <- r | P ) F" :=
  (\big[max/bottom]_(i <- r | P%B) F%O) : order_scope.
Notation "\max^p_ ( i < r ) F" :=
  (\big[max/bottom]_(i <- r) F%O) : order_scope.
Notation "\max^p_ ( m <= i < n | P ) F" :=
  (\big[max/bottom]_(m i < n | P%B) F%O) : order_scope.
Notation "\max^p_ ( m <= i < n ) F" :=
  (\big[max/bottom]_(m i < n) F%O) : order_scope.
Notation "\max^p_ ( i | P ) F" :=
  (\big[max/bottom]_(i | P%B) F%O) : order_scope.
Notation "\max^p_ ( i : t | P ) F" :=
  (\big[max/bottom]_(i : t | P%B) F%O) (only parsing) : order_scope.
Notation "\max^p_ ( i : t ) F" :=
  (\big[max/bottom]_(i : t) F%O) (only parsing) : order_scope.
Notation "\max^p_ ( i < n | P ) F" :=
  (\big[max/bottom]_(i < n | P%B) F%O) : order_scope.
Notation "\max^p_ ( i < n ) F" :=
  (\big[max/bottom]_(i < n) F%O) : order_scope.
Notation "\max^p_ ( i 'in' A | P ) F" :=
  (\big[max/bottom]_(i in A | P%B) F%O) : order_scope.
Notation "\max^p_ ( i 'in' A ) F" :=
  (\big[max/bottom]_(i in A) F%O) : order_scope.

End ProdSyntax.

Module Import SeqProdSyntax.

Export SeqProdSyntax.

The following Local Notations are here to define the \join^sp_ and \meet^sp_ notations later. Do not remove them.
Local Notation "\bot" := (@bottom (seqprod_display _) _).
Local Notation "\top" := (@top (seqprod_display _) _).
Local Notation meet := (@meet (seqprod_display _) _).
Local Notation join := (@join (seqprod_display _) _).
Local Notation min := (@min (seqprod_display _) _).
Local Notation max := (@max (seqprod_display _) _).

Notation "x `&^sp` y" := (meet x y) : order_scope.
Notation "x `|^sp` y" := (join x y) : order_scope.

Notation "\join^sp_ ( i <- r | P ) F" :=
  (\big[join / \bot]_(i <- r | P%B) F%O) : order_scope.
Notation "\join^sp_ ( i <- r ) F" :=
  (\big[join / \bot]_(i <- r) F%O) : order_scope.
Notation "\join^sp_ ( i | P ) F" :=
  (\big[join / \bot]_(i | P%B) F%O) : order_scope.
Notation "\join^sp_ i F" :=
  (\big[join / \bot]_i F%O) : order_scope.
Notation "\join^sp_ ( i : I | P ) F" :=
  (\big[join / \bot]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\join^sp_ ( i : I ) F" :=
  (\big[join / \bot]_(i : I) F%O) (only parsing) : order_scope.
Notation "\join^sp_ ( m <= i < n | P ) F" :=
 (\big[join / \bot]_(m i < n | P%B) F%O) : order_scope.
Notation "\join^sp_ ( m <= i < n ) F" :=
 (\big[join / \bot]_(m i < n) F%O) : order_scope.
Notation "\join^sp_ ( i < n | P ) F" :=
 (\big[join / \bot]_(i < n | P%B) F%O) : order_scope.
Notation "\join^sp_ ( i < n ) F" :=
 (\big[join / \bot]_(i < n) F%O) : order_scope.
Notation "\join^sp_ ( i 'in' A | P ) F" :=
 (\big[join / \bot]_(i in A | P%B) F%O) : order_scope.
Notation "\join^sp_ ( i 'in' A ) F" :=
 (\big[join / \bot]_(i in A) F%O) : order_scope.

Notation "\meet^sp_ ( i <- r | P ) F" :=
  (\big[meet / \top]_(i <- r | P%B) F%O) : order_scope.
Notation "\meet^sp_ ( i <- r ) F" :=
  (\big[meet / \top]_(i <- r) F%O) : order_scope.
Notation "\meet^sp_ ( i | P ) F" :=
  (\big[meet / \top]_(i | P%B) F%O) : order_scope.
Notation "\meet^sp_ i F" :=
  (\big[meet / \top]_i F%O) : order_scope.
Notation "\meet^sp_ ( i : I | P ) F" :=
  (\big[meet / \top]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\meet^sp_ ( i : I ) F" :=
  (\big[meet / \top]_(i : I) F%O) (only parsing) : order_scope.
Notation "\meet^sp_ ( m <= i < n | P ) F" :=
 (\big[meet / \top]_(m i < n | P%B) F%O) : order_scope.
Notation "\meet^sp_ ( m <= i < n ) F" :=
 (\big[meet / \top]_(m i < n) F%O) : order_scope.
Notation "\meet^sp_ ( i < n | P ) F" :=
 (\big[meet / \top]_(i < n | P%B) F%O) : order_scope.
Notation "\meet^sp_ ( i < n ) F" :=
 (\big[meet / \top]_(i < n) F%O) : order_scope.
Notation "\meet^sp_ ( i 'in' A | P ) F" :=
 (\big[meet / \top]_(i in A | P%B) F%O) : order_scope.
Notation "\meet^sp_ ( i 'in' A ) F" :=
 (\big[meet / \top]_(i in A) F%O) : order_scope.

Notation "\min^sp_ i F" :=
  (\big[min/top]_i F) : order_scope.
Notation "\min^sp_ ( i <- r | P ) F" :=
  (\big[min/top]_(i <- r | P%B) F%O) : order_scope.
Notation "\min^sp_ ( i < r ) F" :=
  (\big[min/top]_(i <- r) F%O) : order_scope.
Notation "\min^sp_ ( m <= i < n | P ) F" :=
  (\big[min/top]_(m i < n | P%B) F%O) : order_scope.
Notation "\min^sp_ ( m <= i < n ) F" :=
  (\big[min/top]_(m i < n) F%O) : order_scope.
Notation "\min^sp_ ( i | P ) F" :=
  (\big[min/top]_(i | P%B) F%O) : order_scope.
Notation "\min^sp_ ( i : t | P ) F" :=
  (\big[min/top]_(i : t | P%B) F%O) (only parsing) : order_scope.
Notation "\min^sp_ ( i : t ) F" :=
  (\big[min/top]_(i : t) F%O) (only parsing) : order_scope.
Notation "\min^sp_ ( i < n | P ) F" :=
  (\big[min/top]_(i < n | P%B) F%O) : order_scope.
Notation "\min^sp_ ( i < n ) F" :=
  (\big[min/top]_(i < n) F%O) : order_scope.
Notation "\min^sp_ ( i 'in' A | P ) F" :=
  (\big[min/top]_(i in A | P%B) F%O) : order_scope.
Notation "\min^sp_ ( i 'in' A ) F" :=
  (\big[min/top]_(i in A) F%O) : order_scope.

Notation "\max^sp_ i F" :=
  (\big[max/bottom]_i F%O) : order_scope.
Notation "\max^sp_ ( i <- r | P ) F" :=
  (\big[max/bottom]_(i <- r | P%B) F%O) : order_scope.
Notation "\max^sp_ ( i < r ) F" :=
  (\big[max/bottom]_(i <- r) F%O) : order_scope.
Notation "\max^sp_ ( m <= i < n | P ) F" :=
  (\big[max/bottom]_(m i < n | P%B) F%O) : order_scope.
Notation "\max^sp_ ( m <= i < n ) F" :=
  (\big[max/bottom]_(m i < n) F%O) : order_scope.
Notation "\max^sp_ ( i | P ) F" :=
  (\big[max/bottom]_(i | P%B) F%O) : order_scope.
Notation "\max^sp_ ( i : t | P ) F" :=
  (\big[max/bottom]_(i : t | P%B) F%O) (only parsing) : order_scope.
Notation "\max^sp_ ( i : t ) F" :=
  (\big[max/bottom]_(i : t) F%O) (only parsing) : order_scope.
Notation "\max^sp_ ( i < n | P ) F" :=
  (\big[max/bottom]_(i < n | P%B) F%O) : order_scope.
Notation "\max^sp_ ( i < n ) F" :=
  (\big[max/bottom]_(i < n) F%O) : order_scope.
Notation "\max^sp_ ( i 'in' A | P ) F" :=
  (\big[max/bottom]_(i in A | P%B) F%O) : order_scope.
Notation "\max^sp_ ( i 'in' A ) F" :=
  (\big[max/bottom]_(i in A) F%O) : order_scope.

End SeqProdSyntax.

Module Import LexiSyntax.

Export LexiSyntax.

Notation meetlexi := (@meet (lexi_display _ _) _).
Notation joinlexi := (@join (lexi_display _ _) _).

Notation "x `&^l` y" := (meetlexi x y) : order_scope.
Notation "x `|^l` y" := (joinlexi x y) : order_scope.

The following Local Notations are here to define the \join^l_ and \meet^l_ notations later. Do not remove them.
Local Notation "\bot" := (@bottom (lexi_display _ _) _).
Local Notation "\top" := (@top (lexi_display _ _) _).
Local Notation meet := (@meet (lexi_display _ _) _).
Local Notation join := (@join (lexi_display _ _) _).
Local Notation min := (@min (lexi_display _ _) _).
Local Notation max := (@max (lexi_display _ _) _).

Notation "\join^l_ ( i <- r | P ) F" :=
  (\big[join / \bot]_(i <- r | P%B) F%O) : order_scope.
Notation "\join^l_ ( i <- r ) F" :=
  (\big[join / \bot]_(i <- r) F%O) : order_scope.
Notation "\join^l_ ( i | P ) F" :=
  (\big[join / \bot]_(i | P%B) F%O) : order_scope.
Notation "\join^l_ i F" :=
  (\big[join / \bot]_i F%O) : order_scope.
Notation "\join^l_ ( i : I | P ) F" :=
  (\big[join / \bot]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\join^l_ ( i : I ) F" :=
  (\big[join / \bot]_(i : I) F%O) (only parsing) : order_scope.
Notation "\join^l_ ( m <= i < n | P ) F" :=
 (\big[join / \bot]_(m i < n | P%B) F%O) : order_scope.
Notation "\join^l_ ( m <= i < n ) F" :=
 (\big[join / \bot]_(m i < n) F%O) : order_scope.
Notation "\join^l_ ( i < n | P ) F" :=
 (\big[join / \bot]_(i < n | P%B) F%O) : order_scope.
Notation "\join^l_ ( i < n ) F" :=
 (\big[join / \bot]_(i < n) F%O) : order_scope.
Notation "\join^l_ ( i 'in' A | P ) F" :=
 (\big[join / \bot]_(i in A | P%B) F%O) : order_scope.
Notation "\join^l_ ( i 'in' A ) F" :=
 (\big[join / \bot]_(i in A) F%O) : order_scope.

Notation "\meet^l_ ( i <- r | P ) F" :=
  (\big[meet / \top]_(i <- r | P%B) F%O) : order_scope.
Notation "\meet^l_ ( i <- r ) F" :=
  (\big[meet / \top]_(i <- r) F%O) : order_scope.
Notation "\meet^l_ ( i | P ) F" :=
  (\big[meet / \top]_(i | P%B) F%O) : order_scope.
Notation "\meet^l_ i F" :=
  (\big[meet / \top]_i F%O) : order_scope.
Notation "\meet^l_ ( i : I | P ) F" :=
  (\big[meet / \top]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\meet^l_ ( i : I ) F" :=
  (\big[meet / \top]_(i : I) F%O) (only parsing) : order_scope.
Notation "\meet^l_ ( m <= i < n | P ) F" :=
 (\big[meet / \top]_(m i < n | P%B) F%O) : order_scope.
Notation "\meet^l_ ( m <= i < n ) F" :=
 (\big[meet / \top]_(m i < n) F%O) : order_scope.
Notation "\meet^l_ ( i < n | P ) F" :=
 (\big[meet / \top]_(i < n | P%B) F%O) : order_scope.
Notation "\meet^l_ ( i < n ) F" :=
 (\big[meet / \top]_(i < n) F%O) : order_scope.
Notation "\meet^l_ ( i 'in' A | P ) F" :=
 (\big[meet / \top]_(i in A | P%B) F%O) : order_scope.
Notation "\meet^l_ ( i 'in' A ) F" :=
 (\big[meet / \top]_(i in A) F%O) : order_scope.

Notation "\min^l_ i F" :=
  (\big[min/top]_i F) : order_scope.
Notation "\min^l_ ( i <- r | P ) F" :=
  (\big[min/top]_(i <- r | P%B) F%O) : order_scope.
Notation "\min^l_ ( i < r ) F" :=
  (\big[min/top]_(i <- r) F%O) : order_scope.
Notation "\min^l_ ( m <= i < n | P ) F" :=
  (\big[min/top]_(m i < n | P%B) F%O) : order_scope.
Notation "\min^l_ ( m <= i < n ) F" :=
  (\big[min/top]_(m i < n) F%O) : order_scope.
Notation "\min^l_ ( i | P ) F" :=
  (\big[min/top]_(i | P%B) F%O) : order_scope.
Notation "\min^l_ ( i : t | P ) F" :=
  (\big[min/top]_(i : t | P%B) F%O) (only parsing) : order_scope.
Notation "\min^l_ ( i : t ) F" :=
  (\big[min/top]_(i : t) F%O) (only parsing) : order_scope.
Notation "\min^l_ ( i < n | P ) F" :=
  (\big[min/top]_(i < n | P%B) F%O) : order_scope.
Notation "\min^l_ ( i < n ) F" :=
  (\big[min/top]_(i < n) F%O) : order_scope.
Notation "\min^l_ ( i 'in' A | P ) F" :=
  (\big[min/top]_(i in A | P%B) F%O) : order_scope.
Notation "\min^l_ ( i 'in' A ) F" :=
  (\big[min/top]_(i in A) F%O) : order_scope.

Notation "\max^l_ i F" :=
  (\big[max/bottom]_i F%O) : order_scope.
Notation "\max^l_ ( i <- r | P ) F" :=
  (\big[max/bottom]_(i <- r | P%B) F%O) : order_scope.
Notation "\max^l_ ( i < r ) F" :=
  (\big[max/bottom]_(i <- r) F%O) : order_scope.
Notation "\max^l_ ( m <= i < n | P ) F" :=
  (\big[max/bottom]_(m i < n | P%B) F%O) : order_scope.
Notation "\max^l_ ( m <= i < n ) F" :=
  (\big[max/bottom]_(m i < n) F%O) : order_scope.
Notation "\max^l_ ( i | P ) F" :=
  (\big[max/bottom]_(i | P%B) F%O) : order_scope.
Notation "\max^l_ ( i : t | P ) F" :=
  (\big[max/bottom]_(i : t | P%B) F%O) (only parsing) : order_scope.
Notation "\max^l_ ( i : t ) F" :=
  (\big[max/bottom]_(i : t) F%O) (only parsing) : order_scope.
Notation "\max^l_ ( i < n | P ) F" :=
  (\big[max/bottom]_(i < n | P%B) F%O) : order_scope.
Notation "\max^l_ ( i < n ) F" :=
  (\big[max/bottom]_(i < n) F%O) : order_scope.
Notation "\max^l_ ( i 'in' A | P ) F" :=
  (\big[max/bottom]_(i in A | P%B) F%O) : order_scope.
Notation "\max^l_ ( i 'in' A ) F" :=
  (\big[max/bottom]_(i in A) F%O) : order_scope.

End LexiSyntax.

Module Import SeqLexiSyntax.

Export SeqLexiSyntax.

Notation meetlexi := (@meet (seqlexi_display _) _).
Notation joinlexi := (@join (seqlexi_display _) _).

Notation "x `&^l` y" := (meetlexi x y) : order_scope.
Notation "x `|^l` y" := (joinlexi x y) : order_scope.

The following Local Notations are here to define the \join^l_ and \meet^l_ notations later. Do not remove them.
Local Notation "\bot" := (@bottom (lexi_display _ _) _).
Local Notation "\top" := (@top (lexi_display _ _) _).
Local Notation meet := (@meet (lexi_display _ _) _).
Local Notation join := (@join (lexi_display _ _) _).
Local Notation min := (@min (lexi_display _ _) _).
Local Notation max := (@max (lexi_display _ _) _).

Notation "\join^l_ ( i <- r | P ) F" :=
  (\big[join / \bot]_(i <- r | P%B) F%O) : order_scope.
Notation "\join^l_ ( i <- r ) F" :=
  (\big[join / \bot]_(i <- r) F%O) : order_scope.
Notation "\join^l_ ( i | P ) F" :=
  (\big[join / \bot]_(i | P%B) F%O) : order_scope.
Notation "\join^l_ i F" :=
  (\big[join / \bot]_i F%O) : order_scope.
Notation "\join^l_ ( i : I | P ) F" :=
  (\big[join / \bot]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\join^l_ ( i : I ) F" :=
  (\big[join / \bot]_(i : I) F%O) (only parsing) : order_scope.
Notation "\join^l_ ( m <= i < n | P ) F" :=
 (\big[join / \bot]_(m i < n | P%B) F%O) : order_scope.
Notation "\join^l_ ( m <= i < n ) F" :=
 (\big[join / \bot]_(m i < n) F%O) : order_scope.
Notation "\join^l_ ( i < n | P ) F" :=
 (\big[join / \bot]_(i < n | P%B) F%O) : order_scope.
Notation "\join^l_ ( i < n ) F" :=
 (\big[join / \bot]_(i < n) F%O) : order_scope.
Notation "\join^l_ ( i 'in' A | P ) F" :=
 (\big[join / \bot]_(i in A | P%B) F%O) : order_scope.
Notation "\join^l_ ( i 'in' A ) F" :=
 (\big[join / \bot]_(i in A) F%O) : order_scope.

Notation "\meet^l_ ( i <- r | P ) F" :=
  (\big[meet / \top]_(i <- r | P%B) F%O) : order_scope.
Notation "\meet^l_ ( i <- r ) F" :=
  (\big[meet / \top]_(i <- r) F%O) : order_scope.
Notation "\meet^l_ ( i | P ) F" :=
  (\big[meet / \top]_(i | P%B) F%O) : order_scope.
Notation "\meet^l_ i F" :=
  (\big[meet / \top]_i F%O) : order_scope.
Notation "\meet^l_ ( i : I | P ) F" :=
  (\big[meet / \top]_(i : I | P%B) F%O) (only parsing) : order_scope.
Notation "\meet^l_ ( i : I ) F" :=
  (\big[meet / \top]_(i : I) F%O) (only parsing) : order_scope.
Notation "\meet^l_ ( m <= i < n | P ) F" :=
 (\big[meet / \top]_(m i < n | P%B) F%O) : order_scope.
Notation "\meet^l_ ( m <= i < n ) F" :=
 (\big[meet / \top]_(m i < n) F%O) : order_scope.
Notation "\meet^l_ ( i < n | P ) F" :=
 (\big[meet / \top]_(i < n | P%B) F%O) : order_scope.
Notation "\meet^l_ ( i < n ) F" :=
 (\big[meet / \top]_(i < n) F%O) : order_scope.
Notation "\meet^l_ ( i 'in' A | P ) F" :=
 (\big[meet / \top]_(i in A | P%B) F%O) : order_scope.
Notation "\meet^l_ ( i 'in' A ) F" :=
 (\big[meet / \top]_(i in A) F%O) : order_scope.

Notation "\min^l_ i F" :=
  (\big[min/top]_i F) : order_scope.
Notation "\min^l_ ( i <- r | P ) F" :=
  (\big[min/top]_(i <- r | P%B) F%O) : order_scope.
Notation "\min^l_ ( i < r ) F" :=
  (\big[min/top]_(i <- r) F%O) : order_scope.
Notation "\min^l_ ( m <= i < n | P ) F" :=
  (\big[min/top]_(m i < n | P%B) F%O) : order_scope.
Notation "\min^l_ ( m <= i < n ) F" :=
  (\big[min/top]_(m i < n) F%O) : order_scope.
Notation "\min^l_ ( i | P ) F" :=
  (\big[min/top]_(i | P%B) F%O) : order_scope.
Notation "\min^l_ ( i : t | P ) F" :=
  (\big[min/top]_(i : t | P%B) F%O) (only parsing) : order_scope.
Notation "\min^l_ ( i : t ) F" :=
  (\big[min/top]_(i : t) F%O) (only parsing) : order_scope.
Notation "\min^l_ ( i < n | P ) F" :=
  (\big[min/top]_(i < n | P%B) F%O) : order_scope.
Notation "\min^l_ ( i < n ) F" :=
  (\big[min/top]_(i < n) F%O) : order_scope.
Notation "\min^l_ ( i 'in' A | P ) F" :=
  (\big[min/top]_(i in A | P%B) F%O) : order_scope.
Notation "\min^l_ ( i 'in' A ) F" :=
  (\big[min/top]_(i in A) F%O) : order_scope.

Notation "\max^l_ i F" :=
  (\big[max/bottom]_i F%O) : order_scope.
Notation "\max^l_ ( i <- r | P ) F" :=
  (\big[max/bottom]_(i <- r | P%B) F%O) : order_scope.
Notation "\max^l_ ( i < r ) F" :=
  (\big[max/bottom]_(i <- r) F%O) : order_scope.
Notation "\max^l_ ( m <= i < n | P ) F" :=
  (\big[max/bottom]_(m i < n | P%B) F%O) : order_scope.
Notation "\max^l_ ( m <= i < n ) F" :=
  (\big[max/bottom]_(m i < n) F%O) : order_scope.
Notation "\max^l_ ( i | P ) F" :=
  (\big[max/bottom]_(i | P%B) F%O) : order_scope.
Notation "\max^l_ ( i : t | P ) F" :=
  (\big[max/bottom]_(i : t | P%B) F%O) (only parsing) : order_scope.
Notation "\max^l_ ( i : t ) F" :=
  (\big[max/bottom]_(i : t) F%O) (only parsing) : order_scope.
Notation "\max^l_ ( i < n | P ) F" :=
  (\big[max/bottom]_(i < n | P%B) F%O) : order_scope.
Notation "\max^l_ ( i < n ) F" :=
  (\big[max/bottom]_(i < n) F%O) : order_scope.
Notation "\max^l_ ( i 'in' A | P ) F" :=
  (\big[max/bottom]_(i in A | P%B) F%O) : order_scope.
Notation "\max^l_ ( i 'in' A ) F" :=
  (\big[max/bottom]_(i in A) F%O) : order_scope.

End SeqLexiSyntax.

We declare an alias of the cartesian product which has canonical product order.

Module ProdOrder.

Export ProdOrder.

Local Open Scope type_scope. (* FIXME *)

Section POrder.
Context (disp1 disp2 disp3 : disp_t).
Context (T1 : porderType disp1) (T2 : porderType disp2).

Fact anti : antisymmetric (@le disp1 disp2 disp2 T1 T2).

End POrder.

Section POrder.
Context (disp1 disp2 disp3 : disp_t).
Context (T1 : porderType disp1) (T2 : porderType disp2).
Local Notation "T1 * T2" := (type disp3 T1 T2) : type_scope.
Implicit Types (x y : T1 × T2).

Let T1' : Type := T1.
Let T2' : Type := T2.

#[export]
HB.instance Definition _ :=
  Preorder_isDuallyPOrder.Build disp3 (T1 × T2)
  (@anti _ _ T1' T2') (@anti _ _ T1^d T2^d).

Lemma ltEprod x y : (x < y) = [&& x != y, x.1 y.1 & x.2 y.2].

Lemma lt_pair (x1 y1 : T1) (x2 y2 : T2) : (x1, x2) < (y1, y2) :> T1 × T2 =
  [&& (x1 != y1) || (x2 != y2), x1 y1 & x2 y2].

End POrder.

Section MeetSemilattice.
Context (disp1 disp2 : disp_t).
Context (T1 : meetSemilatticeType disp1) (T2 : meetSemilatticeType disp2).
Local Notation "T1 * T2" := (type (Disp tt tt) T1 T2) : type_scope.
Implicit Types (x y : T1 × T2).

Let meet x y := (x.1 `&` y.1, x.2 `&` y.2).

Fact lexI x y z : (x meet y z) = (x y) && (x z).

End MeetSemilattice.

Section MeetSemilattice.
Context (disp1 disp2 disp3 : disp_t).
Context (T1 : meetSemilatticeType disp1) (T2 : meetSemilatticeType disp2).
Local Notation "T1 * T2" := (type disp3 T1 T2) : type_scope.
Implicit Types (x y : T1 × T2).

Let T1' : Type := T1.
Let T2' : Type := T2.

Definition meet x y := (x.1 `&` y.1, x.2 `&` y.2).

#[export]
HB.instance Definition _ :=
  @POrder_isMeetSemilattice.Build disp3 (T1 × T2) meet (@lexI _ _ T1' T2').

Lemma meetEprod x y : x `&` y = (x.1 `&` y.1, x.2 `&` y.2).

End MeetSemilattice.

Section JoinSemilattice.
Context (disp1 disp2 disp3 : disp_t).
Context (T1 : joinSemilatticeType disp1) (T2 : joinSemilatticeType disp2).
Local Notation "T1 * T2" := (type disp3 T1 T2) : type_scope.
Implicit Types (x y : T1 × T2).

Definition join x y := (x.1 `|` y.1, x.2 `|` y.2).

#[export]
HB.instance Definition _ :=
  @POrder_isJoinSemilattice.Build disp3 (T1 × T2) join
    (fun x y z @lexI _ _ T1^d T2^d z x y).

Lemma joinEprod x y : x `|` y = (x.1 `|` y.1, x.2 `|` y.2).

End JoinSemilattice.

FIXME: use HB.saturate
#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : bPOrderType disp1) (T2 : bPOrderType disp2) :=
  POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : tPOrderType disp1) (T2 : tPOrderType disp2) :=
  POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : tbPOrderType disp1) (T2 : tbPOrderType disp2) :=
  POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : bMeetSemilatticeType disp1) (T2 : bMeetSemilatticeType disp2) :=
  POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : tMeetSemilatticeType disp1) (T2 : tMeetSemilatticeType disp2) :=
  POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : tbMeetSemilatticeType disp1) (T2 : tbMeetSemilatticeType disp2) :=
  POrder.on (type disp3 T1 T2).

#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : bJoinSemilatticeType disp1) (T2 : bJoinSemilatticeType disp2) :=
  POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : tJoinSemilatticeType disp1) (T2 : tJoinSemilatticeType disp2) :=
  POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : tbJoinSemilatticeType disp1) (T2 : tbJoinSemilatticeType disp2) :=
  POrder.on (type disp3 T1 T2).

#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : latticeType disp1) (T2 : latticeType disp2) :=
  POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : bLatticeType disp1) (T2 : bLatticeType disp2) :=
  POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : tLatticeType disp1) (T2 : tLatticeType disp2) :=
  POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (disp1 disp2 disp3 : disp_t)
  (T1 : tbLatticeType disp1) (T2 : tbLatticeType disp2) :=
  POrder.on (type disp3 T1 T2).
/FIXME

Section DistrLattice.
Context (disp1 disp2 : disp_t).
Context (T1 : distrLatticeType disp1) (T2 : distrLatticeType disp2).
Local Notation "T1 * T2" := (type (Disp tt tt) T1 T2) : type_scope.

Fact meetUl : @left_distributive (T1 × T2) _ Order.meet Order.join.

End DistrLattice.

Section DistrLattice.
Context (disp1 disp2 disp3 : disp_t).
Context (T1 : distrLatticeType disp1) (T2 : distrLatticeType disp2).
Local Notation "T1 * T2" := (type disp3 T1 T2) : type_scope.

Let T1' : Type := T1.
Let T2' : Type := T2.

#[export]
HB.instance Definition _ := Lattice_isDistributive.Build disp3 (T1 × T2)
  (@meetUl _ _ T1' T2') (@meetUl _ _ T1^d T2^d).

End DistrLattice.

FIXME: use HB.saturate
/FIXME

Section CDistrLattice.
Context (disp1 disp2 : disp_t).
Context (T1 : cDistrLatticeType disp1) (T2 : cDistrLatticeType disp2).
Local Notation "T1 * T2" := (type (Disp tt tt) T1 T2) : type_scope.
Implicit Types (x y z : T1 × T2).

Let rcompl x y z := (rcompl x.1 y.1 z.1, rcompl x.2 y.2 z.2).

Fact rcomplPmeet x y z : ((x `&` y) `|` z) `&` rcompl x y z = x `&` y.

End CDistrLattice.

Section CDistrLattice.
Context (disp1 disp2 disp3 : disp_t).
Context (T1 : cDistrLatticeType disp1) (T2 : cDistrLatticeType disp2).
Local Notation "T1 * T2" := (type disp3 T1 T2) : type_scope.
Implicit Types (x y z : T1 × T2).

Let T1' : Type := T1.
Let T2' : Type := T2.

Definition rcompl x y z := (rcompl x.1 y.1 z.1, rcompl x.2 y.2 z.2).

#[export]
HB.instance Definition _ :=
  @DistrLattice_hasRelativeComplement.Build disp3 (T1 × T2)
    rcompl (@rcomplPmeet _ _ T1' T2')
    (fun x y @rcomplPmeet _ _ T1^d T2^d y x).

Lemma rcomplEprod x y z :
  rcompl x y z = (Order.rcompl x.1 y.1 z.1, Order.rcompl x.2 y.2 z.2).

End CDistrLattice.

Section CBDistrLattice.
Context (disp1 disp2 : disp_t).
Context (T1 : cbDistrLatticeType disp1) (T2 : cbDistrLatticeType disp2).
Local Notation "T1 * T2" := (type (Disp tt tt) T1 T2) : type_scope.
Implicit Types (x y : T1 × T2).

Let diff x y := (diff x.1 y.1, diff x.2 y.2).

Fact diffErcompl x y : diff x y = rcompl \bot x y.

End CBDistrLattice.

Section CBDistrLattice.
Context (disp1 disp2 disp3 : disp_t).
Context (T1 : cbDistrLatticeType disp1) (T2 : cbDistrLatticeType disp2).
Local Notation "T1 * T2" := (type disp3 T1 T2) : type_scope.
Implicit Types (x y : T1 × T2).

Let T1' : Type := T1.
Let T2' : Type := T2.

Definition diff x y := (diff x.1 y.1, diff x.2 y.2).

#[export]
HB.instance Definition _ :=
  @CDistrLattice_hasSectionalComplement.Build disp3 (T1 × T2)
    diff (@diffErcompl _ _ T1' T2').

Lemma diffEprod x y : x `\` y = (x.1 `\` y.1, x.2 `\` y.2).

End CBDistrLattice.

Section CTDistrLattice.
Context (disp1 disp2 disp3 : disp_t).
Context (T1 : ctDistrLatticeType disp1) (T2 : ctDistrLatticeType disp2).
Local Notation "T1 * T2" := (type disp3 T1 T2) : type_scope.
Implicit Types (x y : T1 × T2).

Definition codiff x y := (codiff x.1 y.1, codiff x.2 y.2).

#[export]
HB.instance Definition _ :=
  @CDistrLattice_hasDualSectionalComplement.Build disp3 (T1 × T2)
    codiff (@diffErcompl _ _ T1^d T2^d).

Lemma codiffEprod x y :
  codiff x y = (Order.codiff x.1 y.1, Order.codiff x.2 y.2).

End CTDistrLattice.

Section CTBDistrLattice.
Context (disp1 disp2 : disp_t).
Context (T1 : ctbDistrLatticeType disp1) (T2 : ctbDistrLatticeType disp2).
Local Notation "T1 * T2" := (type (Disp tt tt) T1 T2) : type_scope.
Implicit Types (x : T1 × T2).

Let compl x := (~` x.1, ~` x.2).

Fact complEdiff x : compl x = (\top : T1 × T2) `\` x.

End CTBDistrLattice.

Section CTBDistrLattice.
Context (disp1 disp2 disp3 : disp_t).
Context (T1 : ctbDistrLatticeType disp1) (T2 : ctbDistrLatticeType disp2).
Local Notation "T1 * T2" := (type disp3 T1 T2) : type_scope.
Implicit Types (x : T1 × T2).

Let T1' : Type := T1.
Let T2' : Type := T2.

Definition compl x := (~` x.1, ~` x.2).

#[export]
HB.instance Definition _ := @CDistrLattice_hasComplement.Build _ (T1 × T2) compl
  (@complEdiff _ _ T1' T2') (@complEdiff _ _ T1^d T2^d).

Lemma complEprod x : ~` x = (~` x.1, ~` x.2).

End CTBDistrLattice.

FIXME: use HB.saturate
Section FinOrder.
Context (disp1 disp2 disp3 : disp_t).

#[export]
HB.instance Definition _ (T1 : finPOrderType disp1)
  (T2 : finPOrderType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finBPOrderType disp1)
  (T2 : finBPOrderType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finTPOrderType disp1)
  (T2 : finTPOrderType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finTBPOrderType disp1)
  (T2 : finTBPOrderType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finMeetSemilatticeType disp1)
  (T2 : finMeetSemilatticeType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finBMeetSemilatticeType disp1)
  (T2 : finBMeetSemilatticeType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finJoinSemilatticeType disp1)
  (T2 : finJoinSemilatticeType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finTJoinSemilatticeType disp1)
  (T2 : finTJoinSemilatticeType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finLatticeType disp1)
  (T2 : finLatticeType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finTBLatticeType disp1)
  (T2 : finTBLatticeType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finDistrLatticeType disp1)
  (T2 : finDistrLatticeType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finTBDistrLatticeType disp1)
  (T2 : finTBDistrLatticeType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finCDistrLatticeType disp1)
  (T2 : finCDistrLatticeType disp2) := POrder.on (type disp3 T1 T2).
#[export]
HB.instance Definition _ (T1 : finCTBDistrLatticeType disp1)
  (T2 : finCTBDistrLatticeType disp2) := POrder.on (type disp3 T1 T2).

End FinOrder.
/FIXME

Module Exports.
Definition ltEprod := @ltEprod.
Definition lt_pair := @lt_pair.
Definition meetEprod := @meetEprod.
Definition joinEprod := @joinEprod.
Definition rcomplEprod := @rcomplEprod.
Definition diffEprod := @diffEprod.
Definition codiffEprod := @codiffEprod.
Definition complEprod := @complEprod.
End Exports.
End ProdOrder.

Module DefaultProdOrder.

Export DefaultProdOrder.

Section DefaultProdOrder.
Context {disp1 disp2 : disp_t}.

Let prod T1 T2 := T1 ×prod[prod_display disp1 disp2] T2.

FIXME: Scopes of arguments are broken in several places. FIXME: Declaring a bunch of copies is still a bit painful.
/FIXME
We declare lexicographic ordering on dependent pairs.

Module SigmaOrder.
Section SigmaOrder.

Context {disp1 disp2 : disp_t}.

Section POrder.
Context (T : porderType disp1) (T' : T porderType disp2).
Implicit Types (x y : {t : T & T' t}).

Definition le x y := (tag x tag y) &&
  ((tag x tag y) ==> (tagged x tagged_as x y)).
Definition lt x y := (tag x tag y) &&
  ((tag x tag y) ==> (tagged x < tagged_as x y)).

Fact refl : reflexive le.

Fact anti : antisymmetric le.

Fact trans : transitive le.

Fact lt_le_def x y : lt x y = le x y && ~~ le y x.

#[export]
HB.instance Definition _ :=
  isPreorder.Build disp2 {t : T & T' t} lt_le_def refl trans.
#[export]
HB.instance Definition _ :=
  Preorder_isPOrder.Build disp2 {t : T & T' t} anti.

Lemma leEsig x y : x y =
  (tag x tag y) && ((tag x tag y) ==> (tagged x tagged_as x y)).

Lemma ltEsig x y : x < y =
  (tag x tag y) && ((tag x tag y) ==> (tagged x < tagged_as x y)).

Lemma le_Taggedl x (u : T' (tag x)) : (Tagged T' u x) = (u tagged x).

Lemma le_Taggedr x (u : T' (tag x)) : (x Tagged T' u) = (tagged x u).

Lemma lt_Taggedl x (u : T' (tag x)) : (Tagged T' u < x) = (u < tagged x).

Lemma lt_Taggedr x (u : T' (tag x)) : (x < Tagged T' u) = (tagged x < u).

End POrder.

Section BPOrder.
Context (T : bPOrderType disp1) (T' : T bPOrderType disp2).

Fact le0x (x : {t : T & T' t}) : Tagged T' (\bot : T' \bot) x.

#[export]
HB.instance Definition _ := hasBottom.Build _ {t : T & T' t} le0x.

Lemma botEsig : \bot = Tagged T' (\bot : T' \bot).

End BPOrder.

Section TPOrder.
Context (T : tPOrderType disp1) (T' : T tPOrderType disp2).

Fact lex1 (x : {t : T & T' t}) : x Tagged T' (\top : T' \top).

#[export]
HB.instance Definition _ := hasTop.Build _ {t : T & T' t} lex1.

Lemma topEsig : \top = Tagged T' (\top : T' \top).

End TPOrder.

Section Total.
Context (T : orderType disp1) (T' : T orderType disp2).
Implicit Types (x y : {t : T & T' t}).

Fact total : total (<=%O : rel {t : T & T' t}).

#[export]
HB.instance Definition _ := POrder_isTotal.Build _ {t : T & T' t} total.

End Total.

FIXME: use HB.saturate
/FIXME

End SigmaOrder.

Module Exports.
Definition leEsig := @leEsig.
Definition ltEsig := @ltEsig.
Definition le_Taggedl := @le_Taggedl.
Definition lt_Taggedl := @lt_Taggedl.
Definition le_Taggedr := @le_Taggedr.
Definition lt_Taggedr := @lt_Taggedr.
Definition topEsig := @topEsig.
Definition botEsig := @botEsig.
End Exports.
End SigmaOrder.

We declare an alias of the cartesian product, which has canonical lexicographic order.

Module ProdLexiOrder.

Export ProdLexiOrder.

Local Open Scope type_scope. (* FIXME *)

Section POrder.
Context (disp1 disp2 disp3 : disp_t).
Context (T1 : porderType disp1) (T2 : porderType disp2).

Fact anti : antisymmetric (@le disp1 disp2 disp2 T1 T2).

End POrder.

Section POrder.
Context (disp1 disp2 disp3 : disp_t).
Context (T1 : porderType disp1) (T2 : porderType disp2).
Local Notation "T1 * T2" := (type disp3 T1 T2) : type_scope.
Implicit Types (x y : T1 × T2).

Let T1' : Type := T1.
Let T2' : Type := T2.

#[export]
HB.instance Definition _ :=
  Preorder_isDuallyPOrder.Build disp3 (T1 × T2)
  (@anti _ _ T1' T2') (@anti _ _ T1^d T2^d).

End POrder.

Section Total.
Context (disp1 disp2 disp3 : disp_t).
Context (T1 : orderType disp1) (T2 : orderType disp2).
Local Notation "T1 * T2" := (type disp3 T1 T2) : type_scope.

Fact total : total (<=%O : rel (T1 × T2)).

FIXME: In order to dualize this instance, we have to dualize the [POrder_isTotal] factory. However, [min] and max are not definitional dual (while [min x y] and [max y x] are).
FIXME: use HB.saturate
/FIXME
FIXME: Scopes of arguments are broken in several places. FIXME: Declaring a bunch of copies is still a bit painful.
/FIXME
We declare an alias of the sequences, which has canonical product order.

Module SeqProdOrder.

Export SeqProdOrder.

Section SeqProdOrder.

Context {disp disp' : disp_t}.

Local Notation seq := (type disp').

Section POrder.
Variable T : porderType disp.
Implicit Types s : seq T.

Fact anti : antisymmetric (@le disp disp' T).

#[export]
HB.instance Definition _ := Preorder_isPOrder.Build disp' (seq T) anti.

End POrder.

Section MeetSemilattice.
Context (T : meetSemilatticeType disp).
Implicit Types (s : seq T).

Fixpoint meet s1 s2 :=
  match s1, s2 with
    | x1 :: s1', x2 :: s2'(x1 `&` x2) :: meet s1' s2'
    | _, _[::]
  end.

Fact lexI s1 s2 s3 : (s1 meet s2 s3) = (s1 s2) && (s1 s3).

#[export]
HB.instance Definition _ := @POrder_isMeetSemilattice.Build _ (seq T) meet lexI.

Lemma meetEseq s1 s2 : s1 `&` s2 = [seq x.1 `&` x.2 | x <- zip s1 s2].

Lemma meet_cons x1 s1 x2 s2 :
  (x1 :: s1 : seq T) `&` (x2 :: s2) = (x1 `&` x2) :: s1 `&` s2.

End MeetSemilattice.

Section JoinSemilattice.
Context (T : joinSemilatticeType disp).
Implicit Types (s : seq T).

Fixpoint join s1 s2 :=
  match s1, s2 with
    | [::], _s2 | _, [::]s1
    | x1 :: s1', x2 :: s2'(x1 `|` x2) :: join s1' s2'
  end.

Fact leUx s1 s2 s3 : (join s1 s2 s3) = (s1 s3) && (s2 s3).

#[export]
HB.instance Definition _ := @POrder_isJoinSemilattice.Build _ (seq T) join leUx.

Lemma joinEseq s1 s2 : s1 `|` s2 =
  match s1, s2 with
    | [::], _s2 | _, [::]s1
    | x1 :: s1', x2 :: s2'(x1 `|` x2) :: ((s1' : seq _) `|` s2')
  end.

Lemma join_cons x1 s1 x2 s2 :
  (x1 :: s1 : seq T) `|` (x2 :: s2) = (x1 `|` x2) :: s1 `|` s2.

End JoinSemilattice.

FIXME: use HB.saturate
/FIXME
We declare an alias of the sequences, which has canonical lexicographic order.

Module SeqLexiOrder.

Export SeqLexiOrder.

Section SeqLexiOrder.

Context {disp disp' : disp_t}.

Local Notation seq := (type disp').

Section POrder.
Variable T : porderType disp.
Implicit Types s : seq T.

Fact anti: antisymmetric (@le disp disp' T).

#[export]
HB.instance Definition _ := Preorder_isPOrder.Build disp' (seq T) anti.

Lemma neqhead_lexiE (x y : T) s1 s2 : x != y
  (x :: s1 y :: s2 :> seq _) = (x < y).

Lemma neqhead_ltxiE (x y : T) s1 s2 : x != y
  (x :: s1 < y :: s2 :> seq _) = (x < y).

End POrder.

Section Total.
Context (T : orderType disp).

Fact total : total (<=%O : rel (seq T)).

#[export]
HB.instance Definition _ := POrder_isTotal.Build _ (seq T) total.

End Total.

End SeqLexiOrder.

Module Exports.


Definition neqhead_lexiE := @neqhead_lexiE.
Definition neqhead_ltxiE := @neqhead_ltxiE.

End Exports.
End SeqLexiOrder.

Module DefaultSeqLexiOrder.

Export DefaultSeqLexiOrder.

Section DefaultSeqLexiOrder.
Context {disp : disp_t}.


End DefaultSeqLexiOrder.
End DefaultSeqLexiOrder.

We declare an alias of the tuples, which has canonical product order.

Module TupleProdOrder.

Export TupleProdOrder.

Import DefaultSeqProdOrder.

Section TupleProdOrder.

Context {disp disp' : disp_t}.
Local Notation "n .-tuple" := (type disp' n) : type_scope.

Section POrder.
Implicit Types (T : porderType disp).

#[export, warning="-HB.no-new-instance"]
HB.instance Definition _ n T := SubChoice.on (n.-tuple T).
#[export]
HB.instance Definition _ n T :=
  [SubChoice_isSubPOrder of n.-tuple T by <: with disp'].

End POrder.

FIXME: use HB.saturate
FIXME: use HB.saturate
/FIXME
FIXME: use HB.saturate
/FIXME
FIXME: use HB.saturate
/FIXME

Section CDistrLattice.
Context (n : nat) (T : cDistrLatticeType disp).
Implicit Types (t : n.-tuple T).

Definition rcompl t1 t2 t3 :=
  [tuple rcompl (tnth t1 i) (tnth t2 i) (tnth t3 i) | i < n].

Fact rcomplPmeet x y z : ((x `&` y) `|` z) `&` rcompl x y z = x `&` y.

Fact rcomplPjoin x y z : ((y `|` x) `&` z) `|` rcompl x y z = y `|` x.

#[export]
HB.instance Definition _ :=
  @DistrLattice_hasRelativeComplement.Build _ (n.-tuple T)
    rcompl rcomplPmeet rcomplPjoin.

Lemma tnth_rcompl t1 t2 t3 i :
  tnth (Order.rcompl t1 t2 t3) i =
    Order.rcompl (tnth t1 i) (tnth t2 i) (tnth t3 i).

Lemma rcomplEtprod t1 t2 t3 :
  Order.rcompl t1 t2 t3 =
    [tuple Order.rcompl (tnth t1 i) (tnth t2 i) (tnth t3 i) | i < n].

End CDistrLattice.

Section CBDistrLattice.
Context (n : nat) (T : cbDistrLatticeType disp).
Implicit Types (t : n.-tuple T).

Definition diff t1 t2 : n.-tuple T := [tuple tnth t1 i `\` tnth t2 i | i < n].

Fact diffErcompl t1 t2 : diff t1 t2 = rcompl \bot t1 t2.

#[export] HB.instance Definition _ :=
  @CDistrLattice_hasSectionalComplement.Build _ (n.-tuple T) diff diffErcompl.

Lemma tnth_diff t1 t2 i : tnth (diff t1 t2) i = tnth t1 i `\` tnth t2 i.

Lemma diffEtprod t1 t2 : t1 `\` t2 = [tuple tnth t1 i `\` tnth t2 i | i < n].

End CBDistrLattice.

Section CTDistrLattice.
Context (n : nat) (T : ctDistrLatticeType disp).
Implicit Types (t : n.-tuple T).

Definition codiff t1 t2 : n.-tuple T :=
  [tuple Order.codiff (tnth t1 i) (tnth t2 i) | i < n].

Fact codiffErcompl t1 t2 : codiff t1 t2 = rcompl t1 \top t2.

#[export] HB.instance Definition _ :=
  @CDistrLattice_hasDualSectionalComplement.Build _ (n.-tuple T)
    codiff codiffErcompl.

Lemma tnth_codiff t1 t2 i :
  tnth (Order.codiff t1 t2) i = Order.codiff (tnth t1 i) (tnth t2 i).

Lemma codiffEtprod t1 t2 :
  Order.codiff t1 t2 = [tuple Order.codiff (tnth t1 i) (tnth t2 i) | i < n].

End CTDistrLattice.

Section CTBDistrLattice.
Context (n : nat) (T : ctbDistrLatticeType disp).
Implicit Types (t : n.-tuple T).

Definition compl t : n.-tuple T := map_tuple compl t.

Fact complEdiff t : compl t = (\top : n.-tuple T) `\` t.

Fact complEcodiff t : compl t = codiff (\bot : n.-tuple T) t.

#[export] HB.instance Definition _ :=
  @CDistrLattice_hasComplement.Build _ (n.-tuple T)
    compl complEdiff complEcodiff.

Lemma tnth_compl t i : tnth (~` t) i = ~` tnth t i.

Lemma complEtprod t : ~` t = map_tuple Order.compl t.

End CTBDistrLattice.

FIXME: use HB.saturate
/FIXME

End TupleProdOrder.

Module Exports.


Definition tnth_meet := @tnth_meet.
Definition meetEtprod := @meetEtprod.
Definition tnth_join := @tnth_join.
Definition joinEtprod := @joinEtprod.
Definition tnth_rcompl := @tnth_rcompl.
Definition rcomplEtprod := @rcomplEtprod.
Definition tnth_diff := @tnth_diff.
Definition diffEtprod := @diffEtprod.
Definition tnth_codiff := @tnth_codiff.
Definition codiffEtprod := @codiffEtprod.
Definition tnth_compl := @tnth_compl.
Definition complEtprod := @complEtprod.

End Exports.
End TupleProdOrder.

Module DefaultTupleProdOrder.

Export DefaultTupleProdOrder.

Section DefaultTupleProdOrder.
Context {disp : disp_t}.


End DefaultTupleProdOrder.
End DefaultTupleProdOrder.

We declare an alias of the tuples, which has canonical lexicographic order.

Module TupleLexiOrder.

Export TupleLexiOrder.

Section TupleLexiOrder.
Import DefaultSeqLexiOrder.

Context {disp disp' : disp_t}.
Local Notation "n .-tuple" := (type disp' n) : type_scope.

Section POrder.
Implicit Types (n : nat) (T : porderType disp).

#[export] HB.instance Definition _ n T :=
  [SubChoice_isSubPOrder of n.-tuple T by <: with disp'].

Lemma lexi_tupleP n T (t1 t2 : n.-tuple T) :
   reflect ( k : 'I_n.+1, i : 'I_n, (i k)%N
               tnth t1 i tnth t2 i ?= iff (i != k :> nat)) (t1 t2).

Lemma ltxi_tupleP n T (t1 t2 : n.-tuple T) :
   reflect ( k : 'I_n, i : 'I_n, (i k)%N
               tnth t1 i tnth t2 i ?= iff (i != k :> nat)) (t1 < t2).

Lemma ltxi_tuplePlt n T (t1 t2 : n.-tuple T) : reflect
  (exists2 k : 'I_n, i : 'I_n, (i < k)%N tnth t1 i = tnth t2 i
                                                 & tnth t1 k < tnth t2 k)
  (t1 < t2).

End POrder.

#[export] HB.instance Definition _ n (T : orderType disp) :=
  [SubChoice_isSubOrder of n.-tuple T by <: with disp'].

FIXME: use HB.saturate
/FIXME

End TupleLexiOrder.

Module Exports.


Definition lexi_tupleP := @lexi_tupleP.
Arguments lexi_tupleP {disp disp' n T t1 t2}.
Definition ltxi_tupleP := @ltxi_tupleP.
Arguments ltxi_tupleP {disp disp' n T t1 t2}.
Definition ltxi_tuplePlt := @ltxi_tuplePlt.
Arguments ltxi_tuplePlt {disp disp' n T t1 t2}.

End Exports.
End TupleLexiOrder.

Module DefaultTupleLexiOrder.

Export DefaultTupleLexiOrder.

Section DefaultTupleLexiOrder.
Context {disp : disp_t}.


End DefaultTupleLexiOrder.
End DefaultTupleLexiOrder.

We declare an alias of the sets, which is canonically ordered by inclusion

Module SetSubsetOrder.

Export SetSubsetOrder.

Section SetSubsetOrder.

Context {disp : disp_t} {T : finType}.
Local Notation "{ 'subset' T }" := (type disp T).
Implicit Type (A B C : type disp T).

Lemma setKUC B A : A :&: (A :|: B) = A.

Lemma setKIC B A : A :|: (A :&: B) = A.

Fact le_anti : antisymmetric (fun A BA \subset B).

#[export]
HB.instance Definition _ := Preorder_isPOrder.Build disp {subset T} le_anti.
#[export]
HB.instance Definition _ := POrder_Meet_isDistrLattice.Build disp {subset T}
  (@setIC _) (@setUC _) (@setIA _) (@setUA _) setKUC setKIC le_def (@setIUl _).

Lemma setIDv A B : B :&: (A :\: B) = set0.

#[export]
HB.instance Definition _ :=
  @BDistrLattice_hasSectionalComplement.Build disp {subset T}
    (@setD _) setIDv (@setID _).

Lemma setTDsym A : ~: A = setT :\: A.

#[export]
HB.instance Definition _ :=
  CBDistrLattice_hasComplement.Build disp {subset T} setTDsym.

Lemma meetEsubset A B : A `&` B = A :&: B.
Lemma joinEsubset A B : A `|` B = A :|: B.
Lemma botEsubset : \bot = set0 :> {subset T}.
Lemma topEsubset : \top = setT :> {subset T}.
Lemma subEsubset A B : A `\` B = A :\: B.
Lemma complEsubset A : ~` A = ~: A.

End SetSubsetOrder.

Module Exports.


Definition meetEsubset := @meetEsubset.
Definition joinEsubset := @joinEsubset.
Definition botEsubset := @botEsubset.
Definition topEsubset := @topEsubset.
Definition subEsubset := @subEsubset.
Definition complEsubset := @complEsubset.

End Exports.
End SetSubsetOrder.
Export SetSubsetOrder.Exports.

Module DefaultSetSubsetOrder.

Export DefaultSetSubsetOrder.


End DefaultSetSubsetOrder.

Lemma mono_unique d (T T' : finPOrderType d) (f g : T T') :
    total (<=%O : rel T) (#|T'| #|T|)%N
    {mono f : x y / x y} {mono g : x y / x y}
  f =1 g.

This module should be exported on demand, as in module tagnat below
Module Import EnumVal.

Export EnumVal.

Section EnumVal.
Import OrdinalOrder.Exports.
Variables (d : disp_t) (T : finPOrderType d).
Implicit Types (x : T) (A : {pred T}).

Section total.
We circumvent a shortcoming of finOrderType which requires the type to be nonempty and we do not want to rule this out
Hypothesis (leT_total : total (<=%O : rel T)).

Lemma le_enum_val A : {mono @enum_val _ _ A : i j / i j}.

Lemma le_enum_rank_in x0 A (Ax0 : x0 \in A) :
  {in A &, {mono enum_rank_in Ax0 : x y / x y}}.

Lemma le_enum_rank : {mono @enum_rank d T : i j / i j}.

End total.
End EnumVal.
End EnumVal.

Notation le_enum_val := le_enum_val.
Notation le_enum_rank_in := le_enum_rank_in.
Notation le_enum_rank := le_enum_rank.

Module Syntax.
Export PreOSyntax.
Export DualSyntax.
Export DvdSyntax.
Export LatticeSyntax.
Export BLatticeSyntax.
Export TLatticeSyntax.
Export CBDistrLatticeSyntax.
Export CTBDistrLatticeSyntax.
Export DualSyntax.
Export DvdSyntax.
End Syntax.

Module LTheory.
Export PreorderTheory.
Export BPreorderTheory.
Export TPreorderTheory.
Export DualPreorder. (* FIXME? *)
Export PreOCoercions.
Export PreorderTheory.
Export POrderTheory.
Export BPOrderTheory.
Export TPOrderTheory.
Export MeetTheory.
Export BMeetTheory.
Export TMeetTheory.
Export JoinTheory.
Export BJoinTheory.
Export TJoinTheory.
Export LatticeTheory.
Export DistrLatticeTheory.
Export BDistrLatticeTheory.
Export TDistrLatticeTheory.
Export DualTotalTheory. (* FIXME? *)
Export DualOrder. (* FIXME? *)

Export OrderMorphismTheory.
Export LatticeMorphismTheory.
Export BLatticeMorphismTheory.
Export TLatticeMorphismTheory.

Export ClosedPredicates.
Export LatticePred.

Export SubPreorderTheory.
End LTheory.

Module CTheory.
Export LTheory.
Export CDistrLatticeTheory.
Export CBDistrLatticeTheory.
Export CTDistrLatticeTheory.
Export CTBDistrLatticeTheory.
End CTheory.

Module TTheory.
Export LTheory TotalTheory.
End TTheory.

Module Theory.
Export CTheory TotalTheory.
End Theory.

Module Exports.
End Exports.
End Order.

Export Order.Exports.

Export Order.Syntax.

Export order.Order.Exports.
Export Order.POrder.Exports.
Export Order.BPOrder.Exports.
Export Order.TPOrder.Exports.
Export Order.TBPOrder.Exports.
Export Order.MeetSemilattice.Exports.
Export Order.BMeetSemilattice.Exports.
Export Order.TMeetSemilattice.Exports.
Export Order.TBMeetSemilattice.Exports.
Export Order.JoinSemilattice.Exports.
Export Order.BJoinSemilattice.Exports.
Export Order.TJoinSemilattice.Exports.
Export Order.TBJoinSemilattice.Exports.
Export Order.Lattice.Exports.
Export Order.BLattice.Exports.
Export Order.TLattice.Exports.
Export Order.TBLattice.Exports.
Export Order.DistrLattice.Exports.
Export Order.BDistrLattice.Exports.
Export Order.TDistrLattice.Exports.
Export Order.TBDistrLattice.Exports.
Export Order.Total.Exports.
Export Order.BTotal.Exports.
Export Order.TTotal.Exports.
Export Order.TBTotal.Exports.
Export Order.CDistrLattice.Exports.
Export Order.CBDistrLattice.Exports.
Export Order.CTDistrLattice.Exports.
Export Order.CTBDistrLattice.Exports.
Export Order.FinPOrder.Exports.
Export Order.FinBPOrder.Exports.
Export Order.FinTPOrder.Exports.
Export Order.FinTBPOrder.Exports.
Export Order.FinMeetSemilattice.Exports.
Export Order.FinBMeetSemilattice.Exports.
Export Order.FinJoinSemilattice.Exports.
Export Order.FinTJoinSemilattice.Exports.
Export Order.FinLattice.Exports.
Export Order.FinTBLattice.Exports.
Export Order.FinDistrLattice.Exports.
Export Order.FinTBDistrLattice.Exports.
Export Order.FinTotal.Exports.
Export Order.FinTBTotal.Exports.
Export Order.FinCDistrLattice.Exports.
Export Order.FinCTBDistrLattice.Exports.

FIXME: check if covered by Order.Exports Export Order.NatOrder.Exports. Export Order.NatMonotonyTheory. Export Order.NatDvd.Exports. Export Order.OrdinalOrder.Exports. Export Order.BoolOrder.Exports. Export Order.ProdOrder.Exports. Export Order.SigmaOrder.Exports. Export Order.ProdLexiOrder.Exports. Export Order.SeqProdOrder.Exports. Export Order.SeqLexiOrder.Exports. Export Order.TupleProdOrder.Exports. Export Order.TupleLexiOrder.Exports.

Module DefaultProdOrder := Order.DefaultProdOrder.
Module DefaultSeqProdOrder := Order.DefaultSeqProdOrder.
Module DefaultTupleProdOrder := Order.DefaultTupleProdOrder.
Module DefaultProdLexiOrder := Order.DefaultProdLexiOrder.
Module DefaultSeqLexiOrder := Order.DefaultSeqLexiOrder.
Module DefaultTupleLexiOrder := Order.DefaultTupleLexiOrder.

Import Order.Theory.

Module tagnat.
Section tagnat.
Import Order.EnumVal.

Context {n : nat} {p_ : 'I_n nat}.

Local Notation ordsum := 'I_(\sum_i p_ i)%N.
Local Notation T := {i & 'I_(p_ i)}.

Implicit Types (i : 'I_n) (s : ordsum) (p : T).

Lemma card : #|{: T}| = \sum_i p_ i.

Definition sig : ordsum T := enum_val \o cast_ord (esym card).
Definition rank : T ordsum := cast_ord card \o enum_rank.

Lemma sigK : cancel sig rank.
Lemma sig_inj : injective sig.

Lemma rankK : cancel rank sig.
Lemma rank_inj : injective rank.

Definition sig1 s : 'I_n := tag (sig s).
Definition sig2 s : 'I_(p_ (sig1 s)) := tagged (sig s).
Definition Rank i (j : 'I_(p_ i)) := rank (Tagged _ j).

Lemma sigE12 s : sig s = @Tagged _ (sig1 s) _ (sig2 s).

Lemma rankE p : rank p = @Rank (tag p) (tagged p).

Lemma sig2K s : Rank (sig2 s) = s.

Lemma Rank1K i0 (k : 'I_(p_ i0)) : sig1 (Rank k) = i0.

Lemma Rank2K i0 (k : 'I_(p_ i0)) :
  sig2 (Rank k) = cast_ord (congr1 p_ (esym (Rank1K k))) k.
#[local] Hint Resolve sigK rankK : core.

Lemma rank_bij : bijective rank.
Lemma sig_bij : bijective sig.

Lemma rank_bij_on : {on [pred _ | true], bijective rank}.

Lemma sig_bij_on : {on [pred _ | true], bijective sig}.

Lemma le_sig : {mono sig : i j / i j}.

Lemma le_sig1 : {homo sig1 : i j / i j}.

Lemma le_rank : {mono rank : p q / p q}.

Lemma le_Rank i : {mono @Rank i : j k / j k}.

Lemma lt_sig : {mono sig : i j / i < j}.

Lemma lt_rank : {mono rank : p q / p < q}.

Lemma lt_Rank i : {mono @Rank i : j k / j < k}.

Lemma eq_Rank i i' (j : 'I_(p_ i)) (j': 'I_(p_ i')) :
  (Rank j == Rank j' :> nat) = (i == i') && (j == j' :> nat).

Lemma rankEsum p : rank p = \sum_(i < n | (i < tag p)%N) p_ i + tagged p :> nat.

Lemma RankEsum i j : @Rank i j = \sum_(k < n | (k < i)%N) p_ k + j :> nat.

Lemma rect s : s = \sum_(i < n | (i < sig1 s)%N) p_ i + sig2 s :> nat.

Lemma eqRank (i0 j : nat) (li0 : (i0 < n)%N) (lj : (j < p_ (Ordinal li0))%N) :
  (\sum_(i < n | (i < i0)%N) p_ i) + j = Rank (Ordinal lj) :> nat.

End tagnat.
End tagnat.
Arguments tagnat.Rank {n p_}.