diff options
| author | Enrico Tassi | 2015-03-09 11:07:53 +0100 |
|---|---|---|
| committer | Enrico Tassi | 2015-03-09 11:24:38 +0100 |
| commit | fc84c27eac260dffd8f2fb1cb56d599f1e3486d9 (patch) | |
| tree | c16205f1637c80833a4c4598993c29fa0fd8c373 /mathcomp/discrete | |
Initial commit
Diffstat (limited to 'mathcomp/discrete')
| -rw-r--r-- | mathcomp/discrete/all.v | 12 | ||||
| -rw-r--r-- | mathcomp/discrete/bigop.v | 1770 | ||||
| -rw-r--r-- | mathcomp/discrete/binomial.v | 524 | ||||
| -rw-r--r-- | mathcomp/discrete/choice.v | 681 | ||||
| -rw-r--r-- | mathcomp/discrete/div.v | 946 | ||||
| -rw-r--r-- | mathcomp/discrete/finfun.v | 302 | ||||
| -rw-r--r-- | mathcomp/discrete/fingraph.v | 721 | ||||
| -rw-r--r-- | mathcomp/discrete/finset.v | 2214 | ||||
| -rw-r--r-- | mathcomp/discrete/fintype.v | 2037 | ||||
| -rw-r--r-- | mathcomp/discrete/generic_quotient.v | 727 | ||||
| -rw-r--r-- | mathcomp/discrete/path.v | 890 | ||||
| -rw-r--r-- | mathcomp/discrete/prime.v | 1404 | ||||
| -rw-r--r-- | mathcomp/discrete/tuple.v | 412 |
13 files changed, 12640 insertions, 0 deletions
diff --git a/mathcomp/discrete/all.v b/mathcomp/discrete/all.v new file mode 100644 index 0000000..d100add --- /dev/null +++ b/mathcomp/discrete/all.v @@ -0,0 +1,12 @@ +Require Export bigop. +Require Export binomial. +Require Export choice. +Require Export div. +Require Export finfun. +Require Export fingraph. +Require Export finset. +Require Export fintype. +Require Export generic_quotient. +Require Export path. +Require Export prime. +Require Export tuple. diff --git a/mathcomp/discrete/bigop.v b/mathcomp/discrete/bigop.v new file mode 100644 index 0000000..88bf4d2 --- /dev/null +++ b/mathcomp/discrete/bigop.v @@ -0,0 +1,1770 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div fintype. +Require Import tuple finfun. + +(******************************************************************************) +(* This file provides a generic definition for iterating an operator over a *) +(* set of indices (reducebig); this big operator is parametrized by the *) +(* return type (R), the type of indices (I), the operator (op), the default *) +(* value on empty lists (idx), the range of indices (r), the filter applied *) +(* on this range (P) and the expression we are iterating (F). The definition *) +(* is not to be used directly, but via the wide range of notations provided *) +(* and which allows a natural use of big operators. *) +(* The lemmas can be classified according to the operator being iterated: *) +(* 1. results independent of the operator: extensionality with respect to *) +(* the range of indices, to the filtering predicate or to the expression *) +(* being iterated; reindexing, widening or narrowing of the range of *) +(* indices; we provide lemmas for the special cases where indices are *) +(* natural numbers or bounded natural numbers ("ordinals"). We supply *) +(* several "functional" induction principles that can be used with the *) +(* ssreflect 1.3 "elim" tactic to do induction over the index range for *) +(* up to 3 bigops simultaneously. *) +(* 2. results depending on the properties of the operator: *) +(* We distinguish: monoid laws (op is associative, idx is an identity *) +(* element), abelian monoid laws (op is also commutative), and laws with *) +(* a distributive operation (semi-rings). Examples of such results are *) +(* splitting, permuting, and exchanging bigops. *) +(* A special section is dedicated to big operators on natural numbers. *) +(******************************************************************************) +(* Notations: *) +(* The general form for iterated operators is *) +(* <bigop>_<range> <general_term> *) +(* - <bigop> is one of \big[op/idx], \sum, \prod, or \max (see below). *) +(* - <general_term> can be any expression. *) +(* - <range> binds an index variable in <general_term>; <range> is one of *) +(* (i <- s) i ranges over the sequence s *) +(* (m <= i < n) i ranges over the nat interval m, m.+1, ..., n.-1 *) +(* (i < n) i ranges over the (finite) type 'I_n (i.e., ordinal n) *) +(* (i : T) i ranges over the finite type T *) +(* i or (i) i ranges over its (inferred) finite type *) +(* (i in A) i ranges over the elements that satisfy the collective *) +(* predicate A (the domain of A must be a finite type) *) +(* (i <- s | <condition>) limits the range to the i for which <condition> *) +(* holds. <condition> can be any expression that coerces to *) +(* bool, and may mention the bound index i. All six kinds of *) +(* ranges above can have a <condition> part. *) +(* - One can use the "\big[op/idx]" notations for any operator. *) +(* - BIG_F and BIG_P are pattern abbreviations for the <general_term> and *) +(* <condition> part of a \big ... expression; for (i in A) and (i in A | C) *) +(* ranges the term matched by BIG_P will include the i \in A condition. *) +(* - The (locked) head constant of a \big notation is bigop. *) +(* - The "\sum", "\prod" and "\max" notations in the %N scope are used for *) +(* natural numbers with addition, multiplication and maximum (and their *) +(* corresponding neutral elements), respectively. *) +(* - The "\sum" and "\prod" reserved notations are overloaded in ssralg in *) +(* the %R scope, in mxalgebra, vector & falgebra in the %MS and %VS scopes; *) +(* "\prod" is also overloaded in fingroup, the %g and %G scopes. *) +(* - We reserve "\bigcup" and "\bigcap" notations for iterated union and *) +(* intersection (of sets, groups, vector spaces, etc). *) +(******************************************************************************) +(* Tips for using lemmas in this file: *) +(* to apply a lemma for a specific operator: if no special property is *) +(* required for the operator, simply apply the lemma; if the lemma needs *) +(* certain properties for the operator, make sure the appropriate Canonical *) +(* instances are declared. *) +(******************************************************************************) +(* Interfaces for operator properties are packaged in the Monoid submodule: *) +(* Monoid.law idx == interface (keyed on the operator) for associative *) +(* operators with identity element idx. *) +(* Monoid.com_law idx == extension (telescope) of Monoid.law for operators *) +(* that are also commutative. *) +(* Monoid.mul_law abz == interface for operators with absorbing (zero) *) +(* element abz. *) +(* Monoid.add_law idx mop == extension of Monoid.com_law for operators over *) +(* which operation mop distributes (mop will often also *) +(* have a Monoid.mul_law idx structure). *) +(* [law of op], [com_law of op], [mul_law of op], [add_law mop of op] == *) +(* syntax for cloning Monoid structures. *) +(* Monoid.Theory == submodule containing basic generic algebra lemmas *) +(* for operators satisfying the Monoid interfaces. *) +(* Monoid.simpm == generic monoid simplification rewrite multirule. *) +(* Monoid structures are predeclared for many basic operators: (_ && _)%B, *) +(* (_ || _)%B, (_ (+) _)%B (exclusive or) , (_ + _)%N, (_ * _)%N, maxn, *) +(* gcdn, lcmn and (_ ++ _)%SEQ (list concatenation). *) +(******************************************************************************) +(* Additional documentation for this file: *) +(* Y. Bertot, G. Gonthier, S. Ould Biha and I. Pasca. *) +(* Canonical Big Operators. In TPHOLs 2008, LNCS vol. 5170, Springer. *) +(* Article available at: *) +(* http://hal.inria.fr/docs/00/33/11/93/PDF/main.pdf *) +(******************************************************************************) +(* Examples of use in: poly.v, matrix.v *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Reserved Notation "\big [ op / idx ]_ i F" + (at level 36, F at level 36, op, idx at level 10, i at level 0, + right associativity, + format "'[' \big [ op / idx ]_ i '/ ' F ']'"). +Reserved Notation "\big [ op / idx ]_ ( i <- r | P ) F" + (at level 36, F at level 36, op, idx at level 10, i, r at level 50, + format "'[' \big [ op / idx ]_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\big [ op / idx ]_ ( i <- r ) F" + (at level 36, F at level 36, op, idx at level 10, i, r at level 50, + format "'[' \big [ op / idx ]_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\big [ op / idx ]_ ( m <= i < n | P ) F" + (at level 36, F at level 36, op, idx at level 10, m, i, n at level 50, + format "'[' \big [ op / idx ]_ ( m <= i < n | P ) F ']'"). +Reserved Notation "\big [ op / idx ]_ ( m <= i < n ) F" + (at level 36, F at level 36, op, idx at level 10, i, m, n at level 50, + format "'[' \big [ op / idx ]_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\big [ op / idx ]_ ( i | P ) F" + (at level 36, F at level 36, op, idx at level 10, i at level 50, + format "'[' \big [ op / idx ]_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\big [ op / idx ]_ ( i : t | P ) F" + (at level 36, F at level 36, op, idx at level 10, i at level 50, + format "'[' \big [ op / idx ]_ ( i : t | P ) '/ ' F ']'"). +Reserved Notation "\big [ op / idx ]_ ( i : t ) F" + (at level 36, F at level 36, op, idx at level 10, i at level 50, + format "'[' \big [ op / idx ]_ ( i : t ) '/ ' F ']'"). +Reserved Notation "\big [ op / idx ]_ ( i < n | P ) F" + (at level 36, F at level 36, op, idx at level 10, i, n at level 50, + format "'[' \big [ op / idx ]_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\big [ op / idx ]_ ( i < n ) F" + (at level 36, F at level 36, op, idx at level 10, i, n at level 50, + format "'[' \big [ op / idx ]_ ( i < n ) F ']'"). +Reserved Notation "\big [ op / idx ]_ ( i 'in' A | P ) F" + (at level 36, F at level 36, op, idx at level 10, i, A at level 50, + format "'[' \big [ op / idx ]_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\big [ op / idx ]_ ( i 'in' A ) F" + (at level 36, F at level 36, op, idx at level 10, i, A at level 50, + format "'[' \big [ op / idx ]_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\sum_ i F" + (at level 41, F at level 41, i at level 0, + right associativity, + format "'[' \sum_ i '/ ' F ']'"). +Reserved Notation "\sum_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \sum_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\sum_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \sum_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\sum_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \sum_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\sum_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \sum_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\sum_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \sum_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\sum_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\sum_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\sum_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \sum_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\sum_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \sum_ ( i < n ) '/ ' F ']'"). +Reserved Notation "\sum_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \sum_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\sum_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \sum_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\max_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \max_ i '/ ' F ']'"). +Reserved Notation "\max_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \max_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\max_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \max_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\max_ ( m <= i < n | P ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \max_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\max_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \max_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\max_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \max_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\max_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\max_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + only parsing). +Reserved Notation "\max_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \max_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\max_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \max_ ( i < n ) F ']'"). +Reserved Notation "\max_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \max_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\max_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \max_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\prod_ i F" + (at level 36, F at level 36, i at level 0, + format "'[' \prod_ i '/ ' F ']'"). +Reserved Notation "\prod_ ( i <- r | P ) F" + (at level 36, F at level 36, i, r at level 50, + format "'[' \prod_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\prod_ ( i <- r ) F" + (at level 36, F at level 36, i, r at level 50, + format "'[' \prod_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\prod_ ( m <= i < n | P ) F" + (at level 36, F at level 36, i, m, n at level 50, + format "'[' \prod_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\prod_ ( m <= i < n ) F" + (at level 36, F at level 36, i, m, n at level 50, + format "'[' \prod_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\prod_ ( i | P ) F" + (at level 36, F at level 36, i at level 50, + format "'[' \prod_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\prod_ ( i : t | P ) F" + (at level 36, F at level 36, i at level 50, + only parsing). +Reserved Notation "\prod_ ( i : t ) F" + (at level 36, F at level 36, i at level 50, + only parsing). +Reserved Notation "\prod_ ( i < n | P ) F" + (at level 36, F at level 36, i, n at level 50, + format "'[' \prod_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\prod_ ( i < n ) F" + (at level 36, F at level 36, i, n at level 50, + format "'[' \prod_ ( i < n ) '/ ' F ']'"). +Reserved Notation "\prod_ ( i 'in' A | P ) F" + (at level 36, F at level 36, i, A at level 50, + format "'[' \prod_ ( i 'in' A | P ) F ']'"). +Reserved Notation "\prod_ ( i 'in' A ) F" + (at level 36, F at level 36, i, A at level 50, + format "'[' \prod_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\bigcup_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \bigcup_ i '/ ' F ']'"). +Reserved Notation "\bigcup_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \bigcup_ ( i <- r | P ) '/ ' F ']'"). +Reserved Notation "\bigcup_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \bigcup_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\bigcup_ ( m <= i < n | P ) F" + (at level 41, F at level 41, m, i, n at level 50, + format "'[' \bigcup_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\bigcup_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \bigcup_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\bigcup_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \bigcup_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\bigcup_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \bigcup_ ( i : t | P ) '/ ' F ']'"). +Reserved Notation "\bigcup_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \bigcup_ ( i : t ) '/ ' F ']'"). +Reserved Notation "\bigcup_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \bigcup_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\bigcup_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \bigcup_ ( i < n ) '/ ' F ']'"). +Reserved Notation "\bigcup_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \bigcup_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\bigcup_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \bigcup_ ( i 'in' A ) '/ ' F ']'"). + +Reserved Notation "\bigcap_ i F" + (at level 41, F at level 41, i at level 0, + format "'[' \bigcap_ i '/ ' F ']'"). +Reserved Notation "\bigcap_ ( i <- r | P ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \bigcap_ ( i <- r | P ) F ']'"). +Reserved Notation "\bigcap_ ( i <- r ) F" + (at level 41, F at level 41, i, r at level 50, + format "'[' \bigcap_ ( i <- r ) '/ ' F ']'"). +Reserved Notation "\bigcap_ ( m <= i < n | P ) F" + (at level 41, F at level 41, m, i, n at level 50, + format "'[' \bigcap_ ( m <= i < n | P ) '/ ' F ']'"). +Reserved Notation "\bigcap_ ( m <= i < n ) F" + (at level 41, F at level 41, i, m, n at level 50, + format "'[' \bigcap_ ( m <= i < n ) '/ ' F ']'"). +Reserved Notation "\bigcap_ ( i | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \bigcap_ ( i | P ) '/ ' F ']'"). +Reserved Notation "\bigcap_ ( i : t | P ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \bigcap_ ( i : t | P ) '/ ' F ']'"). +Reserved Notation "\bigcap_ ( i : t ) F" + (at level 41, F at level 41, i at level 50, + format "'[' \bigcap_ ( i : t ) '/ ' F ']'"). +Reserved Notation "\bigcap_ ( i < n | P ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \bigcap_ ( i < n | P ) '/ ' F ']'"). +Reserved Notation "\bigcap_ ( i < n ) F" + (at level 41, F at level 41, i, n at level 50, + format "'[' \bigcap_ ( i < n ) '/ ' F ']'"). +Reserved Notation "\bigcap_ ( i 'in' A | P ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \bigcap_ ( i 'in' A | P ) '/ ' F ']'"). +Reserved Notation "\bigcap_ ( i 'in' A ) F" + (at level 41, F at level 41, i, A at level 50, + format "'[' \bigcap_ ( i 'in' A ) '/ ' F ']'"). + +Module Monoid. + +Section Definitions. +Variables (T : Type) (idm : T). + +Structure law := Law { + operator : T -> T -> T; + _ : associative operator; + _ : left_id idm operator; + _ : right_id idm operator +}. +Local Coercion operator : law >-> Funclass. + +Structure com_law := ComLaw { + com_operator : law; + _ : commutative com_operator +}. +Local Coercion com_operator : com_law >-> law. + +Structure mul_law := MulLaw { + mul_operator : T -> T -> T; + _ : left_zero idm mul_operator; + _ : right_zero idm mul_operator +}. +Local Coercion mul_operator : mul_law >-> Funclass. + +Structure add_law (mul : T -> T -> T) := AddLaw { + add_operator : com_law; + _ : left_distributive mul add_operator; + _ : right_distributive mul add_operator +}. +Local Coercion add_operator : add_law >-> com_law. + +Let op_id (op1 op2 : T -> T -> T) := phant_id op1 op2. + +Definition clone_law op := + fun (opL : law) & op_id opL op => + fun opmA op1m opm1 (opL' := @Law op opmA op1m opm1) + & phant_id opL' opL => opL'. + +Definition clone_com_law op := + fun (opL : law) (opC : com_law) & op_id opL op & op_id opC op => + fun opmC (opC' := @ComLaw opL opmC) & phant_id opC' opC => opC'. + +Definition clone_mul_law op := + fun (opM : mul_law) & op_id opM op => + fun op0m opm0 (opM' := @MulLaw op op0m opm0) & phant_id opM' opM => opM'. + +Definition clone_add_law mop aop := + fun (opC : com_law) (opA : add_law mop) & op_id opC aop & op_id opA aop => + fun mopDm mopmD (opA' := @AddLaw mop opC mopDm mopmD) + & phant_id opA' opA => opA'. + +End Definitions. + +Module Import Exports. +Coercion operator : law >-> Funclass. +Coercion com_operator : com_law >-> law. +Coercion mul_operator : mul_law >-> Funclass. +Coercion add_operator : add_law >-> com_law. +Notation "[ 'law' 'of' f ]" := (@clone_law _ _ f _ id _ _ _ id) + (at level 0, format"[ 'law' 'of' f ]") : form_scope. +Notation "[ 'com_law' 'of' f ]" := (@clone_com_law _ _ f _ _ id id _ id) + (at level 0, format "[ 'com_law' 'of' f ]") : form_scope. +Notation "[ 'mul_law' 'of' f ]" := (@clone_mul_law _ _ f _ id _ _ id) + (at level 0, format"[ 'mul_law' 'of' f ]") : form_scope. +Notation "[ 'add_law' m 'of' a ]" := (@clone_add_law _ _ m a _ _ id id _ _ id) + (at level 0, format "[ 'add_law' m 'of' a ]") : form_scope. +End Exports. + +Section CommutativeAxioms. + +Variable (T : Type) (zero one : T) (mul add : T -> T -> T) (inv : T -> T). +Hypothesis mulC : commutative mul. + +Lemma mulC_id : left_id one mul -> right_id one mul. +Proof. by move=> mul1x x; rewrite mulC. Qed. + +Lemma mulC_zero : left_zero zero mul -> right_zero zero mul. +Proof. by move=> mul0x x; rewrite mulC. Qed. + +Lemma mulC_dist : left_distributive mul add -> right_distributive mul add. +Proof. by move=> mul_addl x y z; rewrite !(mulC x). Qed. + +End CommutativeAxioms. + +Module Theory. + +Section Theory. +Variables (T : Type) (idm : T). + +Section Plain. +Variable mul : law idm. +Lemma mul1m : left_id idm mul. Proof. by case mul. Qed. +Lemma mulm1 : right_id idm mul. Proof. by case mul. Qed. +Lemma mulmA : associative mul. Proof. by case mul. Qed. +Lemma iteropE n x : iterop n mul x idm = iter n (mul x) idm. +Proof. by case: n => // n; rewrite iterSr mulm1 iteropS. Qed. +End Plain. + +Section Commutative. +Variable mul : com_law idm. +Lemma mulmC : commutative mul. Proof. by case mul. Qed. +Lemma mulmCA : left_commutative mul. +Proof. by move=> x y z; rewrite !mulmA (mulmC x). Qed. +Lemma mulmAC : right_commutative mul. +Proof. by move=> x y z; rewrite -!mulmA (mulmC y). Qed. +Lemma mulmACA : interchange mul mul. +Proof. by move=> x y z t; rewrite -!mulmA (mulmCA y). Qed. +End Commutative. + +Section Mul. +Variable mul : mul_law idm. +Lemma mul0m : left_zero idm mul. Proof. by case mul. Qed. +Lemma mulm0 : right_zero idm mul. Proof. by case mul. Qed. +End Mul. + +Section Add. +Variables (mul : T -> T -> T) (add : add_law idm mul). +Lemma addmA : associative add. Proof. exact: mulmA. Qed. +Lemma addmC : commutative add. Proof. exact: mulmC. Qed. +Lemma addmCA : left_commutative add. Proof. exact: mulmCA. Qed. +Lemma addmAC : right_commutative add. Proof. exact: mulmAC. Qed. +Lemma add0m : left_id idm add. Proof. exact: mul1m. Qed. +Lemma addm0 : right_id idm add. Proof. exact: mulm1. Qed. +Lemma mulm_addl : left_distributive mul add. Proof. by case add. Qed. +Lemma mulm_addr : right_distributive mul add. Proof. by case add. Qed. +End Add. + +Definition simpm := (mulm1, mulm0, mul1m, mul0m, mulmA). + +End Theory. + +End Theory. +Include Theory. + +End Monoid. +Export Monoid.Exports. + +Section PervasiveMonoids. + +Import Monoid. + +Canonical andb_monoid := Law andbA andTb andbT. +Canonical andb_comoid := ComLaw andbC. + +Canonical andb_muloid := MulLaw andFb andbF. +Canonical orb_monoid := Law orbA orFb orbF. +Canonical orb_comoid := ComLaw orbC. +Canonical orb_muloid := MulLaw orTb orbT. +Canonical addb_monoid := Law addbA addFb addbF. +Canonical addb_comoid := ComLaw addbC. +Canonical orb_addoid := AddLaw andb_orl andb_orr. +Canonical andb_addoid := AddLaw orb_andl orb_andr. +Canonical addb_addoid := AddLaw andb_addl andb_addr. + +Canonical addn_monoid := Law addnA add0n addn0. +Canonical addn_comoid := ComLaw addnC. +Canonical muln_monoid := Law mulnA mul1n muln1. +Canonical muln_comoid := ComLaw mulnC. +Canonical muln_muloid := MulLaw mul0n muln0. +Canonical addn_addoid := AddLaw mulnDl mulnDr. + +Canonical maxn_monoid := Law maxnA max0n maxn0. +Canonical maxn_comoid := ComLaw maxnC. +Canonical maxn_addoid := AddLaw maxn_mull maxn_mulr. + +Canonical gcdn_monoid := Law gcdnA gcd0n gcdn0. +Canonical gcdn_comoid := ComLaw gcdnC. +Canonical gcdnDoid := AddLaw muln_gcdl muln_gcdr. + +Canonical lcmn_monoid := Law lcmnA lcm1n lcmn1. +Canonical lcmn_comoid := ComLaw lcmnC. +Canonical lcmn_addoid := AddLaw muln_lcml muln_lcmr. + +Canonical cat_monoid T := Law (@catA T) (@cat0s T) (@cats0 T). + +End PervasiveMonoids. + +(* Unit test for the [...law of ...] Notations +Definition myp := addn. Definition mym := muln. +Canonical myp_mon := [law of myp]. +Canonical myp_cmon := [com_law of myp]. +Canonical mym_mul := [mul_law of mym]. +Canonical myp_add := [add_law _ of myp]. +Print myp_add. +Print Canonical Projections. +*) + +Delimit Scope big_scope with BIG. +Open Scope big_scope. + +(* The bigbody wrapper is a workaround for a quirk of the Coq pretty-printer, *) +(* which would fail to redisplay the \big notation when the <general_term> or *) +(* <condition> do not depend on the bound index. The BigBody constructor *) +(* packages both in in a term in which i occurs; it also depends on the *) +(* iterated <op>, as this can give more information on the expected type of *) +(* the <general_term>, thus allowing for the insertion of coercions. *) +CoInductive bigbody R I := BigBody of I & (R -> R -> R) & bool & R. + +Definition applybig {R I} (body : bigbody R I) x := + let: BigBody _ op b v := body in if b then op v x else x. + +Definition reducebig R I idx r (body : I -> bigbody R I) := + foldr (applybig \o body) idx r. + +Module Type BigOpSig. +Parameter bigop : forall R I, R -> seq I -> (I -> bigbody R I) -> R. +Axiom bigopE : bigop = reducebig. +End BigOpSig. + +Module BigOp : BigOpSig. +Definition bigop := reducebig. +Lemma bigopE : bigop = reducebig. Proof. by []. Qed. +End BigOp. + +Notation bigop := BigOp.bigop (only parsing). +Canonical bigop_unlock := Unlockable BigOp.bigopE. + +Definition index_iota m n := iota m (n - m). + +Definition index_enum (T : finType) := Finite.enum T. + +Lemma mem_index_iota m n i : i \in index_iota m n = (m <= i < n). +Proof. +rewrite mem_iota; case le_m_i: (m <= i) => //=. +by rewrite -leq_subLR subSn // -subn_gt0 -subnDA subnKC // subn_gt0. +Qed. + +Lemma mem_index_enum T i : i \in index_enum T. +Proof. by rewrite -[index_enum T]enumT mem_enum. Qed. +Hint Resolve mem_index_enum. + +Lemma filter_index_enum T P : filter P (index_enum T) = enum P. +Proof. by []. Qed. + +Notation "\big [ op / idx ]_ ( i <- r | P ) F" := + (bigop idx r (fun i => BigBody i op P%B F)) : big_scope. +Notation "\big [ op / idx ]_ ( i <- r ) F" := + (bigop idx r (fun i => BigBody i op true F)) : big_scope. +Notation "\big [ op / idx ]_ ( m <= i < n | P ) F" := + (bigop idx (index_iota m n) (fun i : nat => BigBody i op P%B F)) + : big_scope. +Notation "\big [ op / idx ]_ ( m <= i < n ) F" := + (bigop idx (index_iota m n) (fun i : nat => BigBody i op true F)) + : big_scope. +Notation "\big [ op / idx ]_ ( i | P ) F" := + (bigop idx (index_enum _) (fun i => BigBody i op P%B F)) : big_scope. +Notation "\big [ op / idx ]_ i F" := + (bigop idx (index_enum _) (fun i => BigBody i op true F)) : big_scope. +Notation "\big [ op / idx ]_ ( i : t | P ) F" := + (bigop idx (index_enum _) (fun i : t => BigBody i op P%B F)) + (only parsing) : big_scope. +Notation "\big [ op / idx ]_ ( i : t ) F" := + (bigop idx (index_enum _) (fun i : t => BigBody i op true F)) + (only parsing) : big_scope. +Notation "\big [ op / idx ]_ ( i < n | P ) F" := + (\big[op/idx]_(i : ordinal n | P%B) F) : big_scope. +Notation "\big [ op / idx ]_ ( i < n ) F" := + (\big[op/idx]_(i : ordinal n) F) : big_scope. +Notation "\big [ op / idx ]_ ( i 'in' A | P ) F" := + (\big[op/idx]_(i | (i \in A) && P) F) : big_scope. +Notation "\big [ op / idx ]_ ( i 'in' A ) F" := + (\big[op/idx]_(i | i \in A) F) : big_scope. + +Notation BIG_F := (F in \big[_/_]_(i <- _ | _) F i)%pattern. +Notation BIG_P := (P in \big[_/_]_(i <- _ | P i) _)%pattern. + +Local Notation "+%N" := addn (at level 0, only parsing). +Notation "\sum_ ( i <- r | P ) F" := + (\big[+%N/0%N]_(i <- r | P%B) F%N) : nat_scope. +Notation "\sum_ ( i <- r ) F" := + (\big[+%N/0%N]_(i <- r) F%N) : nat_scope. +Notation "\sum_ ( m <= i < n | P ) F" := + (\big[+%N/0%N]_(m <= i < n | P%B) F%N) : nat_scope. +Notation "\sum_ ( m <= i < n ) F" := + (\big[+%N/0%N]_(m <= i < n) F%N) : nat_scope. +Notation "\sum_ ( i | P ) F" := + (\big[+%N/0%N]_(i | P%B) F%N) : nat_scope. +Notation "\sum_ i F" := + (\big[+%N/0%N]_i F%N) : nat_scope. +Notation "\sum_ ( i : t | P ) F" := + (\big[+%N/0%N]_(i : t | P%B) F%N) (only parsing) : nat_scope. +Notation "\sum_ ( i : t ) F" := + (\big[+%N/0%N]_(i : t) F%N) (only parsing) : nat_scope. +Notation "\sum_ ( i < n | P ) F" := + (\big[+%N/0%N]_(i < n | P%B) F%N) : nat_scope. +Notation "\sum_ ( i < n ) F" := + (\big[+%N/0%N]_(i < n) F%N) : nat_scope. +Notation "\sum_ ( i 'in' A | P ) F" := + (\big[+%N/0%N]_(i in A | P%B) F%N) : nat_scope. +Notation "\sum_ ( i 'in' A ) F" := + (\big[+%N/0%N]_(i in A) F%N) : nat_scope. + +Local Notation "*%N" := muln (at level 0, only parsing). +Notation "\prod_ ( i <- r | P ) F" := + (\big[*%N/1%N]_(i <- r | P%B) F%N) : nat_scope. +Notation "\prod_ ( i <- r ) F" := + (\big[*%N/1%N]_(i <- r) F%N) : nat_scope. +Notation "\prod_ ( m <= i < n | P ) F" := + (\big[*%N/1%N]_(m <= i < n | P%B) F%N) : nat_scope. +Notation "\prod_ ( m <= i < n ) F" := + (\big[*%N/1%N]_(m <= i < n) F%N) : nat_scope. +Notation "\prod_ ( i | P ) F" := + (\big[*%N/1%N]_(i | P%B) F%N) : nat_scope. +Notation "\prod_ i F" := + (\big[*%N/1%N]_i F%N) : nat_scope. +Notation "\prod_ ( i : t | P ) F" := + (\big[*%N/1%N]_(i : t | P%B) F%N) (only parsing) : nat_scope. +Notation "\prod_ ( i : t ) F" := + (\big[*%N/1%N]_(i : t) F%N) (only parsing) : nat_scope. +Notation "\prod_ ( i < n | P ) F" := + (\big[*%N/1%N]_(i < n | P%B) F%N) : nat_scope. +Notation "\prod_ ( i < n ) F" := + (\big[*%N/1%N]_(i < n) F%N) : nat_scope. +Notation "\prod_ ( i 'in' A | P ) F" := + (\big[*%N/1%N]_(i in A | P%B) F%N) : nat_scope. +Notation "\prod_ ( i 'in' A ) F" := + (\big[*%N/1%N]_(i in A) F%N) : nat_scope. + +Notation "\max_ ( i <- r | P ) F" := + (\big[maxn/0%N]_(i <- r | P%B) F%N) : nat_scope. +Notation "\max_ ( i <- r ) F" := + (\big[maxn/0%N]_(i <- r) F%N) : nat_scope. +Notation "\max_ ( i | P ) F" := + (\big[maxn/0%N]_(i | P%B) F%N) : nat_scope. +Notation "\max_ i F" := + (\big[maxn/0%N]_i F%N) : nat_scope. +Notation "\max_ ( i : I | P ) F" := + (\big[maxn/0%N]_(i : I | P%B) F%N) (only parsing) : nat_scope. +Notation "\max_ ( i : I ) F" := + (\big[maxn/0%N]_(i : I) F%N) (only parsing) : nat_scope. +Notation "\max_ ( m <= i < n | P ) F" := + (\big[maxn/0%N]_(m <= i < n | P%B) F%N) : nat_scope. +Notation "\max_ ( m <= i < n ) F" := + (\big[maxn/0%N]_(m <= i < n) F%N) : nat_scope. +Notation "\max_ ( i < n | P ) F" := + (\big[maxn/0%N]_(i < n | P%B) F%N) : nat_scope. +Notation "\max_ ( i < n ) F" := + (\big[maxn/0%N]_(i < n) F%N) : nat_scope. +Notation "\max_ ( i 'in' A | P ) F" := + (\big[maxn/0%N]_(i in A | P%B) F%N) : nat_scope. +Notation "\max_ ( i 'in' A ) F" := + (\big[maxn/0%N]_(i in A) F%N) : nat_scope. + +(* Induction loading *) +Lemma big_load R (K K' : R -> Type) idx op I r (P : pred I) F : + K (\big[op/idx]_(i <- r | P i) F i) * K' (\big[op/idx]_(i <- r | P i) F i) + -> K' (\big[op/idx]_(i <- r | P i) F i). +Proof. by case. Qed. + +Implicit Arguments big_load [R K' I]. + +Section Elim3. + +Variables (R1 R2 R3 : Type) (K : R1 -> R2 -> R3 -> Type). +Variables (id1 : R1) (op1 : R1 -> R1 -> R1). +Variables (id2 : R2) (op2 : R2 -> R2 -> R2). +Variables (id3 : R3) (op3 : R3 -> R3 -> R3). + +Hypothesis Kid : K id1 id2 id3. + +Lemma big_rec3 I r (P : pred I) F1 F2 F3 + (K_F : forall i y1 y2 y3, P i -> K y1 y2 y3 -> + K (op1 (F1 i) y1) (op2 (F2 i) y2) (op3 (F3 i) y3)) : + K (\big[op1/id1]_(i <- r | P i) F1 i) + (\big[op2/id2]_(i <- r | P i) F2 i) + (\big[op3/id3]_(i <- r | P i) F3 i). +Proof. by rewrite unlock; elim: r => //= i r; case: ifP => //; exact: K_F. Qed. + +Hypothesis Kop : forall x1 x2 x3 y1 y2 y3, + K x1 x2 x3 -> K y1 y2 y3-> K (op1 x1 y1) (op2 x2 y2) (op3 x3 y3). +Lemma big_ind3 I r (P : pred I) F1 F2 F3 + (K_F : forall i, P i -> K (F1 i) (F2 i) (F3 i)) : + K (\big[op1/id1]_(i <- r | P i) F1 i) + (\big[op2/id2]_(i <- r | P i) F2 i) + (\big[op3/id3]_(i <- r | P i) F3 i). +Proof. by apply: big_rec3 => i x1 x2 x3 /K_F; exact: Kop. Qed. + +End Elim3. + +Implicit Arguments big_rec3 [R1 R2 R3 id1 op1 id2 op2 id3 op3 I r P F1 F2 F3]. +Implicit Arguments big_ind3 [R1 R2 R3 id1 op1 id2 op2 id3 op3 I r P F1 F2 F3]. + +Section Elim2. + +Variables (R1 R2 : Type) (K : R1 -> R2 -> Type) (f : R2 -> R1). +Variables (id1 : R1) (op1 : R1 -> R1 -> R1). +Variables (id2 : R2) (op2 : R2 -> R2 -> R2). + +Hypothesis Kid : K id1 id2. + +Lemma big_rec2 I r (P : pred I) F1 F2 + (K_F : forall i y1 y2, P i -> K y1 y2 -> + K (op1 (F1 i) y1) (op2 (F2 i) y2)) : + K (\big[op1/id1]_(i <- r | P i) F1 i) (\big[op2/id2]_(i <- r | P i) F2 i). +Proof. by rewrite unlock; elim: r => //= i r; case: ifP => //; exact: K_F. Qed. + +Hypothesis Kop : forall x1 x2 y1 y2, + K x1 x2 -> K y1 y2 -> K (op1 x1 y1) (op2 x2 y2). +Lemma big_ind2 I r (P : pred I) F1 F2 (K_F : forall i, P i -> K (F1 i) (F2 i)) : + K (\big[op1/id1]_(i <- r | P i) F1 i) (\big[op2/id2]_(i <- r | P i) F2 i). +Proof. by apply: big_rec2 => i x1 x2 /K_F; exact: Kop. Qed. + +Hypotheses (f_op : {morph f : x y / op2 x y >-> op1 x y}) (f_id : f id2 = id1). +Lemma big_morph I r (P : pred I) F : + f (\big[op2/id2]_(i <- r | P i) F i) = \big[op1/id1]_(i <- r | P i) f (F i). +Proof. by rewrite unlock; elim: r => //= i r <-; rewrite -f_op -fun_if. Qed. + +End Elim2. + +Implicit Arguments big_rec2 [R1 R2 id1 op1 id2 op2 I r P F1 F2]. +Implicit Arguments big_ind2 [R1 R2 id1 op1 id2 op2 I r P F1 F2]. +Implicit Arguments big_morph [R1 R2 id1 op1 id2 op2 I]. + +Section Elim1. + +Variables (R : Type) (K : R -> Type) (f : R -> R). +Variables (idx : R) (op op' : R -> R -> R). + +Hypothesis Kid : K idx. + +Lemma big_rec I r (P : pred I) F + (Kop : forall i x, P i -> K x -> K (op (F i) x)) : + K (\big[op/idx]_(i <- r | P i) F i). +Proof. by rewrite unlock; elim: r => //= i r; case: ifP => //; exact: Kop. Qed. + +Hypothesis Kop : forall x y, K x -> K y -> K (op x y). +Lemma big_ind I r (P : pred I) F (K_F : forall i, P i -> K (F i)) : + K (\big[op/idx]_(i <- r | P i) F i). +Proof. by apply: big_rec => // i x /K_F /Kop; exact. Qed. + +Hypothesis Kop' : forall x y, K x -> K y -> op x y = op' x y. +Lemma eq_big_op I r (P : pred I) F (K_F : forall i, P i -> K (F i)) : + \big[op/idx]_(i <- r | P i) F i = \big[op'/idx]_(i <- r | P i) F i. +Proof. +by elim/(big_load K): _; elim/big_rec2: _ => // i _ y Pi [Ky <-]; auto. +Qed. + +Hypotheses (fM : {morph f : x y / op x y}) (f_id : f idx = idx). +Lemma big_endo I r (P : pred I) F : + f (\big[op/idx]_(i <- r | P i) F i) = \big[op/idx]_(i <- r | P i) f (F i). +Proof. exact: big_morph. Qed. + +End Elim1. + +Implicit Arguments big_rec [R idx op I r P F]. +Implicit Arguments big_ind [R idx op I r P F]. +Implicit Arguments eq_big_op [R idx op I]. +Implicit Arguments big_endo [R idx op I]. + +Section Extensionality. + +Variables (R : Type) (idx : R) (op : R -> R -> R). + +Section SeqExtension. + +Variable I : Type. + +Lemma big_filter r (P : pred I) F : + \big[op/idx]_(i <- filter P r) F i = \big[op/idx]_(i <- r | P i) F i. +Proof. by rewrite unlock; elim: r => //= i r <-; case (P i). Qed. + +Lemma big_filter_cond r (P1 P2 : pred I) F : + \big[op/idx]_(i <- filter P1 r | P2 i) F i + = \big[op/idx]_(i <- r | P1 i && P2 i) F i. +Proof. +rewrite -big_filter -(big_filter r); congr bigop. +rewrite -filter_predI; apply: eq_filter => i; exact: andbC. +Qed. + +Lemma eq_bigl r (P1 P2 : pred I) F : + P1 =1 P2 -> + \big[op/idx]_(i <- r | P1 i) F i = \big[op/idx]_(i <- r | P2 i) F i. +Proof. by move=> eqP12; rewrite -!(big_filter r) (eq_filter eqP12). Qed. + +(* A lemma to permute aggregate conditions. *) +Lemma big_andbC r (P Q : pred I) F : + \big[op/idx]_(i <- r | P i && Q i) F i + = \big[op/idx]_(i <- r | Q i && P i) F i. +Proof. by apply: eq_bigl => i; exact: andbC. Qed. + +Lemma eq_bigr r (P : pred I) F1 F2 : (forall i, P i -> F1 i = F2 i) -> + \big[op/idx]_(i <- r | P i) F1 i = \big[op/idx]_(i <- r | P i) F2 i. +Proof. by move=> eqF12; elim/big_rec2: _ => // i x _ /eqF12-> ->. Qed. + +Lemma eq_big r (P1 P2 : pred I) F1 F2 : + P1 =1 P2 -> (forall i, P1 i -> F1 i = F2 i) -> + \big[op/idx]_(i <- r | P1 i) F1 i = \big[op/idx]_(i <- r | P2 i) F2 i. +Proof. by move/eq_bigl <-; move/eq_bigr->. Qed. + +Lemma congr_big r1 r2 (P1 P2 : pred I) F1 F2 : + r1 = r2 -> P1 =1 P2 -> (forall i, P1 i -> F1 i = F2 i) -> + \big[op/idx]_(i <- r1 | P1 i) F1 i = \big[op/idx]_(i <- r2 | P2 i) F2 i. +Proof. by move=> <-{r2}; exact: eq_big. Qed. + +Lemma big_nil (P : pred I) F : \big[op/idx]_(i <- [::] | P i) F i = idx. +Proof. by rewrite unlock. Qed. + +Lemma big_cons i r (P : pred I) F : + let x := \big[op/idx]_(j <- r | P j) F j in + \big[op/idx]_(j <- i :: r | P j) F j = if P i then op (F i) x else x. +Proof. by rewrite unlock. Qed. + +Lemma big_map J (h : J -> I) r (P : pred I) F : + \big[op/idx]_(i <- map h r | P i) F i + = \big[op/idx]_(j <- r | P (h j)) F (h j). +Proof. by rewrite unlock; elim: r => //= j r ->. Qed. + +Lemma big_nth x0 r (P : pred I) F : + \big[op/idx]_(i <- r | P i) F i + = \big[op/idx]_(0 <= i < size r | P (nth x0 r i)) (F (nth x0 r i)). +Proof. by rewrite -{1}(mkseq_nth x0 r) big_map /index_iota subn0. Qed. + +Lemma big_hasC r (P : pred I) F : + ~~ has P r -> \big[op/idx]_(i <- r | P i) F i = idx. +Proof. +by rewrite -big_filter has_count -size_filter -eqn0Ngt unlock => /nilP->. +Qed. + +Lemma big_pred0_eq (r : seq I) F : \big[op/idx]_(i <- r | false) F i = idx. +Proof. by rewrite big_hasC // has_pred0. Qed. + +Lemma big_pred0 r (P : pred I) F : + P =1 xpred0 -> \big[op/idx]_(i <- r | P i) F i = idx. +Proof. by move/eq_bigl->; exact: big_pred0_eq. Qed. + +Lemma big_cat_nested r1 r2 (P : pred I) F : + let x := \big[op/idx]_(i <- r2 | P i) F i in + \big[op/idx]_(i <- r1 ++ r2 | P i) F i = \big[op/x]_(i <- r1 | P i) F i. +Proof. by rewrite unlock /reducebig foldr_cat. Qed. + +Lemma big_catl r1 r2 (P : pred I) F : + ~~ has P r2 -> + \big[op/idx]_(i <- r1 ++ r2 | P i) F i = \big[op/idx]_(i <- r1 | P i) F i. +Proof. by rewrite big_cat_nested => /big_hasC->. Qed. + +Lemma big_catr r1 r2 (P : pred I) F : + ~~ has P r1 -> + \big[op/idx]_(i <- r1 ++ r2 | P i) F i = \big[op/idx]_(i <- r2 | P i) F i. +Proof. +rewrite -big_filter -(big_filter r2) filter_cat. +by rewrite has_count -size_filter; case: filter. +Qed. + +Lemma big_const_seq r (P : pred I) x : + \big[op/idx]_(i <- r | P i) x = iter (count P r) (op x) idx. +Proof. by rewrite unlock; elim: r => //= i r ->; case: (P i). Qed. + +End SeqExtension. + +(* The following lemmas can be used to localise extensionality to a specific *) +(* index sequence. This is done by ssreflect rewriting, before applying *) +(* congruence or induction lemmas. *) +Lemma big_seq_cond (I : eqType) r (P : pred I) F : + \big[op/idx]_(i <- r | P i) F i + = \big[op/idx]_(i <- r | (i \in r) && P i) F i. +Proof. +by rewrite -!(big_filter r); congr bigop; apply: eq_in_filter => i ->. +Qed. + +Lemma big_seq (I : eqType) (r : seq I) F : + \big[op/idx]_(i <- r) F i = \big[op/idx]_(i <- r | i \in r) F i. +Proof. by rewrite big_seq_cond big_andbC. Qed. + +Lemma eq_big_seq (I : eqType) (r : seq I) F1 F2 : + {in r, F1 =1 F2} -> \big[op/idx]_(i <- r) F1 i = \big[op/idx]_(i <- r) F2 i. +Proof. by move=> eqF; rewrite !big_seq (eq_bigr _ eqF). Qed. + +(* Similar lemmas for exposing integer indexing in the predicate. *) +Lemma big_nat_cond m n (P : pred nat) F : + \big[op/idx]_(m <= i < n | P i) F i + = \big[op/idx]_(m <= i < n | (m <= i < n) && P i) F i. +Proof. +by rewrite big_seq_cond; apply: eq_bigl => i; rewrite mem_index_iota. +Qed. + +Lemma big_nat m n F : + \big[op/idx]_(m <= i < n) F i = \big[op/idx]_(m <= i < n | m <= i < n) F i. +Proof. by rewrite big_nat_cond big_andbC. Qed. + +Lemma congr_big_nat m1 n1 m2 n2 P1 P2 F1 F2 : + m1 = m2 -> n1 = n2 -> + (forall i, m1 <= i < n2 -> P1 i = P2 i) -> + (forall i, P1 i && (m1 <= i < n2) -> F1 i = F2 i) -> + \big[op/idx]_(m1 <= i < n1 | P1 i) F1 i + = \big[op/idx]_(m2 <= i < n2 | P2 i) F2 i. +Proof. +move=> <- <- eqP12 eqF12; rewrite big_seq_cond (big_seq_cond _ P2). +apply: eq_big => i; rewrite ?inE /= !mem_index_iota. + by apply: andb_id2l; exact: eqP12. +by rewrite andbC; exact: eqF12. +Qed. + +Lemma eq_big_nat m n F1 F2 : + (forall i, m <= i < n -> F1 i = F2 i) -> + \big[op/idx]_(m <= i < n) F1 i = \big[op/idx]_(m <= i < n) F2 i. +Proof. by move=> eqF; apply: congr_big_nat. Qed. + +Lemma big_geq m n (P : pred nat) F : + m >= n -> \big[op/idx]_(m <= i < n | P i) F i = idx. +Proof. by move=> ge_m_n; rewrite /index_iota (eqnP ge_m_n) big_nil. Qed. + +Lemma big_ltn_cond m n (P : pred nat) F : + m < n -> let x := \big[op/idx]_(m.+1 <= i < n | P i) F i in + \big[op/idx]_(m <= i < n | P i) F i = if P m then op (F m) x else x. +Proof. +by case: n => [//|n] le_m_n; rewrite /index_iota subSn // big_cons. +Qed. + +Lemma big_ltn m n F : + m < n -> + \big[op/idx]_(m <= i < n) F i = op (F m) (\big[op/idx]_(m.+1 <= i < n) F i). +Proof. move=> lt_mn; exact: big_ltn_cond. Qed. + +Lemma big_addn m n a (P : pred nat) F : + \big[op/idx]_(m + a <= i < n | P i) F i = + \big[op/idx]_(m <= i < n - a | P (i + a)) F (i + a). +Proof. +rewrite /index_iota -subnDA addnC iota_addl big_map. +by apply: eq_big => ? *; rewrite addnC. +Qed. + +Lemma big_add1 m n (P : pred nat) F : + \big[op/idx]_(m.+1 <= i < n | P i) F i = + \big[op/idx]_(m <= i < n.-1 | P (i.+1)) F (i.+1). +Proof. +by rewrite -addn1 big_addn subn1; apply: eq_big => ? *; rewrite addn1. +Qed. + +Lemma big_nat_recl n m F : m <= n -> + \big[op/idx]_(m <= i < n.+1) F i = + op (F m) (\big[op/idx]_(m <= i < n) F i.+1). +Proof. by move=> lemn; rewrite big_ltn // big_add1. Qed. + +Lemma big_mkord n (P : pred nat) F : + \big[op/idx]_(0 <= i < n | P i) F i = \big[op/idx]_(i < n | P i) F i. +Proof. +rewrite /index_iota subn0 -(big_map (@nat_of_ord n)). +by congr bigop; rewrite /index_enum unlock val_ord_enum. +Qed. + +Lemma big_nat_widen m n1 n2 (P : pred nat) F : + n1 <= n2 -> + \big[op/idx]_(m <= i < n1 | P i) F i + = \big[op/idx]_(m <= i < n2 | P i && (i < n1)) F i. +Proof. +move=> len12; symmetry; rewrite -big_filter filter_predI big_filter. +have [ltn_trans eq_by_mem] := (ltn_trans, eq_sorted_irr ltn_trans ltnn). +congr bigop; apply: eq_by_mem; rewrite ?sorted_filter ?iota_ltn_sorted // => i. +rewrite mem_filter !mem_index_iota andbCA andbA andb_idr => // /andP[_]. +by move/leq_trans->. +Qed. + +Lemma big_ord_widen_cond n1 n2 (P : pred nat) (F : nat -> R) : + n1 <= n2 -> + \big[op/idx]_(i < n1 | P i) F i + = \big[op/idx]_(i < n2 | P i && (i < n1)) F i. +Proof. by move/big_nat_widen=> len12; rewrite -big_mkord len12 big_mkord. Qed. + +Lemma big_ord_widen n1 n2 (F : nat -> R) : + n1 <= n2 -> + \big[op/idx]_(i < n1) F i = \big[op/idx]_(i < n2 | i < n1) F i. +Proof. by move=> le_n12; exact: (big_ord_widen_cond (predT)). Qed. + +Lemma big_ord_widen_leq n1 n2 (P : pred 'I_(n1.+1)) F : + n1 < n2 -> + \big[op/idx]_(i < n1.+1 | P i) F i + = \big[op/idx]_(i < n2 | P (inord i) && (i <= n1)) F (inord i). +Proof. +move=> len12; pose g G i := G (inord i : 'I_(n1.+1)). +rewrite -(big_ord_widen_cond (g _ P) (g _ F) len12) {}/g. +by apply: eq_big => i *; rewrite inord_val. +Qed. + +Lemma big_ord0 P F : \big[op/idx]_(i < 0 | P i) F i = idx. +Proof. by rewrite big_pred0 => [|[]]. Qed. + +Lemma big_tnth I r (P : pred I) F : + let r_ := tnth (in_tuple r) in + \big[op/idx]_(i <- r | P i) F i + = \big[op/idx]_(i < size r | P (r_ i)) (F (r_ i)). +Proof. +case: r => /= [|x0 r]; first by rewrite big_nil big_ord0. +by rewrite (big_nth x0) big_mkord; apply: eq_big => i; rewrite (tnth_nth x0). +Qed. + +Lemma big_index_uniq (I : eqType) (r : seq I) (E : 'I_(size r) -> R) : + uniq r -> + \big[op/idx]_i E i = \big[op/idx]_(x <- r) oapp E idx (insub (index x r)). +Proof. +move=> Ur; apply/esym; rewrite big_tnth; apply: eq_bigr => i _. +by rewrite index_uniq // valK. +Qed. + +Lemma big_tuple I n (t : n.-tuple I) (P : pred I) F : + \big[op/idx]_(i <- t | P i) F i + = \big[op/idx]_(i < n | P (tnth t i)) F (tnth t i). +Proof. by rewrite big_tnth tvalK; case: _ / (esym _). Qed. + +Lemma big_ord_narrow_cond n1 n2 (P : pred 'I_n2) F (le_n12 : n1 <= n2) : + let w := widen_ord le_n12 in + \big[op/idx]_(i < n2 | P i && (i < n1)) F i + = \big[op/idx]_(i < n1 | P (w i)) F (w i). +Proof. +case: n1 => [|n1] /= in le_n12 *. + by rewrite big_ord0 big_pred0 // => i; rewrite andbF. +rewrite (big_ord_widen_leq _ _ le_n12); apply: eq_big => i. + by apply: andb_id2r => le_i_n1; congr P; apply: val_inj; rewrite /= inordK. +by case/andP=> _ le_i_n1; congr F; apply: val_inj; rewrite /= inordK. +Qed. + +Lemma big_ord_narrow_cond_leq n1 n2 (P : pred _) F (le_n12 : n1 <= n2) : + let w := @widen_ord n1.+1 n2.+1 le_n12 in + \big[op/idx]_(i < n2.+1 | P i && (i <= n1)) F i + = \big[op/idx]_(i < n1.+1 | P (w i)) F (w i). +Proof. exact: (@big_ord_narrow_cond n1.+1 n2.+1). Qed. + +Lemma big_ord_narrow n1 n2 F (le_n12 : n1 <= n2) : + let w := widen_ord le_n12 in + \big[op/idx]_(i < n2 | i < n1) F i = \big[op/idx]_(i < n1) F (w i). +Proof. exact: (big_ord_narrow_cond (predT)). Qed. + +Lemma big_ord_narrow_leq n1 n2 F (le_n12 : n1 <= n2) : + let w := @widen_ord n1.+1 n2.+1 le_n12 in + \big[op/idx]_(i < n2.+1 | i <= n1) F i = \big[op/idx]_(i < n1.+1) F (w i). +Proof. exact: (big_ord_narrow_cond_leq (predT)). Qed. + +Lemma big_ord_recl n F : + \big[op/idx]_(i < n.+1) F i = + op (F ord0) (\big[op/idx]_(i < n) F (@lift n.+1 ord0 i)). +Proof. +pose G i := F (inord i); have eqFG i: F i = G i by rewrite /G inord_val. +rewrite (eq_bigr _ (fun i _ => eqFG i)) -(big_mkord _ (fun _ => _) G) eqFG. +rewrite big_ltn // big_add1 /= big_mkord; congr op. +by apply: eq_bigr => i _; rewrite eqFG. +Qed. + +Lemma big_const (I : finType) (A : pred I) x : + \big[op/idx]_(i in A) x = iter #|A| (op x) idx. +Proof. by rewrite big_const_seq -size_filter cardE. Qed. + +Lemma big_const_nat m n x : + \big[op/idx]_(m <= i < n) x = iter (n - m) (op x) idx. +Proof. by rewrite big_const_seq count_predT size_iota. Qed. + +Lemma big_const_ord n x : + \big[op/idx]_(i < n) x = iter n (op x) idx. +Proof. by rewrite big_const card_ord. Qed. + +Lemma big_nseq_cond I n a (P : pred I) F : + \big[op/idx]_(i <- nseq n a | P i) F i = if P a then iter n (op (F a)) idx else idx. +Proof. by rewrite unlock; elim: n => /= [|n ->]; case: (P a). Qed. + +Lemma big_nseq I n a (F : I -> R): + \big[op/idx]_(i <- nseq n a) F i = iter n (op (F a)) idx. +Proof. exact: big_nseq_cond. Qed. + +End Extensionality. + +Section MonoidProperties. + +Import Monoid.Theory. + +Variable R : Type. + +Variable idx : R. +Notation Local "1" := idx. + +Section Plain. + +Variable op : Monoid.law 1. + +Notation Local "*%M" := op (at level 0). +Notation Local "x * y" := (op x y). + +Lemma eq_big_idx_seq idx' I r (P : pred I) F : + right_id idx' *%M -> has P r -> + \big[*%M/idx']_(i <- r | P i) F i =\big[*%M/1]_(i <- r | P i) F i. +Proof. +move=> op_idx'; rewrite -!(big_filter _ _ r) has_count -size_filter. +case/lastP: (filter P r) => {r}// r i _. +by rewrite -cats1 !(big_cat_nested, big_cons, big_nil) op_idx' mulm1. +Qed. + +Lemma eq_big_idx idx' (I : finType) i0 (P : pred I) F : + P i0 -> right_id idx' *%M -> + \big[*%M/idx']_(i | P i) F i =\big[*%M/1]_(i | P i) F i. +Proof. +by move=> Pi0 op_idx'; apply: eq_big_idx_seq => //; apply/hasP; exists i0. +Qed. + +Lemma big1_eq I r (P : pred I) : \big[*%M/1]_(i <- r | P i) 1 = 1. +Proof. +by rewrite big_const_seq; elim: (count _ _) => //= n ->; exact: mul1m. +Qed. + +Lemma big1 I r (P : pred I) F : + (forall i, P i -> F i = 1) -> \big[*%M/1]_(i <- r | P i) F i = 1. +Proof. by move/(eq_bigr _)->; exact: big1_eq. Qed. + +Lemma big1_seq (I : eqType) r (P : pred I) F : + (forall i, P i && (i \in r) -> F i = 1) -> + \big[*%M/1]_(i <- r | P i) F i = 1. +Proof. by move=> eqF1; rewrite big_seq_cond big_andbC big1. Qed. + +Lemma big_seq1 I (i : I) F : \big[*%M/1]_(j <- [:: i]) F j = F i. +Proof. by rewrite unlock /= mulm1. Qed. + +Lemma big_mkcond I r (P : pred I) F : + \big[*%M/1]_(i <- r | P i) F i = + \big[*%M/1]_(i <- r) (if P i then F i else 1). +Proof. by rewrite unlock; elim: r => //= i r ->; case P; rewrite ?mul1m. Qed. + +Lemma big_mkcondr I r (P Q : pred I) F : + \big[*%M/1]_(i <- r | P i && Q i) F i = + \big[*%M/1]_(i <- r | P i) (if Q i then F i else 1). +Proof. by rewrite -big_filter_cond big_mkcond big_filter. Qed. + +Lemma big_mkcondl I r (P Q : pred I) F : + \big[*%M/1]_(i <- r | P i && Q i) F i = + \big[*%M/1]_(i <- r | Q i) (if P i then F i else 1). +Proof. by rewrite big_andbC big_mkcondr. Qed. + +Lemma big_cat I r1 r2 (P : pred I) F : + \big[*%M/1]_(i <- r1 ++ r2 | P i) F i = + \big[*%M/1]_(i <- r1 | P i) F i * \big[*%M/1]_(i <- r2 | P i) F i. +Proof. +rewrite !(big_mkcond _ P) unlock. +by elim: r1 => /= [|i r1 ->]; rewrite (mul1m, mulmA). +Qed. + +Lemma big_pred1_eq (I : finType) (i : I) F : + \big[*%M/1]_(j | j == i) F j = F i. +Proof. by rewrite -big_filter filter_index_enum enum1 big_seq1. Qed. + +Lemma big_pred1 (I : finType) i (P : pred I) F : + P =1 pred1 i -> \big[*%M/1]_(j | P j) F j = F i. +Proof. by move/(eq_bigl _ _)->; exact: big_pred1_eq. Qed. + +Lemma big_cat_nat n m p (P : pred nat) F : m <= n -> n <= p -> + \big[*%M/1]_(m <= i < p | P i) F i = + (\big[*%M/1]_(m <= i < n | P i) F i) * (\big[*%M/1]_(n <= i < p | P i) F i). +Proof. +move=> le_mn le_np; rewrite -big_cat -{2}(subnKC le_mn) -iota_add subnDA. +by rewrite subnKC // leq_sub. +Qed. + +Lemma big_nat1 n F : \big[*%M/1]_(n <= i < n.+1) F i = F n. +Proof. by rewrite big_ltn // big_geq // mulm1. Qed. + +Lemma big_nat_recr n m F : m <= n -> + \big[*%M/1]_(m <= i < n.+1) F i = (\big[*%M/1]_(m <= i < n) F i) * F n. +Proof. by move=> lemn; rewrite (@big_cat_nat n) ?leqnSn // big_nat1. Qed. + +Lemma big_ord_recr n F : + \big[*%M/1]_(i < n.+1) F i = + (\big[*%M/1]_(i < n) F (widen_ord (leqnSn n) i)) * F ord_max. +Proof. +transitivity (\big[*%M/1]_(0 <= i < n.+1) F (inord i)). + by rewrite big_mkord; apply: eq_bigr=> i _; rewrite inord_val. +rewrite big_nat_recr // big_mkord; congr (_ * F _); last first. + by apply: val_inj; rewrite /= inordK. +by apply: eq_bigr => [] i _; congr F; apply: ord_inj; rewrite inordK //= leqW. +Qed. + +Lemma big_sumType (I1 I2 : finType) (P : pred (I1 + I2)) F : + \big[*%M/1]_(i | P i) F i = + (\big[*%M/1]_(i | P (inl _ i)) F (inl _ i)) + * (\big[*%M/1]_(i | P (inr _ i)) F (inr _ i)). +Proof. +by rewrite /index_enum {1}[@Finite.enum]unlock /= big_cat !big_map. +Qed. + +Lemma big_split_ord m n (P : pred 'I_(m + n)) F : + \big[*%M/1]_(i | P i) F i = + (\big[*%M/1]_(i | P (lshift n i)) F (lshift n i)) + * (\big[*%M/1]_(i | P (rshift m i)) F (rshift m i)). +Proof. +rewrite -(big_map _ _ (lshift n) _ P F) -(big_map _ _ (@rshift m _) _ P F). +rewrite -big_cat; congr bigop; apply: (inj_map val_inj). +rewrite /index_enum -!enumT val_enum_ord map_cat -map_comp val_enum_ord. +rewrite -map_comp (map_comp (addn m)) val_enum_ord. +by rewrite -iota_addl addn0 iota_add. +Qed. + +Lemma big_flatten I rr (P : pred I) F : + \big[*%M/1]_(i <- flatten rr | P i) F i + = \big[*%M/1]_(r <- rr) \big[*%M/1]_(i <- r | P i) F i. +Proof. +by elim: rr => [|r rr IHrr]; rewrite ?big_nil //= big_cat big_cons -IHrr. +Qed. + +End Plain. + +Section Abelian. + +Variable op : Monoid.com_law 1. + +Notation Local "'*%M'" := op (at level 0). +Notation Local "x * y" := (op x y). + +Lemma eq_big_perm (I : eqType) r1 r2 (P : pred I) F : + perm_eq r1 r2 -> + \big[*%M/1]_(i <- r1 | P i) F i = \big[*%M/1]_(i <- r2 | P i) F i. +Proof. +move/perm_eqP; rewrite !(big_mkcond _ _ P). +elim: r1 r2 => [|i r1 IHr1] r2 eq_r12. + by case: r2 eq_r12 => // i r2; move/(_ (pred1 i)); rewrite /= eqxx. +have r2i: i \in r2 by rewrite -has_pred1 has_count -eq_r12 /= eqxx. +case/splitPr: r2 / r2i => [r3 r4] in eq_r12 *; rewrite big_cat /= !big_cons. +rewrite mulmCA; congr (_ * _); rewrite -big_cat; apply: IHr1 => a. +move/(_ a): eq_r12; rewrite !count_cat /= addnCA; exact: addnI. +Qed. + +Lemma big_uniq (I : finType) (r : seq I) F : + uniq r -> \big[*%M/1]_(i <- r) F i = \big[*%M/1]_(i in r) F i. +Proof. +move=> uniq_r; rewrite -(big_filter _ _ _ (mem r)); apply: eq_big_perm. +by rewrite filter_index_enum uniq_perm_eq ?enum_uniq // => i; rewrite mem_enum. +Qed. + +Lemma big_rem (I : eqType) r x (P : pred I) F : + x \in r -> + \big[*%M/1]_(y <- r | P y) F y + = (if P x then F x else 1) * \big[*%M/1]_(y <- rem x r | P y) F y. +Proof. +by move/perm_to_rem/(eq_big_perm _)->; rewrite !(big_mkcond _ _ P) big_cons. +Qed. + +Lemma big_undup (I : eqType) (r : seq I) (P : pred I) F : + idempotent *%M -> + \big[*%M/1]_(i <- undup r | P i) F i = \big[*%M/1]_(i <- r | P i) F i. +Proof. +move=> idM; rewrite -!(big_filter _ _ _ P) filter_undup. +elim: {P r}(filter P r) => //= i r IHr. +case: ifP => [r_i | _]; rewrite !big_cons {}IHr //. +by rewrite (big_rem _ _ r_i) mulmA idM. +Qed. + +Lemma eq_big_idem (I : eqType) (r1 r2 : seq I) (P : pred I) F : + idempotent *%M -> r1 =i r2 -> + \big[*%M/1]_(i <- r1 | P i) F i = \big[*%M/1]_(i <- r2 | P i) F i. +Proof. +move=> idM eq_r; rewrite -big_undup // -(big_undup r2) //; apply/eq_big_perm. +by rewrite uniq_perm_eq ?undup_uniq // => i; rewrite !mem_undup eq_r. +Qed. + +Lemma big_undup_iterop_count (I : eqType) (r : seq I) (P : pred I) F : + \big[*%M/1]_(i <- undup r | P i) iterop (count_mem i r) *%M (F i) 1 + = \big[*%M/1]_(i <- r | P i) F i. +Proof. +rewrite -[RHS](eq_big_perm _ F (perm_undup_count _)) big_flatten big_map. +by rewrite big_mkcond; apply: eq_bigr => i _; rewrite big_nseq_cond iteropE. +Qed. + +Lemma big_split I r (P : pred I) F1 F2 : + \big[*%M/1]_(i <- r | P i) (F1 i * F2 i) = + \big[*%M/1]_(i <- r | P i) F1 i * \big[*%M/1]_(i <- r | P i) F2 i. +Proof. +by elim/big_rec3: _ => [|i x y _ _ ->]; rewrite ?mulm1 // mulmCA -!mulmA mulmCA. +Qed. + +Lemma bigID I r (a P : pred I) F : + \big[*%M/1]_(i <- r | P i) F i = + \big[*%M/1]_(i <- r | P i && a i) F i * + \big[*%M/1]_(i <- r | P i && ~~ a i) F i. +Proof. +rewrite !(big_mkcond _ _ _ F) -big_split. +by apply: eq_bigr => i; case: (a i); rewrite !simpm. +Qed. +Implicit Arguments bigID [I r]. + +Lemma bigU (I : finType) (A B : pred I) F : + [disjoint A & B] -> + \big[*%M/1]_(i in [predU A & B]) F i = + (\big[*%M/1]_(i in A) F i) * (\big[*%M/1]_(i in B) F i). +Proof. +move=> dAB; rewrite (bigID (mem A)). +congr (_ * _); apply: eq_bigl => i; first by rewrite orbK. +by have:= pred0P dAB i; rewrite andbC /= !inE; case: (i \in A). +Qed. + +Lemma bigD1 (I : finType) j (P : pred I) F : + P j -> \big[*%M/1]_(i | P i) F i + = F j * \big[*%M/1]_(i | P i && (i != j)) F i. +Proof. +move=> Pj; rewrite (bigID (pred1 j)); congr (_ * _). +by apply: big_pred1 => i; rewrite /= andbC; case: eqP => // ->. +Qed. +Implicit Arguments bigD1 [I P F]. + +Lemma bigD1_seq (I : eqType) (r : seq I) j F : + j \in r -> uniq r -> + \big[*%M/1]_(i <- r) F i = F j * \big[*%M/1]_(i <- r | i != j) F i. +Proof. by move=> /big_rem-> /rem_filter->; rewrite big_filter. Qed. + +Lemma cardD1x (I : finType) (A : pred I) j : + A j -> #|SimplPred A| = 1 + #|[pred i | A i & i != j]|. +Proof. +move=> Aj; rewrite (cardD1 j) [j \in A]Aj; congr (_ + _). +by apply: eq_card => i; rewrite inE /= andbC. +Qed. +Implicit Arguments cardD1x [I A]. + +Lemma partition_big (I J : finType) (P : pred I) p (Q : pred J) F : + (forall i, P i -> Q (p i)) -> + \big[*%M/1]_(i | P i) F i = + \big[*%M/1]_(j | Q j) \big[*%M/1]_(i | P i && (p i == j)) F i. +Proof. +move=> Qp; transitivity (\big[*%M/1]_(i | P i && Q (p i)) F i). + by apply: eq_bigl => i; case Pi: (P i); rewrite // Qp. +elim: {Q Qp}_.+1 {-2}Q (ltnSn #|Q|) => // n IHn Q. +case: (pickP Q) => [j Qj | Q0 _]; last first. + by rewrite !big_pred0 // => i; rewrite Q0 andbF. +rewrite ltnS (cardD1x j Qj) (bigD1 j) //; move/IHn=> {n IHn} <-. +rewrite (bigID (fun i => p i == j)); congr (_ * _); apply: eq_bigl => i. + by case: eqP => [-> | _]; rewrite !(Qj, simpm). +by rewrite andbA. +Qed. + +Implicit Arguments partition_big [I J P F]. + +Lemma reindex_onto (I J : finType) (h : J -> I) h' (P : pred I) F : + (forall i, P i -> h (h' i) = i) -> + \big[*%M/1]_(i | P i) F i = + \big[*%M/1]_(j | P (h j) && (h' (h j) == j)) F (h j). +Proof. +move=> h'K; elim: {P}_.+1 {-3}P h'K (ltnSn #|P|) => //= n IHn P h'K. +case: (pickP P) => [i Pi | P0 _]; last first. + by rewrite !big_pred0 // => j; rewrite P0. +rewrite ltnS (cardD1x i Pi); move/IHn {n IHn} => IH. +rewrite (bigD1 i Pi) (bigD1 (h' i)) h'K ?Pi ?eq_refl //=; congr (_ * _). +rewrite {}IH => [|j]; [apply: eq_bigl => j | by case/andP; auto]. +rewrite andbC -andbA (andbCA (P _)); case: eqP => //= hK; congr (_ && ~~ _). +by apply/eqP/eqP=> [<-|->] //; rewrite h'K. +Qed. +Implicit Arguments reindex_onto [I J P F]. + +Lemma reindex (I J : finType) (h : J -> I) (P : pred I) F : + {on [pred i | P i], bijective h} -> + \big[*%M/1]_(i | P i) F i = \big[*%M/1]_(j | P (h j)) F (h j). +Proof. +case=> h' hK h'K; rewrite (reindex_onto h h' h'K). +by apply: eq_bigl => j; rewrite !inE; case Pi: (P _); rewrite //= hK ?eqxx. +Qed. +Implicit Arguments reindex [I J P F]. + +Lemma reindex_inj (I : finType) (h : I -> I) (P : pred I) F : + injective h -> \big[*%M/1]_(i | P i) F i = \big[*%M/1]_(j | P (h j)) F (h j). +Proof. move=> injh; exact: reindex (onW_bij _ (injF_bij injh)). Qed. +Implicit Arguments reindex_inj [I h P F]. + +Lemma big_nat_rev m n P F : + \big[*%M/1]_(m <= i < n | P i) F i + = \big[*%M/1]_(m <= i < n | P (m + n - i.+1)) F (m + n - i.+1). +Proof. +case: (ltnP m n) => ltmn; last by rewrite !big_geq. +rewrite -{3 4}(subnK (ltnW ltmn)) addnA. +do 2!rewrite (big_addn _ _ 0) big_mkord; rewrite (reindex_inj rev_ord_inj) /=. +by apply: eq_big => [i | i _]; rewrite /= -addSn subnDr addnC addnBA. +Qed. + +Lemma pair_big_dep (I J : finType) (P : pred I) (Q : I -> pred J) F : + \big[*%M/1]_(i | P i) \big[*%M/1]_(j | Q i j) F i j = + \big[*%M/1]_(p | P p.1 && Q p.1 p.2) F p.1 p.2. +Proof. +rewrite (partition_big (fun p => p.1) P) => [|j]; last by case/andP. +apply: eq_bigr => i /= Pi; rewrite (reindex_onto (pair i) (fun p => p.2)). + by apply: eq_bigl => j; rewrite !eqxx [P i]Pi !andbT. +by case=> i' j /=; case/andP=> _ /=; move/eqP->. +Qed. + +Lemma pair_big (I J : finType) (P : pred I) (Q : pred J) F : + \big[*%M/1]_(i | P i) \big[*%M/1]_(j | Q j) F i j = + \big[*%M/1]_(p | P p.1 && Q p.2) F p.1 p.2. +Proof. exact: pair_big_dep. Qed. + +Lemma pair_bigA (I J : finType) (F : I -> J -> R) : + \big[*%M/1]_i \big[*%M/1]_j F i j = \big[*%M/1]_p F p.1 p.2. +Proof. exact: pair_big_dep. Qed. + +Lemma exchange_big_dep I J rI rJ (P : pred I) (Q : I -> pred J) + (xQ : pred J) F : + (forall i j, P i -> Q i j -> xQ j) -> + \big[*%M/1]_(i <- rI | P i) \big[*%M/1]_(j <- rJ | Q i j) F i j = + \big[*%M/1]_(j <- rJ | xQ j) \big[*%M/1]_(i <- rI | P i && Q i j) F i j. +Proof. +move=> PQxQ; pose p u := (u.2, u.1). +rewrite (eq_bigr _ _ _ (fun _ _ => big_tnth _ _ rI _ _)) (big_tnth _ _ rJ). +rewrite (eq_bigr _ _ _ (fun _ _ => (big_tnth _ _ rJ _ _))) big_tnth. +rewrite !pair_big_dep (reindex_onto (p _ _) (p _ _)) => [|[]] //=. +apply: eq_big => [] [j i] //=; symmetry; rewrite eqxx andbT andb_idl //. +by case/andP; exact: PQxQ. +Qed. +Implicit Arguments exchange_big_dep [I J rI rJ P Q F]. + +Lemma exchange_big I J rI rJ (P : pred I) (Q : pred J) F : + \big[*%M/1]_(i <- rI | P i) \big[*%M/1]_(j <- rJ | Q j) F i j = + \big[*%M/1]_(j <- rJ | Q j) \big[*%M/1]_(i <- rI | P i) F i j. +Proof. +rewrite (exchange_big_dep Q) //; apply: eq_bigr => i /= Qi. +by apply: eq_bigl => j; rewrite Qi andbT. +Qed. + +Lemma exchange_big_dep_nat m1 n1 m2 n2 (P : pred nat) (Q : rel nat) + (xQ : pred nat) F : + (forall i j, m1 <= i < n1 -> m2 <= j < n2 -> P i -> Q i j -> xQ j) -> + \big[*%M/1]_(m1 <= i < n1 | P i) \big[*%M/1]_(m2 <= j < n2 | Q i j) F i j = + \big[*%M/1]_(m2 <= j < n2 | xQ j) + \big[*%M/1]_(m1 <= i < n1 | P i && Q i j) F i j. +Proof. +move=> PQxQ; rewrite (eq_bigr _ _ _ (fun _ _ => big_seq_cond _ _ _ _ _)). +rewrite big_seq_cond /= (exchange_big_dep xQ) => [|i j]; last first. + by rewrite !mem_index_iota => /andP[mn_i Pi] /andP[mn_j /PQxQ->]. +rewrite 2!(big_seq_cond _ _ _ xQ); apply: eq_bigr => j /andP[-> _] /=. +by rewrite [rhs in _ = rhs]big_seq_cond; apply: eq_bigl => i; rewrite -andbA. +Qed. +Implicit Arguments exchange_big_dep_nat [m1 n1 m2 n2 P Q F]. + +Lemma exchange_big_nat m1 n1 m2 n2 (P Q : pred nat) F : + \big[*%M/1]_(m1 <= i < n1 | P i) \big[*%M/1]_(m2 <= j < n2 | Q j) F i j = + \big[*%M/1]_(m2 <= j < n2 | Q j) \big[*%M/1]_(m1 <= i < n1 | P i) F i j. +Proof. +rewrite (exchange_big_dep_nat Q) //. +by apply: eq_bigr => i /= Qi; apply: eq_bigl => j; rewrite Qi andbT. +Qed. + +End Abelian. + +End MonoidProperties. + +Implicit Arguments big_filter [R op idx I]. +Implicit Arguments big_filter_cond [R op idx I]. +Implicit Arguments congr_big [R op idx I r1 P1 F1]. +Implicit Arguments eq_big [R op idx I r P1 F1]. +Implicit Arguments eq_bigl [R op idx I r P1]. +Implicit Arguments eq_bigr [R op idx I r P F1]. +Implicit Arguments eq_big_idx [R op idx idx' I P F]. +Implicit Arguments big_seq_cond [R op idx I r]. +Implicit Arguments eq_big_seq [R op idx I r F1]. +Implicit Arguments congr_big_nat [R op idx m1 n1 P1 F1]. +Implicit Arguments big_map [R op idx I J r]. +Implicit Arguments big_nth [R op idx I r]. +Implicit Arguments big_catl [R op idx I r1 r2 P F]. +Implicit Arguments big_catr [R op idx I r1 r2 P F]. +Implicit Arguments big_geq [R op idx m n P F]. +Implicit Arguments big_ltn_cond [R op idx m n P F]. +Implicit Arguments big_ltn [R op idx m n F]. +Implicit Arguments big_addn [R op idx]. +Implicit Arguments big_mkord [R op idx n]. +Implicit Arguments big_nat_widen [R op idx] . +Implicit Arguments big_ord_widen_cond [R op idx n1]. +Implicit Arguments big_ord_widen [R op idx n1]. +Implicit Arguments big_ord_widen_leq [R op idx n1]. +Implicit Arguments big_ord_narrow_cond [R op idx n1 n2 P F]. +Implicit Arguments big_ord_narrow_cond_leq [R op idx n1 n2 P F]. +Implicit Arguments big_ord_narrow [R op idx n1 n2 F]. +Implicit Arguments big_ord_narrow_leq [R op idx n1 n2 F]. +Implicit Arguments big_mkcond [R op idx I r]. +Implicit Arguments big1_eq [R op idx I]. +Implicit Arguments big1_seq [R op idx I]. +Implicit Arguments big1 [R op idx I]. +Implicit Arguments big_pred1 [R op idx I P F]. +Implicit Arguments eq_big_perm [R op idx I r1 P F]. +Implicit Arguments big_uniq [R op idx I F]. +Implicit Arguments big_rem [R op idx I r P F]. +Implicit Arguments bigID [R op idx I r]. +Implicit Arguments bigU [R op idx I]. +Implicit Arguments bigD1 [R op idx I P F]. +Implicit Arguments bigD1_seq [R op idx I r F]. +Implicit Arguments partition_big [R op idx I J P F]. +Implicit Arguments reindex_onto [R op idx I J P F]. +Implicit Arguments reindex [R op idx I J P F]. +Implicit Arguments reindex_inj [R op idx I h P F]. +Implicit Arguments pair_big_dep [R op idx I J]. +Implicit Arguments pair_big [R op idx I J]. +Implicit Arguments exchange_big_dep [R op idx I J rI rJ P Q F]. +Implicit Arguments exchange_big_dep_nat [R op idx m1 n1 m2 n2 P Q F]. +Implicit Arguments big_ord_recl [R op idx]. +Implicit Arguments big_ord_recr [R op idx]. +Implicit Arguments big_nat_recl [R op idx]. +Implicit Arguments big_nat_recr [R op idx]. + +Section Distributivity. + +Import Monoid.Theory. + +Variable R : Type. +Variables zero one : R. +Notation Local "0" := zero. +Notation Local "1" := one. +Variable times : Monoid.mul_law 0. +Notation Local "*%M" := times (at level 0). +Notation Local "x * y" := (times x y). +Variable plus : Monoid.add_law 0 *%M. +Notation Local "+%M" := plus (at level 0). +Notation Local "x + y" := (plus x y). + +Lemma big_distrl I r a (P : pred I) F : + \big[+%M/0]_(i <- r | P i) F i * a = \big[+%M/0]_(i <- r | P i) (F i * a). +Proof. by rewrite (big_endo ( *%M^~ a)) ?mul0m // => x y; exact: mulm_addl. Qed. + +Lemma big_distrr I r a (P : pred I) F : + a * \big[+%M/0]_(i <- r | P i) F i = \big[+%M/0]_(i <- r | P i) (a * F i). +Proof. by rewrite big_endo ?mulm0 // => x y; exact: mulm_addr. Qed. + +Lemma big_distrlr I J rI rJ (pI : pred I) (pJ : pred J) F G : + (\big[+%M/0]_(i <- rI | pI i) F i) * (\big[+%M/0]_(j <- rJ | pJ j) G j) + = \big[+%M/0]_(i <- rI | pI i) \big[+%M/0]_(j <- rJ | pJ j) (F i * G j). +Proof. by rewrite big_distrl; apply: eq_bigr => i _; rewrite big_distrr. Qed. + +Lemma big_distr_big_dep (I J : finType) j0 (P : pred I) (Q : I -> pred J) F : + \big[*%M/1]_(i | P i) \big[+%M/0]_(j | Q i j) F i j = + \big[+%M/0]_(f in pfamily j0 P Q) \big[*%M/1]_(i | P i) F i (f i). +Proof. +pose fIJ := {ffun I -> J}; pose Pf := pfamily j0 (_ : seq I) Q. +rewrite -big_filter filter_index_enum; set r := enum P; symmetry. +transitivity (\big[+%M/0]_(f in Pf r) \big[*%M/1]_(i <- r) F i (f i)). + apply: eq_big => f; last by rewrite -big_filter filter_index_enum. + by apply: eq_forallb => i; rewrite /= mem_enum. +have: uniq r by exact: enum_uniq. +elim: {P}r => /= [_ | i r IHr]. + rewrite (big_pred1 [ffun => j0]) ?big_nil //= => f. + apply/familyP/eqP=> /= [Df |->{f} i]; last by rewrite ffunE !inE. + by apply/ffunP=> i; rewrite ffunE; exact/eqP/Df. +case/andP=> /negbTE nri; rewrite big_cons big_distrl => {IHr}/IHr <-. +rewrite (partition_big (fun f : fIJ => f i) (Q i)) => [|f]; last first. + by move/familyP/(_ i); rewrite /= inE /= eqxx. +pose seti j (f : fIJ) := [ffun k => if k == i then j else f k]. +apply: eq_bigr => j Qij. +rewrite (reindex_onto (seti j) (seti j0)) => [|f /andP[_ /eqP fi]]; last first. + by apply/ffunP=> k; rewrite !ffunE; case: eqP => // ->. +rewrite big_distrr; apply: eq_big => [f | f eq_f]; last first. + rewrite big_cons ffunE eqxx !big_seq; congr (_ * _). + by apply: eq_bigr => k; rewrite ffunE; case: eqP nri => // -> ->. +rewrite !ffunE !eqxx andbT; apply/andP/familyP=> /= [[Pjf fij0] k | Pff]. + have:= familyP Pjf k; rewrite /= ffunE inE; case: eqP => // -> _. + by rewrite nri -(eqP fij0) !ffunE !inE !eqxx. +split; [apply/familyP | apply/eqP/ffunP] => k; have:= Pff k; rewrite !ffunE. + by rewrite inE; case: eqP => // ->. +by case: eqP => // ->; rewrite nri /= => /eqP. +Qed. + +Lemma big_distr_big (I J : finType) j0 (P : pred I) (Q : pred J) F : + \big[*%M/1]_(i | P i) \big[+%M/0]_(j | Q j) F i j = + \big[+%M/0]_(f in pffun_on j0 P Q) \big[*%M/1]_(i | P i) F i (f i). +Proof. +rewrite (big_distr_big_dep j0); apply: eq_bigl => f. +by apply/familyP/familyP=> Pf i; case: ifP (Pf i). +Qed. + +Lemma bigA_distr_big_dep (I J : finType) (Q : I -> pred J) F : + \big[*%M/1]_i \big[+%M/0]_(j | Q i j) F i j + = \big[+%M/0]_(f in family Q) \big[*%M/1]_i F i (f i). +Proof. +case: (pickP J) => [j0 _ | J0]; first exact: (big_distr_big_dep j0). +rewrite {1 4}/index_enum -enumT; case: (enum I) (mem_enum I) => [I0 | i r _]. + have f0: I -> J by move=> i; have:= I0 i. + rewrite (big_pred1 (finfun f0)) ?big_nil // => g. + by apply/familyP/eqP=> _; first apply/ffunP; move=> i; have:= I0 i. +have Q0 i': Q i' =1 pred0 by move=> j; have:= J0 j. +rewrite big_cons /= big_pred0 // mul0m big_pred0 // => f. +by apply/familyP=> /(_ i); rewrite [_ \in _]Q0. +Qed. + +Lemma bigA_distr_big (I J : finType) (Q : pred J) (F : I -> J -> R) : + \big[*%M/1]_i \big[+%M/0]_(j | Q j) F i j + = \big[+%M/0]_(f in ffun_on Q) \big[*%M/1]_i F i (f i). +Proof. exact: bigA_distr_big_dep. Qed. + +Lemma bigA_distr_bigA (I J : finType) F : + \big[*%M/1]_(i : I) \big[+%M/0]_(j : J) F i j + = \big[+%M/0]_(f : {ffun I -> J}) \big[*%M/1]_i F i (f i). +Proof. by rewrite bigA_distr_big; apply: eq_bigl => ?; exact/familyP. Qed. + +End Distributivity. + +Implicit Arguments big_distrl [R zero times plus I r]. +Implicit Arguments big_distrr [R zero times plus I r]. +Implicit Arguments big_distr_big_dep [R zero one times plus I J]. +Implicit Arguments big_distr_big [R zero one times plus I J]. +Implicit Arguments bigA_distr_big_dep [R zero one times plus I J]. +Implicit Arguments bigA_distr_big [R zero one times plus I J]. +Implicit Arguments bigA_distr_bigA [R zero one times plus I J]. + +Section BigBool. + +Section Seq. + +Variables (I : Type) (r : seq I) (P B : pred I). + +Lemma big_has : \big[orb/false]_(i <- r) B i = has B r. +Proof. by rewrite unlock. Qed. + +Lemma big_all : \big[andb/true]_(i <- r) B i = all B r. +Proof. by rewrite unlock. Qed. + +Lemma big_has_cond : \big[orb/false]_(i <- r | P i) B i = has (predI P B) r. +Proof. by rewrite big_mkcond unlock. Qed. + +Lemma big_all_cond : + \big[andb/true]_(i <- r | P i) B i = all [pred i | P i ==> B i] r. +Proof. by rewrite big_mkcond unlock. Qed. + +End Seq. + +Section FinType. + +Variables (I : finType) (P B : pred I). + +Lemma big_orE : \big[orb/false]_(i | P i) B i = [exists (i | P i), B i]. +Proof. by rewrite big_has_cond; apply/hasP/existsP=> [] [i]; exists i. Qed. + +Lemma big_andE : \big[andb/true]_(i | P i) B i = [forall (i | P i), B i]. +Proof. +rewrite big_all_cond; apply/allP/forallP=> /= allB i; rewrite allB //. +exact: mem_index_enum. +Qed. + +End FinType. + +End BigBool. + +Section NatConst. + +Variables (I : finType) (A : pred I). + +Lemma sum_nat_const n : \sum_(i in A) n = #|A| * n. +Proof. by rewrite big_const iter_addn_0 mulnC. Qed. + +Lemma sum1_card : \sum_(i in A) 1 = #|A|. +Proof. by rewrite sum_nat_const muln1. Qed. + +Lemma sum1_count J (r : seq J) (a : pred J) : \sum_(j <- r | a j) 1 = count a r. +Proof. by rewrite big_const_seq iter_addn_0 mul1n. Qed. + +Lemma sum1_size J (r : seq J) : \sum_(j <- r) 1 = size r. +Proof. by rewrite sum1_count count_predT. Qed. + +Lemma prod_nat_const n : \prod_(i in A) n = n ^ #|A|. +Proof. by rewrite big_const -Monoid.iteropE. Qed. + +Lemma sum_nat_const_nat n1 n2 n : \sum_(n1 <= i < n2) n = (n2 - n1) * n. +Proof. by rewrite big_const_nat; elim: (_ - _) => //= ? ->. Qed. + +Lemma prod_nat_const_nat n1 n2 n : \prod_(n1 <= i < n2) n = n ^ (n2 - n1). +Proof. by rewrite big_const_nat -Monoid.iteropE. Qed. + +End NatConst. + +Lemma leqif_sum (I : finType) (P C : pred I) (E1 E2 : I -> nat) : + (forall i, P i -> E1 i <= E2 i ?= iff C i) -> + \sum_(i | P i) E1 i <= \sum_(i | P i) E2 i ?= iff [forall (i | P i), C i]. +Proof. +move=> leE12; rewrite -big_andE. +by elim/big_rec3: _ => // i Ci m1 m2 /leE12; exact: leqif_add. +Qed. + +Lemma leq_sum I r (P : pred I) (E1 E2 : I -> nat) : + (forall i, P i -> E1 i <= E2 i) -> + \sum_(i <- r | P i) E1 i <= \sum_(i <- r | P i) E2 i. +Proof. by move=> leE12; elim/big_ind2: _ => // m1 m2 n1 n2; exact: leq_add. Qed. + +Lemma sum_nat_eq0 (I : finType) (P : pred I) (E : I -> nat) : + (\sum_(i | P i) E i == 0)%N = [forall (i | P i), E i == 0%N]. +Proof. by rewrite eq_sym -(@leqif_sum I P _ (fun _ => 0%N) E) ?big1_eq. Qed. + +Lemma prodn_cond_gt0 I r (P : pred I) F : + (forall i, P i -> 0 < F i) -> 0 < \prod_(i <- r | P i) F i. +Proof. by move=> Fpos; elim/big_ind: _ => // n1 n2; rewrite muln_gt0 => ->. Qed. + +Lemma prodn_gt0 I r (P : pred I) F : + (forall i, 0 < F i) -> 0 < \prod_(i <- r | P i) F i. +Proof. move=> Fpos; exact: prodn_cond_gt0. Qed. + +Lemma leq_bigmax_cond (I : finType) (P : pred I) F i0 : + P i0 -> F i0 <= \max_(i | P i) F i. +Proof. by move=> Pi0; rewrite (bigD1 i0) ?leq_maxl. Qed. +Implicit Arguments leq_bigmax_cond [I P F]. + +Lemma leq_bigmax (I : finType) F (i0 : I) : F i0 <= \max_i F i. +Proof. exact: leq_bigmax_cond. Qed. +Implicit Arguments leq_bigmax [I F]. + +Lemma bigmax_leqP (I : finType) (P : pred I) m F : + reflect (forall i, P i -> F i <= m) (\max_(i | P i) F i <= m). +Proof. +apply: (iffP idP) => leFm => [i Pi|]. + by apply: leq_trans leFm; exact: leq_bigmax_cond. +by elim/big_ind: _ => // m1 m2; rewrite geq_max => ->. +Qed. + +Lemma bigmax_sup (I : finType) i0 (P : pred I) m F : + P i0 -> m <= F i0 -> m <= \max_(i | P i) F i. +Proof. by move=> Pi0 le_m_Fi0; exact: leq_trans (leq_bigmax_cond i0 Pi0). Qed. +Implicit Arguments bigmax_sup [I P m F]. + +Lemma bigmax_eq_arg (I : finType) i0 (P : pred I) F : + P i0 -> \max_(i | P i) F i = F [arg max_(i > i0 | P i) F i]. +Proof. +move=> Pi0; case: arg_maxP => //= i Pi maxFi. +by apply/eqP; rewrite eqn_leq leq_bigmax_cond // andbT; exact/bigmax_leqP. +Qed. +Implicit Arguments bigmax_eq_arg [I P F]. + +Lemma eq_bigmax_cond (I : finType) (A : pred I) F : + #|A| > 0 -> {i0 | i0 \in A & \max_(i in A) F i = F i0}. +Proof. +case: (pickP A) => [i0 Ai0 _ | ]; last by move/eq_card0->. +by exists [arg max_(i > i0 in A) F i]; [case: arg_maxP | exact: bigmax_eq_arg]. +Qed. + +Lemma eq_bigmax (I : finType) F : #|I| > 0 -> {i0 : I | \max_i F i = F i0}. +Proof. by case/(eq_bigmax_cond F) => x _ ->; exists x. Qed. + +Lemma expn_sum m I r (P : pred I) F : + (m ^ (\sum_(i <- r | P i) F i) = \prod_(i <- r | P i) m ^ F i)%N. +Proof. exact: (big_morph _ (expnD m)). Qed. + +Lemma dvdn_biglcmP (I : finType) (P : pred I) F m : + reflect (forall i, P i -> F i %| m) (\big[lcmn/1%N]_(i | P i) F i %| m). +Proof. +apply: (iffP idP) => [dvFm i Pi | dvFm]. + by rewrite (bigD1 i) // dvdn_lcm in dvFm; case/andP: dvFm. +by elim/big_ind: _ => // p q p_m; rewrite dvdn_lcm p_m. +Qed. + +Lemma biglcmn_sup (I : finType) i0 (P : pred I) F m : + P i0 -> m %| F i0 -> m %| \big[lcmn/1%N]_(i | P i) F i. +Proof. +by move=> Pi0 m_Fi0; rewrite (dvdn_trans m_Fi0) // (bigD1 i0) ?dvdn_lcml. +Qed. +Implicit Arguments biglcmn_sup [I P F m]. + +Lemma dvdn_biggcdP (I : finType) (P : pred I) F m : + reflect (forall i, P i -> m %| F i) (m %| \big[gcdn/0]_(i | P i) F i). +Proof. +apply: (iffP idP) => [dvmF i Pi | dvmF]. + by rewrite (bigD1 i) // dvdn_gcd in dvmF; case/andP: dvmF. +by elim/big_ind: _ => // p q m_p; rewrite dvdn_gcd m_p. +Qed. + +Lemma biggcdn_inf (I : finType) i0 (P : pred I) F m : + P i0 -> F i0 %| m -> \big[gcdn/0]_(i | P i) F i %| m. +Proof. by move=> Pi0; apply: dvdn_trans; rewrite (bigD1 i0) ?dvdn_gcdl. Qed. +Implicit Arguments biggcdn_inf [I P F m]. + +Unset Implicit Arguments. diff --git a/mathcomp/discrete/binomial.v b/mathcomp/discrete/binomial.v new file mode 100644 index 0000000..756a8f9 --- /dev/null +++ b/mathcomp/discrete/binomial.v @@ -0,0 +1,524 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq path div. +Require Import fintype tuple finfun bigop prime finset. + +(******************************************************************************) +(* This files contains the definition of: *) +(* n ^_ m == the falling (or lower) factorial of n with m terms, i.e., *) +(* the product n * (n - 1) * ... * (n - m + 1) *) +(* Note that n ^_ m = 0 if m > n. *) +(* 'C(n, m) == the binomial coeficient n choose m *) +(* := n ^_ m %/ fact m *) +(* *) +(* In additions to the properties of these functions, triangular_sum, Wilson *) +(* and Pascal are examples of how to manipulate expressions with bigops. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(** More properties of the factorial **) + +Lemma fact_smonotone m n : 0 < m -> m < n -> m`! < n`!. +Proof. +case: m => // m _; elim: n m => // n IHn [|m] lt_m_n. + by rewrite -[_.+1]muln1 leq_mul ?fact_gt0. +by rewrite ltn_mul ?IHn. +Qed. + +Lemma fact_prod n : n`! = \prod_(1 <= i < n.+1) i. +Proof. +elim: n => [|n IHn] //; first by rewrite big_nil. +by apply sym_equal; rewrite factS IHn // !big_add1 big_nat_recr //= mulnC. +Qed. + +Lemma logn_fact p n : prime p -> logn p n`! = \sum_(1 <= k < n.+1) n %/ p ^ k. +Proof. +move=> p_prime; transitivity (\sum_(1 <= i < n.+1) logn p i). + rewrite big_add1; elim: n => /= [|n IHn]; first by rewrite logn1 big_geq. + by rewrite big_nat_recr // -IHn /= factS mulnC lognM ?fact_gt0. +transitivity (\sum_(1 <= i < n.+1) \sum_(1 <= k < n.+1) (p ^ k %| i)). + apply: eq_big_nat => i /andP[i_gt0 le_i_n]; rewrite logn_count_dvd //. + rewrite -!big_mkcond (big_nat_widen _ _ n.+1) 1?ltnW //; apply: eq_bigl => k. + by apply: andb_idr => /dvdn_leq/(leq_trans (ltn_expl _ (prime_gt1 _)))->. +by rewrite exchange_big_nat; apply: eq_bigr => i _; rewrite divn_count_dvd. +Qed. + +Theorem Wilson p : p > 1 -> prime p = (p %| ((p.-1)`!).+1). +Proof. +have dFact n: 0 < n -> (n.-1)`! = \prod_(0 <= i < n | i != 0) i. + move=> n_gt0; rewrite -big_filter fact_prod; symmetry; apply: congr_big => //. + rewrite /index_iota subn1 -[n]prednK //=; apply/all_filterP. + by rewrite all_predC has_pred1 mem_iota. +move=> lt1p; have p_gt0 := ltnW lt1p. +apply/idP/idP=> [pr_p | dv_pF]; last first. + apply/primeP; split=> // d dv_dp; have: d <= p by exact: dvdn_leq. + rewrite orbC leq_eqVlt => /orP[-> // | ltdp]. + have:= dvdn_trans dv_dp dv_pF; rewrite dFact // big_mkord. + rewrite (bigD1 (Ordinal ltdp)) /=; last by rewrite -lt0n (dvdn_gt0 p_gt0). + by rewrite orbC -addn1 dvdn_addr ?dvdn_mulr // dvdn1 => ->. +pose Fp1 := Ordinal lt1p; pose Fp0 := Ordinal p_gt0. +have ltp1p: p.-1 < p by [rewrite prednK]; pose Fpn1 := Ordinal ltp1p. +case eqF1n1: (Fp1 == Fpn1); first by rewrite -{1}[p]prednK -1?((1 =P p.-1) _). +have toFpP m: m %% p < p by rewrite ltn_mod. +pose toFp := Ordinal (toFpP _); pose mFp (i j : 'I_p) := toFp (i * j). +have Fp_mod (i : 'I_p) : i %% p = i by exact: modn_small. +have mFpA: associative mFp. + by move=> i j k; apply: val_inj; rewrite /= modnMml modnMmr mulnA. +have mFpC: commutative mFp by move=> i j; apply: val_inj; rewrite /= mulnC. +have mFp1: left_id Fp1 mFp by move=> i; apply: val_inj; rewrite /= mul1n. +have mFp1r: right_id Fp1 mFp by move=> i; apply: val_inj; rewrite /= muln1. +pose mFpLaw := Monoid.Law mFpA mFp1 mFp1r. +pose mFpM := Monoid.operator (@Monoid.ComLaw _ _ mFpLaw mFpC). +pose vFp (i : 'I_p) := toFp (egcdn i p).1. +have vFpV i: i != Fp0 -> mFp (vFp i) i = Fp1. + rewrite -val_eqE /= -lt0n => i_gt0; apply: val_inj => /=. + rewrite modnMml; case: egcdnP => //= _ km -> _; rewrite {km}modnMDl. + suffices: coprime i p by move/eqnP->; rewrite modn_small. + rewrite coprime_sym prime_coprime //; apply/negP=> /(dvdn_leq i_gt0). + by rewrite leqNgt ltn_ord. +have vFp0 i: i != Fp0 -> vFp i != Fp0. + move/vFpV=> inv_i; apply/eqP=> vFp0. + by have:= congr1 val inv_i; rewrite vFp0 /= mod0n. +have vFpK: {in predC1 Fp0, involutive vFp}. + move=> i n0i; rewrite /= -[vFp _]mFp1r -(vFpV _ n0i) mFpA. + by rewrite vFpV (vFp0, mFp1). +have le_pmFp (i : 'I_p) m: i <= p + m. + by apply: leq_trans (ltnW _) (leq_addr _ _). +have eqFp (i j : 'I_p): (i == j) = (p %| p + i - j). + by rewrite -eqn_mod_dvd ?(modnDl, Fp_mod). +have vFpId i: (vFp i == i :> nat) = xpred2 Fp1 Fpn1 i. + symmetry; have [->{i} | /eqP ni0] := i =P Fp0. + by rewrite /= -!val_eqE /= -{2}[p]prednK //= modn_small //= -(subnKC lt1p). + rewrite 2!eqFp -Euclid_dvdM //= -[_ - p.-1]subSS prednK //. + have lt0i: 0 < i by rewrite lt0n. + rewrite -addnS addKn -addnBA // mulnDl -{2}(addn1 i) -subn_sqr. + rewrite addnBA ?leq_sqr // mulnS -addnA -mulnn -mulnDl. + rewrite -(subnK (le_pmFp (vFp i) i)) mulnDl addnCA. + rewrite -[1 ^ 2]/(Fp1 : nat) -addnBA // dvdn_addl. + by rewrite Euclid_dvdM // -eqFp eq_sym orbC /dvdn Fp_mod eqn0Ngt lt0i. + by rewrite -eqn_mod_dvd // Fp_mod modnDl -(vFpV _ ni0) eqxx. +suffices [mod_fact]: toFp (p.-1)`! = Fpn1. + by rewrite /dvdn -addn1 -modnDml mod_fact addn1 prednK // modnn. +rewrite dFact //; rewrite ((big_morph toFp) Fp1 mFpM) //; first last. +- by apply: val_inj; rewrite /= modn_small. +- by move=> i j; apply: val_inj; rewrite /= modnMm. +rewrite big_mkord (eq_bigr id) => [|i _]; last by apply: val_inj => /=. +pose ltv i := vFp i < i; rewrite (bigID ltv) -/mFpM [mFpM _ _]mFpC. +rewrite (bigD1 Fp1) -/mFpM; last by rewrite [ltv _]ltn_neqAle vFpId. +rewrite [mFpM _ _]mFp1 (bigD1 Fpn1) -?mFpA -/mFpM; last first. + rewrite -lt0n -ltnS prednK // lt1p. + by rewrite [ltv _]ltn_neqAle vFpId eqxx orbT eq_sym eqF1n1. +rewrite (reindex_onto vFp vFp) -/mFpM => [|i]; last by do 3!case/andP; auto. +rewrite (eq_bigl (xpredD1 ltv Fp0)) => [|i]; last first. + rewrite andbC -!andbA -2!negb_or -vFpId orbC -leq_eqVlt. + rewrite andbA -ltnNge; symmetry; case: (altP eqP) => [->|ni0]. + by case: eqP => // E; rewrite ?E !andbF. + by rewrite vFpK //eqxx vFp0. +rewrite -{2}[mFp]/mFpM -[mFpM _ _]big_split -/mFpM. +by rewrite big1 ?mFp1r //= => i /andP[]; auto. +Qed. + +(** The falling factorial *) + +Fixpoint ffact_rec n m := if m is m'.+1 then n * ffact_rec n.-1 m' else 1. + +Definition falling_factorial := nosimpl ffact_rec. + +Notation "n ^_ m" := (falling_factorial n m) + (at level 30, right associativity) : nat_scope. + +Lemma ffactE : falling_factorial = ffact_rec. Proof. by []. Qed. + +Lemma ffactn0 n : n ^_ 0 = 1. Proof. by []. Qed. + +Lemma ffact0n m : 0 ^_ m = (m == 0). Proof. by case: m. Qed. + +Lemma ffactnS n m : n ^_ m.+1 = n * n.-1 ^_ m. Proof. by []. Qed. + +Lemma ffactSS n m : n.+1 ^_ m.+1 = n.+1 * n ^_ m. Proof. by []. Qed. + +Lemma ffactn1 n : n ^_ 1 = n. Proof. exact: muln1. Qed. + +Lemma ffactnSr n m : n ^_ m.+1 = n ^_ m * (n - m). +Proof. +elim: n m => [|n IHn] [|m] //=; first by rewrite ffactn1 mul1n. +by rewrite !ffactSS IHn mulnA. +Qed. + +Lemma ffact_gt0 n m : (0 < n ^_ m) = (m <= n). +Proof. by elim: n m => [|n IHn] [|m] //=; rewrite ffactSS muln_gt0 IHn. Qed. + +Lemma ffact_small n m : n < m -> n ^_ m = 0. +Proof. by rewrite ltnNge -ffact_gt0; case: posnP. Qed. + +Lemma ffactnn n : n ^_ n = n`!. +Proof. by elim: n => [|n IHn] //; rewrite ffactnS IHn. Qed. + +Lemma ffact_fact n m : m <= n -> n ^_ m * (n - m)`! = n`!. +Proof. +by elim: n m => [|n IHn] [|m] //= le_m_n; rewrite ?mul1n // -mulnA IHn. +Qed. + +Lemma ffact_factd n m : m <= n -> n ^_ m = n`! %/ (n - m)`!. +Proof. by move/ffact_fact <-; rewrite mulnK ?fact_gt0. Qed. + +(** Binomial coefficients *) + +Fixpoint binomial_rec n m := + match n, m with + | n'.+1, m'.+1 => binomial_rec n' m + binomial_rec n' m' + | _, 0 => 1 + | 0, _.+1 => 0 + end. + +Definition binomial := nosimpl binomial_rec. + +Notation "''C' ( n , m )" := (binomial n m) + (at level 8, format "''C' ( n , m )") : nat_scope. + +Lemma binE : binomial = binomial_rec. Proof. by []. Qed. + +Lemma bin0 n : 'C(n, 0) = 1. Proof. by case: n. Qed. + +Lemma bin0n m : 'C(0, m) = (m == 0). Proof. by case: m. Qed. + +Lemma binS n m : 'C(n.+1, m.+1) = 'C(n, m.+1) + 'C(n, m). Proof. by []. Qed. + +Lemma bin1 n : 'C(n, 1) = n. +Proof. by elim: n => //= n IHn; rewrite binS bin0 IHn addn1. Qed. + +Lemma bin_gt0 m n : (0 < 'C(m, n)) = (n <= m). +Proof. +elim: m n => [|m IHm] [|n] //. +by rewrite binS addn_gt0 !IHm orbC ltn_neqAle andKb. +Qed. + +Lemma leq_bin2l m1 m2 n : m1 <= m2 -> 'C(m1, n) <= 'C(m2, n). +Proof. +elim: m1 m2 n => [m2 | m1 IHm [|m2] //] [|n] le_m12; rewrite ?bin0 //. +by rewrite !binS leq_add // IHm. +Qed. + +Lemma bin_small n m : n < m -> 'C(n, m) = 0. +Proof. by rewrite ltnNge -bin_gt0; case: posnP. Qed. + +Lemma binn n : 'C(n, n) = 1. +Proof. by elim: n => [|n IHn] //; rewrite binS bin_small. Qed. + +Lemma mul_Sm_binm m n : m.+1 * 'C(m, n) = n.+1 * 'C(m.+1, n.+1). +Proof. +elim: m n => [|m IHm] [|n] //; first by rewrite bin0 bin1 muln1 mul1n. +by rewrite mulSn {2}binS mulnDr addnCA !IHm -mulnDr. +Qed. + +Lemma bin_fact m n : n <= m -> 'C(m, n) * (n`! * (m - n)`!) = m`!. +Proof. +move/subnKC; move: (m - n) => m0 <-{m}. +elim: n => [|n IHn]; first by rewrite bin0 !mul1n. +by rewrite -mulnA mulnCA mulnA -mul_Sm_binm -mulnA IHn. +Qed. + +(* In fact the only exception is n = 0 and m = 1 *) +Lemma bin_factd n m : 0 < n -> 'C(n, m) = n`! %/ (m`! * (n - m)`!). +Proof. +move=> n_gt0; have [/bin_fact <-|lt_n_m] := leqP m n. + by rewrite mulnK // muln_gt0 !fact_gt0. +by rewrite bin_small // divnMA !divn_small ?fact_gt0 // fact_smonotone. +Qed. + +Lemma bin_ffact n m : 'C(n, m) * m`! = n ^_ m. +Proof. +apply/eqP; have [lt_n_m | le_m_n] := ltnP n m. + by rewrite bin_small ?ffact_small. +by rewrite -(eqn_pmul2r (fact_gt0 (n - m))) ffact_fact // -mulnA bin_fact. +Qed. + +Lemma bin_ffactd n m : 'C(n, m) = n ^_ m %/ m`!. +Proof. by rewrite -bin_ffact mulnK ?fact_gt0. Qed. + +Lemma bin_sub n m : m <= n -> 'C(n, n - m) = 'C(n, m). +Proof. +move=> le_m_n; apply/eqP; move/eqP: (bin_fact (leq_subr m n)). +by rewrite subKn // -(bin_fact le_m_n) !mulnA mulnAC !eqn_pmul2r // fact_gt0. +Qed. + +Lemma binSn n : 'C(n.+1, n) = n.+1. +Proof. by rewrite -bin_sub ?leqnSn // subSnn bin1. Qed. + +Lemma bin2 n : 'C(n, 2) = (n * n.-1)./2. +Proof. +by case: n => //= n; rewrite -{3}[n]bin1 mul_Sm_binm mul2n half_double. +Qed. + +Lemma bin2odd n : odd n -> 'C(n, 2) = n * n.-1./2. +Proof. by case: n => // n oddn; rewrite bin2 -!divn2 muln_divA ?dvdn2. Qed. + +Lemma prime_dvd_bin k p : prime p -> 0 < k < p -> p %| 'C(p, k). +Proof. +move=> p_pr /andP[k_gt0 lt_k_p]; have def_p := ltn_predK lt_k_p. +have: p %| p * 'C(p.-1, k.-1) by rewrite dvdn_mulr. +by rewrite -def_p mul_Sm_binm def_p prednK // Euclid_dvdM // gtnNdvd. +Qed. + +Lemma triangular_sum n : \sum_(0 <= i < n) i = 'C(n, 2). +Proof. +elim: n => [|n IHn]; first by rewrite big_geq. +by rewrite big_nat_recr // IHn binS bin1. +Qed. + +Lemma textbook_triangular_sum n : \sum_(0 <= i < n) i = 'C(n, 2). +Proof. +rewrite bin2; apply: canRL half_double _. +rewrite -addnn {1}big_nat_rev -big_split big_mkord /= ?add0n. +rewrite (eq_bigr (fun _ => n.-1)); first by rewrite sum_nat_const card_ord. +by case: n => [|n] [i le_i_n] //=; rewrite subSS subnK. +Qed. + +Theorem Pascal a b n : + (a + b) ^ n = \sum_(i < n.+1) 'C(n, i) * (a ^ (n - i) * b ^ i). +Proof. +elim: n => [|n IHn]; rewrite big_ord_recl muln1 ?big_ord0 //. +rewrite expnS {}IHn /= mulnDl !big_distrr /= big_ord_recl muln1 subn0. +rewrite !big_ord_recr /= !binn !subnn bin0 !subn0 !mul1n -!expnS -addnA. +congr (_ + _); rewrite addnA -big_split /=; congr (_ + _). +apply: eq_bigr => i _; rewrite mulnCA (mulnA a) -expnS subnSK //=. +by rewrite (mulnC b) -2!mulnA -expnSr -mulnDl. +Qed. +Definition expnDn := Pascal. + +Lemma Vandermonde k l i : + \sum_(j < i.+1) 'C(k, j) * 'C(l, i - j) = 'C(k + l , i). +Proof. +pose f k i := \sum_(j < i.+1) 'C(k, j) * 'C(l, i - j). +suffices{k i} fxx k i: f k.+1 i.+1 = f k i.+1 + f k i. + elim: k i => [i | k IHk [|i]]; last by rewrite -/(f _ _) fxx /f !IHk -binS. + by rewrite big_ord_recl big1_eq addn0 mul1n subn0. + by rewrite big_ord_recl big_ord0 addn0 !bin0 muln1. +rewrite {}/f big_ord_recl (big_ord_recl (i.+1)) !bin0 !mul1n. +rewrite -addnA -big_split /=; congr (_ + _). +by apply: eq_bigr => j _ ; rewrite -mulnDl. +Qed. + +Lemma subn_exp m n k : + m ^ k - n ^ k = (m - n) * (\sum_(i < k) m ^ (k.-1 -i) * n ^ i). +Proof. +case: k => [|k]; first by rewrite big_ord0. +rewrite mulnBl !big_distrr big_ord_recl big_ord_recr /= subn0 muln1. +rewrite subnn mul1n -!expnS subnDA; congr (_ - _); apply: canRL (addnK _) _. +congr (_ + _); apply: eq_bigr => i _. +by rewrite (mulnCA n) -expnS mulnA -expnS subnSK /=. +Qed. + +Lemma predn_exp m k : (m ^ k).-1 = m.-1 * (\sum_(i < k) m ^ i). +Proof. +rewrite -!subn1 -{1}(exp1n k) subn_exp; congr (_ * _). +symmetry; rewrite (reindex_inj rev_ord_inj); apply: eq_bigr => i _ /=. +by rewrite -subn1 -subnDA exp1n muln1. +Qed. + +Lemma dvdn_pred_predX n e : (n.-1 %| (n ^ e).-1)%N. +Proof. by rewrite predn_exp dvdn_mulr. Qed. + +Lemma modn_summ I r (P : pred I) F d : + \sum_(i <- r | P i) F i %% d = \sum_(i <- r | P i) F i %[mod d]. +Proof. +by apply/eqP; elim/big_rec2: _ => // i m n _; rewrite modnDml eqn_modDl. +Qed. + +(* Combinatorial characterizations. *) + +Section Combinations. + +Implicit Types T D : finType. + +Lemma card_uniq_tuples T n (A : pred T) : + #|[set t : n.-tuple T | all A t & uniq t]| = #|A| ^_ n. +Proof. +elim: n A => [|n IHn] A. + by rewrite (@eq_card1 _ [tuple]) // => t; rewrite [t]tuple0 inE. +rewrite -sum1dep_card (partition_big (@thead _ _) A) /= => [|t]; last first. + by case/tupleP: t => x t; do 2!case/andP. +transitivity (#|A| * #|A|.-1 ^_ n)%N; last by case: #|A|. +rewrite -sum_nat_const; apply: eq_bigr => x Ax. +rewrite (cardD1 x) [x \in A]Ax /= -(IHn [predD1 A & x]) -sum1dep_card. +rewrite (reindex (fun t : n.-tuple T => [tuple of x :: t])) /=; last first. + pose ttail (t : n.+1.-tuple T) := [tuple of behead t]. + exists ttail => [t _ | t /andP[_ /eqP <-]]; first exact: val_inj. + by rewrite -tuple_eta. +apply: eq_bigl=> t; rewrite Ax theadE eqxx andbT /= andbA; congr (_ && _). +by rewrite all_predI all_predC has_pred1 andbC. +Qed. + +Lemma card_inj_ffuns_on D T (R : pred T) : + #|[set f : {ffun D -> T} in ffun_on R | injectiveb f]| = #|R| ^_ #|D|. +Proof. +rewrite -card_uniq_tuples. +have bijFF: {on (_ : pred _), bijective (@Finfun D T)}. + by exists val => // x _; exact: val_inj. +rewrite -(on_card_preimset (bijFF _)); apply: eq_card => t. +rewrite !inE -(codom_ffun (Finfun t)); congr (_ && _); apply: negb_inj. +by rewrite -has_predC has_map enumT has_filter -size_eq0 -cardE. +Qed. + +Lemma card_inj_ffuns D T : + #|[set f : {ffun D -> T} | injectiveb f]| = #|T| ^_ #|D|. +Proof. +rewrite -card_inj_ffuns_on; apply: eq_card => f. +by rewrite 2!inE; case: ffun_onP => // []. +Qed. + +Lemma card_draws T k : #|[set A : {set T} | #|A| == k]| = 'C(#|T|, k). +Proof. +have [ltTk | lekT] := ltnP #|T| k. + rewrite bin_small // eq_card0 // => A. + by rewrite inE eqn_leq andbC leqNgt (leq_ltn_trans (max_card _)). +apply/eqP; rewrite -(eqn_pmul2r (fact_gt0 k)) bin_ffact // eq_sym. +rewrite -sum_nat_dep_const -{1 3}(card_ord k) -card_inj_ffuns -sum1dep_card. +pose imIk (f : {ffun 'I_k -> T}) := f @: 'I_k. +rewrite (partition_big imIk (fun A => #|A| == k)) /= => [|f]; last first. + by move/injectiveP=> inj_f; rewrite card_imset ?card_ord. +apply/eqP; apply: eq_bigr => A /eqP cardAk. +have [f0 inj_f0 im_f0]: exists2 f, injective f & f @: 'I_k = A. + rewrite -cardAk; exists enum_val; first exact: enum_val_inj. + apply/setP=> a; apply/imsetP/idP=> [[i _ ->] | Aa]; first exact: enum_valP. + by exists (enum_rank_in Aa a); rewrite ?enum_rankK_in. +rewrite (reindex (fun p : {ffun _} => [ffun i => f0 (p i)])) /=; last first. + pose ff0' f i := odflt i [pick j | f i == f0 j]. + exists (fun f => [ffun i => ff0' f i]) => [p _ | f]. + apply/ffunP=> i; rewrite ffunE /ff0'; case: pickP => [j | /(_ (p i))]. + by rewrite ffunE (inj_eq inj_f0) => /eqP. + by rewrite ffunE eqxx. + rewrite -im_f0 => /andP[/injectiveP injf /eqP im_f]. + apply/ffunP=> i; rewrite !ffunE /ff0'; case: pickP => [y /eqP //|]. + have /imsetP[j _ eq_f0j_fi]: f i \in f0 @: 'I_k by rewrite -im_f mem_imset. + by move/(_ j)=> /eqP[]. +rewrite -ffactnn -card_inj_ffuns -sum1dep_card; apply: eq_bigl => p. +apply/andP/injectiveP=> [[/injectiveP inj_f0p _] i j eq_pij | inj_p]. + by apply: inj_f0p; rewrite !ffunE eq_pij. +set f := finfun _. +have injf: injective f by move=> i j; rewrite !ffunE => /inj_f0; exact: inj_p. +split; first exact/injectiveP. +rewrite eqEcard card_imset // cardAk card_ord leqnn andbT -im_f0. +by apply/subsetP=> x /imsetP[i _ ->]; rewrite ffunE mem_imset. +Qed. + +Lemma card_ltn_sorted_tuples m n : + #|[set t : m.-tuple 'I_n | sorted ltn (map val t)]| = 'C(n, m). +Proof. +have [-> | n_gt0] := posnP n; last pose i0 := Ordinal n_gt0. + case: m => [|m]; last by apply: eq_card0; case/tupleP=> [[]]. + by apply: (@eq_card1 _ [tuple]) => t; rewrite [t]tuple0 inE. +rewrite -{12}[n]card_ord -card_draws. +pose f_t (t : m.-tuple 'I_n) := [set i in t]. +pose f_A (A : {set 'I_n}) := [tuple of mkseq (nth i0 (enum A)) m]. +have val_fA (A : {set 'I_n}) : #|A| = m -> val (f_A A) = enum A. + by move=> Am; rewrite -[enum _](mkseq_nth i0) -cardE Am. +have inc_A (A : {set 'I_n}) : sorted ltn (map val (enum A)). + rewrite -[enum _](eq_filter (mem_enum _)). + rewrite -(eq_filter (mem_map val_inj _)) -filter_map. + by rewrite (sorted_filter ltn_trans) // unlock val_ord_enum iota_ltn_sorted. +rewrite -!sum1dep_card (reindex_onto f_t f_A) /= => [|A]; last first. + by move/eqP=> cardAm; apply/setP=> x; rewrite inE -(mem_enum (mem A)) -val_fA. +apply: eq_bigl => t; apply/idP/idP=> [inc_t|]; last first. + by case/andP; move/eqP=> t_m; move/eqP=> <-; rewrite val_fA. +have ft_m: #|f_t t| = m. + rewrite cardsE (card_uniqP _) ?size_tuple // -(map_inj_uniq val_inj). + exact: (sorted_uniq ltn_trans ltnn). +rewrite ft_m eqxx -val_eqE val_fA // -(inj_eq (inj_map val_inj)) /=. +apply/eqP; apply: (eq_sorted_irr ltn_trans ltnn) => // y. +by apply/mapP/mapP=> [] [x t_x ->]; exists x; rewrite // mem_enum inE in t_x *. +Qed. + +Lemma card_sorted_tuples m n : + #|[set t : m.-tuple 'I_n.+1 | sorted leq (map val t)]| = 'C(m + n, m). +Proof. +set In1 := 'I_n.+1; pose x0 : In1 := ord0. +have add_mnP (i : 'I_m) (x : In1) : i + x < m + n. + by rewrite -ltnS -addSn -!addnS leq_add. +pose add_mn t i := Ordinal (add_mnP i (tnth t i)). +pose add_mn_nat (t : m.-tuple In1) i := i + nth x0 t i. +have add_mnC t: val \o add_mn t =1 add_mn_nat t \o val. + by move=> i; rewrite /= (tnth_nth x0). +pose f_add t := [tuple of map (add_mn t) (ord_tuple m)]. +rewrite -card_ltn_sorted_tuples -!sum1dep_card (reindex f_add) /=. + apply: eq_bigl => t; rewrite -map_comp (eq_map (add_mnC t)) map_comp. + rewrite enumT unlock val_ord_enum -{1}(drop0 t). + have [m0 | m_gt0] := posnP m. + by rewrite {2}m0 /= drop_oversize // size_tuple m0. + have def_m := subnK m_gt0; rewrite -{2}def_m addn1 /= {1}/add_mn_nat. + move: 0 (m - 1) def_m => i k; rewrite -{1}(size_tuple t) => def_m. + rewrite (drop_nth x0) /=; last by rewrite -def_m leq_addl. + elim: k i (nth x0 t i) def_m => [|k IHk] i x /=. + by rewrite add0n => ->; rewrite drop_size. + rewrite addSnnS => def_m; rewrite -addSn leq_add2l -IHk //. + by rewrite (drop_nth x0) // -def_m leq_addl. +pose sub_mn (t : m.-tuple 'I_(m + n)) i : In1 := inord (tnth t i - i). +exists (fun t => [tuple of map (sub_mn t) (ord_tuple m)]) => [t _ | t]. + apply: eq_from_tnth => i; apply: val_inj. + by rewrite /sub_mn !(tnth_ord_tuple, tnth_map) addKn inord_val. +rewrite inE /= => inc_t; apply: eq_from_tnth => i; apply: val_inj. +rewrite tnth_map tnth_ord_tuple /= tnth_map tnth_ord_tuple. +suffices [le_i_ti le_ti_ni]: i <= tnth t i /\ tnth t i <= i + n. + by rewrite /sub_mn inordK ?subnKC // ltnS leq_subLR. +pose y0 := tnth t i; rewrite (tnth_nth y0) -(nth_map _ (val i)) ?size_tuple //. +case def_e: (map _ _) => [|x e] /=; first by rewrite nth_nil ?leq_addr. +rewrite def_e in inc_t; split. + case: {-2}i; rewrite /= -{1}(size_tuple t) -(size_map val) def_e. + elim=> //= j IHj lt_j_t; apply: leq_trans (pathP (val i) inc_t _ lt_j_t). + by rewrite ltnS IHj 1?ltnW. +move: (_ - _) (subnK (valP i)) => k /=. +elim: k {-2}(val i) => /= [|k IHk] j def_m; rewrite -ltnS -addSn. + by rewrite [j.+1]def_m -def_e (nth_map y0) ?ltn_ord // size_tuple -def_m. +rewrite (leq_trans _ (IHk _ _)) -1?addSnnS //; apply: (pathP _ inc_t). +rewrite -ltnS (leq_trans (leq_addl k _)) // -addSnnS def_m. +by rewrite -(size_tuple t) -(size_map val) def_e. +Qed. + +Lemma card_partial_ord_partitions m n : + #|[set t : m.-tuple 'I_n.+1 | \sum_(i <- t) i <= n]| = 'C(m + n, m). +Proof. +symmetry; set In1 := 'I_n.+1; pose x0 : In1 := ord0. +pose add_mn (i j : In1) : In1 := inord (i + j). +pose f_add (t : m.-tuple In1) := [tuple of scanl add_mn x0 t]. +rewrite -card_sorted_tuples -!sum1dep_card (reindex f_add) /=. + apply: eq_bigl => t; rewrite -[\sum_(i <- t) i]add0n. + transitivity (path leq x0 (map val (f_add t))) => /=; first by case: map. + rewrite -{1 2}[0]/(val x0); elim: {t}(val t) (x0) => /= [|x t IHt] s. + by rewrite big_nil addn0 -ltnS ltn_ord. + rewrite big_cons addnA IHt /= val_insubd ltnS. + have [_ | ltn_n_sx] := leqP (s + x) n; first by rewrite leq_addr. + rewrite -(leq_add2r x) leqNgt (leq_trans (valP x)) //=. + by rewrite leqNgt (leq_trans ltn_n_sx) ?leq_addr. +pose sub_mn (i j : In1) := Ordinal (leq_ltn_trans (leq_subr i j) (valP j)). +exists (fun t : m.-tuple In1 => [tuple of pairmap sub_mn x0 t]) => /= t inc_t. + apply: val_inj => /=; have{inc_t}: path leq x0 (map val (f_add t)). + by move: inc_t; rewrite inE /=; case: map. + rewrite [map _ _]/=; elim: {t}(val t) (x0) => //= x t IHt s. + case/andP=> le_s_sx /IHt->; congr (_ :: _); apply: val_inj => /=. + move: le_s_sx; rewrite val_insubd. + case le_sx_n: (_ < n.+1); first by rewrite addKn. + by case: (val s) le_sx_n; rewrite ?ltn_ord. +apply: val_inj => /=; have{inc_t}: path leq x0 (map val t). + by move: inc_t; rewrite inE /=; case: map. +elim: {t}(val t) (x0) => //= x t IHt s /andP[le_s_sx inc_t]. +suffices ->: add_mn s (sub_mn s x) = x by rewrite IHt. +by apply: val_inj; rewrite /add_mn /= subnKC ?inord_val. +Qed. + +Lemma card_ord_partitions m n : + #|[set t : m.+1.-tuple 'I_n.+1 | \sum_(i <- t) i == n]| = 'C(m + n, m). +Proof. +symmetry; set In1 := 'I_n.+1; pose x0 : In1 := ord0. +pose f_add (t : m.-tuple In1) := [tuple of sub_ord (\sum_(x <- t) x) :: t]. +rewrite -card_partial_ord_partitions -!sum1dep_card (reindex f_add) /=. + by apply: eq_bigl => t; rewrite big_cons /= addnC (sameP maxn_idPr eqP) maxnE. +exists (fun t : m.+1.-tuple In1 => [tuple of behead t]) => [t _|]. + exact: val_inj. +case/tupleP=> x t; rewrite inE /= big_cons => /eqP def_n. +by apply: val_inj; congr (_ :: _); apply: val_inj; rewrite /= -{1}def_n addnK. +Qed. + +End Combinations. + diff --git a/mathcomp/discrete/choice.v b/mathcomp/discrete/choice.v new file mode 100644 index 0000000..96b549c --- /dev/null +++ b/mathcomp/discrete/choice.v @@ -0,0 +1,681 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. + +(******************************************************************************) +(* This file contains the definitions of: *) +(* choiceType == interface for types with a choice operator *) +(* countType == interface for countable types *) +(* subCountType == interface for types that are both subType and countType. *) +(* xchoose exP == a standard x such that P x, given exP : exists x : T, P x *) +(* when T is a choiceType. The choice depends only on the *) +(* extent of P (in particular, it is independent of exP). *) +(* choose P x0 == if P x0, a standard x such that P x. *) +(* pickle x == a nat encoding the value x : T, where T is a countType. *) +(* unpickle n == a partial inverse to pickle: unpickle (pickle x) = Some x *) +(* pickle_inv n == a sharp partial inverse to pickle pickle_inv n = Some x *) +(* if and only if pickle x = n. *) +(* [choiceType of T for cT] == clone for T of the choiceType cT. *) +(* [choiceType of T] == clone for T of the choiceType inferred for T. *) +(* [countType of T for cT] == clone for T of the countType cT. *) +(* [count Type of T] == clone for T of the countType inferred for T. *) +(* [choiceMixin of T by <:] == Choice mixin for T when T has a subType p *) +(* structure with p : pred cT and cT has a Choice *) +(* structure; the corresponding structure is Canonical.*) +(* [countMixin of T by <:] == Count mixin for a subType T of a countType. *) +(* PcanChoiceMixin fK == Choice mixin for T, given f : T -> cT where cT has *) +(* a Choice structure, a left inverse partial function *) +(* g and fK : pcancel f g. *) +(* CanChoiceMixin fK == Choice mixin for T, given f : T -> cT, g and *) +(* fK : cancel f g. *) +(* PcanCountMixin fK == Count mixin for T, given f : T -> cT where cT has *) +(* a Countable structure, a left inverse partial *) +(* function g and fK : pcancel f g. *) +(* CanCountMixin fK == Count mixin for T, given f : T -> cT, g and *) +(* fK : cancel f g. *) +(* GenTree.tree T == generic n-ary tree type with nat-labeled nodes and *) +(* T-labeled nodes. It is equipped with canonical *) +(* eqType, choiceType, and countType instances, and so *) +(* can be used to similarly equip simple datatypes *) +(* by using the mixins above. *) +(* In addition to the lemmas relevant to these definitions, this file also *) +(* contains definitions of a Canonical choiceType and countType instances for *) +(* all basic datatypes (e.g., nat, bool, subTypes, pairs, sums, etc.). *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(* Technical definitions about coding and decoding of nat sequances, which *) +(* are used below to define various Canonical instances of the choice and *) +(* countable interfaces. *) + +Module CodeSeq. + +(* Goedel-style one-to-one encoding of seq nat into nat. *) +(* The code for [:: n1; ...; nk] has binary representation *) +(* 1 0 ... 0 1 ... 1 0 ... 0 1 0 ... 0 *) +(* <-----> <-----> <-----> *) +(* nk 0s n2 0s n1 0s *) + +Definition code := foldr (fun n m => 2 ^ n * m.*2.+1) 0. + +Fixpoint decode_rec (v q r : nat) {struct q} := + match q, r with + | 0, _ => [:: v] + | q'.+1, 0 => v :: [rec 0, q', q'] + | q'.+1, 1 => [rec v.+1, q', q'] + | q'.+1, r'.+2 => [rec v, q', r'] + end where "[ 'rec' v , q , r ]" := (decode_rec v q r). + +Definition decode n := if n is 0 then [::] else [rec 0, n.-1, n.-1]. + +Lemma decodeK : cancel decode code. +Proof. +have m2s: forall n, n.*2 - n = n by move=> n; rewrite -addnn addnK. +case=> //= n; rewrite -[n.+1]mul1n -(expn0 2) -{3}[n]m2s. +elim: n {2 4}n {1 3}0 => [|q IHq] [|[|r]] v //=; rewrite {}IHq ?mul1n ?m2s //. +by rewrite expnSr -mulnA mul2n. +Qed. + +Lemma codeK : cancel code decode. +Proof. +elim=> //= v s IHs; rewrite -[_ * _]prednK ?muln_gt0 ?expn_gt0 //=. +rewrite -{3}[v]addn0; elim: v {1 4}0 => [|v IHv {IHs}] q. + rewrite mul1n /= -{1}addnn -{4}IHs; move: (_ s) {IHs} => n. + by elim: {1 3}n => //=; case: n. +rewrite expnS -mulnA mul2n -{1}addnn -[_ * _]prednK ?muln_gt0 ?expn_gt0 //. +by rewrite doubleS addSn /= addSnnS; elim: {-2}_.-1 => //=. +Qed. + +Lemma ltn_code s : all (fun j => j < code s) s. +Proof. +elim: s => //= i s IHs; rewrite -[_.+1]muln1 leq_mul 1?ltn_expl //=. +apply: sub_all IHs => j /leqW lejs; rewrite -[j.+1]mul1n leq_mul ?expn_gt0 //. +by rewrite ltnS -[j]mul1n -mul2n leq_mul. +Qed. + +Lemma gtn_decode n : all (ltn^~ n) (decode n). +Proof. by rewrite -{1}[n]decodeK ltn_code. Qed. + +End CodeSeq. + +Section OtherEncodings. +(* Miscellaneous encodings: option T -c-> seq T, T1 * T2 -c-> {i : T1 & T2} *) +(* T1 + T2 -c-> option T1 * option T2, unit -c-> bool; bool -c-> nat is *) +(* already covered in ssrnat by the nat_of_bool coercion, the odd predicate, *) +(* and their "cancellation" lemma oddb. We use these encodings to propagate *) +(* canonical structures through these type constructors so that ultimately *) +(* all Choice and Countable instanced derive from nat and the seq and sigT *) +(* constructors. *) + +Variables T T1 T2 : Type. + +Definition seq_of_opt := @oapp T _ (nseq 1) [::]. +Lemma seq_of_optK : cancel seq_of_opt ohead. Proof. by case. Qed. + +Definition tag_of_pair (p : T1 * T2) := @Tagged T1 p.1 (fun _ => T2) p.2. +Definition pair_of_tag (u : {i : T1 & T2}) := (tag u, tagged u). +Lemma tag_of_pairK : cancel tag_of_pair pair_of_tag. Proof. by case. Qed. +Lemma pair_of_tagK : cancel pair_of_tag tag_of_pair. Proof. by case. Qed. + +Definition opair_of_sum (s : T1 + T2) := + match s with inl x => (Some x, None) | inr y => (None, Some y) end. +Definition sum_of_opair p := + oapp (some \o @inr T1 T2) (omap (@inl _ T2) p.1) p.2. +Lemma opair_of_sumK : pcancel opair_of_sum sum_of_opair. Proof. by case. Qed. + +Lemma bool_of_unitK : cancel (fun _ => true) (fun _ => tt). +Proof. by case. Qed. + +End OtherEncodings. + +(* Generic variable-arity tree type, providing an encoding target for *) +(* miscellaneous user datatypes. The GenTree.tree type can be combined with *) +(* a sigT type to model multi-sorted concrete datatypes. *) +Module GenTree. + +Section Def. + +Variable T : Type. + +Unset Elimination Schemes. +Inductive tree := Leaf of T | Node of nat & seq tree. + +Definition tree_rect K IH_leaf IH_node := + fix loop t : K t := match t with + | Leaf x => IH_leaf x + | Node n f0 => + let fix iter_pair f : foldr (fun t => prod (K t)) unit f := + if f is t :: f' then (loop t, iter_pair f') else tt in + IH_node n f0 (iter_pair f0) + end. +Definition tree_rec (K : tree -> Set) := @tree_rect K. +Definition tree_ind K IH_leaf IH_node := + fix loop t : K t : Prop := match t with + | Leaf x => IH_leaf x + | Node n f0 => + let fix iter_conj f : foldr (fun t => and (K t)) True f := + if f is t :: f' then conj (loop t) (iter_conj f') else Logic.I + in IH_node n f0 (iter_conj f0) + end. + +Fixpoint encode t : seq (nat + T) := + match t with + | Leaf x => [:: inr _ x] + | Node n f => inl _ n.+1 :: rcons (flatten (map encode f)) (inl _ 0) + end. + +Definition decode_step c fs := + match c with + | inr x => (Leaf x :: fs.1, fs.2) + | inl 0 => ([::], fs.1 :: fs.2) + | inl n.+1 => (Node n fs.1 :: head [::] fs.2, behead fs.2) + end. + +Definition decode c := ohead (foldr decode_step ([::], [::]) c).1. + +Lemma codeK : pcancel encode decode. +Proof. +move=> t; rewrite /decode; set fs := (_, _). +suffices ->: foldr decode_step fs (encode t) = (t :: fs.1, fs.2) by []. +elim: t => //= n f IHt in (fs) *; elim: f IHt => //= t f IHf []. +by rewrite rcons_cat foldr_cat => -> /= /IHf[-> -> ->]. +Qed. + +End Def. + +End GenTree. +Implicit Arguments GenTree.codeK []. + +Definition tree_eqMixin (T : eqType) := PcanEqMixin (GenTree.codeK T). +Canonical tree_eqType (T : eqType) := EqType (GenTree.tree T) (tree_eqMixin T). + +(* Structures for Types with a choice function, and for Types with countably *) +(* many elements. The two concepts are closely linked: we indeed make *) +(* Countable a subclass of Choice, as countable choice is valid in CiC. This *) +(* apparent redundancy is needed to ensure the consistency of the Canonical *) +(* inference, as the canonical Choice for a given type may differ from the *) +(* countable choice for its canonical Countable structure, e.g., for options. *) +(* The Choice interface exposes two choice functions; for T : choiceType *) +(* and P : pred T, we provide: *) +(* xchoose : (exists x, P x) -> T *) +(* choose : pred T -> T -> T *) +(* While P (xchoose exP) will always hold, P (choose P x0) will be true if *) +(* and only if P x0 holds. Both xchoose and choose are extensional in P and *) +(* do not depend on the witness exP or x0 (provided P x0 holds). Note that *) +(* xchoose is slightly more powerful, but less convenient to use. *) +(* However, neither choose nor xchoose are composable: it would not be *) +(* be possible to extend the Choice structure to arbitrary pairs using only *) +(* these functions, for instance. Internally, the interfaces provides a *) +(* subtly stronger operation, Choice.InternalTheory.find, which performs a *) +(* limited search using an integer parameter only rather than a full value as *) +(* [x]choose does. This is not a restriction in the constructive setting *) +(* (where all types are concrete and hence countable). In the case of *) +(* axiomatizations, like for the Coq reals library, postulating a suitable *) +(* axiom of choice suppresses the need for guidance. Nevertheless this *) +(* operation is just what is needed to make the Choice interface compose. *) +(* The Countable interface provides three functions; for T : countType we *) +(* geth pickle : T -> nat, and unpickle, pickle_inv : nat -> option T. *) +(* The functions provide an effective embedding of T in nat: unpickle is a *) +(* left inverse to pickle, which satisfies pcancel pickle unpickle, i.e., *) +(* unpickle \o pickle =1 some; pickle_inv is a more precise inverse for which *) +(* we also have ocancel pickle_inv pickle. Both unpickle and pickle need to *) +(* partial functions, to allow for possibly empty types such as {x | P x}. *) +(* The names of these functions underline the correspondence with the *) +(* notion of "Serializable" types in programming languages. *) +(* Finally, we need to provide a join class to let type inference unify *) +(* subType and countType class constraints, e.g., for a countable subType of *) +(* an uncountable choiceType (the issue does not arise earlier with eqType or *) +(* choiceType because in practice the base type of an Equality/Choice subType *) +(* is always an Equality/Choice Type). *) + +Module Choice. + +Section ClassDef. + +Record mixin_of T := Mixin { + find : pred T -> nat -> option T; + _ : forall P n x, find P n = Some x -> P x; + _ : forall P : pred T, (exists x, P x) -> exists n, find P n; + _ : forall P Q : pred T, P =1 Q -> find P =1 find Q +}. + +Record class_of T := Class {base : Equality.class_of T; mixin : mixin_of T}. +Local Coercion base : class_of >-> Equality.class_of. + +Structure type := Pack {sort; _ : class_of sort; _ : Type}. +Local Coercion sort : type >-> Sortclass. +Variables (T : Type) (cT : type). +Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack T c T. +Let xT := let: Pack T _ _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack m := + fun b bT & phant_id (Equality.class bT) b => Pack (@Class T b m) T. + +(* Inheritance *) +Definition eqType := @Equality.Pack cT xclass xT. + +End ClassDef. + +Module Import Exports. +Coercion base : class_of >-> Equality.class_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Canonical eqType. +Notation choiceType := type. +Notation choiceMixin := mixin_of. +Notation ChoiceType T m := (@pack T m _ _ id). +Notation "[ 'choiceType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) + (at level 0, format "[ 'choiceType' 'of' T 'for' cT ]") : form_scope. +Notation "[ 'choiceType' 'of' T ]" := (@clone T _ _ id) + (at level 0, format "[ 'choiceType' 'of' T ]") : form_scope. + +End Exports. + +Module InternalTheory. +Section InternalTheory. +(* Inner choice function. *) +Definition find T := find (mixin (class T)). + +Variable T : choiceType. +Implicit Types P Q : pred T. + +Lemma correct P n x : find P n = Some x -> P x. +Proof. by case: T => _ [_ []] //= in P n x *. Qed. + +Lemma complete P : (exists x, P x) -> (exists n, find P n). +Proof. by case: T => _ [_ []] //= in P *. Qed. + +Lemma extensional P Q : P =1 Q -> find P =1 find Q. +Proof. by case: T => _ [_ []] //= in P Q *. Qed. + +Fact xchoose_subproof P exP : {x | find P (ex_minn (@complete P exP)) = Some x}. +Proof. +by case: (ex_minnP (complete exP)) => n; case: (find P n) => // x; exists x. +Qed. + +End InternalTheory. +End InternalTheory. + +End Choice. +Export Choice.Exports. + +Section ChoiceTheory. + +Implicit Type T : choiceType. +Import Choice.InternalTheory CodeSeq. +Local Notation dc := decode. + +Section OneType. + +Variable T : choiceType. +Implicit Types P Q : pred T. + +Definition xchoose P exP := sval (@xchoose_subproof T P exP). + +Lemma xchooseP P exP : P (@xchoose P exP). +Proof. by rewrite /xchoose; case: (xchoose_subproof exP) => x /= /correct. Qed. + +Lemma eq_xchoose P Q exP exQ : P =1 Q -> @xchoose P exP = @xchoose Q exQ. +Proof. +rewrite /xchoose => eqPQ. +case: (xchoose_subproof exP) => x; case: (xchoose_subproof exQ) => y /=. +case: ex_minnP => n; case: ex_minnP => m. +rewrite -(extensional eqPQ) {1}(extensional eqPQ). +move=> Qm minPm Pn minQn; suffices /eqP->: m == n by move=> -> []. +by rewrite eqn_leq minQn ?minPm. +Qed. + +Lemma sigW P : (exists x, P x) -> {x | P x}. +Proof. by move=> exP; exists (xchoose exP); exact: xchooseP. Qed. + +Lemma sig2W P Q : (exists2 x, P x & Q x) -> {x | P x & Q x}. +Proof. +move=> exPQ; have [|x /andP[]] := @sigW (predI P Q); last by exists x. +by have [x Px Qx] := exPQ; exists x; exact/andP. +Qed. + +Lemma sig_eqW (vT : eqType) (lhs rhs : T -> vT) : + (exists x, lhs x = rhs x) -> {x | lhs x = rhs x}. +Proof. +move=> exP; suffices [x /eqP Ex]: {x | lhs x == rhs x} by exists x. +by apply: sigW; have [x /eqP Ex] := exP; exists x. +Qed. + +Lemma sig2_eqW (vT : eqType) (P : pred T) (lhs rhs : T -> vT) : + (exists2 x, P x & lhs x = rhs x) -> {x | P x & lhs x = rhs x}. +Proof. +move=> exP; suffices [x Px /eqP Ex]: {x | P x & lhs x == rhs x} by exists x. +by apply: sig2W; have [x Px /eqP Ex] := exP; exists x. +Qed. + +Definition choose P x0 := + if insub x0 : {? x | P x} is Some (exist x Px) then + xchoose (ex_intro [eta P] x Px) + else x0. + +Lemma chooseP P x0 : P x0 -> P (choose P x0). +Proof. by move=> Px0; rewrite /choose insubT xchooseP. Qed. + +Lemma choose_id P x0 y0 : P x0 -> P y0 -> choose P x0 = choose P y0. +Proof. by move=> Px0 Py0; rewrite /choose !insubT /=; exact: eq_xchoose. Qed. + +Lemma eq_choose P Q : P =1 Q -> choose P =1 choose Q. +Proof. +rewrite /choose => eqPQ x0. +do [case: insubP; rewrite eqPQ] => [[x Px] Qx0 _| ?]; last by rewrite insubN. +by rewrite insubT; exact: eq_xchoose. +Qed. + +Section CanChoice. + +Variables (sT : Type) (f : sT -> T). + +Lemma PcanChoiceMixin f' : pcancel f f' -> choiceMixin sT. +Proof. +move=> fK; pose liftP sP := [pred x | oapp sP false (f' x)]. +pose sf sP := [fun n => obind f' (find (liftP sP) n)]. +exists sf => [sP n x | sP [y sPy] | sP sQ eqPQ n] /=. +- by case Df: (find _ n) => //= [?] Dx; have:= correct Df; rewrite /= Dx. +- have [|n Pn] := @complete T (liftP sP); first by exists (f y); rewrite /= fK. + exists n; case Df: (find _ n) Pn => //= [x] _. + by have:= correct Df => /=; case: (f' x). +by congr (obind _ _); apply: extensional => x /=; case: (f' x) => /=. +Qed. + +Definition CanChoiceMixin f' (fK : cancel f f') := + PcanChoiceMixin (can_pcan fK). + +End CanChoice. + +Section SubChoice. + +Variables (P : pred T) (sT : subType P). + +Definition sub_choiceMixin := PcanChoiceMixin (@valK T P sT). +Definition sub_choiceClass := @Choice.Class sT (sub_eqMixin sT) sub_choiceMixin. +Canonical sub_choiceType := Choice.Pack sub_choiceClass sT. + +End SubChoice. + +Fact seq_choiceMixin : choiceMixin (seq T). +Proof. +pose r f := [fun xs => fun x : T => f (x :: xs) : option (seq T)]. +pose fix f sP ns xs {struct ns} := + if ns is n :: ns1 then let fr := r (f sP ns1) xs in obind fr (find fr n) + else if sP xs then Some xs else None. +exists (fun sP nn => f sP (dc nn) nil) => [sP n ys | sP [ys] | sP sQ eqPQ n]. +- elim: {n}(dc n) nil => [|n ns IHs] xs /=; first by case: ifP => // sPxs [<-]. + by case: (find _ n) => //= [x]; apply: IHs. +- rewrite -(cats0 ys); elim/last_ind: ys nil => [|ys y IHs] xs /=. + by move=> sPxs; exists 0; rewrite /= sPxs. + rewrite cat_rcons => /IHs[n1 sPn1] {IHs}. + have /complete[n]: exists z, f sP (dc n1) (z :: xs) by exists y. + case Df: (find _ n)=> // [x] _; exists (code (n :: dc n1)). + by rewrite codeK /= Df /= (correct Df). +elim: {n}(dc n) nil => [|n ns IHs] xs /=; first by rewrite eqPQ. +rewrite (@extensional _ _ (r (f sQ ns) xs)) => [|x]; last by rewrite IHs. +by case: find => /=. +Qed. +Canonical seq_choiceType := Eval hnf in ChoiceType (seq T) seq_choiceMixin. + +End OneType. + +Section TagChoice. + +Variables (I : choiceType) (T_ : I -> choiceType). + +Fact tagged_choiceMixin : choiceMixin {i : I & T_ i}. +Proof. +pose mkT i (x : T_ i) := Tagged T_ x. +pose ft tP n i := omap (mkT i) (find (tP \o mkT i) n). +pose fi tP ni nt := obind (ft tP nt) (find (ft tP nt) ni). +pose f tP n := if dc n is [:: ni; nt] then fi tP ni nt else None. +exists f => [tP n u | tP [[i x] tPxi] | sP sQ eqPQ n]. +- rewrite /f /fi; case: (dc n) => [|ni [|nt []]] //=. + case: (find _ _) => //= [i]; rewrite /ft. + by case Df: (find _ _) => //= [x] [<-]; have:= correct Df. +- have /complete[nt tPnt]: exists y, (tP \o mkT i) y by exists x. + have{tPnt}: exists j, ft tP nt j by exists i; rewrite /ft; case: find tPnt. + case/complete=> ni tPn; exists (code [:: ni; nt]); rewrite /f codeK /fi. + by case Df: find tPn => //= [j] _; have:= correct Df. +rewrite /f /fi; case: (dc n) => [|ni [|nt []]] //=. +rewrite (@extensional _ _ (ft sQ nt)) => [|i]. + by case: find => //= i; congr (omap _ _); apply: extensional => x /=. +by congr (omap _ _); apply: extensional => x /=. +Qed. +Canonical tagged_choiceType := + Eval hnf in ChoiceType {i : I & T_ i} tagged_choiceMixin. + +End TagChoice. + +Fact nat_choiceMixin : choiceMixin nat. +Proof. +pose f := [fun (P : pred nat) n => if P n then Some n else None]. +exists f => [P n m | P [n Pn] | P Q eqPQ n] /=; last by rewrite eqPQ. + by case: ifP => // Pn [<-]. +by exists n; rewrite Pn. +Qed. +Canonical nat_choiceType := Eval hnf in ChoiceType nat nat_choiceMixin. + +Definition bool_choiceMixin := CanChoiceMixin oddb. +Canonical bool_choiceType := Eval hnf in ChoiceType bool bool_choiceMixin. +Canonical bitseq_choiceType := Eval hnf in [choiceType of bitseq]. + +Definition unit_choiceMixin := CanChoiceMixin bool_of_unitK. +Canonical unit_choiceType := Eval hnf in ChoiceType unit unit_choiceMixin. + +Definition option_choiceMixin T := CanChoiceMixin (@seq_of_optK T). +Canonical option_choiceType T := + Eval hnf in ChoiceType (option T) (option_choiceMixin T). + +Definition sig_choiceMixin T (P : pred T) : choiceMixin {x | P x} := + sub_choiceMixin _. +Canonical sig_choiceType T (P : pred T) := + Eval hnf in ChoiceType {x | P x} (sig_choiceMixin P). + +Definition prod_choiceMixin T1 T2 := CanChoiceMixin (@tag_of_pairK T1 T2). +Canonical prod_choiceType T1 T2 := + Eval hnf in ChoiceType (T1 * T2) (prod_choiceMixin T1 T2). + +Definition sum_choiceMixin T1 T2 := PcanChoiceMixin (@opair_of_sumK T1 T2). +Canonical sum_choiceType T1 T2 := + Eval hnf in ChoiceType (T1 + T2) (sum_choiceMixin T1 T2). + +Definition tree_choiceMixin T := PcanChoiceMixin (GenTree.codeK T). +Canonical tree_choiceType T := ChoiceType (GenTree.tree T) (tree_choiceMixin T). + +End ChoiceTheory. + +Prenex Implicits xchoose choose. +Notation "[ 'choiceMixin' 'of' T 'by' <: ]" := + (sub_choiceMixin _ : choiceMixin T) + (at level 0, format "[ 'choiceMixin' 'of' T 'by' <: ]") : form_scope. + +Module Countable. + +Record mixin_of (T : Type) : Type := Mixin { + pickle : T -> nat; + unpickle : nat -> option T; + pickleK : pcancel pickle unpickle +}. + +Definition EqMixin T m := PcanEqMixin (@pickleK T m). +Definition ChoiceMixin T m := PcanChoiceMixin (@pickleK T m). + +Section ClassDef. + +Record class_of T := Class { base : Choice.class_of T; mixin : mixin_of T }. +Local Coercion base : class_of >-> Choice.class_of. + +Structure type : Type := Pack {sort : Type; _ : class_of sort; _ : Type}. +Local Coercion sort : type >-> Sortclass. +Variables (T : Type) (cT : type). +Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack T c T. +Let xT := let: Pack T _ _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack m := + fun bT b & phant_id (Choice.class bT) b => Pack (@Class T b m) T. + +Definition eqType := @Equality.Pack cT xclass xT. +Definition choiceType := @Choice.Pack cT xclass xT. + +End ClassDef. + +Module Exports. +Coercion base : class_of >-> Choice.class_of. +Coercion mixin : class_of >-> mixin_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Canonical eqType. +Coercion choiceType : type >-> Choice.type. +Canonical choiceType. +Notation countType := type. +Notation CountType T m := (@pack T m _ _ id). +Notation CountMixin := Mixin. +Notation CountChoiceMixin := ChoiceMixin. +Notation "[ 'countType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) + (at level 0, format "[ 'countType' 'of' T 'for' cT ]") : form_scope. +Notation "[ 'countType' 'of' T ]" := (@clone T _ _ id) + (at level 0, format "[ 'countType' 'of' T ]") : form_scope. + +End Exports. + +End Countable. +Export Countable.Exports. + +Definition unpickle T := Countable.unpickle (Countable.class T). +Definition pickle T := Countable.pickle (Countable.class T). +Implicit Arguments unpickle [T]. +Prenex Implicits pickle unpickle. + +Section CountableTheory. + +Variable T : countType. + +Lemma pickleK : @pcancel nat T pickle unpickle. +Proof. exact: Countable.pickleK. Qed. + +Definition pickle_inv n := + obind (fun x : T => if pickle x == n then Some x else None) (unpickle n). + +Lemma pickle_invK : ocancel pickle_inv pickle. +Proof. +by rewrite /pickle_inv => n; case def_x: (unpickle n) => //= [x]; case: eqP. +Qed. + +Lemma pickleK_inv : pcancel pickle pickle_inv. +Proof. by rewrite /pickle_inv => x; rewrite pickleK /= eqxx. Qed. + +Lemma pcan_pickleK sT f f' : + @pcancel T sT f f' -> pcancel (pickle \o f) (pcomp f' unpickle). +Proof. by move=> fK x; rewrite /pcomp pickleK /= fK. Qed. + +Definition PcanCountMixin sT f f' (fK : pcancel f f') := + @CountMixin sT _ _ (pcan_pickleK fK). + +Definition CanCountMixin sT f f' (fK : cancel f f') := + @PcanCountMixin sT _ _ (can_pcan fK). + +Definition sub_countMixin P sT := PcanCountMixin (@valK T P sT). + +Definition pickle_seq s := CodeSeq.code (map (@pickle T) s). +Definition unpickle_seq n := Some (pmap (@unpickle T) (CodeSeq.decode n)). +Lemma pickle_seqK : pcancel pickle_seq unpickle_seq. +Proof. by move=> s; rewrite /unpickle_seq CodeSeq.codeK (map_pK pickleK). Qed. + +Definition seq_countMixin := CountMixin pickle_seqK. +Canonical seq_countType := Eval hnf in CountType (seq T) seq_countMixin. + +End CountableTheory. + +Notation "[ 'countMixin' 'of' T 'by' <: ]" := + (sub_countMixin _ : Countable.mixin_of T) + (at level 0, format "[ 'countMixin' 'of' T 'by' <: ]") : form_scope. + +Section SubCountType. + +Variables (T : choiceType) (P : pred T). +Import Countable. + +Structure subCountType : Type := + SubCountType {subCount_sort :> subType P; _ : mixin_of subCount_sort}. + +Coercion sub_countType (sT : subCountType) := + Eval hnf in pack (let: SubCountType _ m := sT return mixin_of sT in m) id. +Canonical sub_countType. + +Definition pack_subCountType U := + fun sT cT & sub_sort sT * sort cT -> U * U => + fun b m & phant_id (Class b m) (class cT) => @SubCountType sT m. + +End SubCountType. + +(* This assumes that T has both countType and subType structures. *) +Notation "[ 'subCountType' 'of' T ]" := + (@pack_subCountType _ _ T _ _ id _ _ id) + (at level 0, format "[ 'subCountType' 'of' T ]") : form_scope. + +Section TagCountType. + +Variables (I : countType) (T_ : I -> countType). + +Definition pickle_tagged (u : {i : I & T_ i}) := + CodeSeq.code [:: pickle (tag u); pickle (tagged u)]. +Definition unpickle_tagged s := + if CodeSeq.decode s is [:: ni; nx] then + obind (fun i => omap (@Tagged I i T_) (unpickle nx)) (unpickle ni) + else None. +Lemma pickle_taggedK : pcancel pickle_tagged unpickle_tagged. +Proof. +by case=> i x; rewrite /unpickle_tagged CodeSeq.codeK /= pickleK /= pickleK. +Qed. + +Definition tag_countMixin := CountMixin pickle_taggedK. +Canonical tag_countType := Eval hnf in CountType {i : I & T_ i} tag_countMixin. + +End TagCountType. + +(* The remaining Canonicals for standard datatypes. *) +Section CountableDataTypes. + +Implicit Type T : countType. + +Lemma nat_pickleK : pcancel id (@Some nat). Proof. by []. Qed. +Definition nat_countMixin := CountMixin nat_pickleK. +Canonical nat_countType := Eval hnf in CountType nat nat_countMixin. + +Definition bool_countMixin := CanCountMixin oddb. +Canonical bool_countType := Eval hnf in CountType bool bool_countMixin. +Canonical bitseq_countType := Eval hnf in [countType of bitseq]. + +Definition unit_countMixin := CanCountMixin bool_of_unitK. +Canonical unit_countType := Eval hnf in CountType unit unit_countMixin. + +Definition option_countMixin T := CanCountMixin (@seq_of_optK T). +Canonical option_countType T := + Eval hnf in CountType (option T) (option_countMixin T). + +Definition sig_countMixin T (P : pred T) := [countMixin of {x | P x} by <:]. +Canonical sig_countType T (P : pred T) := + Eval hnf in CountType {x | P x} (sig_countMixin P). +Canonical sig_subCountType T (P : pred T) := + Eval hnf in [subCountType of {x | P x}]. + +Definition prod_countMixin T1 T2 := CanCountMixin (@tag_of_pairK T1 T2). +Canonical prod_countType T1 T2 := + Eval hnf in CountType (T1 * T2) (prod_countMixin T1 T2). + +Definition sum_countMixin T1 T2 := PcanCountMixin (@opair_of_sumK T1 T2). +Canonical sum_countType T1 T2 := + Eval hnf in CountType (T1 + T2) (sum_countMixin T1 T2). + +Definition tree_countMixin T := PcanCountMixin (GenTree.codeK T). +Canonical tree_countType T := CountType (GenTree.tree T) (tree_countMixin T). + +End CountableDataTypes. diff --git a/mathcomp/discrete/div.v b/mathcomp/discrete/div.v new file mode 100644 index 0000000..d06a8e3 --- /dev/null +++ b/mathcomp/discrete/div.v @@ -0,0 +1,946 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. + +(******************************************************************************) +(* This file deals with divisibility for natural numbers. *) +(* It contains the definitions of: *) +(* edivn m d == the pair composed of the quotient and remainder *) +(* of the Euclidean division of m by d. *) +(* m %/ d == quotient of m by d. *) +(* m %% d == remainder of m by d. *) +(* m = n %[mod d] <-> m equals n modulo d. *) +(* m == n %[mod d] <=> m equals n modulo d (boolean version). *) +(* m <> n %[mod d] <-> m differs from n modulo d. *) +(* m != n %[mod d] <=> m differs from n modulo d (boolean version). *) +(* d %| m <=> d divides m. *) +(* gcdn m n == the GCD of m and n. *) +(* egcdn m n == the extended GCD of m and n. *) +(* lcmn m n == the LCM of m and n. *) +(* coprime m n <=> m and n are coprime (:= gcdn m n == 1). *) +(* chinese m n r s == witness of the chinese remainder theorem. *) +(* We adjoin an m to operator suffixes to indicate a nested %% (modn), as in *) +(* modnDml : m %% d + n = m + n %[mod d]. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(** Euclidean division *) + +Definition edivn_rec d := + fix loop m q := if m - d is m'.+1 then loop m' q.+1 else (q, m). + +Definition edivn m d := if d > 0 then edivn_rec d.-1 m 0 else (0, m). + +CoInductive edivn_spec m d : nat * nat -> Type := + EdivnSpec q r of m = q * d + r & (d > 0) ==> (r < d) : edivn_spec m d (q, r). + +Lemma edivnP m d : edivn_spec m d (edivn m d). +Proof. +rewrite -{1}[m]/(0 * d + m) /edivn; case: d => //= d. +elim: m {-2}m 0 (leqnn m) => [|n IHn] [|m] q //= le_mn. +have le_m'n: m - d <= n by rewrite (leq_trans (leq_subr d m)). +rewrite subn_if_gt; case: ltnP => [// | le_dm]. +by rewrite -{1}(subnKC le_dm) -addSn addnA -mulSnr; exact: IHn. +Qed. + +Lemma edivn_eq d q r : r < d -> edivn (q * d + r) d = (q, r). +Proof. +move=> lt_rd; have d_gt0: 0 < d by exact: leq_trans lt_rd. +case: edivnP lt_rd => q' r'; rewrite d_gt0 /=. +wlog: q q' r r' / q <= q' by case/orP: (leq_total q q'); last symmetry; eauto. +rewrite leq_eqVlt; case/predU1P => [-> /addnI-> |] //=. +rewrite -(leq_pmul2r d_gt0) => /leq_add lt_qr eq_qr _ /lt_qr {lt_qr}. +by rewrite addnS ltnNge mulSn -addnA eq_qr addnCA addnA leq_addr. +Qed. + +Definition divn m d := (edivn m d).1. + +Notation "m %/ d" := (divn m d) : nat_scope. + +(* We redefine modn so that it is structurally decreasing. *) + +Definition modn_rec d := fix loop m := if m - d is m'.+1 then loop m' else m. + +Definition modn m d := if d > 0 then modn_rec d.-1 m else m. + +Notation "m %% d" := (modn m d) : nat_scope. +Notation "m = n %[mod d ]" := (m %% d = n %% d) : nat_scope. +Notation "m == n %[mod d ]" := (m %% d == n %% d) : nat_scope. +Notation "m <> n %[mod d ]" := (m %% d <> n %% d) : nat_scope. +Notation "m != n %[mod d ]" := (m %% d != n %% d) : nat_scope. + +Lemma modn_def m d : m %% d = (edivn m d).2. +Proof. +case: d => //= d; rewrite /modn /edivn /=. +elim: m {-2}m 0 (leqnn m) => [|n IHn] [|m] q //=. +rewrite ltnS !subn_if_gt; case: (d <= m) => // le_mn. +by apply: IHn; apply: leq_trans le_mn; exact: leq_subr. +Qed. + +Lemma edivn_def m d : edivn m d = (m %/ d, m %% d). +Proof. by rewrite /divn modn_def; case: (edivn m d). Qed. + +Lemma divn_eq m d : m = m %/ d * d + m %% d. +Proof. by rewrite /divn modn_def; case: edivnP. Qed. + +Lemma div0n d : 0 %/ d = 0. Proof. by case: d. Qed. +Lemma divn0 m : m %/ 0 = 0. Proof. by []. Qed. +Lemma mod0n d : 0 %% d = 0. Proof. by case: d. Qed. +Lemma modn0 m : m %% 0 = m. Proof. by []. Qed. + +Lemma divn_small m d : m < d -> m %/ d = 0. +Proof. by move=> lt_md; rewrite /divn (edivn_eq 0). Qed. + +Lemma divnMDl q m d : 0 < d -> (q * d + m) %/ d = q + m %/ d. +Proof. +move=> d_gt0; rewrite {1}(divn_eq m d) addnA -mulnDl. +by rewrite /divn edivn_eq // modn_def; case: edivnP; rewrite d_gt0. +Qed. + +Lemma mulnK m d : 0 < d -> m * d %/ d = m. +Proof. by move=> d_gt0; rewrite -[m * d]addn0 divnMDl // div0n addn0. Qed. + +Lemma mulKn m d : 0 < d -> d * m %/ d = m. +Proof. by move=> d_gt0; rewrite mulnC mulnK. Qed. + +Lemma expnB p m n : p > 0 -> m >= n -> p ^ (m - n) = p ^ m %/ p ^ n. +Proof. +by move=> p_gt0 /subnK{2}<-; rewrite expnD mulnK // expn_gt0 p_gt0. +Qed. + +Lemma modn1 m : m %% 1 = 0. +Proof. by rewrite modn_def; case: edivnP => ? []. Qed. + +Lemma divn1 m : m %/ 1 = m. +Proof. by rewrite {2}(@divn_eq m 1) // modn1 addn0 muln1. Qed. + +Lemma divnn d : d %/ d = (0 < d). +Proof. by case: d => // d; rewrite -{1}[d.+1]muln1 mulKn. Qed. + +Lemma divnMl p m d : p > 0 -> p * m %/ (p * d) = m %/ d. +Proof. +move=> p_gt0; case: (posnP d) => [-> | d_gt0]; first by rewrite muln0. +rewrite {2}/divn; case: edivnP; rewrite d_gt0 /= => q r ->{m} lt_rd. +rewrite mulnDr mulnCA divnMDl; last by rewrite muln_gt0 p_gt0. +by rewrite addnC divn_small // ltn_pmul2l. +Qed. +Implicit Arguments divnMl [p m d]. + +Lemma divnMr p m d : p > 0 -> m * p %/ (d * p) = m %/ d. +Proof. by move=> p_gt0; rewrite -!(mulnC p) divnMl. Qed. +Implicit Arguments divnMr [p m d]. + +Lemma ltn_mod m d : (m %% d < d) = (0 < d). +Proof. by case: d => // d; rewrite modn_def; case: edivnP. Qed. + +Lemma ltn_pmod m d : 0 < d -> m %% d < d. +Proof. by rewrite ltn_mod. Qed. + +Lemma leq_trunc_div m d : m %/ d * d <= m. +Proof. by rewrite {2}(divn_eq m d) leq_addr. Qed. + +Lemma leq_mod m d : m %% d <= m. +Proof. by rewrite {2}(divn_eq m d) leq_addl. Qed. + +Lemma leq_div m d : m %/ d <= m. +Proof. +by case: d => // d; apply: leq_trans (leq_pmulr _ _) (leq_trunc_div _ _). +Qed. + +Lemma ltn_ceil m d : 0 < d -> m < (m %/ d).+1 * d. +Proof. +by move=> d_gt0; rewrite {1}(divn_eq m d) -addnS mulSnr leq_add2l ltn_mod. +Qed. + +Lemma ltn_divLR m n d : d > 0 -> (m %/ d < n) = (m < n * d). +Proof. +move=> d_gt0; apply/idP/idP. + by rewrite -(leq_pmul2r d_gt0); apply: leq_trans (ltn_ceil _ _). +rewrite !ltnNge -(@leq_pmul2r d n) //; apply: contra => le_nd_floor. +exact: leq_trans le_nd_floor (leq_trunc_div _ _). +Qed. + +Lemma leq_divRL m n d : d > 0 -> (m <= n %/ d) = (m * d <= n). +Proof. by move=> d_gt0; rewrite leqNgt ltn_divLR // -leqNgt. Qed. + +Lemma ltn_Pdiv m d : 1 < d -> 0 < m -> m %/ d < m. +Proof. by move=> d_gt1 m_gt0; rewrite ltn_divLR ?ltn_Pmulr // ltnW. Qed. + +Lemma divn_gt0 d m : 0 < d -> (0 < m %/ d) = (d <= m). +Proof. by move=> d_gt0; rewrite leq_divRL ?mul1n. Qed. + +Lemma leq_div2r d m n : m <= n -> m %/ d <= n %/ d. +Proof. +have [-> //| d_gt0 le_mn] := posnP d. +by rewrite leq_divRL // (leq_trans _ le_mn) -?leq_divRL. +Qed. + +Lemma leq_div2l m d e : 0 < d -> d <= e -> m %/ e <= m %/ d. +Proof. +move/leq_divRL=> -> le_de. +by apply: leq_trans (leq_trunc_div m e); apply: leq_mul. +Qed. + +Lemma leq_divDl p m n : (m + n) %/ p <= m %/ p + n %/ p + 1. +Proof. +have [-> //| p_gt0] := posnP p; rewrite -ltnS -addnS ltn_divLR // ltnW //. +rewrite {1}(divn_eq n p) {1}(divn_eq m p) addnACA !mulnDl -3!addnS leq_add2l. +by rewrite mul2n -addnn -addSn leq_add // ltn_mod. +Qed. + +Lemma geq_divBl k m p : k %/ p - m %/ p <= (k - m) %/ p + 1. +Proof. +rewrite leq_subLR addnA; apply: leq_trans (leq_divDl _ _ _). +by rewrite -maxnE leq_div2r ?leq_maxr. +Qed. + +Lemma divnMA m n p : m %/ (n * p) = m %/ n %/ p. +Proof. +case: n p => [|n] [|p]; rewrite ?muln0 ?div0n //. +rewrite {2}(divn_eq m (n.+1 * p.+1)) mulnA mulnAC !divnMDl //. +by rewrite [_ %/ p.+1]divn_small ?addn0 // ltn_divLR // mulnC ltn_mod. +Qed. + +Lemma divnAC m n p : m %/ n %/ p = m %/ p %/ n. +Proof. by rewrite -!divnMA mulnC. Qed. + +Lemma modn_small m d : m < d -> m %% d = m. +Proof. by move=> lt_md; rewrite {2}(divn_eq m d) divn_small. Qed. + +Lemma modn_mod m d : m %% d = m %[mod d]. +Proof. by case: d => // d; apply: modn_small; rewrite ltn_mod. Qed. + +Lemma modnMDl p m d : p * d + m = m %[mod d]. +Proof. +case: (posnP d) => [-> | d_gt0]; first by rewrite muln0. +by rewrite {1}(divn_eq m d) addnA -mulnDl modn_def edivn_eq // ltn_mod. +Qed. + +Lemma muln_modr {p m d} : 0 < p -> p * (m %% d) = (p * m) %% (p * d). +Proof. +move=> p_gt0; apply: (@addnI (p * (m %/ d * d))). +by rewrite -mulnDr -divn_eq mulnCA -(divnMl p_gt0) -divn_eq. +Qed. + +Lemma muln_modl {p m d} : 0 < p -> (m %% d) * p = (m * p) %% (d * p). +Proof. by rewrite -!(mulnC p); apply: muln_modr. Qed. + +Lemma modnDl m d : d + m = m %[mod d]. +Proof. by rewrite -{1}[d]mul1n modnMDl. Qed. + +Lemma modnDr m d : m + d = m %[mod d]. +Proof. by rewrite addnC modnDl. Qed. + +Lemma modnn d : d %% d = 0. +Proof. by rewrite -{1}[d]addn0 modnDl mod0n. Qed. + +Lemma modnMl p d : p * d %% d = 0. +Proof. by rewrite -[p * d]addn0 modnMDl mod0n. Qed. + +Lemma modnMr p d : d * p %% d = 0. +Proof. by rewrite mulnC modnMl. Qed. + +Lemma modnDml m n d : m %% d + n = m + n %[mod d]. +Proof. by rewrite {2}(divn_eq m d) -addnA modnMDl. Qed. + +Lemma modnDmr m n d : m + n %% d = m + n %[mod d]. +Proof. by rewrite !(addnC m) modnDml. Qed. + +Lemma modnDm m n d : m %% d + n %% d = m + n %[mod d]. +Proof. by rewrite modnDml modnDmr. Qed. + +Lemma eqn_modDl p m n d : (p + m == p + n %[mod d]) = (m == n %[mod d]). +Proof. +case: d => [|d]; first by rewrite !modn0 eqn_add2l. +apply/eqP/eqP=> eq_mn; last by rewrite -modnDmr eq_mn modnDmr. +rewrite -(modnMDl p m) -(modnMDl p n) !mulnSr -!addnA. +by rewrite -modnDmr eq_mn modnDmr. +Qed. + +Lemma eqn_modDr p m n d : (m + p == n + p %[mod d]) = (m == n %[mod d]). +Proof. by rewrite -!(addnC p) eqn_modDl. Qed. + +Lemma modnMml m n d : m %% d * n = m * n %[mod d]. +Proof. by rewrite {2}(divn_eq m d) mulnDl mulnAC modnMDl. Qed. + +Lemma modnMmr m n d : m * (n %% d) = m * n %[mod d]. +Proof. by rewrite !(mulnC m) modnMml. Qed. + +Lemma modnMm m n d : m %% d * (n %% d) = m * n %[mod d]. +Proof. by rewrite modnMml modnMmr. Qed. + +Lemma modn2 m : m %% 2 = odd m. +Proof. by elim: m => //= m IHm; rewrite -addn1 -modnDml IHm; case odd. Qed. + +Lemma divn2 m : m %/ 2 = m./2. +Proof. by rewrite {2}(divn_eq m 2) modn2 muln2 addnC half_bit_double. Qed. + +Lemma odd_mod m d : odd d = false -> odd (m %% d) = odd m. +Proof. +by move=> d_even; rewrite {2}(divn_eq m d) odd_add odd_mul d_even andbF. +Qed. + +Lemma modnXm m n a : (a %% n) ^ m = a ^ m %[mod n]. +Proof. +by elim: m => // m IHm; rewrite !expnS -modnMmr IHm modnMml modnMmr. +Qed. + +(** Divisibility **) + +Definition dvdn d m := m %% d == 0. + +Notation "m %| d" := (dvdn m d) : nat_scope. + +Lemma dvdnP d m : reflect (exists k, m = k * d) (d %| m). +Proof. +apply: (iffP eqP) => [md0 | [k ->]]; last by rewrite modnMl. +by exists (m %/ d); rewrite {1}(divn_eq m d) md0 addn0. +Qed. +Implicit Arguments dvdnP [d m]. +Prenex Implicits dvdnP. + +Lemma dvdn0 d : d %| 0. +Proof. by case: d. Qed. + +Lemma dvd0n n : (0 %| n) = (n == 0). +Proof. by case: n. Qed. + +Lemma dvdn1 d : (d %| 1) = (d == 1). +Proof. by case: d => [|[|d]] //; rewrite /dvdn modn_small. Qed. + +Lemma dvd1n m : 1 %| m. +Proof. by rewrite /dvdn modn1. Qed. + +Lemma dvdn_gt0 d m : m > 0 -> d %| m -> d > 0. +Proof. by case: d => // /prednK <-. Qed. + +Lemma dvdnn m : m %| m. +Proof. by rewrite /dvdn modnn. Qed. + +Lemma dvdn_mull d m n : d %| n -> d %| m * n. +Proof. by case/dvdnP=> n' ->; rewrite /dvdn mulnA modnMl. Qed. + +Lemma dvdn_mulr d m n : d %| m -> d %| m * n. +Proof. by move=> d_m; rewrite mulnC dvdn_mull. Qed. +Hint Resolve dvdn0 dvd1n dvdnn dvdn_mull dvdn_mulr. + +Lemma dvdn_mul d1 d2 m1 m2 : d1 %| m1 -> d2 %| m2 -> d1 * d2 %| m1 * m2. +Proof. +by move=> /dvdnP[q1 ->] /dvdnP[q2 ->]; rewrite mulnCA -mulnA 2?dvdn_mull. +Qed. + +Lemma dvdn_trans n d m : d %| n -> n %| m -> d %| m. +Proof. by move=> d_dv_n /dvdnP[n1 ->]; exact: dvdn_mull. Qed. + +Lemma dvdn_eq d m : (d %| m) = (m %/ d * d == m). +Proof. +apply/eqP/eqP=> [modm0 | <-]; last exact: modnMl. +by rewrite {2}(divn_eq m d) modm0 addn0. +Qed. + +Lemma dvdn2 n : (2 %| n) = ~~ odd n. +Proof. by rewrite /dvdn modn2; case (odd n). Qed. + +Lemma dvdn_odd m n : m %| n -> odd n -> odd m. +Proof. +by move=> m_dv_n; apply: contraTT; rewrite -!dvdn2 => /dvdn_trans->. +Qed. + +Lemma divnK d m : d %| m -> m %/ d * d = m. +Proof. by rewrite dvdn_eq; move/eqP. Qed. + +Lemma leq_divLR d m n : d %| m -> (m %/ d <= n) = (m <= n * d). +Proof. by case: d m => [|d] [|m] ///divnK=> {2}<-; rewrite leq_pmul2r. Qed. + +Lemma ltn_divRL d m n : d %| m -> (n < m %/ d) = (n * d < m). +Proof. by move=> dv_d_m; rewrite !ltnNge leq_divLR. Qed. + +Lemma eqn_div d m n : d > 0 -> d %| m -> (n == m %/ d) = (n * d == m). +Proof. by move=> d_gt0 dv_d_m; rewrite -(eqn_pmul2r d_gt0) divnK. Qed. + +Lemma eqn_mul d m n : d > 0 -> d %| m -> (m == n * d) = (m %/ d == n). +Proof. by move=> d_gt0 dv_d_m; rewrite eq_sym -eqn_div // eq_sym. Qed. + +Lemma divn_mulAC d m n : d %| m -> m %/ d * n = m * n %/ d. +Proof. +case: d m => [[] //| d m] dv_d_m; apply/eqP. +by rewrite eqn_div ?dvdn_mulr // mulnAC divnK. +Qed. + +Lemma muln_divA d m n : d %| n -> m * (n %/ d) = m * n %/ d. +Proof. by move=> dv_d_m; rewrite !(mulnC m) divn_mulAC. Qed. + +Lemma muln_divCA d m n : d %| m -> d %| n -> m * (n %/ d) = n * (m %/ d). +Proof. by move=> dv_d_m dv_d_n; rewrite mulnC divn_mulAC ?muln_divA. Qed. + +Lemma divnA m n p : p %| n -> m %/ (n %/ p) = m * p %/ n. +Proof. by case: p => [|p] dv_n; rewrite -{2}(divnK dv_n) // divnMr. Qed. + +Lemma modn_dvdm m n d : d %| m -> n %% m = n %[mod d]. +Proof. +by case/dvdnP=> q def_m; rewrite {2}(divn_eq n m) {3}def_m mulnA modnMDl. +Qed. + +Lemma dvdn_leq d m : 0 < m -> d %| m -> d <= m. +Proof. by move=> m_gt0 /dvdnP[[|k] Dm]; rewrite Dm // leq_addr in m_gt0 *. Qed. + +Lemma gtnNdvd n d : 0 < n -> n < d -> (d %| n) = false. +Proof. by move=> n_gt0 lt_nd; rewrite /dvdn eqn0Ngt modn_small ?n_gt0. Qed. + +Lemma eqn_dvd m n : (m == n) = (m %| n) && (n %| m). +Proof. +case: m n => [|m] [|n] //; apply/idP/andP; first by move/eqP->; auto. +rewrite eqn_leq => [[Hmn Hnm]]; apply/andP; have:= dvdn_leq; auto. +Qed. + +Lemma dvdn_pmul2l p d m : 0 < p -> (p * d %| p * m) = (d %| m). +Proof. by case: p => // p _; rewrite /dvdn -muln_modr // muln_eq0. Qed. +Implicit Arguments dvdn_pmul2l [p m d]. + +Lemma dvdn_pmul2r p d m : 0 < p -> (d * p %| m * p) = (d %| m). +Proof. by move=> p_gt0; rewrite -!(mulnC p) dvdn_pmul2l. Qed. +Implicit Arguments dvdn_pmul2r [p m d]. + +Lemma dvdn_divLR p d m : 0 < p -> p %| d -> (d %/ p %| m) = (d %| m * p). +Proof. by move=> /(@dvdn_pmul2r p _ m) <- /divnK->. Qed. + +Lemma dvdn_divRL p d m : p %| m -> (d %| m %/ p) = (d * p %| m). +Proof. +have [-> | /(@dvdn_pmul2r p d) <- /divnK-> //] := posnP p. +by rewrite divn0 muln0 dvdn0. +Qed. + +Lemma dvdn_div d m : d %| m -> m %/ d %| m. +Proof. by move/divnK=> {2}<-; apply: dvdn_mulr. Qed. + +Lemma dvdn_exp2l p m n : m <= n -> p ^ m %| p ^ n. +Proof. by move/subnK <-; rewrite expnD dvdn_mull. Qed. + +Lemma dvdn_Pexp2l p m n : p > 1 -> (p ^ m %| p ^ n) = (m <= n). +Proof. +move=> p_gt1; case: leqP => [|gt_n_m]; first exact: dvdn_exp2l. +by rewrite gtnNdvd ?ltn_exp2l ?expn_gt0 // ltnW. +Qed. + +Lemma dvdn_exp2r m n k : m %| n -> m ^ k %| n ^ k. +Proof. by case/dvdnP=> q ->; rewrite expnMn dvdn_mull. Qed. + +Lemma dvdn_addr m d n : d %| m -> (d %| m + n) = (d %| n). +Proof. by case/dvdnP=> q ->; rewrite /dvdn modnMDl. Qed. + +Lemma dvdn_addl n d m : d %| n -> (d %| m + n) = (d %| m). +Proof. by rewrite addnC; exact: dvdn_addr. Qed. + +Lemma dvdn_add d m n : d %| m -> d %| n -> d %| m + n. +Proof. by move/dvdn_addr->. Qed. + +Lemma dvdn_add_eq d m n : d %| m + n -> (d %| m) = (d %| n). +Proof. by move=> dv_d_mn; apply/idP/idP => [/dvdn_addr | /dvdn_addl] <-. Qed. + +Lemma dvdn_subr d m n : n <= m -> d %| m -> (d %| m - n) = (d %| n). +Proof. by move=> le_n_m dv_d_m; apply: dvdn_add_eq; rewrite subnK. Qed. + +Lemma dvdn_subl d m n : n <= m -> d %| n -> (d %| m - n) = (d %| m). +Proof. by move=> le_n_m dv_d_m; rewrite -(dvdn_addl _ dv_d_m) subnK. Qed. + +Lemma dvdn_sub d m n : d %| m -> d %| n -> d %| m - n. +Proof. +by case: (leqP n m) => [le_nm /dvdn_subr <- // | /ltnW/eqnP ->]; rewrite dvdn0. +Qed. + +Lemma dvdn_exp k d m : 0 < k -> d %| m -> d %| (m ^ k). +Proof. by case: k => // k _ d_dv_m; rewrite expnS dvdn_mulr. Qed. + +Hint Resolve dvdn_add dvdn_sub dvdn_exp. + +Lemma eqn_mod_dvd d m n : n <= m -> (m == n %[mod d]) = (d %| m - n). +Proof. +by move=> le_mn; rewrite -{1}[n]add0n -{1}(subnK le_mn) eqn_modDr mod0n. +Qed. + +Lemma divnDl m n d : d %| m -> (m + n) %/ d = m %/ d + n %/ d. +Proof. by case: d => // d /divnK{1}<-; rewrite divnMDl. Qed. + +Lemma divnDr m n d : d %| n -> (m + n) %/ d = m %/ d + n %/ d. +Proof. by move=> dv_n; rewrite addnC divnDl // addnC. Qed. + +(***********************************************************************) +(* A function that computes the gcd of 2 numbers *) +(***********************************************************************) + +Fixpoint gcdn_rec m n := + let n' := n %% m in if n' is 0 then m else + if m - n'.-1 is m'.+1 then gcdn_rec (m' %% n') n' else n'. + +Definition gcdn := nosimpl gcdn_rec. + +Lemma gcdnE m n : gcdn m n = if m == 0 then n else gcdn (n %% m) m. +Proof. +rewrite /gcdn; elim: m {-2}m (leqnn m) n => [|s IHs] [|m] le_ms [|n] //=. +case def_n': (_ %% _) => // [n']. +have{def_n'} lt_n'm: n' < m by rewrite -def_n' -ltnS ltn_pmod. +rewrite {}IHs ?(leq_trans lt_n'm) // subn_if_gt ltnW //=; congr gcdn_rec. +by rewrite -{2}(subnK (ltnW lt_n'm)) -addnS modnDr. +Qed. + +Lemma gcdnn : idempotent gcdn. +Proof. by case=> // n; rewrite gcdnE modnn. Qed. + +Lemma gcdnC : commutative gcdn. +Proof. +move=> m n; wlog lt_nm: m n / n < m. + by case: (ltngtP n m) => [||-> //]; last symmetry; auto. +by rewrite gcdnE -{1}(ltn_predK lt_nm) modn_small. +Qed. + +Lemma gcd0n : left_id 0 gcdn. Proof. by case. Qed. +Lemma gcdn0 : right_id 0 gcdn. Proof. by case. Qed. + +Lemma gcd1n : left_zero 1 gcdn. +Proof. by move=> n; rewrite gcdnE modn1. Qed. + +Lemma gcdn1 : right_zero 1 gcdn. +Proof. by move=> n; rewrite gcdnC gcd1n. Qed. + +Lemma dvdn_gcdr m n : gcdn m n %| n. +Proof. +elim: m {-2}m (leqnn m) n => [|s IHs] [|m] le_ms [|n] //. +rewrite gcdnE; case def_n': (_ %% _) => [|n']; first by rewrite /dvdn def_n'. +have lt_n's: n' < s by rewrite -ltnS (leq_trans _ le_ms) // -def_n' ltn_pmod. +rewrite /= (divn_eq n.+1 m.+1) def_n' dvdn_addr ?dvdn_mull //; last exact: IHs. +by rewrite gcdnE /= IHs // (leq_trans _ lt_n's) // ltnW // ltn_pmod. +Qed. + +Lemma dvdn_gcdl m n : gcdn m n %| m. +Proof. by rewrite gcdnC dvdn_gcdr. Qed. + +Lemma gcdn_gt0 m n : (0 < gcdn m n) = (0 < m) || (0 < n). +Proof. +by case: m n => [|m] [|n] //; apply: (@dvdn_gt0 _ m.+1) => //; exact: dvdn_gcdl. +Qed. + +Lemma gcdnMDl k m n : gcdn m (k * m + n) = gcdn m n. +Proof. by rewrite !(gcdnE m) modnMDl mulnC; case: m. Qed. + +Lemma gcdnDl m n : gcdn m (m + n) = gcdn m n. +Proof. by rewrite -{2}(mul1n m) gcdnMDl. Qed. + +Lemma gcdnDr m n : gcdn m (n + m) = gcdn m n. +Proof. by rewrite addnC gcdnDl. Qed. + +Lemma gcdnMl n m : gcdn n (m * n) = n. +Proof. by case: n => [|n]; rewrite gcdnE modnMl gcd0n. Qed. + +Lemma gcdnMr n m : gcdn n (n * m) = n. +Proof. by rewrite mulnC gcdnMl. Qed. + +Lemma gcdn_idPl {m n} : reflect (gcdn m n = m) (m %| n). +Proof. +by apply: (iffP idP) => [/dvdnP[q ->] | <-]; rewrite (gcdnMl, dvdn_gcdr). +Qed. + +Lemma gcdn_idPr {m n} : reflect (gcdn m n = n) (n %| m). +Proof. by rewrite gcdnC; apply: gcdn_idPl. Qed. + +Lemma expn_min e m n : e ^ minn m n = gcdn (e ^ m) (e ^ n). +Proof. +rewrite /minn; case: leqP; [rewrite gcdnC | move/ltnW]; + by move/(dvdn_exp2l e)/gcdn_idPl. +Qed. + +Lemma gcdn_modr m n : gcdn m (n %% m) = gcdn m n. +Proof. by rewrite {2}(divn_eq n m) gcdnMDl. Qed. + +Lemma gcdn_modl m n : gcdn (m %% n) n = gcdn m n. +Proof. by rewrite !(gcdnC _ n) gcdn_modr. Qed. + +(* Extended gcd, which computes Bezout coefficients. *) + +Fixpoint Bezout_rec km kn qs := + if qs is q :: qs' then Bezout_rec kn (NatTrec.add_mul q kn km) qs' + else (km, kn). + +Fixpoint egcdn_rec m n s qs := + if s is s'.+1 then + let: (q, r) := edivn m n in + if r > 0 then egcdn_rec n r s' (q :: qs) else + if odd (size qs) then qs else q.-1 :: qs + else [::0]. + +Definition egcdn m n := Bezout_rec 0 1 (egcdn_rec m n n [::]). + +CoInductive egcdn_spec m n : nat * nat -> Type := + EgcdnSpec km kn of km * m = kn * n + gcdn m n & kn * gcdn m n < m : + egcdn_spec m n (km, kn). + +Lemma egcd0n n : egcdn 0 n = (1, 0). +Proof. by case: n. Qed. + +Lemma egcdnP m n : m > 0 -> egcdn_spec m n (egcdn m n). +Proof. +rewrite /egcdn; have: (n, m) = Bezout_rec n m [::] by []. +case: (posnP n) => [-> /=|]; first by split; rewrite // mul1n gcdn0. +move: {2 6}n {4 6}n {1 4}m [::] (ltnSn n) => s n0 m0. +elim: s n m => [[]//|s IHs] n m qs /= le_ns n_gt0 def_mn0 m_gt0. +case: edivnP => q r def_m; rewrite n_gt0 /= => lt_rn. +case: posnP => [r0 {s le_ns IHs lt_rn}|r_gt0]; last first. + by apply: IHs => //=; [rewrite (leq_trans lt_rn) | rewrite natTrecE -def_m]. +rewrite {r}r0 addn0 in def_m; set b := odd _; pose d := gcdn m n. +pose km := ~~ b : nat; pose kn := if b then 1 else q.-1. +rewrite (_ : Bezout_rec _ _ _ = Bezout_rec km kn qs); last first. + by rewrite /kn /km; case: (b) => //=; rewrite natTrecE addn0 muln1. +have def_d: d = n by rewrite /d def_m gcdnC gcdnE modnMl gcd0n -[n]prednK. +have: km * m + 2 * b * d = kn * n + d. + rewrite {}/kn {}/km def_m def_d -mulSnr; case: b; rewrite //= addn0 mul1n. + by rewrite prednK //; apply: dvdn_gt0 m_gt0 _; rewrite def_m dvdn_mulr. +have{def_m}: kn * d <= m. + have q_gt0 : 0 < q by rewrite def_m muln_gt0 n_gt0 ?andbT in m_gt0. + by rewrite /kn; case b; rewrite def_d def_m leq_pmul2r // leq_pred. +have{def_d}: km * d <= n by rewrite -[n]mul1n def_d leq_pmul2r // leq_b1. +move: km {q}kn m_gt0 n_gt0 def_mn0; rewrite {}/d {}/b. +elim: qs m n => [|q qs IHq] n r kn kr n_gt0 r_gt0 /=. + case=> -> -> {m0 n0}; rewrite !addn0 => le_kn_r _ def_d; split=> //. + have d_gt0: 0 < gcdn n r by rewrite gcdn_gt0 n_gt0. + have: 0 < kn * n by rewrite def_d addn_gt0 d_gt0 orbT. + rewrite muln_gt0 n_gt0 andbT; move/ltn_pmul2l <-. + by rewrite def_d -addn1 leq_add // mulnCA leq_mul2l le_kn_r orbT. +rewrite !natTrecE; set m:= _ + r; set km := _ * _ + kn; pose d := gcdn m n. +have ->: gcdn n r = d by rewrite [d]gcdnC gcdnMDl. +have m_gt0: 0 < m by rewrite addn_gt0 r_gt0 orbT. +have d_gt0: 0 < d by rewrite gcdn_gt0 m_gt0. +move/IHq=> {IHq} IHq le_kn_r le_kr_n def_d; apply: IHq => //; rewrite -/d. + by rewrite mulnDl leq_add // -mulnA leq_mul2l le_kr_n orbT. +apply: (@addIn d); rewrite -!addnA addnn addnCA mulnDr -addnA addnCA. +rewrite /km mulnDl mulnCA mulnA -addnA; congr (_ + _). +by rewrite -def_d addnC -addnA -mulnDl -mulnDr addn_negb -mul2n. +Qed. + +Lemma Bezoutl m n : m > 0 -> {a | a < m & m %| gcdn m n + a * n}. +Proof. +move=> m_gt0; case: (egcdnP n m_gt0) => km kn def_d lt_kn_m. +exists kn; last by rewrite addnC -def_d dvdn_mull. +apply: leq_ltn_trans lt_kn_m. +by rewrite -{1}[kn]muln1 leq_mul2l gcdn_gt0 m_gt0 orbT. +Qed. + +Lemma Bezoutr m n : n > 0 -> {a | a < n & n %| gcdn m n + a * m}. +Proof. by rewrite gcdnC; exact: Bezoutl. Qed. + +(* Back to the gcd. *) + +Lemma dvdn_gcd p m n : p %| gcdn m n = (p %| m) && (p %| n). +Proof. +apply/idP/andP=> [dv_pmn | [dv_pm dv_pn]]. + by rewrite !(dvdn_trans dv_pmn) ?dvdn_gcdl ?dvdn_gcdr. +case (posnP n) => [->|n_gt0]; first by rewrite gcdn0. +case: (Bezoutr m n_gt0) => // km _ /(dvdn_trans dv_pn). +by rewrite dvdn_addl // dvdn_mull. +Qed. + +Lemma gcdnAC : right_commutative gcdn. +Proof. +suffices dvd m n p: gcdn (gcdn m n) p %| gcdn (gcdn m p) n. + by move=> m n p; apply/eqP; rewrite eqn_dvd !dvd. +rewrite !dvdn_gcd dvdn_gcdr. +by rewrite !(dvdn_trans (dvdn_gcdl _ p)) ?dvdn_gcdl ?dvdn_gcdr. +Qed. + +Lemma gcdnA : associative gcdn. +Proof. by move=> m n p; rewrite !(gcdnC m) gcdnAC. Qed. + +Lemma gcdnCA : left_commutative gcdn. +Proof. by move=> m n p; rewrite !gcdnA (gcdnC m). Qed. + +Lemma gcdnACA : interchange gcdn gcdn. +Proof. by move=> m n p q; rewrite -!gcdnA (gcdnCA n). Qed. + +Lemma muln_gcdr : right_distributive muln gcdn. +Proof. +move=> p m n; case: (posnP p) => [-> //| p_gt0]. +elim: {m}m.+1 {-2}m n (ltnSn m) => // s IHs m n; rewrite ltnS => le_ms. +rewrite gcdnE [rhs in _ = rhs]gcdnE muln_eq0 (gtn_eqF p_gt0) -muln_modr //=. +by case: posnP => // m_gt0; apply: IHs; apply: leq_trans le_ms; apply: ltn_pmod. +Qed. + +Lemma muln_gcdl : left_distributive muln gcdn. +Proof. by move=> m n p; rewrite -!(mulnC p) muln_gcdr. Qed. + +Lemma gcdn_def d m n : + d %| m -> d %| n -> (forall d', d' %| m -> d' %| n -> d' %| d) -> + gcdn m n = d. +Proof. +move=> dv_dm dv_dn gdv_d; apply/eqP. +by rewrite eqn_dvd dvdn_gcd dv_dm dv_dn gdv_d ?dvdn_gcdl ?dvdn_gcdr. +Qed. + +Lemma muln_divCA_gcd n m : n * (m %/ gcdn n m) = m * (n %/ gcdn n m). +Proof. by rewrite muln_divCA ?dvdn_gcdl ?dvdn_gcdr. Qed. + +(* We derive the lcm directly. *) + +Definition lcmn m n := m * n %/ gcdn m n. + +Lemma lcmnC : commutative lcmn. +Proof. by move=> m n; rewrite /lcmn mulnC gcdnC. Qed. + +Lemma lcm0n : left_zero 0 lcmn. Proof. by move=> n; exact: div0n. Qed. +Lemma lcmn0 : right_zero 0 lcmn. Proof. by move=> n; rewrite lcmnC lcm0n. Qed. + +Lemma lcm1n : left_id 1 lcmn. +Proof. by move=> n; rewrite /lcmn gcd1n mul1n divn1. Qed. + +Lemma lcmn1 : right_id 1 lcmn. +Proof. by move=> n; rewrite lcmnC lcm1n. Qed. + +Lemma muln_lcm_gcd m n : lcmn m n * gcdn m n = m * n. +Proof. by apply/eqP; rewrite divnK ?dvdn_mull ?dvdn_gcdr. Qed. + +Lemma lcmn_gt0 m n : (0 < lcmn m n) = (0 < m) && (0 < n). +Proof. by rewrite -muln_gt0 ltn_divRL ?dvdn_mull ?dvdn_gcdr. Qed. + +Lemma muln_lcmr : right_distributive muln lcmn. +Proof. +case=> // m n p; rewrite /lcmn -muln_gcdr -!mulnA divnMl // mulnCA. +by rewrite muln_divA ?dvdn_mull ?dvdn_gcdr. +Qed. + +Lemma muln_lcml : left_distributive muln lcmn. +Proof. by move=> m n p; rewrite -!(mulnC p) muln_lcmr. Qed. + +Lemma lcmnA : associative lcmn. +Proof. +move=> m n p; rewrite {1 3}/lcmn mulnC !divn_mulAC ?dvdn_mull ?dvdn_gcdr //. +rewrite -!divnMA ?dvdn_mulr ?dvdn_gcdl // mulnC mulnA !muln_gcdr. +by rewrite ![_ * lcmn _ _]mulnC !muln_lcm_gcd !muln_gcdl -!(mulnC m) gcdnA. +Qed. + +Lemma lcmnCA : left_commutative lcmn. +Proof. by move=> m n p; rewrite !lcmnA (lcmnC m). Qed. + +Lemma lcmnAC : right_commutative lcmn. +Proof. by move=> m n p; rewrite -!lcmnA (lcmnC n). Qed. + +Lemma lcmnACA : interchange lcmn lcmn. +Proof. by move=> m n p q; rewrite -!lcmnA (lcmnCA n). Qed. + +Lemma dvdn_lcml d1 d2 : d1 %| lcmn d1 d2. +Proof. by rewrite /lcmn -muln_divA ?dvdn_gcdr ?dvdn_mulr. Qed. + +Lemma dvdn_lcmr d1 d2 : d2 %| lcmn d1 d2. +Proof. by rewrite lcmnC dvdn_lcml. Qed. + +Lemma dvdn_lcm d1 d2 m : lcmn d1 d2 %| m = (d1 %| m) && (d2 %| m). +Proof. +case: d1 d2 => [|d1] [|d2]; try by case: m => [|m]; rewrite ?lcmn0 ?andbF. +rewrite -(@dvdn_pmul2r (gcdn d1.+1 d2.+1)) ?gcdn_gt0 // muln_lcm_gcd. +by rewrite muln_gcdr dvdn_gcd {1}mulnC andbC !dvdn_pmul2r. +Qed. + +Lemma lcmnMl m n : lcmn m (m * n) = m * n. +Proof. by case: m => // m; rewrite /lcmn gcdnMr mulKn. Qed. + +Lemma lcmnMr m n : lcmn n (m * n) = m * n. +Proof. by rewrite mulnC lcmnMl. Qed. + +Lemma lcmn_idPr {m n} : reflect (lcmn m n = n) (m %| n). +Proof. +by apply: (iffP idP) => [/dvdnP[q ->] | <-]; rewrite (lcmnMr, dvdn_lcml). +Qed. + +Lemma lcmn_idPl {m n} : reflect (lcmn m n = m) (n %| m). +Proof. by rewrite lcmnC; apply: lcmn_idPr. Qed. + +Lemma expn_max e m n : e ^ maxn m n = lcmn (e ^ m) (e ^ n). +Proof. +rewrite /maxn; case: leqP; [rewrite lcmnC | move/ltnW]; + by move/(dvdn_exp2l e)/lcmn_idPr. +Qed. + +(* Coprime factors *) + +Definition coprime m n := gcdn m n == 1. + +Lemma coprime1n n : coprime 1 n. +Proof. by rewrite /coprime gcd1n. Qed. + +Lemma coprimen1 n : coprime n 1. +Proof. by rewrite /coprime gcdn1. Qed. + +Lemma coprime_sym m n : coprime m n = coprime n m. +Proof. by rewrite /coprime gcdnC. Qed. + +Lemma coprime_modl m n : coprime (m %% n) n = coprime m n. +Proof. by rewrite /coprime gcdn_modl. Qed. + +Lemma coprime_modr m n : coprime m (n %% m) = coprime m n. +Proof. by rewrite /coprime gcdn_modr. Qed. + +Lemma coprime2n n : coprime 2 n = odd n. +Proof. by rewrite -coprime_modr modn2; case: (odd n). Qed. + +Lemma coprimen2 n : coprime n 2 = odd n. +Proof. by rewrite coprime_sym coprime2n. Qed. + +Lemma coprimeSn n : coprime n.+1 n. +Proof. by rewrite -coprime_modl (modnDr 1) coprime_modl coprime1n. Qed. + +Lemma coprimenS n : coprime n n.+1. +Proof. by rewrite coprime_sym coprimeSn. Qed. + +Lemma coprimePn n : n > 0 -> coprime n.-1 n. +Proof. by case: n => // n _; rewrite coprimenS. Qed. + +Lemma coprimenP n : n > 0 -> coprime n n.-1. +Proof. by case: n => // n _; rewrite coprimeSn. Qed. + +Lemma coprimeP n m : + n > 0 -> reflect (exists u, u.1 * n - u.2 * m = 1) (coprime n m). +Proof. +move=> n_gt0; apply: (iffP eqP) => [<-| [[kn km] /= kn_km_1]]. + by have [kn km kg _] := egcdnP m n_gt0; exists (kn, km); rewrite kg addKn. +apply gcdn_def; rewrite ?dvd1n // => d dv_d_n dv_d_m. +by rewrite -kn_km_1 dvdn_subr ?dvdn_mull // ltnW // -subn_gt0 kn_km_1. +Qed. + +Lemma modn_coprime k n : 0 < k -> (exists u, (k * u) %% n = 1) -> coprime k n. +Proof. +move=> k_gt0 [u Hu]; apply/coprimeP=> //. +by exists (u, k * u %/ n); rewrite /= mulnC {1}(divn_eq (k * u) n) addKn. +Qed. + +Lemma Gauss_dvd m n p : coprime m n -> (m * n %| p) = (m %| p) && (n %| p). +Proof. by move=> co_mn; rewrite -muln_lcm_gcd (eqnP co_mn) muln1 dvdn_lcm. Qed. + +Lemma Gauss_dvdr m n p : coprime m n -> (m %| n * p) = (m %| p). +Proof. +case: n => [|n] co_mn; first by case: m co_mn => [|[]] // _; rewrite !dvd1n. +by symmetry; rewrite mulnC -(@dvdn_pmul2r n.+1) ?Gauss_dvd // andbC dvdn_mull. +Qed. + +Lemma Gauss_dvdl m n p : coprime m p -> (m %| n * p) = (m %| n). +Proof. by rewrite mulnC; apply: Gauss_dvdr. Qed. + +Lemma dvdn_double_leq m n : m %| n -> odd m -> ~~ odd n -> 0 < n -> m.*2 <= n. +Proof. +move=> m_dv_n odd_m even_n n_gt0. +by rewrite -muln2 dvdn_leq // Gauss_dvd ?coprimen2 ?m_dv_n ?dvdn2. +Qed. + +Lemma dvdn_double_ltn m n : m %| n.-1 -> odd m -> odd n -> 1 < n -> m.*2 < n. +Proof. by case: n => //; apply: dvdn_double_leq. Qed. + +Lemma Gauss_gcdr p m n : coprime p m -> gcdn p (m * n) = gcdn p n. +Proof. +move=> co_pm; apply/eqP; rewrite eqn_dvd !dvdn_gcd !dvdn_gcdl /=. +rewrite andbC dvdn_mull ?dvdn_gcdr //= -(@Gauss_dvdr _ m) ?dvdn_gcdr //. +by rewrite /coprime gcdnAC (eqnP co_pm) gcd1n. +Qed. + +Lemma Gauss_gcdl p m n : coprime p n -> gcdn p (m * n) = gcdn p m. +Proof. by move=> co_pn; rewrite mulnC Gauss_gcdr. Qed. + +Lemma coprime_mulr p m n : coprime p (m * n) = coprime p m && coprime p n. +Proof. +case co_pm: (coprime p m) => /=; first by rewrite /coprime Gauss_gcdr. +apply/eqP=> co_p_mn; case/eqnP: co_pm; apply gcdn_def => // d dv_dp dv_dm. +by rewrite -co_p_mn dvdn_gcd dv_dp dvdn_mulr. +Qed. + +Lemma coprime_mull p m n : coprime (m * n) p = coprime m p && coprime n p. +Proof. by rewrite -!(coprime_sym p) coprime_mulr. Qed. + +Lemma coprime_pexpl k m n : 0 < k -> coprime (m ^ k) n = coprime m n. +Proof. +case: k => // k _; elim: k => [|k IHk]; first by rewrite expn1. +by rewrite expnS coprime_mull -IHk; case coprime. +Qed. + +Lemma coprime_pexpr k m n : 0 < k -> coprime m (n ^ k) = coprime m n. +Proof. by move=> k_gt0; rewrite !(coprime_sym m) coprime_pexpl. Qed. + +Lemma coprime_expl k m n : coprime m n -> coprime (m ^ k) n. +Proof. by case: k => [|k] co_pm; rewrite ?coprime1n // coprime_pexpl. Qed. + +Lemma coprime_expr k m n : coprime m n -> coprime m (n ^ k). +Proof. by rewrite !(coprime_sym m); exact: coprime_expl. Qed. + +Lemma coprime_dvdl m n p : m %| n -> coprime n p -> coprime m p. +Proof. by case/dvdnP=> d ->; rewrite coprime_mull => /andP[]. Qed. + +Lemma coprime_dvdr m n p : m %| n -> coprime p n -> coprime p m. +Proof. by rewrite !(coprime_sym p); exact: coprime_dvdl. Qed. + +Lemma coprime_egcdn n m : n > 0 -> coprime (egcdn n m).1 (egcdn n m).2. +Proof. +move=> n_gt0; case: (egcdnP m n_gt0) => kn km /= /eqP. +have [/dvdnP[u defn] /dvdnP[v defm]] := (dvdn_gcdl n m, dvdn_gcdr n m). +rewrite -[gcdn n m]mul1n {1}defm {1}defn !mulnA -mulnDl addnC. +rewrite eqn_pmul2r ?gcdn_gt0 ?n_gt0 //; case: kn => // kn /eqP def_knu _. +by apply/coprimeP=> //; exists (u, v); rewrite mulnC def_knu mulnC addnK. +Qed. + +Lemma dvdn_pexp2r m n k : k > 0 -> (m ^ k %| n ^ k) = (m %| n). +Proof. +move=> k_gt0; apply/idP/idP=> [dv_mn_k|]; last exact: dvdn_exp2r. +case: (posnP n) => [-> | n_gt0]; first by rewrite dvdn0. +have [n' def_n] := dvdnP (dvdn_gcdr m n); set d := gcdn m n in def_n. +have [m' def_m] := dvdnP (dvdn_gcdl m n); rewrite -/d in def_m. +have d_gt0: d > 0 by rewrite gcdn_gt0 n_gt0 orbT. +rewrite def_m def_n !expnMn dvdn_pmul2r ?expn_gt0 ?d_gt0 // in dv_mn_k. +have: coprime (m' ^ k) (n' ^ k). + rewrite coprime_pexpl // coprime_pexpr // /coprime -(eqn_pmul2r d_gt0) mul1n. + by rewrite muln_gcdl -def_m -def_n. +rewrite /coprime -gcdn_modr (eqnP dv_mn_k) gcdn0 -(exp1n k). +by rewrite (inj_eq (expIn k_gt0)) def_m; move/eqP->; rewrite mul1n dvdn_gcdr. +Qed. + +Section Chinese. + +(***********************************************************************) +(* The chinese remainder theorem *) +(***********************************************************************) + +Variables m1 m2 : nat. +Hypothesis co_m12 : coprime m1 m2. + +Lemma chinese_remainder x y : + (x == y %[mod m1 * m2]) = (x == y %[mod m1]) && (x == y %[mod m2]). +Proof. +wlog le_yx : x y / y <= x; last by rewrite !eqn_mod_dvd // Gauss_dvd. +by case/orP: (leq_total y x); last rewrite !(eq_sym (x %% _)); auto. +Qed. + +(***********************************************************************) +(* A function that solves the chinese remainder problem *) +(***********************************************************************) + +Definition chinese r1 r2 := + r1 * m2 * (egcdn m2 m1).1 + r2 * m1 * (egcdn m1 m2).1. + +Lemma chinese_modl r1 r2 : chinese r1 r2 = r1 %[mod m1]. +Proof. +rewrite /chinese; case: (posnP m2) co_m12 => [-> /eqnP | m2_gt0 _]. + by rewrite gcdn0 => ->; rewrite !modn1. +case: egcdnP => // k2 k1 def_m1 _. +rewrite mulnAC -mulnA def_m1 gcdnC (eqnP co_m12) mulnDr mulnA muln1. +by rewrite addnAC (mulnAC _ m1) -mulnDl modnMDl. +Qed. + +Lemma chinese_modr r1 r2 : chinese r1 r2 = r2 %[mod m2]. +Proof. +rewrite /chinese; case: (posnP m1) co_m12 => [-> /eqnP | m1_gt0 _]. + by rewrite gcd0n => ->; rewrite !modn1. +case: (egcdnP m2) => // k1 k2 def_m2 _. +rewrite addnC mulnAC -mulnA def_m2 (eqnP co_m12) mulnDr mulnA muln1. +by rewrite addnAC (mulnAC _ m2) -mulnDl modnMDl. +Qed. + +Lemma chinese_mod x : x = chinese (x %% m1) (x %% m2) %[mod m1 * m2]. +Proof. +apply/eqP; rewrite chinese_remainder //. +by rewrite chinese_modl chinese_modr !modn_mod !eqxx. +Qed. + +End Chinese. diff --git a/mathcomp/discrete/finfun.v b/mathcomp/discrete/finfun.v new file mode 100644 index 0000000..f880260 --- /dev/null +++ b/mathcomp/discrete/finfun.v @@ -0,0 +1,302 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype tuple. + +(******************************************************************************) +(* This file implements a type for functions with a finite domain: *) +(* {ffun aT -> rT} where aT should have a finType structure. *) +(* Any eqType, choiceType, countType and finType structures on rT extend to *) +(* {ffun aT -> rT} as Leibnitz equality and extensional equalities coincide. *) +(* (T ^ n)%type is notation for {ffun 'I_n -> T}, which is isomorphic *) +(* ot n.-tuple T. *) +(* For f : {ffun aT -> rT}, we define *) +(* f x == the image of x under f (f coerces to a CiC function) *) +(* fgraph f == the graph of f, i.e., the #|aT|.-tuple rT of the *) +(* values of f over enum aT. *) +(* finfun lam == the f such that f =1 lam; this is the RECOMMENDED *) +(* interface to build an element of {ffun aT -> rT}. *) +(* [ffun x => expr] == finfun (fun x => expr) *) +(* [ffun => expr] == finfun (fun _ => expr) *) +(* f \in ffun_on R == the range of f is a subset of R *) +(* f \in family F == f belongs to the family F (f x \in F x for all x) *) +(* y.-support f == the y-support of f, i.e., [pred x | f x != y]. *) +(* Thus, y.-support f \subset D means f has y-support D. *) +(* We will put Notation support := 0.-support in ssralg. *) +(* f \in pffun_on y D R == f is a y-partial function from D to R: *) +(* f has y-support D and f x \in R for all x \in D. *) +(* f \in pfamily y D F == f belongs to the y-partial family from D to F: *) +(* f has y-support D and f x \in F x for all x \in D. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Section Def. + +Variables (aT : finType) (rT : Type). + +Inductive finfun_type : predArgType := Finfun of #|aT|.-tuple rT. + +Definition finfun_of of phant (aT -> rT) := finfun_type. + +Identity Coercion type_of_finfun : finfun_of >-> finfun_type. + +Definition fgraph f := let: Finfun t := f in t. + +Canonical finfun_subType := Eval hnf in [newType for fgraph]. + +End Def. + +Notation "{ 'ffun' fT }" := (finfun_of (Phant fT)) + (at level 0, format "{ 'ffun' '[hv' fT ']' }") : type_scope. +Definition finexp_domFinType n := ordinal_finType n. +Notation "T ^ n" := (@finfun_of (finexp_domFinType n) T (Phant _)) : type_scope. + +Notation Local fun_of_fin_def := + (fun aT rT f x => tnth (@fgraph aT rT f) (enum_rank x)). + +Notation Local finfun_def := (fun aT rT f => @Finfun aT rT (codom_tuple f)). + +Module Type FunFinfunSig. +Parameter fun_of_fin : forall aT rT, finfun_type aT rT -> aT -> rT. +Parameter finfun : forall (aT : finType) rT, (aT -> rT) -> {ffun aT -> rT}. +Axiom fun_of_finE : fun_of_fin = fun_of_fin_def. +Axiom finfunE : finfun = finfun_def. +End FunFinfunSig. + +Module FunFinfun : FunFinfunSig. +Definition fun_of_fin := fun_of_fin_def. +Definition finfun := finfun_def. +Lemma fun_of_finE : fun_of_fin = fun_of_fin_def. Proof. by []. Qed. +Lemma finfunE : finfun = finfun_def. Proof. by []. Qed. +End FunFinfun. + +Notation fun_of_fin := FunFinfun.fun_of_fin. +Notation finfun := FunFinfun.finfun. +Coercion fun_of_fin : finfun_type >-> Funclass. +Canonical fun_of_fin_unlock := Unlockable FunFinfun.fun_of_finE. +Canonical finfun_unlock := Unlockable FunFinfun.finfunE. + +Notation "[ 'ffun' x : aT => F ]" := (finfun (fun x : aT => F)) + (at level 0, x ident, only parsing) : fun_scope. + +Notation "[ 'ffun' : aT => F ]" := (finfun (fun _ : aT => F)) + (at level 0, only parsing) : fun_scope. + +Notation "[ 'ffun' x => F ]" := [ffun x : _ => F] + (at level 0, x ident, format "[ 'ffun' x => F ]") : fun_scope. + +Notation "[ 'ffun' => F ]" := [ffun : _ => F] + (at level 0, format "[ 'ffun' => F ]") : fun_scope. + +(* Helper for defining notation for function families. *) +Definition fmem aT rT (pT : predType rT) (f : aT -> pT) := [fun x => mem (f x)]. + +(* Lemmas on the correspondance between finfun_type and CiC functions. *) +Section PlainTheory. + +Variables (aT : finType) (rT : Type). +Notation fT := {ffun aT -> rT}. +Implicit Types (f : fT) (R : pred rT). + +Canonical finfun_of_subType := Eval hnf in [subType of fT]. + +Lemma tnth_fgraph f i : tnth (fgraph f) i = f (enum_val i). +Proof. by rewrite [@fun_of_fin]unlock enum_valK. Qed. + +Lemma ffunE (g : aT -> rT) : finfun g =1 g. +Proof. +move=> x; rewrite [@finfun]unlock unlock tnth_map. +by rewrite -[tnth _ _]enum_val_nth enum_rankK. +Qed. + +Lemma fgraph_codom f : fgraph f = codom_tuple f. +Proof. +apply: eq_from_tnth => i; rewrite [@fun_of_fin]unlock tnth_map. +by congr tnth; rewrite -[tnth _ _]enum_val_nth enum_valK. +Qed. + +Lemma codom_ffun f : codom f = val f. +Proof. by rewrite /= fgraph_codom. Qed. + +Lemma ffunP f1 f2 : f1 =1 f2 <-> f1 = f2. +Proof. +split=> [eq_f12 | -> //]; do 2!apply: val_inj => /=. +by rewrite !fgraph_codom /= (eq_codom eq_f12). +Qed. + +Lemma ffunK : cancel (@fun_of_fin aT rT) (@finfun aT rT). +Proof. by move=> f; apply/ffunP/ffunE. Qed. + +Definition family_mem mF := [pred f : fT | [forall x, in_mem (f x) (mF x)]]. + +Lemma familyP (pT : predType rT) (F : aT -> pT) f : + reflect (forall x, f x \in F x) (f \in family_mem (fmem F)). +Proof. exact: forallP. Qed. + +Definition ffun_on_mem mR := family_mem (fun _ => mR). + +Lemma ffun_onP R f : reflect (forall x, f x \in R) (f \in ffun_on_mem (mem R)). +Proof. exact: forallP. Qed. + +End PlainTheory. + +Notation family F := (family_mem (fun_of_simpl (fmem F))). +Notation ffun_on R := (ffun_on_mem _ (mem R)). + +Implicit Arguments familyP [aT rT pT F f]. +Implicit Arguments ffun_onP [aT rT R f]. + +(*****************************************************************************) + +Lemma nth_fgraph_ord T n (x0 : T) (i : 'I_n) f : nth x0 (fgraph f) i = f i. +Proof. +by rewrite -{2}(enum_rankK i) -tnth_fgraph (tnth_nth x0) enum_rank_ord. +Qed. + +Section Support. + +Variables (aT : Type) (rT : eqType). + +Definition support_for y (f : aT -> rT) := [pred x | f x != y]. + +Lemma supportE x y f : (x \in support_for y f) = (f x != y). Proof. by []. Qed. + +End Support. + +Notation "y .-support" := (support_for y) + (at level 2, format "y .-support") : fun_scope. + +Section EqTheory. + +Variables (aT : finType) (rT : eqType). +Notation fT := {ffun aT -> rT}. +Implicit Types (y : rT) (D : pred aT) (R : pred rT) (f : fT). + +Lemma supportP y D g : + reflect (forall x, x \notin D -> g x = y) (y.-support g \subset D). +Proof. +by apply: (iffP subsetP) => Dg x; [apply: contraNeq | apply: contraR] => /Dg->. +Qed. + +Definition finfun_eqMixin := + Eval hnf in [eqMixin of finfun_type aT rT by <:]. +Canonical finfun_eqType := Eval hnf in EqType _ finfun_eqMixin. +Canonical finfun_of_eqType := Eval hnf in [eqType of fT]. + +Definition pfamily_mem y mD (mF : aT -> mem_pred rT) := + family (fun i : aT => if in_mem i mD then pred_of_simpl (mF i) else pred1 y). + +Lemma pfamilyP (pT : predType rT) y D (F : aT -> pT) f : + reflect (y.-support f \subset D /\ {in D, forall x, f x \in F x}) + (f \in pfamily_mem y (mem D) (fmem F)). +Proof. +apply: (iffP familyP) => [/= f_pfam | [/supportP f_supp f_fam] x]. + split=> [|x Ax]; last by have:= f_pfam x; rewrite Ax. + by apply/subsetP=> x; case: ifP (f_pfam x) => //= _ fx0 /negP[]. +by case: ifPn => Ax /=; rewrite inE /= (f_fam, f_supp). +Qed. + +Definition pffun_on_mem y mD mR := pfamily_mem y mD (fun _ => mR). + +Lemma pffun_onP y D R f : + reflect (y.-support f \subset D /\ {subset image f D <= R}) + (f \in pffun_on_mem y (mem D) (mem R)). +Proof. +apply: (iffP (pfamilyP y D (fun _ => R) f)) => [] [-> f_fam]; split=> //. + by move=> _ /imageP[x Ax ->]; exact: f_fam. +by move=> x Ax; apply: f_fam; apply/imageP; exists x. +Qed. + +End EqTheory. +Canonical exp_eqType (T : eqType) n := [eqType of T ^ n]. + +Implicit Arguments supportP [aT rT y D g]. +Notation pfamily y D F := (pfamily_mem y (mem D) (fun_of_simpl (fmem F))). +Notation pffun_on y D R := (pffun_on_mem y (mem D) (mem R)). + +Definition finfun_choiceMixin aT (rT : choiceType) := + [choiceMixin of finfun_type aT rT by <:]. +Canonical finfun_choiceType aT rT := + Eval hnf in ChoiceType _ (finfun_choiceMixin aT rT). +Canonical finfun_of_choiceType (aT : finType) (rT : choiceType) := + Eval hnf in [choiceType of {ffun aT -> rT}]. +Canonical exp_choiceType (T : choiceType) n := [choiceType of T ^ n]. + +Definition finfun_countMixin aT (rT : countType) := + [countMixin of finfun_type aT rT by <:]. +Canonical finfun_countType aT (rT : countType) := + Eval hnf in CountType _ (finfun_countMixin aT rT). +Canonical finfun_of_countType (aT : finType) (rT : countType) := + Eval hnf in [countType of {ffun aT -> rT}]. +Canonical finfun_subCountType aT (rT : countType) := + Eval hnf in [subCountType of finfun_type aT rT]. +Canonical finfun_of_subCountType (aT : finType) (rT : countType) := + Eval hnf in [subCountType of {ffun aT -> rT}]. + +(*****************************************************************************) + +Section FinTheory. + +Variables aT rT : finType. +Notation fT := {ffun aT -> rT}. +Notation ffT := (finfun_type aT rT). +Implicit Types (D : pred aT) (R : pred rT) (F : aT -> pred rT). + +Definition finfun_finMixin := [finMixin of ffT by <:]. +Canonical finfun_finType := Eval hnf in FinType ffT finfun_finMixin. +Canonical finfun_subFinType := Eval hnf in [subFinType of ffT]. +Canonical finfun_of_finType := Eval hnf in [finType of fT for finfun_finType]. +Canonical finfun_of_subFinType := Eval hnf in [subFinType of fT]. + +Lemma card_pfamily y0 D F : + #|pfamily y0 D F| = foldr muln 1 [seq #|F x| | x in D]. +Proof. +rewrite /image_mem; transitivity #|pfamily y0 (enum D) F|. + by apply/eq_card=> f; apply/eq_forallb=> x /=; rewrite mem_enum. +elim: {D}(enum D) (enum_uniq D) => /= [_|x0 s IHs /andP[s'x0 /IHs<-{IHs}]]. + apply: eq_card1 [ffun=> y0] _ _ => f. + apply/familyP/eqP=> [y0_f|-> x]; last by rewrite ffunE inE. + by apply/ffunP=> x; rewrite ffunE (eqP (y0_f x)). +pose g (xf : rT * fT) := finfun [eta xf.2 with x0 |-> xf.1]. +have gK: cancel (fun f : fT => (f x0, g (y0, f))) g. + by move=> f; apply/ffunP=> x; do !rewrite ffunE /=; case: eqP => // ->. +rewrite -cardX -(card_image (can_inj gK)); apply: eq_card => [] [y f] /=. +apply/imageP/andP=> [[f0 /familyP/=Ff0] [{f}-> ->]| [Fy /familyP/=Ff]]. + split; first by have:= Ff0 x0; rewrite /= mem_head. + apply/familyP=> x; have:= Ff0 x; rewrite ffunE inE /=. + by case: eqP => //= -> _; rewrite ifN ?inE. +exists (g (y, f)). + by apply/familyP=> x; have:= Ff x; rewrite ffunE /= inE; case: eqP => // ->. +congr (_, _); last apply/ffunP=> x; do !rewrite ffunE /= ?eqxx //. +by case: eqP => // ->{x}; apply/eqP; have:= Ff x0; rewrite ifN. +Qed. + +Lemma card_family F : #|family F| = foldr muln 1 [seq #|F x| | x : aT]. +Proof. +have [y0 _ | rT0] := pickP rT; first exact: (card_pfamily y0 aT). +rewrite /image_mem; case DaT: (enum aT) => [{rT0}|x0 e] /=; last first. + by rewrite !eq_card0 // => [f | y]; [have:= rT0 (f x0) | have:= rT0 y]. +have{DaT} no_aT P (x : aT) : P by have:= mem_enum aT x; rewrite DaT. +apply: eq_card1 [ffun x => no_aT rT x] _ _ => f. +by apply/familyP/eqP=> _; [apply/ffunP | ] => x; apply: no_aT. +Qed. + +Lemma card_pffun_on y0 D R : #|pffun_on y0 D R| = #|R| ^ #|D|. +Proof. +rewrite (cardE D) card_pfamily /image_mem. +by elim: (enum D) => //= _ e ->; rewrite expnS. +Qed. + +Lemma card_ffun_on R : #|ffun_on R| = #|R| ^ #|aT|. +Proof. +rewrite card_family /image_mem cardT. +by elim: (enum aT) => //= _ e ->; rewrite expnS. +Qed. + +Lemma card_ffun : #|fT| = #|rT| ^ #|aT|. +Proof. by rewrite -card_ffun_on; apply/esym/eq_card=> f; apply/forallP. Qed. + +End FinTheory. +Canonical exp_finType (T : finType) n := [finType of T ^ n]. + diff --git a/mathcomp/discrete/fingraph.v b/mathcomp/discrete/fingraph.v new file mode 100644 index 0000000..403d1ca --- /dev/null +++ b/mathcomp/discrete/fingraph.v @@ -0,0 +1,721 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path fintype. + +(******************************************************************************) +(* This file develops the theory of finite graphs represented by an "edge" *) +(* relation over a finType T; this mainly amounts to the theory of the *) +(* transitive closure of such relations. *) +(* For g : T -> seq T, e : rel T and f : T -> T we define: *) +(* grel g == the adjacency relation y \in g x of the graph g. *) +(* rgraph e == the graph (x |-> enum (e x)) of the relation e. *) +(* dfs g n v x == the list of points traversed by a depth-first search of *) +(* the g, at depth n, starting from x, and avoiding v. *) +(* dfs_path g v x y <-> there is a path from x to y in g \ v. *) +(* connect e == the transitive closure of e (computed by dfs). *) +(* connect_sym e <-> connect e is symmetric, hence an equivalence relation. *) +(* root e x == a representative of connect e x, which is the component *) +(* of x in the transitive closure of e. *) +(* roots e == the codomain predicate of root e. *) +(* n_comp e a == the number of e-connected components of a, when a is *) +(* e-closed and connect e is symmetric. *) +(* equivalence classes of connect e if connect_sym e holds. *) +(* closed e a == the collective predicate a is e-invariant. *) +(* closure e a == the e-closure of a (the image of a under connect e). *) +(* rel_adjunction h e e' a <-> in the e-closed domain a, h is the left part *) +(* of an adjunction from e to another relation e'. *) +(* fconnect f == connect (frel f), i.e., "connected under f iteration". *) +(* froot f x == root (frel f) x, the root of the orbit of x under f. *) +(* froots f == roots (frel f) == orbit representatives for f. *) +(* orbit f x == lists the f-orbit of x. *) +(* findex f x y == index of y in the f-orbit of x. *) +(* order f x == size (cardinal) of the f-orbit of x. *) +(* order_set f n == elements of f-order n. *) +(* finv f == the inverse of f, if f is injective. *) +(* := finv f x := iter (order x).-1 f x. *) +(* fcard f a == number of orbits of f in a, provided a is f-invariant *) +(* f is one-to-one. *) +(* fclosed f a == the collective predicate a is f-invariant. *) +(* fclosure f a == the closure of a under f iteration. *) +(* fun_adjunction == rel_adjunction (frel f). *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Definition grel (T : eqType) (g : T -> seq T) := [rel x y | y \in g x]. + +(* Decidable connectivity in finite types. *) +Section Connect. + +Variable T : finType. + +Section Dfs. + +Variable g : T -> seq T. +Implicit Type v w a : seq T. + +Fixpoint dfs n v x := + if x \in v then v else + if n is n'.+1 then foldl (dfs n') (x :: v) (g x) else v. + +Lemma subset_dfs n v a : v \subset foldl (dfs n) v a. +Proof. +elim: n a v => [|n IHn]; first by elim=> //= *; rewrite if_same. +elim=> //= x a IHa v; apply: subset_trans {IHa}(IHa _); case: ifP => // _. +by apply: subset_trans (IHn _ _); apply/subsetP=> y; exact: predU1r. +Qed. + +Inductive dfs_path v x y : Prop := + DfsPath p of path (grel g) x p & y = last x p & [disjoint x :: p & v]. + +Lemma dfs_pathP n x y v : + #|T| <= #|v| + n -> y \notin v -> reflect (dfs_path v x y) (y \in dfs n v x). +Proof. +have dfs_id w z: z \notin w -> dfs_path w z z. + by exists [::]; rewrite ?disjoint_has //= orbF. +elim: n => [|n IHn] /= in x y v * => le_v'_n not_vy. + rewrite addn0 (geq_leqif (subset_leqif_card (subset_predT _))) in le_v'_n. + by rewrite predT_subset in not_vy. +have [v_x | not_vx] := ifPn. + by rewrite (negPf not_vy); right=> [] [p _ _]; rewrite disjoint_has /= v_x. +set v1 := x :: v; set a := g x; have sub_dfs := subsetP (subset_dfs n _ _). +have [-> | neq_yx] := eqVneq y x. + by rewrite sub_dfs ?mem_head //; left; exact: dfs_id. +apply: (@equivP (exists2 x1, x1 \in a & dfs_path v1 x1 y)); last first. + split=> {IHn} [[x1 a_x1 [p g_p p_y]] | [p /shortenP[]]]. + rewrite disjoint_has has_sym /= has_sym /= => /norP[_ not_pv]. + by exists (x1 :: p); rewrite /= ?a_x1 // disjoint_has negb_or not_vx. + case=> [_ _ _ eq_yx | x1 p1 /=]; first by case/eqP: neq_yx. + case/andP=> a_x1 g_p1 /andP[not_p1x _] /subsetP p_p1 p1y not_pv. + exists x1 => //; exists p1 => //. + rewrite disjoint_sym disjoint_cons not_p1x disjoint_sym. + by move: not_pv; rewrite disjoint_cons => /andP[_ /disjoint_trans->]. +have{neq_yx not_vy}: y \notin v1 by exact/norP. +have{le_v'_n not_vx}: #|T| <= #|v1| + n by rewrite cardU1 not_vx addSnnS. +elim: {x v}a v1 => [|x a IHa] v /= le_v'_n not_vy. + by rewrite (negPf not_vy); right=> [] []. +set v2 := dfs n v x; have v2v: v \subset v2 := subset_dfs n v [:: x]. +have [v2y | not_v2y] := boolP (y \in v2). + by rewrite sub_dfs //; left; exists x; [exact: mem_head | exact: IHn]. +apply: {IHa}(equivP (IHa _ _ not_v2y)). + by rewrite (leq_trans le_v'_n) // leq_add2r subset_leq_card. +split=> [] [x1 a_x1 [p g_p p_y not_pv]]. + exists x1; [exact: predU1r | exists p => //]. + by rewrite disjoint_sym (disjoint_trans v2v) // disjoint_sym. +suffices not_p1v2: [disjoint x1 :: p & v2]. + case/predU1P: a_x1 => [def_x1 | ]; last by exists x1; last exists p. + case/pred0Pn: not_p1v2; exists x; rewrite /= def_x1 mem_head /=. + suffices not_vx: x \notin v by apply/IHn; last exact: dfs_id. + by move: not_pv; rewrite disjoint_cons def_x1 => /andP[]. +apply: contraR not_v2y => /pred0Pn[x2 /andP[/= p_x2 v2x2]]. +case/splitPl: p_x2 p_y g_p not_pv => p0 p2 p0x2. +rewrite last_cat cat_path -cat_cons lastI cat_rcons {}p0x2 => p2y /andP[_ g_p2]. +rewrite disjoint_cat disjoint_cons => /and3P[{p0}_ not_vx2 not_p2v]. +have{not_vx2 v2x2} [p1 g_p1 p1_x2 not_p1v] := IHn _ _ v le_v'_n not_vx2 v2x2. +apply/IHn=> //; exists (p1 ++ p2); rewrite ?cat_path ?last_cat -?p1_x2 ?g_p1 //. +by rewrite -cat_cons disjoint_cat not_p1v. +Qed. + +Lemma dfsP x y : + reflect (exists2 p, path (grel g) x p & y = last x p) (y \in dfs #|T| [::] x). +Proof. +apply: (iffP (dfs_pathP _ _ _)); rewrite ?card0 // => [] [p]; exists p => //. +by rewrite disjoint_sym disjoint0. +Qed. + +End Dfs. + +Variable e : rel T. + +Definition rgraph x := enum (e x). + +Lemma rgraphK : grel rgraph =2 e. +Proof. by move=> x y; rewrite /= mem_enum. Qed. + +Definition connect : rel T := fun x y => y \in dfs rgraph #|T| [::] x. +Canonical connect_app_pred x := ApplicativePred (connect x). + +Lemma connectP x y : + reflect (exists2 p, path e x p & y = last x p) (connect x y). +Proof. +apply: (equivP (dfsP _ x y)). +by split=> [] [p e_p ->]; exists p => //; rewrite (eq_path rgraphK) in e_p *. +Qed. + +Lemma connect_trans : transitive connect. +Proof. +move=> x y z /connectP[p e_p ->] /connectP[q e_q ->]; apply/connectP. +by exists (p ++ q); rewrite ?cat_path ?e_p ?last_cat. +Qed. + +Lemma connect0 x : connect x x. +Proof. by apply/connectP; exists [::]. Qed. + +Lemma eq_connect0 x y : x = y -> connect x y. +Proof. move->; exact: connect0. Qed. + +Lemma connect1 x y : e x y -> connect x y. +Proof. by move=> e_xy; apply/connectP; exists [:: y]; rewrite /= ?e_xy. Qed. + +Lemma path_connect x p : path e x p -> subpred (mem (x :: p)) (connect x). +Proof. +move=> e_p y p_y; case/splitPl: p / p_y e_p => p q <-. +by rewrite cat_path => /andP[e_p _]; apply/connectP; exists p. +Qed. + +Definition root x := odflt x (pick (connect x)). + +Definition roots : pred T := fun x => root x == x. +Canonical roots_pred := ApplicativePred roots. + +Definition n_comp_mem (m_a : mem_pred T) := #|predI roots m_a|. + +Lemma connect_root x : connect x (root x). +Proof. by rewrite /root; case: pickP; rewrite ?connect0. Qed. + +Definition connect_sym := symmetric connect. + +Hypothesis sym_e : connect_sym. + +Lemma same_connect : left_transitive connect. +Proof. exact: sym_left_transitive connect_trans. Qed. + +Lemma same_connect_r : right_transitive connect. +Proof. exact: sym_right_transitive connect_trans. Qed. + +Lemma same_connect1 x y : e x y -> connect x =1 connect y. +Proof. by move/connect1; exact: same_connect. Qed. + +Lemma same_connect1r x y : e x y -> connect^~ x =1 connect^~ y. +Proof. by move/connect1; exact: same_connect_r. Qed. + +Lemma rootP x y : reflect (root x = root y) (connect x y). +Proof. +apply: (iffP idP) => e_xy. + by rewrite /root -(eq_pick (same_connect e_xy)); case: pickP e_xy => // ->. +by apply: (connect_trans (connect_root x)); rewrite e_xy sym_e connect_root. +Qed. + +Lemma root_root x : root (root x) = root x. +Proof. exact/esym/rootP/connect_root. Qed. + +Lemma roots_root x : roots (root x). +Proof. exact/eqP/root_root. Qed. + +Lemma root_connect x y : (root x == root y) = connect x y. +Proof. exact: sameP eqP (rootP x y). Qed. + +Definition closed_mem m_a := forall x y, e x y -> in_mem x m_a = in_mem y m_a. + +Definition closure_mem m_a : pred T := + fun x => ~~ disjoint (mem (connect x)) m_a. + +End Connect. + +Hint Resolve connect0. + +Notation n_comp e a := (n_comp_mem e (mem a)). +Notation closed e a := (closed_mem e (mem a)). +Notation closure e a := (closure_mem e (mem a)). + +Prenex Implicits connect root roots. + +Implicit Arguments dfsP [T g x y]. +Implicit Arguments connectP [T e x y]. +Implicit Arguments rootP [T e x y]. + +Notation fconnect f := (connect (coerced_frel f)). +Notation froot f := (root (coerced_frel f)). +Notation froots f := (roots (coerced_frel f)). +Notation fcard_mem f := (n_comp_mem (coerced_frel f)). +Notation fcard f a := (fcard_mem f (mem a)). +Notation fclosed f a := (closed (coerced_frel f) a). +Notation fclosure f a := (closure (coerced_frel f) a). + +Section EqConnect. + +Variable T : finType. +Implicit Types (e : rel T) (a : pred T). + +Lemma connect_sub e e' : + subrel e (connect e') -> subrel (connect e) (connect e'). +Proof. +move=> e'e x _ /connectP[p e_p ->]; elim: p x e_p => //= y p IHp x /andP[exy]. +by move/IHp; apply: connect_trans; exact: e'e. +Qed. + +Lemma relU_sym e e' : + connect_sym e -> connect_sym e' -> connect_sym (relU e e'). +Proof. +move=> sym_e sym_e'; apply: symmetric_from_pre => x _ /connectP[p e_p ->]. +elim: p x e_p => //= y p IHp x /andP[e_xy /IHp{IHp}/connect_trans]; apply. +case/orP: e_xy => /connect1; rewrite (sym_e, sym_e'); + by apply: connect_sub y x => x y e_xy; rewrite connect1 //= e_xy ?orbT. +Qed. + +Lemma eq_connect e e' : e =2 e' -> connect e =2 connect e'. +Proof. +move=> eq_e x y; apply/connectP/connectP=> [] [p e_p ->]; + by exists p; rewrite // (eq_path eq_e) in e_p *. +Qed. + +Lemma eq_n_comp e e' : connect e =2 connect e' -> n_comp_mem e =1 n_comp_mem e'. +Proof. +move=> eq_e [a]; apply: eq_card => x /=. +by rewrite !inE /= /roots /root /= (eq_pick (eq_e x)). +Qed. + +Lemma eq_n_comp_r {e} a a' : a =i a' -> n_comp e a = n_comp e a'. +Proof. by move=> eq_a; apply: eq_card => x; rewrite inE /= eq_a. Qed. + +Lemma n_compC a e : n_comp e T = n_comp e a + n_comp e [predC a]. +Proof. +rewrite /n_comp_mem (eq_card (fun _ => andbT _)) -(cardID a); congr (_ + _). +by apply: eq_card => x; rewrite !inE andbC. +Qed. + +Lemma eq_root e e' : e =2 e' -> root e =1 root e'. +Proof. by move=> eq_e x; rewrite /root (eq_pick (eq_connect eq_e x)). Qed. + +Lemma eq_roots e e' : e =2 e' -> roots e =1 roots e'. +Proof. by move=> eq_e x; rewrite /roots (eq_root eq_e). Qed. + +End EqConnect. + +Section Closure. + +Variables (T : finType) (e : rel T). +Hypothesis sym_e : connect_sym e. +Implicit Type a : pred T. + +Lemma same_connect_rev : connect e =2 connect (fun x y => e y x). +Proof. +suff crev e': subrel (connect (fun x : T => e'^~ x)) (fun x => (connect e')^~x). + by move=> x y; rewrite sym_e; apply/idP/idP; exact: crev. +move=> x y /connectP[p e_p p_y]; apply/connectP. +exists (rev (belast x p)); first by rewrite p_y rev_path. +by rewrite -(last_cons x) -rev_rcons p_y -lastI rev_cons last_rcons. +Qed. + +Lemma intro_closed a : (forall x y, e x y -> x \in a -> y \in a) -> closed e a. +Proof. +move=> cl_a x y e_xy; apply/idP/idP=> [|a_y]; first exact: cl_a. +have{x e_xy} /connectP[p e_p ->]: connect e y x by rewrite sym_e connect1. +by elim: p y a_y e_p => //= y p IHp x a_x /andP[/cl_a/(_ a_x)]; exact: IHp. +Qed. + +Lemma closed_connect a : + closed e a -> forall x y, connect e x y -> (x \in a) = (y \in a). +Proof. +move=> cl_a x _ /connectP[p e_p ->]. +by elim: p x e_p => //= y p IHp x /andP[/cl_a->]; exact: IHp. +Qed. + +Lemma connect_closed x : closed e (connect e x). +Proof. by move=> y z /connect1/same_connect_r; exact. Qed. + +Lemma predC_closed a : closed e a -> closed e [predC a]. +Proof. by move=> cl_a x y /cl_a; rewrite !inE => ->. Qed. + +Lemma closure_closed a : closed e (closure e a). +Proof. +apply: intro_closed => x y /connect1 e_xy; congr (~~ _). +by apply: eq_disjoint; exact: same_connect. +Qed. + +Lemma mem_closure a : {subset a <= closure e a}. +Proof. by move=> x a_x; apply/existsP; exists x; rewrite !inE connect0. Qed. + +Lemma subset_closure a : a \subset closure e a. +Proof. by apply/subsetP; exact: mem_closure. Qed. + +Lemma n_comp_closure2 x y : + n_comp e (closure e (pred2 x y)) = (~~ connect e x y).+1. +Proof. +rewrite -(root_connect sym_e) -card2; apply: eq_card => z. +apply/idP/idP=> [/andP[/eqP {2}<- /pred0Pn[t /andP[/= ezt exyt]]] |]. + by case/pred2P: exyt => <-; rewrite (rootP sym_e ezt) !inE eqxx ?orbT. +by case/pred2P=> ->; rewrite !inE roots_root //; apply/existsP; + [exists x | exists y]; rewrite !inE eqxx ?orbT sym_e connect_root. +Qed. + +Lemma n_comp_connect x : n_comp e (connect e x) = 1. +Proof. +rewrite -(card1 (root e x)); apply: eq_card => y. +apply/andP/eqP => [[/eqP r_y /rootP-> //] | ->] /=. +by rewrite inE connect_root roots_root. +Qed. + +End Closure. + +Section Orbit. + +Variables (T : finType) (f : T -> T). + +Definition order x := #|fconnect f x|. + +Definition orbit x := traject f x (order x). + +Definition findex x y := index y (orbit x). + +Definition finv x := iter (order x).-1 f x. + +Lemma fconnect_iter n x : fconnect f x (iter n f x). +Proof. +apply/connectP. +by exists (traject f (f x) n); [ exact: fpath_traject | rewrite last_traject ]. +Qed. + +Lemma fconnect1 x : fconnect f x (f x). +Proof. exact: (fconnect_iter 1). Qed. + +Lemma fconnect_finv x : fconnect f x (finv x). +Proof. exact: fconnect_iter. Qed. + +Lemma orderSpred x : (order x).-1.+1 = order x. +Proof. by rewrite /order (cardD1 x) [_ x _]connect0. Qed. + +Lemma size_orbit x : size (orbit x) = order x. +Proof. exact: size_traject. Qed. + +Lemma looping_order x : looping f x (order x). +Proof. +apply: contraFT (ltnn (order x)); rewrite -looping_uniq => /card_uniqP. +rewrite size_traject => <-; apply: subset_leq_card. +by apply/subsetP=> _ /trajectP[i _ ->]; exact: fconnect_iter. +Qed. + +Lemma fconnect_orbit x y : fconnect f x y = (y \in orbit x). +Proof. +apply/idP/idP=> [/connectP[_ /fpathP[m ->] ->] | /trajectP[i _ ->]]. + by rewrite last_traject; exact/loopingP/looping_order. +exact: fconnect_iter. +Qed. + +Lemma orbit_uniq x : uniq (orbit x). +Proof. +rewrite /orbit -orderSpred looping_uniq; set n := (order x).-1. +apply: contraFN (ltnn n) => /trajectP[i lt_i_n eq_fnx_fix]. +rewrite {1}/n orderSpred /order -(size_traject f x n). +apply: (leq_trans (subset_leq_card _) (card_size _)); apply/subsetP=> z. +rewrite inE fconnect_orbit => /trajectP[j le_jn ->{z}]. +rewrite -orderSpred -/n ltnS leq_eqVlt in le_jn. +by apply/trajectP; case/predU1P: le_jn => [->|]; [exists i | exists j]. +Qed. + +Lemma findex_max x y : fconnect f x y -> findex x y < order x. +Proof. by rewrite [_ y]fconnect_orbit -index_mem size_orbit. Qed. + +Lemma findex_iter x i : i < order x -> findex x (iter i f x) = i. +Proof. +move=> lt_ix; rewrite -(nth_traject f lt_ix) /findex index_uniq ?orbit_uniq //. +by rewrite size_orbit. +Qed. + +Lemma iter_findex x y : fconnect f x y -> iter (findex x y) f x = y. +Proof. +rewrite [_ y]fconnect_orbit => fxy; pose i := index y (orbit x). +have lt_ix: i < order x by rewrite -size_orbit index_mem. +by rewrite -(nth_traject f lt_ix) nth_index. +Qed. + +Lemma findex0 x : findex x x = 0. +Proof. by rewrite /findex /orbit -orderSpred /= eqxx. Qed. + +Lemma fconnect_invariant (T' : eqType) (k : T -> T') : + invariant f k =1 xpredT -> forall x y, fconnect f x y -> k x = k y. +Proof. +move=> eq_k_f x y /iter_findex <-; elim: {y}(findex x y) => //= n ->. +by rewrite (eqP (eq_k_f _)). +Qed. + +Section Loop. + +Variable p : seq T. +Hypotheses (f_p : fcycle f p) (Up : uniq p). +Variable x : T. +Hypothesis p_x : x \in p. + +(* This lemma does not depend on Up : (uniq p) *) +Lemma fconnect_cycle y : fconnect f x y = (y \in p). +Proof. +have [i q def_p] := rot_to p_x; rewrite -(mem_rot i p) def_p. +have{i def_p} /andP[/eqP q_x f_q]: (f (last x q) == x) && fpath f x q. + by have:= f_p; rewrite -(rot_cycle i) def_p (cycle_path x). +apply/idP/idP=> [/connectP[_ /fpathP[j ->] ->] | ]; last exact: path_connect. +case/fpathP: f_q q_x => n ->; rewrite !last_traject -iterS => def_x. +by apply: (@loopingP _ f x n.+1); rewrite /looping def_x /= mem_head. +Qed. + +Lemma order_cycle : order x = size p. +Proof. by rewrite -(card_uniqP Up); exact (eq_card fconnect_cycle). Qed. + +Lemma orbit_rot_cycle : {i : nat | orbit x = rot i p}. +Proof. +have [i q def_p] := rot_to p_x; exists i. +rewrite /orbit order_cycle -(size_rot i) def_p. +suffices /fpathP[j ->]: fpath f x q by rewrite /= size_traject. +by move: f_p; rewrite -(rot_cycle i) def_p (cycle_path x); case/andP. +Qed. + +End Loop. + +Hypothesis injf : injective f. + +Lemma f_finv : cancel finv f. +Proof. +move=> x; move: (looping_order x) (orbit_uniq x). +rewrite /looping /orbit -orderSpred looping_uniq /= /looping; set n := _.-1. +case/predU1P=> // /trajectP[i lt_i_n]; rewrite -iterSr => /= /injf ->. +by case/trajectP; exists i. +Qed. + +Lemma finv_f : cancel f finv. +Proof. exact (inj_can_sym f_finv injf). Qed. + +Lemma fin_inj_bij : bijective f. +Proof. exists finv; [ exact finv_f | exact f_finv ]. Qed. + +Lemma finv_bij : bijective finv. +Proof. exists f; [ exact f_finv | exact finv_f ]. Qed. + +Lemma finv_inj : injective finv. +Proof. exact (can_inj f_finv). Qed. + +Lemma fconnect_sym x y : fconnect f x y = fconnect f y x. +Proof. +suff{x y} Sf x y: fconnect f x y -> fconnect f y x by apply/idP/idP; auto. +case/connectP=> p f_p -> {y}; elim: p x f_p => //= y p IHp x. +rewrite -{2}(finv_f x) => /andP[/eqP-> /IHp/connect_trans-> //]. +exact: fconnect_finv. +Qed. +Let symf := fconnect_sym. + +Lemma iter_order x : iter (order x) f x = x. +Proof. by rewrite -orderSpred iterS; exact (f_finv x). Qed. + +Lemma iter_finv n x : n <= order x -> iter n finv x = iter (order x - n) f x. +Proof. +rewrite -{2}[x]iter_order => /subnKC {1}<-; move: (_ - n) => m. +by rewrite iter_add; elim: n => // n {2}<-; rewrite iterSr /= finv_f. +Qed. + +Lemma cycle_orbit x : fcycle f (orbit x). +Proof. +rewrite /orbit -orderSpred (cycle_path x) /= last_traject -/(finv x). +by rewrite fpath_traject f_finv andbT /=. +Qed. + +Lemma fpath_finv x p : fpath finv x p = fpath f (last x p) (rev (belast x p)). +Proof. +elim: p x => //= y p IHp x; rewrite rev_cons rcons_path -{}IHp andbC /=. +rewrite (canF_eq finv_f) eq_sym; congr (_ && (_ == _)). +by case: p => //= z p; rewrite rev_cons last_rcons. +Qed. + +Lemma same_fconnect_finv : fconnect finv =2 fconnect f. +Proof. +move=> x y; rewrite (same_connect_rev symf); apply: {x y}eq_connect => x y /=. +by rewrite (canF_eq finv_f) eq_sym. +Qed. + +Lemma fcard_finv : fcard_mem finv =1 fcard_mem f. +Proof. exact: eq_n_comp same_fconnect_finv. Qed. + +Definition order_set n : pred T := [pred x | order x == n]. + +Lemma fcard_order_set n (a : pred T) : + a \subset order_set n -> fclosed f a -> fcard f a * n = #|a|. +Proof. +move=> a_n cl_a; rewrite /n_comp_mem; set b := [predI froots f & a]. +symmetry; transitivity #|preim (froot f) b|. + apply: eq_card => x; rewrite !inE (roots_root fconnect_sym). + by rewrite -(closed_connect cl_a (connect_root _ x)). +have{cl_a a_n} (x): b x -> froot f x = x /\ order x = n. + by case/andP=> /eqP-> /(subsetP a_n)/eqnP->. +elim: {a b}#|b| {1 3 4}b (eqxx #|b|) => [|m IHm] b def_m f_b. + by rewrite eq_card0 // => x; exact: (pred0P def_m). +have [x b_x | b0] := pickP b; last by rewrite (eq_card0 b0) in def_m. +have [r_x ox_n] := f_b x b_x; rewrite (cardD1 x) [x \in b]b_x eqSS in def_m. +rewrite mulSn -{1}ox_n -(IHm _ def_m) => [|_ /andP[_ /f_b //]]. +rewrite -(cardID (fconnect f x)); congr (_ + _); apply: eq_card => y. + by apply: andb_idl => /= fxy; rewrite !inE -(rootP symf fxy) r_x. +by congr (~~ _ && _); rewrite /= /in_mem /= symf -(root_connect symf) r_x. +Qed. + +Lemma fclosed1 (a : pred T) : fclosed f a -> forall x, (x \in a) = (f x \in a). +Proof. by move=> cl_a x; exact: cl_a (eqxx _). Qed. + +Lemma same_fconnect1 x : fconnect f x =1 fconnect f (f x). +Proof. by apply: same_connect1 => /=. Qed. + +Lemma same_fconnect1_r x y : fconnect f x y = fconnect f x (f y). +Proof. by apply: same_connect1r x => /=. Qed. + +End Orbit. + +Prenex Implicits order orbit findex finv order_set. + +Section FconnectId. + +Variable T : finType. + +Lemma fconnect_id (x : T) : fconnect id x =1 xpred1 x. +Proof. by move=> y; rewrite (@fconnect_cycle _ _ [:: x]) //= ?inE ?eqxx. Qed. + +Lemma order_id (x : T) : order id x = 1. +Proof. by rewrite /order (eq_card (fconnect_id x)) card1. Qed. + +Lemma orbit_id (x : T) : orbit id x = [:: x]. +Proof. by rewrite /orbit order_id. Qed. + +Lemma froots_id (x : T) : froots id x. +Proof. by rewrite /roots -fconnect_id connect_root. Qed. + +Lemma froot_id (x : T) : froot id x = x. +Proof. by apply/eqP; exact: froots_id. Qed. + +Lemma fcard_id (a : pred T) : fcard id a = #|a|. +Proof. by apply: eq_card => x; rewrite inE froots_id. Qed. + +End FconnectId. + +Section FconnectEq. + +Variables (T : finType) (f f' : T -> T). + +Lemma finv_eq_can : cancel f f' -> finv f =1 f'. +Proof. +move=> fK; exact: (bij_can_eq (fin_inj_bij (can_inj fK)) (finv_f (can_inj fK))). +Qed. + +Hypothesis eq_f : f =1 f'. +Let eq_rf := eq_frel eq_f. + +Lemma eq_fconnect : fconnect f =2 fconnect f'. +Proof. exact: eq_connect eq_rf. Qed. + +Lemma eq_fcard : fcard_mem f =1 fcard_mem f'. +Proof. exact: eq_n_comp eq_fconnect. Qed. + +Lemma eq_finv : finv f =1 finv f'. +Proof. +by move=> x; rewrite /finv /order (eq_card (eq_fconnect x)) (eq_iter eq_f). +Qed. + +Lemma eq_froot : froot f =1 froot f'. +Proof. exact: eq_root eq_rf. Qed. + +Lemma eq_froots : froots f =1 froots f'. +Proof. exact: eq_roots eq_rf. Qed. + +End FconnectEq. + +Section FinvEq. + +Variables (T : finType) (f : T -> T). +Hypothesis injf : injective f. + +Lemma finv_inv : finv (finv f) =1 f. +Proof. exact: (finv_eq_can (f_finv injf)). Qed. + +Lemma order_finv : order (finv f) =1 order f. +Proof. by move=> x; exact: eq_card (same_fconnect_finv injf x). Qed. + +Lemma order_set_finv n : order_set (finv f) n =i order_set f n. +Proof. by move=> x; rewrite !inE order_finv. Qed. + +End FinvEq. + +Section RelAdjunction. + +Variables (T T' : finType) (h : T' -> T) (e : rel T) (e' : rel T'). +Hypotheses (sym_e : connect_sym e) (sym_e' : connect_sym e'). + +Record rel_adjunction_mem m_a := RelAdjunction { + rel_unit x : in_mem x m_a -> {x' : T' | connect e x (h x')}; + rel_functor x' y' : + in_mem (h x') m_a -> connect e' x' y' = connect e (h x') (h y') +}. + +Variable a : pred T. +Hypothesis cl_a : closed e a. + +Local Notation rel_adjunction := (rel_adjunction_mem (mem a)). + +Lemma intro_adjunction (h' : forall x, x \in a -> T') : + (forall x a_x, + [/\ connect e x (h (h' x a_x)) + & forall y a_y, e x y -> connect e' (h' x a_x) (h' y a_y)]) -> + (forall x' a_x, + [/\ connect e' x' (h' (h x') a_x) + & forall y', e' x' y' -> connect e (h x') (h y')]) -> + rel_adjunction. +Proof. +move=> Aee' Ae'e; split=> [y a_y | x' z' a_x]. + by exists (h' y a_y); case/Aee': (a_y). +apply/idP/idP=> [/connectP[p e'p ->{z'}] | /connectP[p e_p p_z']]. + elim: p x' a_x e'p => //= y' p IHp x' a_x. + case: (Ae'e x' a_x) => _ Ae'x /andP[/Ae'x e_xy /IHp e_yz] {Ae'x}. + by apply: connect_trans (e_yz _); rewrite // -(closed_connect cl_a e_xy). +case: (Ae'e x' a_x) => /connect_trans-> //. +elim: p {x'}(h x') p_z' a_x e_p => /= [|y p IHp] x p_z' a_x. + by rewrite -p_z' in a_x *; case: (Ae'e _ a_x); rewrite sym_e'. +case/andP=> e_xy /(IHp _ p_z') e'yz; have a_y: y \in a by rewrite -(cl_a e_xy). +by apply: connect_trans (e'yz a_y); case: (Aee' _ a_x) => _ ->. +Qed. + +Lemma strict_adjunction : + injective h -> a \subset codom h -> rel_base h e e' [predC a] -> + rel_adjunction. +Proof. +move=> /= injh h_a a_ee'; pose h' x Hx := iinv (subsetP h_a x Hx). +apply: (@intro_adjunction h') => [x a_x | x' a_x]. + rewrite f_iinv connect0; split=> // y a_y e_xy. + by rewrite connect1 // -a_ee' !f_iinv ?negbK. +rewrite [h' _ _]iinv_f //; split=> // y' e'xy. +by rewrite connect1 // a_ee' ?negbK. +Qed. + +Let ccl_a := closed_connect cl_a. + +Lemma adjunction_closed : rel_adjunction -> closed e' [preim h of a]. +Proof. +case=> _ Ae'e; apply: intro_closed => // x' y' /connect1 e'xy a_x. +by rewrite Ae'e // in e'xy; rewrite !inE -(ccl_a e'xy). +Qed. + +Lemma adjunction_n_comp : + rel_adjunction -> n_comp e a = n_comp e' [preim h of a]. +Proof. +case=> Aee' Ae'e. +have inj_h: {in predI (roots e') [preim h of a] &, injective (root e \o h)}. + move=> x' y' /andP[/eqP r_x' /= a_x'] /andP[/eqP r_y' _] /(rootP sym_e). + by rewrite -Ae'e // => /(rootP sym_e'); rewrite r_x' r_y'. +rewrite /n_comp_mem -(card_in_image inj_h); apply: eq_card => x. +apply/andP/imageP=> [[/eqP rx a_x] | [x' /andP[/eqP r_x' a_x'] ->]]; last first. + by rewrite /= -(ccl_a (connect_root _ _)) roots_root. +have [y' e_xy]:= Aee' x a_x; pose x' := root e' y'. +have ay': h y' \in a by rewrite -(ccl_a e_xy). +have e_yx: connect e (h y') (h x') by rewrite -Ae'e ?connect_root. +exists x'; first by rewrite inE /= -(ccl_a e_yx) ?roots_root. +by rewrite /= -(rootP sym_e e_yx) -(rootP sym_e e_xy). +Qed. + +End RelAdjunction. + +Notation rel_adjunction h e e' a := (rel_adjunction_mem h e e' (mem a)). +Notation "@ 'rel_adjunction' T T' h e e' a" := + (@rel_adjunction_mem T T' h e e' (mem a)) + (at level 10, T, T', h, e, e', a at level 8, only parsing) : type_scope. +Notation fun_adjunction h f f' a := (rel_adjunction h (frel f) (frel f') a). +Notation "@ 'fun_adjunction' T T' h f f' a" := + (@rel_adjunction T T' h (frel f) (frel f') a) + (at level 10, T, T', h, f, f', a at level 8, only parsing) : type_scope. + +Implicit Arguments intro_adjunction [T T' h e e' a]. +Implicit Arguments adjunction_n_comp [T T' e e' a]. + +Unset Implicit Arguments. + diff --git a/mathcomp/discrete/finset.v b/mathcomp/discrete/finset.v new file mode 100644 index 0000000..54dacc0 --- /dev/null +++ b/mathcomp/discrete/finset.v @@ -0,0 +1,2214 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrbool ssrfun eqtype ssrnat div seq choice fintype. +Require Import finfun bigop. + +(******************************************************************************) +(* This file defines a type for sets over a finite Type, similar to the type *) +(* of functions over a finite Type defined in finfun.v (indeed, based in it): *) +(* {set T} where T must have a finType structure *) +(* We equip {set T} itself with a finType structure, hence Leibnitz and *) +(* extensional equalities coincide on {set T}, and we can form {set {set T}} *) +(* If A, B : {set T} and P : {set {set T}}, we define: *) +(* x \in A == x belongs to A (i.e., {set T} implements predType, *) +(* by coercion to pred_sort). *) +(* mem A == the predicate corresponding to A. *) +(* finset p == the set corresponding to a predicate p. *) +(* [set x | P] == the set containing the x such that P is true (x may *) +(* appear in P). *) +(* [set x | P & Q] := [set x | P && Q]. *) +(* [set x in A] == the set containing the x in a collective predicate A. *) +(* [set x in A | P] == the set containing the x in A such that P is true. *) +(* [set x in A | P & Q] := [set x in A | P && Q]. *) +(* All these have typed variants [set x : T | P], [set x : T in A], etc. *) +(* set0 == the empty set. *) +(* [set: T] or setT == the full set (the A containing all x : T). *) +(* A :|: B == the union of A and B. *) +(* x |: A == A with the element x added (:= [set x] :| A). *) +(* A :&: B == the intersection of A and B. *) +(* ~: A == the complement of A. *) +(* A :\: B == the difference A minus B. *) +(* A :\ x == A with the element x removed (:= A :\: [set x]). *) +(* \bigcup_<range> A == the union of all A, for i in <range> (i is bound in *) +(* A, see bigop.v). *) +(* \bigcap_<range> A == the intersection of all A, for i in <range>. *) +(* cover P == the union of the set of sets P. *) +(* trivIset P <=> the elements of P are pairwise disjoint. *) +(* partition P A <=> P is a partition of A. *) +(* pblock P x == a block of P containing x, or else set0. *) +(* equivalence_partition R D == the partition induced on D by the relation R *) +(* (provided R is an equivalence relation in D). *) +(* preim_partition f D == the partition induced on D by the equivalence *) +(* [rel x y | f x == f y]. *) +(* is_transversal X P D <=> X is a transversal of the partition P of D. *) +(* transversal P D == a transversal of P, provided P is a partition of D. *) +(* transversal_repr x0 X B == a representative of B \in P selected by the *) +(* tranversal X of P, or else x0. *) +(* powerset A == the set of all subset of the set A. *) +(* P ::&: A == those sets in P that are subsets of the set A. *) +(* f @^-1: A == the preimage of the collective predicate A under f. *) +(* f @: A == the image set of the collective predicate A by f. *) +(* f @2:(A, B) == the image set of A x B by the binary function f. *) +(* [set E | x in A] == the set of all the values of the expression E, for x *) +(* drawn from the collective predicate A. *) +(* [set E | x in A & P] == the set of values of E for x drawn from A, such *) +(* that P is true. *) +(* [set E | x in A, y in B] == the set of values of E for x drawn from A and *) +(* and y drawn from B; B may depend on x. *) +(* [set E | x <- A, y <- B & P] == the set of values of E for x drawn from A *) +(* y drawn from B, such that P is trye. *) +(* [set E | x : T] == the set of all values of E, with x in type T. *) +(* [set E | x : T & P] == the set of values of E for x : T s.t. P is true. *) +(* [set E | x : T, y : U in B], [set E | x : T, y : U in B & P], *) +(* [set E | x : T in A, y : U], [set E | x : T in A, y : U & P], *) +(* [set E | x : T, y : U], [set E | x : T, y : U & P] *) +(* == type-ranging versions of the binary comprehensions. *) +(* [set E | x : T in A], [set E | x in A, y], [set E | x, y & P], etc. *) +(* == typed and untyped variants of the comprehensions above. *) +(* The types may be required as type inference processes E *) +(* before considering A or B. Note that type casts in the *) +(* binary comprehension must either be both present or absent *) +(* and that there are no untyped variants for single-type *) +(* comprehension as Coq parsing confuses [x | P] and [E | x]. *) +(* minset p A == A is a minimal set satisfying p. *) +(* maxset p A == A is a maximal set satisfying p. *) +(* We also provide notations A :=: B, A :<>: B, A :==: B, A :!=: B, A :=P: B *) +(* that specialize A = B, A <> B, A == B, etc., to {set _}. This is useful *) +(* for subtypes of {set T}, such as {group T}, that coerce to {set T}. *) +(* We give many lemmas on these operations, on card, and on set inclusion. *) +(* In addition to the standard suffixes described in ssrbool.v, we associate *) +(* the following suffixes to set operations: *) +(* 0 -- the empty set, as in in_set0 : (x \in set0) = false. *) +(* T -- the full set, as in in_setT : x \in [set: T]. *) +(* 1 -- a singleton set, as in in_set1 : (x \in [set a]) = (x == a). *) +(* 2 -- an unordered pair, as in *) +(* in_set2 : (x \in [set a; b]) = (x == a) || (x == b). *) +(* C -- complement, as in setCK : ~: ~: A = A. *) +(* I -- intersection, as in setIid : A :&: A = A. *) +(* U -- union, as in setUid : A :|: A = A. *) +(* D -- difference, as in setDv : A :\: A = set0. *) +(* S -- a subset argument, as in *) +(* setIS: B \subset C -> A :&: B \subset A :&: C *) +(* These suffixes are sometimes preceded with an `s' to distinguish them from *) +(* their basic ssrbool interpretation, e.g., *) +(* card1 : #|pred1 x| = 1 and cards1 : #|[set x]| = 1 *) +(* We also use a trailling `r' to distinguish a right-hand complement from *) +(* commutativity, e.g., *) +(* setIC : A :&: B = B :&: A and setICr : A :&: ~: A = set0. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. +Section SetType. + +Variable T : finType. + +Inductive set_type : predArgType := FinSet of {ffun pred T}. +Definition finfun_of_set A := let: FinSet f := A in f. +Definition set_of of phant T := set_type. +Identity Coercion type_of_set_of : set_of >-> set_type. + +Canonical set_subType := Eval hnf in [newType for finfun_of_set]. +Definition set_eqMixin := Eval hnf in [eqMixin of set_type by <:]. +Canonical set_eqType := Eval hnf in EqType set_type set_eqMixin. +Definition set_choiceMixin := [choiceMixin of set_type by <:]. +Canonical set_choiceType := Eval hnf in ChoiceType set_type set_choiceMixin. +Definition set_countMixin := [countMixin of set_type by <:]. +Canonical set_countType := Eval hnf in CountType set_type set_countMixin. +Canonical set_subCountType := Eval hnf in [subCountType of set_type]. +Definition set_finMixin := [finMixin of set_type by <:]. +Canonical set_finType := Eval hnf in FinType set_type set_finMixin. +Canonical set_subFinType := Eval hnf in [subFinType of set_type]. + +End SetType. + +Delimit Scope set_scope with SET. +Bind Scope set_scope with set_type. +Bind Scope set_scope with set_of. +Open Scope set_scope. +Arguments Scope finfun_of_set [_ set_scope]. + +Notation "{ 'set' T }" := (set_of (Phant T)) + (at level 0, format "{ 'set' T }") : type_scope. + +(* We later define several subtypes that coerce to set; for these it is *) +(* preferable to state equalities at the {set _} level, even when comparing *) +(* subtype values, because the primitive "injection" tactic tends to diverge *) +(* on complex types (e.g., quotient groups). We provide some parse-only *) +(* notation to make this technicality less obstrusive. *) +Notation "A :=: B" := (A = B :> {set _}) + (at level 70, no associativity, only parsing) : set_scope. +Notation "A :<>: B" := (A <> B :> {set _}) + (at level 70, no associativity, only parsing) : set_scope. +Notation "A :==: B" := (A == B :> {set _}) + (at level 70, no associativity, only parsing) : set_scope. +Notation "A :!=: B" := (A != B :> {set _}) + (at level 70, no associativity, only parsing) : set_scope. +Notation "A :=P: B" := (A =P B :> {set _}) + (at level 70, no associativity, only parsing) : set_scope. + +Notation Local finset_def := (fun T P => @FinSet T (finfun P)). + +Notation Local pred_of_set_def := (fun T (A : set_type T) => val A : _ -> _). + +Module Type SetDefSig. +Parameter finset : forall T : finType, pred T -> {set T}. +Parameter pred_of_set : forall T, set_type T -> fin_pred_sort (predPredType T). +(* The weird type of pred_of_set is imposed by the syntactic restrictions on *) +(* coercion declarations; it is unfortunately not possible to use a functor *) +(* to retype the declaration, because this triggers an ugly bug in the Coq *) +(* coercion chaining code. *) +Axiom finsetE : finset = finset_def. +Axiom pred_of_setE : pred_of_set = pred_of_set_def. +End SetDefSig. + +Module SetDef : SetDefSig. +Definition finset := finset_def. +Definition pred_of_set := pred_of_set_def. +Lemma finsetE : finset = finset_def. Proof. by []. Qed. +Lemma pred_of_setE : pred_of_set = pred_of_set_def. Proof. by []. Qed. +End SetDef. + +Notation finset := SetDef.finset. +Notation pred_of_set := SetDef.pred_of_set. +Canonical finset_unlock := Unlockable SetDef.finsetE. +Canonical pred_of_set_unlock := Unlockable SetDef.pred_of_setE. + +Notation "[ 'set' x : T | P ]" := (finset (fun x : T => P%B)) + (at level 0, x at level 99, only parsing) : set_scope. +Notation "[ 'set' x | P ]" := [set x : _ | P] + (at level 0, x, P at level 99, format "[ 'set' x | P ]") : set_scope. +Notation "[ 'set' x 'in' A ]" := [set x | x \in A] + (at level 0, x at level 99, format "[ 'set' x 'in' A ]") : set_scope. +Notation "[ 'set' x : T 'in' A ]" := [set x : T | x \in A] + (at level 0, x at level 99, only parsing) : set_scope. +Notation "[ 'set' x : T | P & Q ]" := [set x : T | P && Q] + (at level 0, x at level 99, only parsing) : set_scope. +Notation "[ 'set' x | P & Q ]" := [set x | P && Q ] + (at level 0, x, P at level 99, format "[ 'set' x | P & Q ]") : set_scope. +Notation "[ 'set' x : T 'in' A | P ]" := [set x : T | x \in A & P] + (at level 0, x at level 99, only parsing) : set_scope. +Notation "[ 'set' x 'in' A | P ]" := [set x | x \in A & P] + (at level 0, x at level 99, format "[ 'set' x 'in' A | P ]") : set_scope. +Notation "[ 'set' x 'in' A | P & Q ]" := [set x in A | P && Q] + (at level 0, x at level 99, + format "[ 'set' x 'in' A | P & Q ]") : set_scope. +Notation "[ 'set' x : T 'in' A | P & Q ]" := [set x : T in A | P && Q] + (at level 0, x at level 99, only parsing) : set_scope. + +(* This lets us use set and subtypes of set, like group or coset_of, both as *) +(* collective predicates and as arguments of the \pi(_) notation. *) +Coercion pred_of_set: set_type >-> fin_pred_sort. + +(* Declare pred_of_set as a canonical instance of topred, but use the *) +(* coercion to resolve mem A to @mem (predPredType T) (pred_of_set A). *) +Canonical set_predType T := + Eval hnf in @mkPredType _ (unkeyed (set_type T)) (@pred_of_set T). + +Section BasicSetTheory. + +Variable T : finType. +Implicit Types (x : T) (A B : {set T}) (pA : pred T). + +Canonical set_of_subType := Eval hnf in [subType of {set T}]. +Canonical set_of_eqType := Eval hnf in [eqType of {set T}]. +Canonical set_of_choiceType := Eval hnf in [choiceType of {set T}]. +Canonical set_of_countType := Eval hnf in [countType of {set T}]. +Canonical set_of_subCountType := Eval hnf in [subCountType of {set T}]. +Canonical set_of_finType := Eval hnf in [finType of {set T}]. +Canonical set_of_subFinType := Eval hnf in [subFinType of {set T}]. + +Lemma in_set pA x : x \in finset pA = pA x. +Proof. by rewrite [@finset]unlock unlock [x \in _]ffunE. Qed. + +Lemma setP A B : A =i B <-> A = B. +Proof. +by split=> [eqAB|-> //]; apply/val_inj/ffunP=> x; have:= eqAB x; rewrite unlock. +Qed. + +Definition set0 := [set x : T | false]. +Definition setTfor (phT : phant T) := [set x : T | true]. + +Lemma in_setT x : x \in setTfor (Phant T). +Proof. by rewrite in_set. Qed. + +Lemma eqsVneq A B : {A = B} + {A != B}. +Proof. exact: eqVneq. Qed. + +End BasicSetTheory. + +Definition inE := (in_set, inE). + +Implicit Arguments set0 [T]. +Prenex Implicits set0. +Hint Resolve in_setT. + +Notation "[ 'set' : T ]" := (setTfor (Phant T)) + (at level 0, format "[ 'set' : T ]") : set_scope. + +Notation setT := [set: _] (only parsing). + +Section setOpsDefs. + +Variable T : finType. +Implicit Types (a x : T) (A B D : {set T}) (P : {set {set T}}). + +Definition set1 a := [set x | x == a]. +Definition setU A B := [set x | (x \in A) || (x \in B)]. +Definition setI A B := [set x in A | x \in B]. +Definition setC A := [set x | x \notin A]. +Definition setD A B := [set x | x \notin B & x \in A]. +Definition ssetI P D := [set A in P | A \subset D]. +Definition powerset D := [set A : {set T} | A \subset D]. + +End setOpsDefs. + +Notation "[ 'set' a ]" := (set1 a) + (at level 0, a at level 99, format "[ 'set' a ]") : set_scope. +Notation "[ 'set' a : T ]" := [set (a : T)] + (at level 0, a at level 99, format "[ 'set' a : T ]") : set_scope. +Notation "A :|: B" := (setU A B) : set_scope. +Notation "a |: A" := ([set a] :|: A) : set_scope. +(* This is left-associative due to historical limitations of the .. Notation. *) +Notation "[ 'set' a1 ; a2 ; .. ; an ]" := (setU .. (a1 |: [set a2]) .. [set an]) + (at level 0, a1 at level 99, + format "[ 'set' a1 ; a2 ; .. ; an ]") : set_scope. +Notation "A :&: B" := (setI A B) : set_scope. +Notation "~: A" := (setC A) (at level 35, right associativity) : set_scope. +Notation "[ 'set' ~ a ]" := (~: [set a]) + (at level 0, format "[ 'set' ~ a ]") : set_scope. +Notation "A :\: B" := (setD A B) : set_scope. +Notation "A :\ a" := (A :\: [set a]) : set_scope. +Notation "P ::&: D" := (ssetI P D) (at level 48) : set_scope. + +Section setOps. + +Variable T : finType. +Implicit Types (a x : T) (A B C D : {set T}) (pA pB pC : pred T). + +Lemma eqEsubset A B : (A == B) = (A \subset B) && (B \subset A). +Proof. by apply/eqP/subset_eqP=> /setP. Qed. + +Lemma subEproper A B : A \subset B = (A == B) || (A \proper B). +Proof. by rewrite eqEsubset -andb_orr orbN andbT. Qed. + +Lemma eqVproper A B : A \subset B -> A = B \/ A \proper B. +Proof. by rewrite subEproper => /predU1P. Qed. + +Lemma properEneq A B : A \proper B = (A != B) && (A \subset B). +Proof. by rewrite andbC eqEsubset negb_and andb_orr andbN. Qed. + +Lemma proper_neq A B : A \proper B -> A != B. +Proof. by rewrite properEneq; case/andP. Qed. + +Lemma eqEproper A B : (A == B) = (A \subset B) && ~~ (A \proper B). +Proof. by rewrite negb_and negbK andb_orr andbN eqEsubset. Qed. + +Lemma eqEcard A B : (A == B) = (A \subset B) && (#|B| <= #|A|). +Proof. +rewrite eqEsubset; apply: andb_id2l => sAB. +by rewrite (geq_leqif (subset_leqif_card sAB)). +Qed. + +Lemma properEcard A B : (A \proper B) = (A \subset B) && (#|A| < #|B|). +Proof. by rewrite properEneq ltnNge andbC eqEcard; case: (A \subset B). Qed. + +Lemma subset_leqif_cards A B : A \subset B -> (#|A| <= #|B| ?= iff (A == B)). +Proof. by move=> sAB; rewrite eqEsubset sAB; exact: subset_leqif_card. Qed. + +Lemma in_set0 x : x \in set0 = false. +Proof. by rewrite inE. Qed. + +Lemma sub0set A : set0 \subset A. +Proof. by apply/subsetP=> x; rewrite inE. Qed. + +Lemma subset0 A : (A \subset set0) = (A == set0). +Proof. by rewrite eqEsubset sub0set andbT. Qed. + +Lemma proper0 A : (set0 \proper A) = (A != set0). +Proof. by rewrite properE sub0set subset0. Qed. + +Lemma subset_neq0 A B : A \subset B -> A != set0 -> B != set0. +Proof. by rewrite -!proper0 => sAB /proper_sub_trans->. Qed. + +Lemma set_0Vmem A : (A = set0) + {x : T | x \in A}. +Proof. +case: (pickP (mem A)) => [x Ax | A0]; [by right; exists x | left]. +apply/setP=> x; rewrite inE; exact: A0. +Qed. + +Lemma enum_set0 : enum set0 = [::] :> seq T. +Proof. by rewrite (eq_enum (in_set _)) enum0. Qed. + +Lemma subsetT A : A \subset setT. +Proof. by apply/subsetP=> x; rewrite inE. Qed. + +Lemma subsetT_hint mA : subset mA (mem [set: T]). +Proof. by rewrite unlock; apply/pred0P=> x; rewrite !inE. Qed. +Hint Resolve subsetT_hint. + +Lemma subTset A : (setT \subset A) = (A == setT). +Proof. by rewrite eqEsubset subsetT. Qed. + +Lemma properT A : (A \proper setT) = (A != setT). +Proof. by rewrite properEneq subsetT andbT. Qed. + +Lemma set1P x a : reflect (x = a) (x \in [set a]). +Proof. by rewrite inE; exact: eqP. Qed. + +Lemma enum_setT : enum [set: T] = Finite.enum T. +Proof. by rewrite (eq_enum (in_set _)) enumT. Qed. + +Lemma in_set1 x a : (x \in [set a]) = (x == a). +Proof. exact: in_set. Qed. + +Lemma set11 x : x \in [set x]. +Proof. by rewrite inE. Qed. + +Lemma set1_inj : injective (@set1 T). +Proof. by move=> a b eqsab; apply/set1P; rewrite -eqsab set11. Qed. + +Lemma enum_set1 a : enum [set a] = [:: a]. +Proof. by rewrite (eq_enum (in_set _)) enum1. Qed. + +Lemma setU1P x a B : reflect (x = a \/ x \in B) (x \in a |: B). +Proof. by rewrite !inE; exact: predU1P. Qed. + +Lemma in_setU1 x a B : (x \in a |: B) = (x == a) || (x \in B). +Proof. by rewrite !inE. Qed. + +Lemma set_cons a s : [set x in a :: s] = a |: [set x in s]. +Proof. by apply/setP=> x; rewrite !inE. Qed. + +Lemma setU11 x B : x \in x |: B. +Proof. by rewrite !inE eqxx. Qed. + +Lemma setU1r x a B : x \in B -> x \in a |: B. +Proof. by move=> Bx; rewrite !inE predU1r. Qed. + +(* We need separate lemmas for the explicit enumerations since they *) +(* associate on the left. *) +Lemma set1Ul x A b : x \in A -> x \in A :|: [set b]. +Proof. by move=> Ax; rewrite !inE Ax. Qed. + +Lemma set1Ur A b : b \in A :|: [set b]. +Proof. by rewrite !inE eqxx orbT. Qed. + +Lemma in_setC1 x a : (x \in [set~ a]) = (x != a). +Proof. by rewrite !inE. Qed. + +Lemma setC11 x : (x \in [set~ x]) = false. +Proof. by rewrite !inE eqxx. Qed. + +Lemma setD1P x A b : reflect (x != b /\ x \in A) (x \in A :\ b). +Proof. rewrite !inE; exact: andP. Qed. + +Lemma in_setD1 x A b : (x \in A :\ b) = (x != b) && (x \in A) . +Proof. by rewrite !inE. Qed. + +Lemma setD11 b A : (b \in A :\ b) = false. +Proof. by rewrite !inE eqxx. Qed. + +Lemma setD1K a A : a \in A -> a |: (A :\ a) = A. +Proof. by move=> Aa; apply/setP=> x; rewrite !inE; case: eqP => // ->. Qed. + +Lemma setU1K a B : a \notin B -> (a |: B) :\ a = B. +Proof. +by move/negPf=> nBa; apply/setP=> x; rewrite !inE; case: eqP => // ->. +Qed. + +Lemma set2P x a b : reflect (x = a \/ x = b) (x \in [set a; b]). +Proof. rewrite !inE; exact: pred2P. Qed. + +Lemma in_set2 x a b : (x \in [set a; b]) = (x == a) || (x == b). +Proof. by rewrite !inE. Qed. + +Lemma set21 a b : a \in [set a; b]. +Proof. by rewrite !inE eqxx. Qed. + +Lemma set22 a b : b \in [set a; b]. +Proof. by rewrite !inE eqxx orbT. Qed. + +Lemma setUP x A B : reflect (x \in A \/ x \in B) (x \in A :|: B). +Proof. by rewrite !inE; exact: orP. Qed. + +Lemma in_setU x A B : (x \in A :|: B) = (x \in A) || (x \in B). +Proof. exact: in_set. Qed. + +Lemma setUC A B : A :|: B = B :|: A. +Proof. by apply/setP => x; rewrite !inE orbC. Qed. + +Lemma setUS A B C : A \subset B -> C :|: A \subset C :|: B. +Proof. +move=> sAB; apply/subsetP=> x; rewrite !inE. +by case: (x \in C) => //; exact: (subsetP sAB). +Qed. + +Lemma setSU A B C : A \subset B -> A :|: C \subset B :|: C. +Proof. by move=> sAB; rewrite -!(setUC C) setUS. Qed. + +Lemma setUSS A B C D : A \subset C -> B \subset D -> A :|: B \subset C :|: D. +Proof. by move=> /(setSU B) /subset_trans sAC /(setUS C)/sAC. Qed. + +Lemma set0U A : set0 :|: A = A. +Proof. by apply/setP => x; rewrite !inE orFb. Qed. + +Lemma setU0 A : A :|: set0 = A. +Proof. by rewrite setUC set0U. Qed. + +Lemma setUA A B C : A :|: (B :|: C) = A :|: B :|: C. +Proof. by apply/setP => x; rewrite !inE orbA. Qed. + +Lemma setUCA A B C : A :|: (B :|: C) = B :|: (A :|: C). +Proof. by rewrite !setUA (setUC B). Qed. + +Lemma setUAC A B C : A :|: B :|: C = A :|: C :|: B. +Proof. by rewrite -!setUA (setUC B). Qed. + +Lemma setUACA A B C D : (A :|: B) :|: (C :|: D) = (A :|: C) :|: (B :|: D). +Proof. by rewrite -!setUA (setUCA B). Qed. + +Lemma setTU A : setT :|: A = setT. +Proof. by apply/setP => x; rewrite !inE orTb. Qed. + +Lemma setUT A : A :|: setT = setT. +Proof. by rewrite setUC setTU. Qed. + +Lemma setUid A : A :|: A = A. +Proof. by apply/setP=> x; rewrite inE orbb. Qed. + +Lemma setUUl A B C : A :|: B :|: C = (A :|: C) :|: (B :|: C). +Proof. by rewrite setUA !(setUAC _ C) -(setUA _ C) setUid. Qed. + +Lemma setUUr A B C : A :|: (B :|: C) = (A :|: B) :|: (A :|: C). +Proof. by rewrite !(setUC A) setUUl. Qed. + +(* intersection *) + +(* setIdP is a generalisation of setIP that applies to comprehensions. *) +Lemma setIdP x pA pB : reflect (pA x /\ pB x) (x \in [set y | pA y & pB y]). +Proof. by rewrite !inE; exact: andP. Qed. + +Lemma setId2P x pA pB pC : + reflect [/\ pA x, pB x & pC x] (x \in [set y | pA y & pB y && pC y]). +Proof. by rewrite !inE; exact: and3P. Qed. + +Lemma setIdE A pB : [set x in A | pB x] = A :&: [set x | pB x]. +Proof. by apply/setP=> x; rewrite !inE. Qed. + +Lemma setIP x A B : reflect (x \in A /\ x \in B) (x \in A :&: B). +Proof. exact: (iffP (@setIdP _ _ _)). Qed. + +Lemma in_setI x A B : (x \in A :&: B) = (x \in A) && (x \in B). +Proof. exact: in_set. Qed. + +Lemma setIC A B : A :&: B = B :&: A. +Proof. by apply/setP => x; rewrite !inE andbC. Qed. + +Lemma setIS A B C : A \subset B -> C :&: A \subset C :&: B. +Proof. +move=> sAB; apply/subsetP=> x; rewrite !inE. +by case: (x \in C) => //; exact: (subsetP sAB). +Qed. + +Lemma setSI A B C : A \subset B -> A :&: C \subset B :&: C. +Proof. by move=> sAB; rewrite -!(setIC C) setIS. Qed. + +Lemma setISS A B C D : A \subset C -> B \subset D -> A :&: B \subset C :&: D. +Proof. by move=> /(setSI B) /subset_trans sAC /(setIS C) /sAC. Qed. + +Lemma setTI A : setT :&: A = A. +Proof. by apply/setP => x; rewrite !inE andTb. Qed. + +Lemma setIT A : A :&: setT = A. +Proof. by rewrite setIC setTI. Qed. + +Lemma set0I A : set0 :&: A = set0. +Proof. by apply/setP => x; rewrite !inE andFb. Qed. + +Lemma setI0 A : A :&: set0 = set0. + +Proof. by rewrite setIC set0I. Qed. + +Lemma setIA A B C : A :&: (B :&: C) = A :&: B :&: C. +Proof. by apply/setP=> x; rewrite !inE andbA. Qed. + +Lemma setICA A B C : A :&: (B :&: C) = B :&: (A :&: C). +Proof. by rewrite !setIA (setIC A). Qed. + +Lemma setIAC A B C : A :&: B :&: C = A :&: C :&: B. +Proof. by rewrite -!setIA (setIC B). Qed. + +Lemma setIACA A B C D : (A :&: B) :&: (C :&: D) = (A :&: C) :&: (B :&: D). +Proof. by rewrite -!setIA (setICA B). Qed. + +Lemma setIid A : A :&: A = A. +Proof. by apply/setP=> x; rewrite inE andbb. Qed. + +Lemma setIIl A B C : A :&: B :&: C = (A :&: C) :&: (B :&: C). +Proof. by rewrite setIA !(setIAC _ C) -(setIA _ C) setIid. Qed. + +Lemma setIIr A B C : A :&: (B :&: C) = (A :&: B) :&: (A :&: C). +Proof. by rewrite !(setIC A) setIIl. Qed. + +(* distribute /cancel *) + +Lemma setIUr A B C : A :&: (B :|: C) = (A :&: B) :|: (A :&: C). +Proof. by apply/setP=> x; rewrite !inE andb_orr. Qed. + +Lemma setIUl A B C : (A :|: B) :&: C = (A :&: C) :|: (B :&: C). +Proof. by apply/setP=> x; rewrite !inE andb_orl. Qed. + +Lemma setUIr A B C : A :|: (B :&: C) = (A :|: B) :&: (A :|: C). +Proof. by apply/setP=> x; rewrite !inE orb_andr. Qed. + +Lemma setUIl A B C : (A :&: B) :|: C = (A :|: C) :&: (B :|: C). +Proof. by apply/setP=> x; rewrite !inE orb_andl. Qed. + +Lemma setUK A B : (A :|: B) :&: A = A. +Proof. by apply/setP=> x; rewrite !inE orbK. Qed. + +Lemma setKU A B : A :&: (B :|: A) = A. +Proof. by apply/setP=> x; rewrite !inE orKb. Qed. + +Lemma setIK A B : (A :&: B) :|: A = A. +Proof. by apply/setP=> x; rewrite !inE andbK. Qed. + +Lemma setKI A B : A :|: (B :&: A) = A. +Proof. by apply/setP=> x; rewrite !inE andKb. Qed. + +(* complement *) + +Lemma setCP x A : reflect (~ x \in A) (x \in ~: A). +Proof. by rewrite !inE; exact: negP. Qed. + +Lemma in_setC x A : (x \in ~: A) = (x \notin A). +Proof. exact: in_set. Qed. + +Lemma setCK : involutive (@setC T). +Proof. by move=> A; apply/setP=> x; rewrite !inE negbK. Qed. + +Lemma setC_inj : injective (@setC T). +Proof. exact: can_inj setCK. Qed. + +Lemma subsets_disjoint A B : (A \subset B) = [disjoint A & ~: B]. +Proof. by rewrite subset_disjoint; apply: eq_disjoint_r => x; rewrite !inE. Qed. + +Lemma disjoints_subset A B : [disjoint A & B] = (A \subset ~: B). +Proof. by rewrite subsets_disjoint setCK. Qed. + +Lemma powersetCE A B : (A \in powerset (~: B)) = [disjoint A & B]. +Proof. by rewrite inE disjoints_subset. Qed. + +Lemma setCS A B : (~: A \subset ~: B) = (B \subset A). +Proof. by rewrite !subsets_disjoint setCK disjoint_sym. Qed. + +Lemma setCU A B : ~: (A :|: B) = ~: A :&: ~: B. +Proof. by apply/setP=> x; rewrite !inE negb_or. Qed. + +Lemma setCI A B : ~: (A :&: B) = ~: A :|: ~: B. +Proof. by apply/setP=> x; rewrite !inE negb_and. Qed. + +Lemma setUCr A : A :|: ~: A = setT. +Proof. by apply/setP=> x; rewrite !inE orbN. Qed. + +Lemma setICr A : A :&: ~: A = set0. +Proof. by apply/setP=> x; rewrite !inE andbN. Qed. + +Lemma setC0 : ~: set0 = [set: T]. +Proof. by apply/setP=> x; rewrite !inE. Qed. + +Lemma setCT : ~: [set: T] = set0. +Proof. by rewrite -setC0 setCK. Qed. + +(* difference *) + +Lemma setDP A B x : reflect (x \in A /\ x \notin B) (x \in A :\: B). +Proof. by rewrite inE andbC; exact: andP. Qed. + +Lemma in_setD A B x : (x \in A :\: B) = (x \notin B) && (x \in A). +Proof. exact: in_set. Qed. + +Lemma setDE A B : A :\: B = A :&: ~: B. +Proof. by apply/setP => x; rewrite !inE andbC. Qed. + +Lemma setSD A B C : A \subset B -> A :\: C \subset B :\: C. +Proof. by rewrite !setDE; exact: setSI. Qed. + +Lemma setDS A B C : A \subset B -> C :\: B \subset C :\: A. +Proof. by rewrite !setDE -setCS; exact: setIS. Qed. + +Lemma setDSS A B C D : A \subset C -> D \subset B -> A :\: B \subset C :\: D. +Proof. by move=> /(setSD B) /subset_trans sAC /(setDS C) /sAC. Qed. + +Lemma setD0 A : A :\: set0 = A. +Proof. by apply/setP=> x; rewrite !inE. Qed. + +Lemma set0D A : set0 :\: A = set0. +Proof. by apply/setP=> x; rewrite !inE andbF. Qed. + +Lemma setDT A : A :\: setT = set0. +Proof. by apply/setP=> x; rewrite !inE. Qed. + +Lemma setTD A : setT :\: A = ~: A. +Proof. by apply/setP=> x; rewrite !inE andbT. Qed. + +Lemma setDv A : A :\: A = set0. +Proof. by apply/setP=> x; rewrite !inE andNb. Qed. + +Lemma setCD A B : ~: (A :\: B) = ~: A :|: B. +Proof. by rewrite !setDE setCI setCK. Qed. + +Lemma setID A B : A :&: B :|: A :\: B = A. +Proof. by rewrite setDE -setIUr setUCr setIT. Qed. + +Lemma setDUl A B C : (A :|: B) :\: C = (A :\: C) :|: (B :\: C). +Proof. by rewrite !setDE setIUl. Qed. + +Lemma setDUr A B C : A :\: (B :|: C) = (A :\: B) :&: (A :\: C). +Proof. by rewrite !setDE setCU setIIr. Qed. + +Lemma setDIl A B C : (A :&: B) :\: C = (A :\: C) :&: (B :\: C). +Proof. by rewrite !setDE setIIl. Qed. + +Lemma setIDA A B C : A :&: (B :\: C) = (A :&: B) :\: C. +Proof. by rewrite !setDE setIA. Qed. + +Lemma setIDAC A B C : (A :\: B) :&: C = (A :&: C) :\: B. +Proof. by rewrite !setDE setIAC. Qed. + +Lemma setDIr A B C : A :\: (B :&: C) = (A :\: B) :|: (A :\: C). +Proof. by rewrite !setDE setCI setIUr. Qed. + +Lemma setDDl A B C : (A :\: B) :\: C = A :\: (B :|: C). +Proof. by rewrite !setDE setCU setIA. Qed. + +Lemma setDDr A B C : A :\: (B :\: C) = (A :\: B) :|: (A :&: C). +Proof. by rewrite !setDE setCI setIUr setCK. Qed. + +(* powerset *) + +Lemma powersetE A B : (A \in powerset B) = (A \subset B). +Proof. by rewrite inE. Qed. + +Lemma powersetS A B : (powerset A \subset powerset B) = (A \subset B). +Proof. +apply/subsetP/idP=> [sAB | sAB C]; last by rewrite !inE => /subset_trans ->. +by rewrite -powersetE sAB // inE. +Qed. + +Lemma powerset0 : powerset set0 = [set set0] :> {set {set T}}. +Proof. by apply/setP=> A; rewrite !inE subset0. Qed. + +Lemma powersetT : powerset [set: T] = [set: {set T}]. +Proof. by apply/setP=> A; rewrite !inE subsetT. Qed. + +Lemma setI_powerset P A : P :&: powerset A = P ::&: A. +Proof. by apply/setP=> B; rewrite !inE. Qed. + +(* cardinal lemmas for sets *) + +Lemma cardsE pA : #|[set x in pA]| = #|pA|. +Proof. by apply: eq_card; exact: in_set. Qed. + +Lemma sum1dep_card pA : \sum_(x | pA x) 1 = #|[set x | pA x]|. +Proof. by rewrite sum1_card cardsE. Qed. + +Lemma sum_nat_dep_const pA n : \sum_(x | pA x) n = #|[set x | pA x]| * n. +Proof. by rewrite sum_nat_const cardsE. Qed. + +Lemma cards0 : #|@set0 T| = 0. +Proof. by rewrite cardsE card0. Qed. + +Lemma cards_eq0 A : (#|A| == 0) = (A == set0). +Proof. by rewrite (eq_sym A) eqEcard sub0set cards0 leqn0. Qed. + +Lemma set0Pn A : reflect (exists x, x \in A) (A != set0). +Proof. by rewrite -cards_eq0; exact: existsP. Qed. + +Lemma card_gt0 A : (0 < #|A|) = (A != set0). +Proof. by rewrite lt0n cards_eq0. Qed. + +Lemma cards0_eq A : #|A| = 0 -> A = set0. +Proof. by move=> A_0; apply/setP=> x; rewrite inE (card0_eq A_0). Qed. + +Lemma cards1 x : #|[set x]| = 1. +Proof. by rewrite cardsE card1. Qed. + +Lemma cardsUI A B : #|A :|: B| + #|A :&: B| = #|A| + #|B|. +Proof. by rewrite !cardsE cardUI. Qed. + +Lemma cardsU A B : #|A :|: B| = (#|A| + #|B| - #|A :&: B|)%N. +Proof. by rewrite -cardsUI addnK. Qed. + +Lemma cardsI A B : #|A :&: B| = (#|A| + #|B| - #|A :|: B|)%N. +Proof. by rewrite -cardsUI addKn. Qed. + +Lemma cardsT : #|[set: T]| = #|T|. +Proof. by rewrite cardsE. Qed. + +Lemma cardsID B A : #|A :&: B| + #|A :\: B| = #|A|. +Proof. by rewrite !cardsE cardID. Qed. + +Lemma cardsD A B : #|A :\: B| = (#|A| - #|A :&: B|)%N. +Proof. by rewrite -(cardsID B A) addKn. Qed. + +Lemma cardsC A : #|A| + #|~: A| = #|T|. +Proof. by rewrite cardsE cardC. Qed. + +Lemma cardsCs A : #|A| = #|T| - #|~: A|. +Proof. by rewrite -(cardsC A) addnK. Qed. + +Lemma cardsU1 a A : #|a |: A| = (a \notin A) + #|A|. +Proof. by rewrite -cardU1; apply: eq_card=> x; rewrite !inE. Qed. + +Lemma cards2 a b : #|[set a; b]| = (a != b).+1. +Proof. by rewrite -card2; apply: eq_card=> x; rewrite !inE. Qed. + +Lemma cardsC1 a : #|[set~ a]| = #|T|.-1. +Proof. by rewrite -(cardC1 a); apply: eq_card=> x; rewrite !inE. Qed. + +Lemma cardsD1 a A : #|A| = (a \in A) + #|A :\ a|. +Proof. +by rewrite (cardD1 a); congr (_ + _); apply: eq_card => x; rewrite !inE. +Qed. + +(* other inclusions *) + +Lemma subsetIl A B : A :&: B \subset A. +Proof. by apply/subsetP=> x; rewrite inE; case/andP. Qed. + +Lemma subsetIr A B : A :&: B \subset B. +Proof. by apply/subsetP=> x; rewrite inE; case/andP. Qed. + +Lemma subsetUl A B : A \subset A :|: B. +Proof. by apply/subsetP=> x; rewrite inE => ->. Qed. + +Lemma subsetUr A B : B \subset A :|: B. +Proof. by apply/subsetP=> x; rewrite inE orbC => ->. Qed. + +Lemma subsetU1 x A : A \subset x |: A. +Proof. exact: subsetUr. Qed. + +Lemma subsetDl A B : A :\: B \subset A. +Proof. by rewrite setDE subsetIl. Qed. + +Lemma subD1set A x : A :\ x \subset A. +Proof. by rewrite subsetDl. Qed. + +Lemma subsetDr A B : A :\: B \subset ~: B. +Proof. by rewrite setDE subsetIr. Qed. + +Lemma sub1set A x : ([set x] \subset A) = (x \in A). +Proof. by rewrite -subset_pred1; apply: eq_subset=> y; rewrite !inE. Qed. + +Lemma cards1P A : reflect (exists x, A = [set x]) (#|A| == 1). +Proof. +apply: (iffP idP) => [|[x ->]]; last by rewrite cards1. +rewrite eq_sym eqn_leq card_gt0 => /andP[/set0Pn[x Ax] leA1]. +by exists x; apply/eqP; rewrite eq_sym eqEcard sub1set Ax cards1 leA1. +Qed. + +Lemma subset1 A x : (A \subset [set x]) = (A == [set x]) || (A == set0). +Proof. +rewrite eqEcard cards1 -cards_eq0 orbC andbC. +by case: posnP => // A0; rewrite (cards0_eq A0) sub0set. +Qed. + +Lemma powerset1 x : powerset [set x] = [set set0; [set x]]. +Proof. by apply/setP=> A; rewrite !inE subset1 orbC. Qed. + +Lemma setIidPl A B : reflect (A :&: B = A) (A \subset B). +Proof. +apply: (iffP subsetP) => [sAB | <- x /setIP[] //]. +by apply/setP=> x; rewrite inE; apply: andb_idr; exact: sAB. +Qed. +Implicit Arguments setIidPl [A B]. + +Lemma setIidPr A B : reflect (A :&: B = B) (B \subset A). +Proof. rewrite setIC; exact: setIidPl. Qed. + +Lemma cardsDS A B : B \subset A -> #|A :\: B| = (#|A| - #|B|)%N. +Proof. by rewrite cardsD => /setIidPr->. Qed. + +Lemma setUidPl A B : reflect (A :|: B = A) (B \subset A). +Proof. +by rewrite -setCS (sameP setIidPl eqP) -setCU (inj_eq setC_inj); exact: eqP. +Qed. + +Lemma setUidPr A B : reflect (A :|: B = B) (A \subset B). +Proof. rewrite setUC; exact: setUidPl. Qed. + +Lemma setDidPl A B : reflect (A :\: B = A) [disjoint A & B]. +Proof. rewrite setDE disjoints_subset; exact: setIidPl. Qed. + +Lemma subIset A B C : (B \subset A) || (C \subset A) -> (B :&: C \subset A). +Proof. by case/orP; apply: subset_trans; rewrite (subsetIl, subsetIr). Qed. + +Lemma subsetI A B C : (A \subset B :&: C) = (A \subset B) && (A \subset C). +Proof. +rewrite !(sameP setIidPl eqP) setIA; have [-> //| ] := altP (A :&: B =P A). +by apply: contraNF => /eqP <-; rewrite -setIA -setIIl setIAC. +Qed. + +Lemma subsetIP A B C : reflect (A \subset B /\ A \subset C) (A \subset B :&: C). +Proof. by rewrite subsetI; exact: andP. Qed. + +Lemma subsetIidl A B : (A \subset A :&: B) = (A \subset B). +Proof. by rewrite subsetI subxx. Qed. + +Lemma subsetIidr A B : (B \subset A :&: B) = (B \subset A). +Proof. by rewrite setIC subsetIidl. Qed. + +Lemma powersetI A B : powerset (A :&: B) = powerset A :&: powerset B. +Proof. by apply/setP=> C; rewrite !inE subsetI. Qed. + +Lemma subUset A B C : (B :|: C \subset A) = (B \subset A) && (C \subset A). +Proof. by rewrite -setCS setCU subsetI !setCS. Qed. + +Lemma subsetU A B C : (A \subset B) || (A \subset C) -> A \subset B :|: C. +Proof. by rewrite -!(setCS _ A) setCU; exact: subIset. Qed. + +Lemma subUsetP A B C : reflect (A \subset C /\ B \subset C) (A :|: B \subset C). +Proof. by rewrite subUset; exact: andP. Qed. + +Lemma subsetC A B : (A \subset ~: B) = (B \subset ~: A). +Proof. by rewrite -setCS setCK. Qed. + +Lemma subCset A B : (~: A \subset B) = (~: B \subset A). +Proof. by rewrite -setCS setCK. Qed. + +Lemma subsetD A B C : (A \subset B :\: C) = (A \subset B) && [disjoint A & C]. +Proof. by rewrite setDE subsetI -disjoints_subset. Qed. + +Lemma subDset A B C : (A :\: B \subset C) = (A \subset B :|: C). +Proof. +apply/subsetP/subsetP=> sABC x; rewrite !inE. + by case Bx: (x \in B) => // Ax; rewrite sABC ?inE ?Bx. +by case Bx: (x \in B) => //; move/sABC; rewrite inE Bx. +Qed. + +Lemma subsetDP A B C : + reflect (A \subset B /\ [disjoint A & C]) (A \subset B :\: C). +Proof. by rewrite subsetD; exact: andP. Qed. + +Lemma setU_eq0 A B : (A :|: B == set0) = (A == set0) && (B == set0). +Proof. by rewrite -!subset0 subUset. Qed. + +Lemma setD_eq0 A B : (A :\: B == set0) = (A \subset B). +Proof. by rewrite -subset0 subDset setU0. Qed. + +Lemma setI_eq0 A B : (A :&: B == set0) = [disjoint A & B]. +Proof. by rewrite disjoints_subset -setD_eq0 setDE setCK. Qed. + +Lemma disjoint_setI0 A B : [disjoint A & B] -> A :&: B = set0. +Proof. by rewrite -setI_eq0; move/eqP. Qed. + +Lemma subsetD1 A B x : (A \subset B :\ x) = (A \subset B) && (x \notin A). +Proof. by rewrite setDE subsetI subsetC sub1set inE. Qed. + +Lemma subsetD1P A B x : reflect (A \subset B /\ x \notin A) (A \subset B :\ x). +Proof. by rewrite subsetD1; exact: andP. Qed. + +Lemma properD1 A x : x \in A -> A :\ x \proper A. +Proof. +move=> Ax; rewrite properE subsetDl; apply/subsetPn; exists x=> //. +by rewrite in_setD1 Ax eqxx. +Qed. + +Lemma properIr A B : ~~ (B \subset A) -> A :&: B \proper B. +Proof. by move=> nsAB; rewrite properE subsetIr subsetI negb_and nsAB. Qed. + +Lemma properIl A B : ~~ (A \subset B) -> A :&: B \proper A. +Proof. by move=> nsBA; rewrite properE subsetIl subsetI negb_and nsBA orbT. Qed. + +Lemma properUr A B : ~~ (A \subset B) -> B \proper A :|: B. +Proof. by rewrite properE subsetUr subUset subxx /= andbT. Qed. + +Lemma properUl A B : ~~ (B \subset A) -> A \proper A :|: B. +Proof. by move=> not_sBA; rewrite setUC properUr. Qed. + +Lemma proper1set A x : ([set x] \proper A) -> (x \in A). +Proof. by move/proper_sub; rewrite sub1set. Qed. + +Lemma properIset A B C : (B \proper A) || (C \proper A) -> (B :&: C \proper A). +Proof. by case/orP; apply: sub_proper_trans; rewrite (subsetIl, subsetIr). Qed. + +Lemma properI A B C : (A \proper B :&: C) -> (A \proper B) && (A \proper C). +Proof. +move=> pAI; apply/andP. +by split; apply: (proper_sub_trans pAI); rewrite (subsetIl, subsetIr). +Qed. + +Lemma properU A B C : (B :|: C \proper A) -> (B \proper A) && (C \proper A). +Proof. +move=> pUA; apply/andP. +by split; apply: sub_proper_trans pUA; rewrite (subsetUr, subsetUl). +Qed. + +Lemma properD A B C : (A \proper B :\: C) -> (A \proper B) && [disjoint A & C]. +Proof. by rewrite setDE disjoints_subset => /properI/andP[-> /proper_sub]. Qed. + +End setOps. + +Implicit Arguments set1P [T x a]. +Implicit Arguments set1_inj [T]. +Implicit Arguments set2P [T x a b]. +Implicit Arguments setIdP [T x pA pB]. +Implicit Arguments setIP [T x A B]. +Implicit Arguments setU1P [T x a B]. +Implicit Arguments setD1P [T x A b]. +Implicit Arguments setUP [T x A B]. +Implicit Arguments setDP [T x A B]. +Implicit Arguments cards1P [T A]. +Implicit Arguments setCP [T x A]. +Implicit Arguments setIidPl [T A B]. +Implicit Arguments setIidPr [T A B]. +Implicit Arguments setUidPl [T A B]. +Implicit Arguments setUidPr [T A B]. +Implicit Arguments setDidPl [T A B]. +Implicit Arguments subsetIP [T A B C]. +Implicit Arguments subUsetP [T A B C]. +Implicit Arguments subsetDP [T A B C]. +Implicit Arguments subsetD1P [T A B x]. +Prenex Implicits set1 set1_inj. +Prenex Implicits set1P set2P setU1P setD1P setIdP setIP setUP setDP. +Prenex Implicits cards1P setCP setIidPl setIidPr setUidPl setUidPr setDidPl. +Hint Resolve subsetT_hint. + +Section setOpsAlgebra. + +Import Monoid. + +Variable T : finType. + +Canonical setI_monoid := Law (@setIA T) (@setTI T) (@setIT T). + +Canonical setI_comoid := ComLaw (@setIC T). +Canonical setI_muloid := MulLaw (@set0I T) (@setI0 T). + +Canonical setU_monoid := Law (@setUA T) (@set0U T) (@setU0 T). +Canonical setU_comoid := ComLaw (@setUC T). +Canonical setU_muloid := MulLaw (@setTU T) (@setUT T). + +Canonical setI_addoid := AddLaw (@setUIl T) (@setUIr T). +Canonical setU_addoid := AddLaw (@setIUl T) (@setIUr T). + +End setOpsAlgebra. + +Section CartesianProd. + +Variables fT1 fT2 : finType. +Variables (A1 : {set fT1}) (A2 : {set fT2}). + +Definition setX := [set u | u.1 \in A1 & u.2 \in A2]. + +Lemma in_setX x1 x2 : ((x1, x2) \in setX) = (x1 \in A1) && (x2 \in A2). +Proof. by rewrite inE. Qed. + +Lemma setXP x1 x2 : reflect (x1 \in A1 /\ x2 \in A2) ((x1, x2) \in setX). +Proof. by rewrite inE; exact: andP. Qed. + +Lemma cardsX : #|setX| = #|A1| * #|A2|. +Proof. by rewrite cardsE cardX. Qed. + +End CartesianProd. + +Implicit Arguments setXP [x1 x2 fT1 fT2 A1 A2]. +Prenex Implicits setXP. + +Notation Local imset_def := + (fun (aT rT : finType) f mD => [set y in @image_mem aT rT f mD]). +Notation Local imset2_def := + (fun (aT1 aT2 rT : finType) f (D1 : mem_pred aT1) (D2 : _ -> mem_pred aT2) => + [set y in @image_mem _ rT (prod_curry f) + (mem [pred u | D1 u.1 & D2 u.1 u.2])]). + +Module Type ImsetSig. +Parameter imset : forall aT rT : finType, + (aT -> rT) -> mem_pred aT -> {set rT}. +Parameter imset2 : forall aT1 aT2 rT : finType, + (aT1 -> aT2 -> rT) -> mem_pred aT1 -> (aT1 -> mem_pred aT2) -> {set rT}. +Axiom imsetE : imset = imset_def. +Axiom imset2E : imset2 = imset2_def. +End ImsetSig. + +Module Imset : ImsetSig. +Definition imset := imset_def. +Definition imset2 := imset2_def. +Lemma imsetE : imset = imset_def. Proof. by []. Qed. +Lemma imset2E : imset2 = imset2_def. Proof. by []. Qed. +End Imset. + +Notation imset := Imset.imset. +Notation imset2 := Imset.imset2. +Canonical imset_unlock := Unlockable Imset.imsetE. +Canonical imset2_unlock := Unlockable Imset.imset2E. +Definition preimset (aT : finType) rT f (R : mem_pred rT) := + [set x : aT | in_mem (f x) R]. + +Notation "f @^-1: A" := (preimset f (mem A)) (at level 24) : set_scope. +Notation "f @: A" := (imset f (mem A)) (at level 24) : set_scope. +Notation "f @2: ( A , B )" := (imset2 f (mem A) (fun _ => mem B)) + (at level 24, format "f @2: ( A , B )") : set_scope. + +(* Comprehensions *) +Notation "[ 'set' E | x 'in' A ]" := ((fun x => E) @: A) + (at level 0, E, x at level 99, + format "[ '[hv' 'set' E '/ ' | x 'in' A ] ']'") : set_scope. +Notation "[ 'set' E | x 'in' A & P ]" := [set E | x in [set x in A | P]] + (at level 0, E, x at level 99, + format "[ '[hv' 'set' E '/ ' | x 'in' A '/ ' & P ] ']'") : set_scope. +Notation "[ 'set' E | x 'in' A , y 'in' B ]" := + (imset2 (fun x y => E) (mem A) (fun x => (mem B))) + (at level 0, E, x, y at level 99, format + "[ '[hv' 'set' E '/ ' | x 'in' A , '/ ' y 'in' B ] ']'" + ) : set_scope. +Notation "[ 'set' E | x 'in' A , y 'in' B & P ]" := + [set E | x in A, y in [set y in B | P]] + (at level 0, E, x, y at level 99, format + "[ '[hv' 'set' E '/ ' | x 'in' A , '/ ' y 'in' B '/ ' & P ] ']'" + ) : set_scope. + +(* Typed variants. *) +Notation "[ 'set' E | x : T 'in' A ]" := ((fun x : T => E) @: A) + (at level 0, E, x at level 99, only parsing) : set_scope. +Notation "[ 'set' E | x : T 'in' A & P ]" := + [set E | x : T in [set x : T in A | P]] + (at level 0, E, x at level 99, only parsing) : set_scope. +Notation "[ 'set' E | x : T 'in' A , y : U 'in' B ]" := + (imset2 (fun (x : T) (y : U) => E) (mem A) (fun (x : T) => (mem B))) + (at level 0, E, x, y at level 99, only parsing) : set_scope. +Notation "[ 'set' E | x : T 'in' A , y : U 'in' B & P ]" := + [set E | x : T in A, y : U in [set y : U in B | P]] + (at level 0, E, x, y at level 99, only parsing) : set_scope. + +(* Comprehensions over a type. *) +Local Notation predOfType T := (sort_of_simpl_pred (@pred_of_argType T)). +Notation "[ 'set' E | x : T ]" := [set E | x : T in predOfType T] + (at level 0, E, x at level 99, + format "[ '[hv' 'set' E '/ ' | x : T ] ']'") : set_scope. +Notation "[ 'set' E | x : T & P ]" := [set E | x : T in [set x : T | P]] + (at level 0, E, x at level 99, + format "[ '[hv' 'set' E '/ ' | x : T '/ ' & P ] ']'") : set_scope. +Notation "[ 'set' E | x : T , y : U 'in' B ]" := + [set E | x : T in predOfType T, y : U in B] + (at level 0, E, x, y at level 99, format + "[ '[hv' 'set' E '/ ' | x : T , '/ ' y : U 'in' B ] ']'") + : set_scope. +Notation "[ 'set' E | x : T , y : U 'in' B & P ]" := + [set E | x : T, y : U in [set y in B | P]] + (at level 0, E, x, y at level 99, format + "[ '[hv ' 'set' E '/' | x : T , '/ ' y : U 'in' B '/' & P ] ']'" + ) : set_scope. +Notation "[ 'set' E | x : T 'in' A , y : U ]" := + [set E | x : T in A, y : U in predOfType U] + (at level 0, E, x, y at level 99, format + "[ '[hv' 'set' E '/ ' | x : T 'in' A , '/ ' y : U ] ']'") + : set_scope. +Notation "[ 'set' E | x : T 'in' A , y : U & P ]" := + [set E | x : T in A, y : U in [set y in P]] + (at level 0, E, x, y at level 99, format + "[ '[hv' 'set' E '/ ' | x : T 'in' A , '/ ' y : U & P ] ']'") + : set_scope. +Notation "[ 'set' E | x : T , y : U ]" := + [set E | x : T, y : U in predOfType U] + (at level 0, E, x, y at level 99, format + "[ '[hv' 'set' E '/ ' | x : T , '/ ' y : U ] ']'") + : set_scope. +Notation "[ 'set' E | x : T , y : U & P ]" := + [set E | x : T, y : U in [set y in P]] + (at level 0, E, x, y at level 99, format + "[ '[hv' 'set' E '/ ' | x : T , '/ ' y : U & P ] ']'") + : set_scope. + +(* Untyped variants. *) +Notation "[ 'set' E | x , y 'in' B ]" := [set E | x : _, y : _ in B] + (at level 0, E, x, y at level 99, only parsing) : set_scope. +Notation "[ 'set' E | x , y 'in' B & P ]" := [set E | x : _, y : _ in B & P] + (at level 0, E, x, y at level 99, only parsing) : set_scope. +Notation "[ 'set' E | x 'in' A , y ]" := [set E | x : _ in A, y : _] + (at level 0, E, x, y at level 99, only parsing) : set_scope. +Notation "[ 'set' E | x 'in' A , y & P ]" := [set E | x : _ in A, y : _ & P] + (at level 0, E, x, y at level 99, only parsing) : set_scope. +Notation "[ 'set' E | x , y ]" := [set E | x : _, y : _] + (at level 0, E, x, y at level 99, only parsing) : set_scope. +Notation "[ 'set' E | x , y & P ]" := [set E | x : _, y : _ & P ] + (at level 0, E, x, y at level 99, only parsing) : set_scope. + +(* Print-only variants to work around the Coq pretty-printer K-term kink. *) +Notation "[ 'se' 't' E | x 'in' A , y 'in' B ]" := + (imset2 (fun x y => E) (mem A) (fun _ => mem B)) + (at level 0, E, x, y at level 99, format + "[ '[hv' 'se' 't' E '/ ' | x 'in' A , '/ ' y 'in' B ] ']'") + : set_scope. +Notation "[ 'se' 't' E | x 'in' A , y 'in' B & P ]" := + [se t E | x in A, y in [set y in B | P]] + (at level 0, E, x, y at level 99, format + "[ '[hv ' 'se' 't' E '/' | x 'in' A , '/ ' y 'in' B '/' & P ] ']'" + ) : set_scope. +Notation "[ 'se' 't' E | x : T , y : U 'in' B ]" := + (imset2 (fun x (y : U) => E) (mem (predOfType T)) (fun _ => mem B)) + (at level 0, E, x, y at level 99, format + "[ '[hv ' 'se' 't' E '/' | x : T , '/ ' y : U 'in' B ] ']'") + : set_scope. +Notation "[ 'se' 't' E | x : T , y : U 'in' B & P ]" := + [se t E | x : T, y : U in [set y in B | P]] + (at level 0, E, x, y at level 99, format +"[ '[hv ' 'se' 't' E '/' | x : T , '/ ' y : U 'in' B '/' & P ] ']'" + ) : set_scope. +Notation "[ 'se' 't' E | x : T 'in' A , y : U ]" := + (imset2 (fun x y => E) (mem A) (fun _ : T => mem (predOfType U))) + (at level 0, E, x, y at level 99, format + "[ '[hv' 'se' 't' E '/ ' | x : T 'in' A , '/ ' y : U ] ']'") + : set_scope. +Notation "[ 'se' 't' E | x : T 'in' A , y : U & P ]" := + (imset2 (fun x (y : U) => E) (mem A) (fun _ : T => mem [set y \in P])) + (at level 0, E, x, y at level 99, format +"[ '[hv ' 'se' 't' E '/' | x : T 'in' A , '/ ' y : U '/' & P ] ']'" + ) : set_scope. +Notation "[ 'se' 't' E | x : T , y : U ]" := + [se t E | x : T, y : U in predOfType U] + (at level 0, E, x, y at level 99, format + "[ '[hv' 'se' 't' E '/ ' | x : T , '/ ' y : U ] ']'") + : set_scope. +Notation "[ 'se' 't' E | x : T , y : U & P ]" := + [se t E | x : T, y : U in [set y in P]] + (at level 0, E, x, y at level 99, format + "[ '[hv' 'se' 't' E '/' | x : T , '/ ' y : U '/' & P ] ']'") + : set_scope. + +Section FunImage. + +Variables aT aT2 : finType. + +Section ImsetTheory. + +Variable rT : finType. + +Section ImsetProp. + +Variables (f : aT -> rT) (f2 : aT -> aT2 -> rT). + +Lemma imsetP D y : reflect (exists2 x, in_mem x D & y = f x) (y \in imset f D). +Proof. rewrite [@imset]unlock inE; exact: imageP. Qed. + +CoInductive imset2_spec D1 D2 y : Prop := + Imset2spec x1 x2 of in_mem x1 D1 & in_mem x2 (D2 x1) & y = f2 x1 x2. + +Lemma imset2P D1 D2 y : reflect (imset2_spec D1 D2 y) (y \in imset2 f2 D1 D2). +Proof. +rewrite [@imset2]unlock inE. +apply: (iffP imageP) => [[[x1 x2] Dx12] | [x1 x2 Dx1 Dx2]] -> {y}. + by case/andP: Dx12; exists x1 x2. +by exists (x1, x2); rewrite //= !inE Dx1. +Qed. + +Lemma mem_imset (D : pred aT) x : x \in D -> f x \in f @: D. +Proof. by move=> Dx; apply/imsetP; exists x. Qed. + +Lemma imset0 : f @: set0 = set0. +Proof. by apply/setP => y; rewrite inE; apply/imsetP=> [[x]]; rewrite inE. Qed. + +Lemma imset_eq0 (A : {set aT}) : (f @: A == set0) = (A == set0). +Proof. +have [-> | [x Ax]] := set_0Vmem A; first by rewrite imset0 !eqxx. +by rewrite -!cards_eq0 (cardsD1 x) Ax (cardsD1 (f x)) mem_imset. +Qed. + +Lemma imset_set1 x : f @: [set x] = [set f x]. +Proof. +apply/setP => y. +by apply/imsetP/set1P=> [[x' /set1P-> //]| ->]; exists x; rewrite ?set11. +Qed. + +Lemma mem_imset2 (D : pred aT) (D2 : aT -> pred aT2) x x2 : + x \in D -> x2 \in D2 x -> + f2 x x2 \in imset2 f2 (mem D) (fun x1 => mem (D2 x1)). +Proof. by move=> Dx Dx2; apply/imset2P; exists x x2. Qed. + +Lemma sub_imset_pre (A : pred aT) (B : pred rT) : + (f @: A \subset B) = (A \subset f @^-1: B). +Proof. +apply/subsetP/subsetP=> [sfAB x Ax | sAf'B fx]. + by rewrite inE sfAB ?mem_imset. +by case/imsetP=> x Ax ->; move/sAf'B: Ax; rewrite inE. +Qed. + +Lemma preimsetS (A B : pred rT) : + A \subset B -> (f @^-1: A) \subset (f @^-1: B). +Proof. move=> sAB; apply/subsetP=> y; rewrite !inE; exact: (subsetP sAB). Qed. + +Lemma preimset0 : f @^-1: set0 = set0. +Proof. by apply/setP=> x; rewrite !inE. Qed. + +Lemma preimsetT : f @^-1: setT = setT. +Proof. by apply/setP=> x; rewrite !inE. Qed. + +Lemma preimsetI (A B : {set rT}) : + f @^-1: (A :&: B) = (f @^-1: A) :&: (f @^-1: B). +Proof. by apply/setP=> y; rewrite !inE. Qed. + +Lemma preimsetU (A B : {set rT}) : + f @^-1: (A :|: B) = (f @^-1: A) :|: (f @^-1: B). +Proof. by apply/setP=> y; rewrite !inE. Qed. + +Lemma preimsetD (A B : {set rT}) : + f @^-1: (A :\: B) = (f @^-1: A) :\: (f @^-1: B). +Proof. by apply/setP=> y; rewrite !inE. Qed. + +Lemma preimsetC (A : {set rT}) : f @^-1: (~: A) = ~: f @^-1: A. +Proof. by apply/setP=> y; rewrite !inE. Qed. + +Lemma imsetS (A B : pred aT) : A \subset B -> f @: A \subset f @: B. +Proof. +move=> sAB; apply/subsetP=> _ /imsetP[x Ax ->]. +by apply/imsetP; exists x; rewrite ?(subsetP sAB). +Qed. + +Lemma imset_proper (A B : {set aT}) : + {in B &, injective f} -> A \proper B -> f @: A \proper f @: B. +Proof. +move=> injf /properP[sAB [x Bx nAx]]; rewrite properE imsetS //=. +apply: contra nAx => sfBA. +have: f x \in f @: A by rewrite (subsetP sfBA) ?mem_imset. +by case/imsetP=> y Ay /injf-> //; exact: subsetP sAB y Ay. +Qed. + +Lemma preimset_proper (A B : {set rT}) : + B \subset codom f -> A \proper B -> (f @^-1: A) \proper (f @^-1: B). +Proof. +move=> sBc /properP[sAB [u Bu nAu]]; rewrite properE preimsetS //=. +by apply/subsetPn; exists (iinv (subsetP sBc _ Bu)); rewrite inE /= f_iinv. +Qed. + +Lemma imsetU (A B : {set aT}) : f @: (A :|: B) = (f @: A) :|: (f @: B). +Proof. +apply/eqP; rewrite eqEsubset subUset. +rewrite 2?imsetS (andbT, subsetUl, subsetUr) // andbT. +apply/subsetP=> _ /imsetP[x ABx ->]; apply/setUP. +by case/setUP: ABx => [Ax | Bx]; [left | right]; apply/imsetP; exists x. +Qed. + +Lemma imsetU1 a (A : {set aT}) : f @: (a |: A) = f a |: (f @: A). +Proof. by rewrite imsetU imset_set1. Qed. + +Lemma imsetI (A B : {set aT}) : + {in A & B, injective f} -> f @: (A :&: B) = f @: A :&: f @: B. +Proof. +move=> injf; apply/eqP; rewrite eqEsubset subsetI. +rewrite 2?imsetS (andTb, subsetIl, subsetIr) //=. +apply/subsetP=> _ /setIP[/imsetP[x Ax ->] /imsetP[z Bz /injf eqxz]]. +by rewrite mem_imset // inE Ax eqxz. +Qed. + +Lemma imset2Sl (A B : pred aT) (C : pred aT2) : + A \subset B -> f2 @2: (A, C) \subset f2 @2: (B, C). +Proof. +move=> sAB; apply/subsetP=> _ /imset2P[x y Ax Cy ->]. +by apply/imset2P; exists x y; rewrite ?(subsetP sAB). +Qed. + +Lemma imset2Sr (A B : pred aT2) (C : pred aT) : + A \subset B -> f2 @2: (C, A) \subset f2 @2: (C, B). +Proof. +move=> sAB; apply/subsetP=> _ /imset2P[x y Ax Cy ->]. +by apply/imset2P; exists x y; rewrite ?(subsetP sAB). +Qed. + +Lemma imset2S (A B : pred aT) (A2 B2 : pred aT2) : + A \subset B -> A2 \subset B2 -> f2 @2: (A, A2) \subset f2 @2: (B, B2). +Proof. by move=> /(imset2Sl B2) sBA /(imset2Sr A)/subset_trans->. Qed. + +End ImsetProp. + +Implicit Types (f g : aT -> rT) (D : {set aT}) (R : pred rT). + +Lemma eq_preimset f g R : f =1 g -> f @^-1: R = g @^-1: R. +Proof. by move=> eqfg; apply/setP => y; rewrite !inE eqfg. Qed. + +Lemma eq_imset f g D : f =1 g -> f @: D = g @: D. +Proof. +move=> eqfg; apply/setP=> y. +by apply/imsetP/imsetP=> [] [x Dx ->]; exists x; rewrite ?eqfg. +Qed. + +Lemma eq_in_imset f g D : {in D, f =1 g} -> f @: D = g @: D. +Proof. +move=> eqfg; apply/setP => y. +by apply/imsetP/imsetP=> [] [x Dx ->]; exists x; rewrite ?eqfg. +Qed. + +Lemma eq_in_imset2 (f g : aT -> aT2 -> rT) (D : pred aT) (D2 : pred aT2) : + {in D & D2, f =2 g} -> f @2: (D, D2) = g @2: (D, D2). +Proof. +move=> eqfg; apply/setP => y. +by apply/imset2P/imset2P=> [] [x x2 Dx Dx2 ->]; exists x x2; rewrite ?eqfg. +Qed. + +End ImsetTheory. + +Lemma imset2_pair (A : {set aT}) (B : {set aT2}) : + [set (x, y) | x in A, y in B] = setX A B. +Proof. +apply/setP=> [[x y]]; rewrite !inE /=. +by apply/imset2P/andP=> [[_ _ _ _ [-> ->]//]| []]; exists x y. +Qed. + +Lemma setXS (A1 B1 : {set aT}) (A2 B2 : {set aT2}) : + A1 \subset B1 -> A2 \subset B2 -> setX A1 A2 \subset setX B1 B2. +Proof. by move=> sAB1 sAB2; rewrite -!imset2_pair imset2S. Qed. + +End FunImage. + +Implicit Arguments imsetP [aT rT f D y]. +Implicit Arguments imset2P [aT aT2 rT f2 D1 D2 y]. +Prenex Implicits imsetP imset2P. + +Section BigOps. + +Variables (R : Type) (idx : R). +Variables (op : Monoid.law idx) (aop : Monoid.com_law idx). +Variables I J : finType. +Implicit Type A B : {set I}. +Implicit Type h : I -> J. +Implicit Type P : pred I. +Implicit Type F : I -> R. + +Lemma big_set0 F : \big[op/idx]_(i in set0) F i = idx. +Proof. by apply: big_pred0 => i; rewrite inE. Qed. + +Lemma big_set1 a F : \big[op/idx]_(i in [set a]) F i = F a. +Proof. by apply: big_pred1 => i; rewrite !inE. Qed. + +Lemma big_setIDdep A B P F : + \big[aop/idx]_(i in A | P i) F i = + aop (\big[aop/idx]_(i in A :&: B | P i) F i) + (\big[aop/idx]_(i in A :\: B | P i) F i). +Proof. +rewrite (bigID (mem B)) setDE. +by congr (aop _ _); apply: eq_bigl => i; rewrite !inE andbAC. +Qed. + +Lemma big_setID A B F : + \big[aop/idx]_(i in A) F i = + aop (\big[aop/idx]_(i in A :&: B) F i) + (\big[aop/idx]_(i in A :\: B) F i). +Proof. +rewrite (bigID (mem B)) !(eq_bigl _ _ (in_set _)) //=. +by congr (aop _); apply: eq_bigl => i; rewrite andbC. +Qed. + +Lemma big_setD1 a A F : a \in A -> + \big[aop/idx]_(i in A) F i = aop (F a) (\big[aop/idx]_(i in A :\ a) F i). +Proof. +move=> Aa; rewrite (bigD1 a Aa); congr (aop _). +by apply: eq_bigl => x; rewrite !inE andbC. +Qed. + +Lemma big_setU1 a A F : a \notin A -> + \big[aop/idx]_(i in a |: A) F i = aop (F a) (\big[aop/idx]_(i in A) F i). +Proof. by move=> notAa; rewrite (@big_setD1 a) ?setU11 //= setU1K. Qed. + +Lemma big_imset h (A : pred I) G : + {in A &, injective h} -> + \big[aop/idx]_(j in h @: A) G j = \big[aop/idx]_(i in A) G (h i). +Proof. +move=> injh; pose hA := mem (image h A). +have [x0 Ax0 | A0] := pickP A; last first. + by rewrite !big_pred0 // => x; apply/imsetP=> [[i]]; rewrite unfold_in A0. +rewrite (eq_bigl hA) => [|j]; last by exact/imsetP/imageP. +pose h' j := if insub j : {? j | hA j} is Some u then iinv (svalP u) else x0. +rewrite (reindex_onto h h') => [|j hAj]; rewrite {}/h'; last first. + by rewrite (insubT hA hAj) f_iinv. +apply: eq_bigl => i; case: insubP => [u -> /= def_u | nhAhi]. + set i' := iinv _; have Ai' : i' \in A := mem_iinv (svalP u). + by apply/eqP/idP=> [<- // | Ai]; apply: injh; rewrite ?f_iinv. +symmetry; rewrite (negbTE nhAhi); apply/idP=> Ai. +by case/imageP: nhAhi; exists i. +Qed. + +Lemma partition_big_imset h (A : pred I) F : + \big[aop/idx]_(i in A) F i = + \big[aop/idx]_(j in h @: A) \big[aop/idx]_(i in A | h i == j) F i. +Proof. by apply: partition_big => i Ai; apply/imsetP; exists i. Qed. + +End BigOps. + +Implicit Arguments big_setID [R idx aop I A]. +Implicit Arguments big_setD1 [R idx aop I A F]. +Implicit Arguments big_setU1 [R idx aop I A F]. +Implicit Arguments big_imset [R idx aop h I J A]. +Implicit Arguments partition_big_imset [R idx aop I J]. + +Section Fun2Set1. + +Variables aT1 aT2 rT : finType. +Variables (f : aT1 -> aT2 -> rT). + +Lemma imset2_set1l x1 (D2 : pred aT2) : f @2: ([set x1], D2) = f x1 @: D2. +Proof. +apply/setP=> y; apply/imset2P/imsetP=> [[x x2 /set1P->]| [x2 Dx2 ->]]. + by exists x2. +by exists x1 x2; rewrite ?set11. +Qed. + +Lemma imset2_set1r x2 (D1 : pred aT1) : f @2: (D1, [set x2]) = f^~ x2 @: D1. +Proof. +apply/setP=> y; apply/imset2P/imsetP=> [[x1 x Dx1 /set1P->]| [x1 Dx1 ->]]. + by exists x1. +by exists x1 x2; rewrite ?set11. +Qed. + +End Fun2Set1. + +Section CardFunImage. + +Variables aT aT2 rT : finType. +Variables (f : aT -> rT) (g : rT -> aT) (f2 : aT -> aT2 -> rT). +Variables (D : pred aT) (D2 : pred aT). + +Lemma imset_card : #|f @: D| = #|image f D|. +Proof. by rewrite [@imset]unlock cardsE. Qed. + +Lemma leq_imset_card : #|f @: D| <= #|D|. +Proof. by rewrite imset_card leq_image_card. Qed. + +Lemma card_in_imset : {in D &, injective f} -> #|f @: D| = #|D|. +Proof. by move=> injf; rewrite imset_card card_in_image. Qed. + +Lemma card_imset : injective f -> #|f @: D| = #|D|. +Proof. by move=> injf; rewrite imset_card card_image. Qed. + +Lemma imset_injP : reflect {in D &, injective f} (#|f @: D| == #|D|). +Proof. by rewrite [@imset]unlock cardsE; exact: image_injP. Qed. + +Lemma can2_in_imset_pre : + {in D, cancel f g} -> {on D, cancel g & f} -> f @: D = g @^-1: D. +Proof. +move=> fK gK; apply/setP=> y; rewrite inE. +by apply/imsetP/idP=> [[x Ax ->] | Agy]; last exists (g y); rewrite ?(fK, gK). +Qed. + +Lemma can2_imset_pre : cancel f g -> cancel g f -> f @: D = g @^-1: D. +Proof. by move=> fK gK; apply: can2_in_imset_pre; exact: in1W. Qed. + +End CardFunImage. + +Implicit Arguments imset_injP [aT rT f D]. + +Lemma on_card_preimset (aT rT : finType) (f : aT -> rT) (R : pred rT) : + {on R, bijective f} -> #|f @^-1: R| = #|R|. +Proof. +case=> g fK gK; rewrite -(can2_in_imset_pre gK) // card_in_imset //. +exact: can_in_inj gK. +Qed. + +Lemma can_imset_pre (T : finType) f g (A : {set T}) : + cancel f g -> f @: A = g @^-1: A :> {set T}. +Proof. +move=> fK; apply: can2_imset_pre => // x. +suffices fx: x \in codom f by rewrite -(f_iinv fx) fK. +move: x; apply/(subset_cardP (card_codom (can_inj fK))); exact/subsetP. +Qed. + +Lemma imset_id (T : finType) (A : {set T}) : [set x | x in A] = A. +Proof. by apply/setP=> x; rewrite (@can_imset_pre _ _ id) ?inE. Qed. + +Lemma card_preimset (T : finType) (f : T -> T) (A : {set T}) : + injective f -> #|f @^-1: A| = #|A|. +Proof. +move=> injf; apply: on_card_preimset; apply: onW_bij. +have ontof: _ \in codom f by exact/(subset_cardP (card_codom injf))/subsetP. +by exists (fun x => iinv (ontof x)) => x; rewrite (f_iinv, iinv_f). +Qed. + +Lemma card_powerset (T : finType) (A : {set T}) : #|powerset A| = 2 ^ #|A|. +Proof. +rewrite -card_bool -(card_pffun_on false) -(card_imset _ val_inj). +apply: eq_card => f; pose sf := false.-support f; pose D := finset sf. +have sDA: (D \subset A) = (sf \subset A) by apply: eq_subset; exact: in_set. +have eq_sf x : sf x = f x by rewrite /= negb_eqb addbF. +have valD: val D = f by rewrite /D unlock; apply/ffunP=> x; rewrite ffunE eq_sf. +apply/imsetP/pffun_onP=> [[B] | [sBA _]]; last by exists D; rewrite // inE ?sDA. +by rewrite inE -sDA -valD => sBA /val_inj->. +Qed. + +Section FunImageComp. + +Variables T T' U : finType. + +Lemma imset_comp (f : T' -> U) (g : T -> T') (H : pred T) : + (f \o g) @: H = f @: (g @: H). +Proof. +apply/setP/subset_eqP/andP. +split; apply/subsetP=> _ /imsetP[x0 Hx0 ->]; apply/imsetP. + by exists (g x0); first apply: mem_imset. +by move/imsetP: Hx0 => [x1 Hx1 ->]; exists x1. +Qed. + +End FunImageComp. + +Notation "\bigcup_ ( i <- r | P ) F" := + (\big[@setU _/set0]_(i <- r | P) F%SET) : set_scope. +Notation "\bigcup_ ( i <- r ) F" := + (\big[@setU _/set0]_(i <- r) F%SET) : set_scope. +Notation "\bigcup_ ( m <= i < n | P ) F" := + (\big[@setU _/set0]_(m <= i < n | P%B) F%SET) : set_scope. +Notation "\bigcup_ ( m <= i < n ) F" := + (\big[@setU _/set0]_(m <= i < n) F%SET) : set_scope. +Notation "\bigcup_ ( i | P ) F" := + (\big[@setU _/set0]_(i | P%B) F%SET) : set_scope. +Notation "\bigcup_ i F" := + (\big[@setU _/set0]_i F%SET) : set_scope. +Notation "\bigcup_ ( i : t | P ) F" := + (\big[@setU _/set0]_(i : t | P%B) F%SET) (only parsing): set_scope. +Notation "\bigcup_ ( i : t ) F" := + (\big[@setU _/set0]_(i : t) F%SET) (only parsing) : set_scope. +Notation "\bigcup_ ( i < n | P ) F" := + (\big[@setU _/set0]_(i < n | P%B) F%SET) : set_scope. +Notation "\bigcup_ ( i < n ) F" := + (\big[@setU _/set0]_ (i < n) F%SET) : set_scope. +Notation "\bigcup_ ( i 'in' A | P ) F" := + (\big[@setU _/set0]_(i in A | P%B) F%SET) : set_scope. +Notation "\bigcup_ ( i 'in' A ) F" := + (\big[@setU _/set0]_(i in A) F%SET) : set_scope. + +Notation "\bigcap_ ( i <- r | P ) F" := + (\big[@setI _/setT]_(i <- r | P%B) F%SET) : set_scope. +Notation "\bigcap_ ( i <- r ) F" := + (\big[@setI _/setT]_(i <- r) F%SET) : set_scope. +Notation "\bigcap_ ( m <= i < n | P ) F" := + (\big[@setI _/setT]_(m <= i < n | P%B) F%SET) : set_scope. +Notation "\bigcap_ ( m <= i < n ) F" := + (\big[@setI _/setT]_(m <= i < n) F%SET) : set_scope. +Notation "\bigcap_ ( i | P ) F" := + (\big[@setI _/setT]_(i | P%B) F%SET) : set_scope. +Notation "\bigcap_ i F" := + (\big[@setI _/setT]_i F%SET) : set_scope. +Notation "\bigcap_ ( i : t | P ) F" := + (\big[@setI _/setT]_(i : t | P%B) F%SET) (only parsing): set_scope. +Notation "\bigcap_ ( i : t ) F" := + (\big[@setI _/setT]_(i : t) F%SET) (only parsing) : set_scope. +Notation "\bigcap_ ( i < n | P ) F" := + (\big[@setI _/setT]_(i < n | P%B) F%SET) : set_scope. +Notation "\bigcap_ ( i < n ) F" := + (\big[@setI _/setT]_(i < n) F%SET) : set_scope. +Notation "\bigcap_ ( i 'in' A | P ) F" := + (\big[@setI _/setT]_(i in A | P%B) F%SET) : set_scope. +Notation "\bigcap_ ( i 'in' A ) F" := + (\big[@setI _/setT]_(i in A) F%SET) : set_scope. + +Section BigSetOps. + +Variables T I : finType. +Implicit Types (U : pred T) (P : pred I) (A B : {set I}) (F : I -> {set T}). + +(* It is very hard to use this lemma, because the unification fails to *) +(* defer the F j pattern (even though it's a Miller pattern!). *) +Lemma bigcup_sup j P F : P j -> F j \subset \bigcup_(i | P i) F i. +Proof. by move=> Pj; rewrite (bigD1 j) //= subsetUl. Qed. + +Lemma bigcup_max j U P F : + P j -> U \subset F j -> U \subset \bigcup_(i | P i) F i. +Proof. by move=> Pj sUF; exact: subset_trans (bigcup_sup _ Pj). Qed. + +Lemma bigcupP x P F : + reflect (exists2 i, P i & x \in F i) (x \in \bigcup_(i | P i) F i). +Proof. +apply: (iffP idP) => [|[i Pi]]; last first. + apply: subsetP x; exact: bigcup_sup. +by elim/big_rec: _ => [|i _ Pi _ /setUP[|//]]; [rewrite inE | exists i]. +Qed. + +Lemma bigcupsP U P F : + reflect (forall i, P i -> F i \subset U) (\bigcup_(i | P i) F i \subset U). +Proof. +apply: (iffP idP) => [sFU i Pi| sFU]. + by apply: subset_trans sFU; exact: bigcup_sup. +by apply/subsetP=> x /bigcupP[i Pi]; exact: (subsetP (sFU i Pi)). +Qed. + +Lemma bigcup_disjoint U P F : + (forall i, P i -> [disjoint U & F i]) -> [disjoint U & \bigcup_(i | P i) F i]. +Proof. +move=> dUF; rewrite disjoint_sym disjoint_subset. +by apply/bigcupsP=> i /dUF; rewrite disjoint_sym disjoint_subset. +Qed. + +Lemma bigcup_setU A B F : + \bigcup_(i in A :|: B) F i = + (\bigcup_(i in A) F i) :|: (\bigcup_ (i in B) F i). +Proof. +apply/setP=> x; apply/bigcupP/setUP=> [[i] | ]. + by case/setUP; [left | right]; apply/bigcupP; exists i. +by case=> /bigcupP[i Pi]; exists i; rewrite // inE Pi ?orbT. +Qed. + +Lemma bigcup_seq r F : \bigcup_(i <- r) F i = \bigcup_(i in r) F i. +Proof. +elim: r => [|i r IHr]; first by rewrite big_nil big_pred0. +rewrite big_cons {}IHr; case r_i: (i \in r). + rewrite (setUidPr _) ?bigcup_sup //. + by apply: eq_bigl => j; rewrite !inE; case: eqP => // ->. +rewrite (bigD1 i (mem_head i r)) /=; congr (_ :|: _). +by apply: eq_bigl => j /=; rewrite andbC; case: eqP => // ->. +Qed. + +(* Unlike its setU counterpart, this lemma is useable. *) +Lemma bigcap_inf j P F : P j -> \bigcap_(i | P i) F i \subset F j. +Proof. by move=> Pj; rewrite (bigD1 j) //= subsetIl. Qed. + +Lemma bigcap_min j U P F : + P j -> F j \subset U -> \bigcap_(i | P i) F i \subset U. +Proof. by move=> Pj; exact: subset_trans (bigcap_inf _ Pj). Qed. + +Lemma bigcapsP U P F : + reflect (forall i, P i -> U \subset F i) (U \subset \bigcap_(i | P i) F i). +Proof. +apply: (iffP idP) => [sUF i Pi | sUF]. + apply: subset_trans sUF _; exact: bigcap_inf. +elim/big_rec: _ => [|i V Pi sUV]; apply/subsetP=> x Ux; rewrite inE //. +by rewrite !(subsetP _ x Ux) ?sUF. +Qed. + +Lemma bigcapP x P F : + reflect (forall i, P i -> x \in F i) (x \in \bigcap_(i | P i) F i). +Proof. +rewrite -sub1set. +by apply: (iffP (bigcapsP _ _ _)) => Fx i /Fx; rewrite sub1set. +Qed. + +Lemma setC_bigcup J r (P : pred J) (F : J -> {set T}) : + ~: (\bigcup_(j <- r | P j) F j) = \bigcap_(j <- r | P j) ~: F j. +Proof. by apply: big_morph => [A B|]; rewrite ?setC0 ?setCU. Qed. + +Lemma setC_bigcap J r (P : pred J) (F : J -> {set T}) : + ~: (\bigcap_(j <- r | P j) F j) = \bigcup_(j <- r | P j) ~: F j. +Proof. by apply: big_morph => [A B|]; rewrite ?setCT ?setCI. Qed. + +Lemma bigcap_setU A B F : + (\bigcap_(i in A :|: B) F i) = + (\bigcap_(i in A) F i) :&: (\bigcap_(i in B) F i). +Proof. by apply: setC_inj; rewrite setCI !setC_bigcap bigcup_setU. Qed. + +Lemma bigcap_seq r F : \bigcap_(i <- r) F i = \bigcap_(i in r) F i. +Proof. by apply: setC_inj; rewrite !setC_bigcap bigcup_seq. Qed. + +End BigSetOps. + +Implicit Arguments bigcup_sup [T I P F]. +Implicit Arguments bigcup_max [T I U P F]. +Implicit Arguments bigcupP [T I x P F]. +Implicit Arguments bigcupsP [T I U P F]. +Implicit Arguments bigcap_inf [T I P F]. +Implicit Arguments bigcap_min [T I U P F]. +Implicit Arguments bigcapP [T I x P F]. +Implicit Arguments bigcapsP [T I U P F]. +Prenex Implicits bigcupP bigcupsP bigcapP bigcapsP. + +Section ImsetCurry. + +Variables (aT1 aT2 rT : finType) (f : aT1 -> aT2 -> rT). + +Section Curry. + +Variables (A1 : {set aT1}) (A2 : {set aT2}). +Variables (D1 : pred aT1) (D2 : pred aT2). + +Lemma curry_imset2X : f @2: (A1, A2) = prod_curry f @: (setX A1 A2). +Proof. +rewrite [@imset]unlock unlock; apply/setP=> x; rewrite !in_set; congr (x \in _). +by apply: eq_image => u //=; rewrite !inE. +Qed. + +Lemma curry_imset2l : f @2: (D1, D2) = \bigcup_(x1 in D1) f x1 @: D2. +Proof. +apply/setP=> y; apply/imset2P/bigcupP => [[x1 x2 Dx1 Dx2 ->{y}] | [x1 Dx1]]. + by exists x1; rewrite // mem_imset. +by case/imsetP=> x2 Dx2 ->{y}; exists x1 x2. +Qed. + +Lemma curry_imset2r : f @2: (D1, D2) = \bigcup_(x2 in D2) f^~ x2 @: D1. +Proof. +apply/setP=> y; apply/imset2P/bigcupP => [[x1 x2 Dx1 Dx2 ->{y}] | [x2 Dx2]]. + by exists x2; rewrite // (mem_imset (f^~ x2)). +by case/imsetP=> x1 Dx1 ->{y}; exists x1 x2. +Qed. + +End Curry. + +Lemma imset2Ul (A B : {set aT1}) (C : {set aT2}) : + f @2: (A :|: B, C) = f @2: (A, C) :|: f @2: (B, C). +Proof. by rewrite !curry_imset2l bigcup_setU. Qed. + +Lemma imset2Ur (A : {set aT1}) (B C : {set aT2}) : + f @2: (A, B :|: C) = f @2: (A, B) :|: f @2: (A, C). +Proof. by rewrite !curry_imset2r bigcup_setU. Qed. + +End ImsetCurry. + +Section Partitions. + +Variables T I : finType. +Implicit Types (x y z : T) (A B D X : {set T}) (P Q : {set {set T}}). +Implicit Types (J : pred I) (F : I -> {set T}). + +Definition cover P := \bigcup_(B in P) B. +Definition pblock P x := odflt set0 (pick [pred B in P | x \in B]). +Definition trivIset P := \sum_(B in P) #|B| == #|cover P|. +Definition partition P D := [&& cover P == D, trivIset P & set0 \notin P]. + +Definition is_transversal X P D := + [&& partition P D, X \subset D & [forall B in P, #|X :&: B| == 1]]. +Definition transversal P D := [set odflt x [pick y in pblock P x] | x in D]. +Definition transversal_repr x0 X B := odflt x0 [pick x in X :&: B]. + +Lemma leq_card_setU A B : #|A :|: B| <= #|A| + #|B| ?= iff [disjoint A & B]. +Proof. +rewrite -(addn0 #|_|) -setI_eq0 -cards_eq0 -cardsUI eq_sym. +by rewrite (mono_leqif (leq_add2l _)). +Qed. + +Lemma leq_card_cover P : #|cover P| <= \sum_(A in P) #|A| ?= iff trivIset P. +Proof. +split; last exact: eq_sym. +rewrite /cover; elim/big_rec2: _ => [|A n U _ leUn]; first by rewrite cards0. +by rewrite (leq_trans (leq_card_setU A U).1) ?leq_add2l. +Qed. + +Lemma trivIsetP P : + reflect {in P &, forall A B, A != B -> [disjoint A & B]} (trivIset P). +Proof. +have->: P = [set x in enum (mem P)] by apply/setP=> x; rewrite inE mem_enum. +elim: {P}(enum _) (enum_uniq (mem P)) => [_ | A e IHe] /=. + by rewrite /trivIset /cover !big_set0 cards0; left=> A; rewrite inE. +case/andP; rewrite set_cons -(in_set (fun B => B \in e)) => PA {IHe}/IHe. +move: {e}[set x in e] PA => P PA IHP. +rewrite /trivIset /cover !big_setU1 //= eq_sym. +have:= leq_card_cover P; rewrite -(mono_leqif (leq_add2l #|A|)). +move/(leqif_trans (leq_card_setU _ _))->; rewrite disjoints_subset setC_bigcup. +case: bigcapsP => [disjA | meetA]; last first. + right=> [tI]; case: meetA => B PB; rewrite -disjoints_subset. + by rewrite tI ?setU11 ?setU1r //; apply: contraNneq PA => ->. +apply: (iffP IHP) => [] tI B C PB PC; last by apply: tI; exact: setU1r. +by case/setU1P: PC PB => [->|PC] /setU1P[->|PB]; try by [exact: tI | case/eqP]; + first rewrite disjoint_sym; rewrite disjoints_subset disjA. +Qed. + +Lemma trivIsetS P Q : P \subset Q -> trivIset Q -> trivIset P. +Proof. by move/subsetP/sub_in2=> sPQ /trivIsetP/sPQ/trivIsetP. Qed. + +Lemma trivIsetI P D : trivIset P -> trivIset (P ::&: D). +Proof. by apply: trivIsetS; rewrite -setI_powerset subsetIl. Qed. + +Lemma cover_setI P D : cover (P ::&: D) \subset cover P :&: D. +Proof. +by apply/bigcupsP=> A /setIdP[PA sAD]; rewrite subsetI sAD andbT (bigcup_max A). +Qed. + +Lemma mem_pblock P x : (x \in pblock P x) = (x \in cover P). +Proof. +rewrite /pblock; apply/esym/bigcupP. +case: pickP => /= [A /andP[PA Ax]| noA]; first by rewrite Ax; exists A. +by rewrite inE => [[A PA Ax]]; case/andP: (noA A). +Qed. + +Lemma pblock_mem P x : x \in cover P -> pblock P x \in P. +Proof. +by rewrite -mem_pblock /pblock; case: pickP => [A /andP[]| _] //=; rewrite inE. +Qed. + +Lemma def_pblock P B x : trivIset P -> B \in P -> x \in B -> pblock P x = B. +Proof. +move/trivIsetP=> tiP PB Bx; have Px: x \in cover P by apply/bigcupP; exists B. +apply: (contraNeq (tiP _ _ _ PB)); first by rewrite pblock_mem. +by apply/pred0Pn; exists x; rewrite /= mem_pblock Px. +Qed. + +Lemma same_pblock P x y : + trivIset P -> x \in pblock P y -> pblock P x = pblock P y. +Proof. +rewrite {1 3}/pblock => tI; case: pickP => [A|]; last by rewrite inE. +by case/andP=> PA _{y} /= Ax; exact: def_pblock. +Qed. + +Lemma eq_pblock P x y : + trivIset P -> x \in cover P -> + (pblock P x == pblock P y) = (y \in pblock P x). +Proof. +move=> tiP Px; apply/eqP/idP=> [eq_xy | /same_pblock-> //]. +move: Px; rewrite -mem_pblock eq_xy /pblock. +by case: pickP => [B /andP[] // | _]; rewrite inE. +Qed. + +Lemma trivIsetU1 A P : + {in P, forall B, [disjoint A & B]} -> trivIset P -> set0 \notin P -> + trivIset (A |: P) /\ A \notin P. +Proof. +move=> tiAP tiP notPset0; split; last first. + apply: contra notPset0 => P_A. + by have:= tiAP A P_A; rewrite -setI_eq0 setIid => /eqP <-. +apply/trivIsetP=> B1 B2 /setU1P[->|PB1] /setU1P[->|PB2]; + by [exact: (trivIsetP _ tiP) | rewrite ?eqxx // ?(tiAP, disjoint_sym)]. +Qed. + +Lemma cover_imset J F : cover (F @: J) = \bigcup_(i in J) F i. +Proof. +apply/setP=> x. +apply/bigcupP/bigcupP=> [[_ /imsetP[i Ji ->]] | [i]]; first by exists i. +by exists (F i); first exact: mem_imset. +Qed. + +Lemma trivIimset J F (P := F @: J) : + {in J &, forall i j, j != i -> [disjoint F i & F j]} -> set0 \notin P -> + trivIset P /\ {in J &, injective F}. +Proof. +move=> tiF notPset0; split=> [|i j Ji Jj /= eqFij]. + apply/trivIsetP=> _ _ /imsetP[i Ji ->] /imsetP[j Jj ->] neqFij. + by rewrite tiF // (contraNneq _ neqFij) // => ->. +apply: contraNeq notPset0 => neq_ij; apply/imsetP; exists i => //; apply/eqP. +by rewrite eq_sym -[F i]setIid setI_eq0 {1}eqFij tiF. +Qed. + +Lemma cover_partition P D : partition P D -> cover P = D. +Proof. by case/and3P=> /eqP. Qed. + +Lemma card_partition P D : partition P D -> #|D| = \sum_(A in P) #|A|. +Proof. by case/and3P=> /eqP <- /eqnP. Qed. + +Lemma card_uniform_partition n P D : + {in P, forall A, #|A| = n} -> partition P D -> #|D| = #|P| * n. +Proof. +by move=> uniP /card_partition->; rewrite -sum_nat_const; exact: eq_bigr. +Qed. + +Section BigOps. + +Variables (R : Type) (idx : R) (op : Monoid.com_law idx). +Let rhs_cond P K E := \big[op/idx]_(A in P) \big[op/idx]_(x in A | K x) E x. +Let rhs P E := \big[op/idx]_(A in P) \big[op/idx]_(x in A) E x. + +Lemma big_trivIset_cond P (K : pred T) (E : T -> R) : + trivIset P -> \big[op/idx]_(x in cover P | K x) E x = rhs_cond P K E. +Proof. +move=> tiP; rewrite (partition_big (pblock P) (mem P)) -/op => /= [|x]. + apply: eq_bigr => A PA; apply: eq_bigl => x; rewrite andbAC; congr (_ && _). + rewrite -mem_pblock; apply/andP/idP=> [[Px /eqP <- //] | Ax]. + by rewrite (def_pblock tiP PA Ax). +by case/andP=> Px _; exact: pblock_mem. +Qed. + +Lemma big_trivIset P (E : T -> R) : + trivIset P -> \big[op/idx]_(x in cover P) E x = rhs P E. +Proof. +have biginT := eq_bigl _ _ (fun _ => andbT _) => tiP. +by rewrite -biginT big_trivIset_cond //; apply: eq_bigr => A _; exact: biginT. +Qed. + +Lemma set_partition_big_cond P D (K : pred T) (E : T -> R) : + partition P D -> \big[op/idx]_(x in D | K x) E x = rhs_cond P K E. +Proof. by case/and3P=> /eqP <- tI_P _; exact: big_trivIset_cond. Qed. + +Lemma set_partition_big P D (E : T -> R) : + partition P D -> \big[op/idx]_(x in D) E x = rhs P E. +Proof. by case/and3P=> /eqP <- tI_P _; exact: big_trivIset. Qed. + +Lemma partition_disjoint_bigcup (F : I -> {set T}) E : + (forall i j, i != j -> [disjoint F i & F j]) -> + \big[op/idx]_(x in \bigcup_i F i) E x = + \big[op/idx]_i \big[op/idx]_(x in F i) E x. +Proof. +move=> disjF; pose P := [set F i | i in I & F i != set0]. +have trivP: trivIset P. + apply/trivIsetP=> _ _ /imsetP[i _ ->] /imsetP[j _ ->] neqFij. + by apply: disjF; apply: contraNneq neqFij => ->. +have ->: \bigcup_i F i = cover P. + apply/esym; rewrite cover_imset big_mkcond; apply: eq_bigr => i _. + by rewrite inE; case: eqP. +rewrite big_trivIset // /rhs big_imset => [|i j _ /setIdP[_ notFj0] eqFij]. + rewrite big_mkcond; apply: eq_bigr => i _; rewrite inE. + by case: eqP => //= ->; rewrite big_set0. +by apply: contraNeq (disjF _ _) _; rewrite -setI_eq0 eqFij setIid. +Qed. + +End BigOps. + +Section Equivalence. + +Variables (R : rel T) (D : {set T}). + +Let Px x := [set y in D | R x y]. +Definition equivalence_partition := [set Px x | x in D]. +Local Notation P := equivalence_partition. +Hypothesis eqiR : {in D & &, equivalence_rel R}. + +Let Pxx x : x \in D -> x \in Px x. +Proof. by move=> Dx; rewrite !inE Dx (eqiR Dx Dx). Qed. +Let PPx x : x \in D -> Px x \in P := fun Dx => mem_imset _ Dx. + +Lemma equivalence_partitionP : partition P D. +Proof. +have defD: cover P == D. + rewrite eqEsubset; apply/andP; split. + by apply/bigcupsP=> _ /imsetP[x Dx ->]; rewrite /Px setIdE subsetIl. + by apply/subsetP=> x Dx; apply/bigcupP; exists (Px x); rewrite (Pxx, PPx). +have tiP: trivIset P. + apply/trivIsetP=> _ _ /imsetP[x Dx ->] /imsetP[y Dy ->]; apply: contraR. + case/pred0Pn=> z /andP[]; rewrite !inE => /andP[Dz Rxz] /andP[_ Ryz]. + apply/eqP/setP=> t; rewrite !inE; apply: andb_id2l => Dt. + by rewrite (eqiR Dx Dz Dt) // (eqiR Dy Dz Dt). +rewrite /partition tiP defD /=. +by apply/imsetP=> [[x /Pxx Px_x Px0]]; rewrite -Px0 inE in Px_x. +Qed. + +Lemma pblock_equivalence_partition : + {in D &, forall x y, (y \in pblock P x) = R x y}. +Proof. +have [_ tiP _] := and3P equivalence_partitionP. +by move=> x y Dx Dy; rewrite /= (def_pblock tiP (PPx Dx) (Pxx Dx)) inE Dy. +Qed. + +End Equivalence. + +Lemma pblock_equivalence P D : + partition P D -> {in D & &, equivalence_rel (fun x y => y \in pblock P x)}. +Proof. +case/and3P=> /eqP <- tiP _ x y z Px Py Pz. +by rewrite mem_pblock; split=> // /same_pblock->. +Qed. + +Lemma equivalence_partition_pblock P D : + partition P D -> equivalence_partition (fun x y => y \in pblock P x) D = P. +Proof. +case/and3P=> /eqP <-{D} tiP notP0; apply/setP=> B /=; set D := cover P. +have defP x: x \in D -> [set y in D | y \in pblock P x] = pblock P x. + by move=> Dx; apply/setIidPr; rewrite (bigcup_max (pblock P x)) ?pblock_mem. +apply/imsetP/idP=> [[x Px ->{B}] | PB]; first by rewrite defP ?pblock_mem. +have /set0Pn[x Bx]: B != set0 := memPn notP0 B PB. +have Px: x \in cover P by apply/bigcupP; exists B. +by exists x; rewrite // defP // (def_pblock tiP PB Bx). +Qed. + +Section Preim. + +Variables (rT : eqType) (f : T -> rT). + +Definition preim_partition := equivalence_partition (fun x y => f x == f y). + +Lemma preim_partitionP D : partition (preim_partition D) D. +Proof. by apply/equivalence_partitionP; split=> // /eqP->. Qed. + +End Preim. + +Lemma preim_partition_pblock P D : + partition P D -> preim_partition (pblock P) D = P. +Proof. +move=> partP; have [/eqP defD tiP _] := and3P partP. +rewrite -{2}(equivalence_partition_pblock partP); apply: eq_in_imset => x Dx. +by apply/setP=> y; rewrite !inE eq_pblock ?defD. +Qed. + +Lemma transversalP P D : partition P D -> is_transversal (transversal P D) P D. +Proof. +case/and3P=> /eqP <- tiP notP0; apply/and3P; split; first exact/and3P. + apply/subsetP=> _ /imsetP[x Px ->]; case: pickP => //= y Pxy. + by apply/bigcupP; exists (pblock P x); rewrite ?pblock_mem //. +apply/forall_inP=> B PB; have /set0Pn[x Bx]: B != set0 := memPn notP0 B PB. +apply/cards1P; exists (odflt x [pick y in pblock P x]); apply/esym/eqP. +rewrite eqEsubset sub1set inE -andbA; apply/andP; split. + by apply/mem_imset/bigcupP; exists B. +rewrite (def_pblock tiP PB Bx); case def_y: _ / pickP => [y By | /(_ x)/idP//]. +rewrite By /=; apply/subsetP=> _ /setIP[/imsetP[z Pz ->]]. +case: {1}_ / pickP => [t zPt Bt | /(_ z)/idP[]]; last by rewrite mem_pblock. +by rewrite -(same_pblock tiP zPt) (def_pblock tiP PB Bt) def_y set11. +Qed. + +Section Transversals. + +Variables (X : {set T}) (P : {set {set T}}) (D : {set T}). +Hypothesis trPX : is_transversal X P D. + +Lemma transversal_sub : X \subset D. Proof. by case/and3P: trPX. Qed. + +Let tiP : trivIset P. Proof. by case/andP: trPX => /and3P[]. Qed. + +Let sXP : {subset X <= cover P}. +Proof. by case/and3P: trPX => /andP[/eqP-> _] /subsetP. Qed. + +Let trX : {in P, forall B, #|X :&: B| == 1}. +Proof. by case/and3P: trPX => _ _ /forall_inP. Qed. + +Lemma setI_transversal_pblock x0 B : + B \in P -> X :&: B = [set transversal_repr x0 X B]. +Proof. +by case/trX/cards1P=> x defXB; rewrite /transversal_repr defXB /pick enum_set1. +Qed. + +Lemma repr_mem_pblock x0 B : B \in P -> transversal_repr x0 X B \in B. +Proof. by move=> PB; rewrite -sub1set -setI_transversal_pblock ?subsetIr. Qed. + +Lemma repr_mem_transversal x0 B : B \in P -> transversal_repr x0 X B \in X. +Proof. by move=> PB; rewrite -sub1set -setI_transversal_pblock ?subsetIl. Qed. + +Lemma transversal_reprK x0 : {in P, cancel (transversal_repr x0 X) (pblock P)}. +Proof. by move=> B PB; rewrite /= (def_pblock tiP PB) ?repr_mem_pblock. Qed. + +Lemma pblockK x0 : {in X, cancel (pblock P) (transversal_repr x0 X)}. +Proof. +move=> x Xx; have /bigcupP[B PB Bx] := sXP Xx; rewrite (def_pblock tiP PB Bx). +by apply/esym/set1P; rewrite -setI_transversal_pblock // inE Xx. +Qed. + +Lemma pblock_inj : {in X &, injective (pblock P)}. +Proof. by move=> x0; exact: (can_in_inj (pblockK x0)). Qed. + +Lemma pblock_transversal : pblock P @: X = P. +Proof. +apply/setP=> B; apply/imsetP/idP=> [[x Xx ->] | PB]. + by rewrite pblock_mem ?sXP. +have /cards1P[x0 _] := trX PB; set x := transversal_repr x0 X B. +by exists x; rewrite ?transversal_reprK ?repr_mem_transversal. +Qed. + +Lemma card_transversal : #|X| = #|P|. +Proof. rewrite -pblock_transversal card_in_imset //; exact: pblock_inj. Qed. + +Lemma im_transversal_repr x0 : transversal_repr x0 X @: P = X. +Proof. +rewrite -{2}[X]imset_id -pblock_transversal -imset_comp. +by apply: eq_in_imset; exact: pblockK. +Qed. + +End Transversals. + +End Partitions. + +Implicit Arguments trivIsetP [T P]. +Implicit Arguments big_trivIset_cond [T R idx op K E]. +Implicit Arguments set_partition_big_cond [T R idx op D K E]. +Implicit Arguments big_trivIset [T R idx op E]. +Implicit Arguments set_partition_big [T R idx op D E]. + +Prenex Implicits cover trivIset partition pblock trivIsetP. + +Lemma partition_partition (T : finType) (D : {set T}) P Q : + partition P D -> partition Q P -> + partition (cover @: Q) D /\ {in Q &, injective cover}. +Proof. +move=> /and3P[/eqP defG tiP notP0] /and3P[/eqP defP tiQ notQ0]. +have sQP E: E \in Q -> {subset E <= P}. + by move=> Q_E; apply/subsetP; rewrite -defP (bigcup_max E). +rewrite /partition cover_imset -(big_trivIset _ tiQ) defP -defG eqxx /= andbC. +have{notQ0} notQ0: set0 \notin cover @: Q. + apply: contra notP0 => /imsetP[E Q_E E0]. + have /set0Pn[/= A E_A] := memPn notQ0 E Q_E. + congr (_ \in P): (sQP E Q_E A E_A). + by apply/eqP; rewrite -subset0 E0 (bigcup_max A). +rewrite notQ0; apply: trivIimset => // E F Q_E Q_F. +apply: contraR => /pred0Pn[x /andP[/bigcupP[A E_A Ax] /bigcupP[B F_B Bx]]]. +rewrite -(def_pblock tiQ Q_E E_A) -(def_pblock tiP _ Ax) ?(sQP E) //. +by rewrite -(def_pblock tiQ Q_F F_B) -(def_pblock tiP _ Bx) ?(sQP F). +Qed. + +(**********************************************************************) +(* *) +(* Maximum and minimun (sub)set with respect to a given pred *) +(* *) +(**********************************************************************) + +Section MaxSetMinSet. + +Variable T : finType. +Notation sT := {set T}. +Implicit Types A B C : sT. +Implicit Type P : pred sT. + +Definition minset P A := [forall (B : sT | B \subset A), (B == A) == P B]. + +Lemma minset_eq P1 P2 A : P1 =1 P2 -> minset P1 A = minset P2 A. +Proof. by move=> eP12; apply: eq_forallb => B; rewrite eP12. Qed. + +Lemma minsetP P A : + reflect ((P A) /\ (forall B, P B -> B \subset A -> B = A)) (minset P A). +Proof. +apply: (iffP forallP) => [minA | [PA minA] B]. + split; first by have:= minA A; rewrite subxx eqxx /= => /eqP. + by move=> B PB sBA; have:= minA B; rewrite PB sBA /= eqb_id => /eqP. +by apply/implyP=> sBA; apply/eqP; apply/eqP/idP=> [-> // | /minA]; exact. +Qed. +Implicit Arguments minsetP [P A]. + +Lemma minsetp P A : minset P A -> P A. +Proof. by case/minsetP. Qed. + +Lemma minsetinf P A B : minset P A -> P B -> B \subset A -> B = A. +Proof. by case/minsetP=> _; exact. Qed. + +Lemma ex_minset P : (exists A, P A) -> {A | minset P A}. +Proof. +move=> exP; pose pS n := [pred B | P B & #|B| == n]. +pose p n := ~~ pred0b (pS n); have{exP}: exists n, p n. + by case: exP => A PA; exists #|A|; apply/existsP; exists A; rewrite /= PA /=. +case/ex_minnP=> n /pred0P; case: (pickP (pS n)) => // A /andP[PA] /eqP <-{n} _. +move=> minA; exists A => //; apply/minsetP; split=> // B PB sBA; apply/eqP. +by rewrite eqEcard sBA minA //; apply/pred0Pn; exists B; rewrite /= PB /=. +Qed. + +Lemma minset_exists P C : P C -> {A | minset P A & A \subset C}. +Proof. +move=> PC; have{PC}: exists A, P A && (A \subset C) by exists C; rewrite PC /=. +case/ex_minset=> A /minsetP[/andP[PA sAC] minA]; exists A => //; apply/minsetP. +by split=> // B PB sBA; rewrite (minA B) // PB (subset_trans sBA). +Qed. + +(* The 'locked_with' allows Coq to find the value of P by unification. *) +Fact maxset_key : unit. Proof. by []. Qed. +Definition maxset P A := + minset (fun B => locked_with maxset_key P (~: B)) (~: A). + +Lemma maxset_eq P1 P2 A : P1 =1 P2 -> maxset P1 A = maxset P2 A. +Proof. by move=> eP12; apply: minset_eq => x /=; rewrite !unlock_with eP12. Qed. + +Lemma maxminset P A : maxset P A = minset [pred B | P (~: B)] (~: A). +Proof. by rewrite /maxset unlock. Qed. + +Lemma minmaxset P A : minset P A = maxset [pred B | P (~: B)] (~: A). +Proof. +by rewrite /maxset unlock setCK; apply: minset_eq => B /=; rewrite setCK. +Qed. + +Lemma maxsetP P A : + reflect ((P A) /\ (forall B, P B -> A \subset B -> B = A)) (maxset P A). +Proof. +apply: (iffP minsetP); rewrite ?setCK unlock_with => [] [PA minA]. + by split=> // B PB sAB; rewrite -[B]setCK [~: B]minA (setCK, setCS). +by split=> // B PB' sBA'; rewrite -(minA _ PB') -1?setCS setCK. +Qed. + +Lemma maxsetp P A : maxset P A -> P A. +Proof. by case/maxsetP. Qed. + +Lemma maxsetsup P A B : maxset P A -> P B -> A \subset B -> B = A. +Proof. by case/maxsetP=> _; exact. Qed. + +Lemma ex_maxset P : (exists A, P A) -> {A | maxset P A}. +Proof. +move=> exP; have{exP}: exists A, P (~: A). + by case: exP => A PA; exists (~: A); rewrite setCK. +by case/ex_minset=> A minA; exists (~: A); rewrite /maxset unlock setCK. +Qed. + +Lemma maxset_exists P C : P C -> {A : sT | maxset P A & C \subset A}. +Proof. +move=> PC; pose P' B := P (~: B); have: P' (~: C) by rewrite /P' setCK. +case/minset_exists=> B; rewrite -[B]setCK setCS. +by exists (~: B); rewrite // /maxset unlock. +Qed. + +End MaxSetMinSet. + +Implicit Arguments minsetP [T P A]. +Implicit Arguments maxsetP [T P A]. +Prenex Implicits minset maxset minsetP maxsetP. + diff --git a/mathcomp/discrete/fintype.v b/mathcomp/discrete/fintype.v new file mode 100644 index 0000000..63d5e84 --- /dev/null +++ b/mathcomp/discrete/fintype.v @@ -0,0 +1,2037 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice. + +(******************************************************************************) +(* The Finite interface describes Types with finitely many elements, *) +(* supplying a duplicate-free sequence of all the elements. It is a subclass *) +(* of Countable and thus of Choice and Equality. As with Countable, the *) +(* interface explicitly includes these somewhat redundant superclasses to *) +(* ensure that Canonical instance inference remains consistent. Finiteness *) +(* could be stated more simply by bounding the range of the pickle function *) +(* supplied by the Countable interface, but this would yield a useless *) +(* computational interpretation due to the wasteful Peano integer encodings. *) +(* Because the Countable interface is closely tied to the Finite interface *) +(* and is not much used on its own, the Countable mixin is included inside *) +(* the Finite mixin; this makes it much easier to derive Finite variants of *) +(* interfaces, in this file for subFinType, and in the finalg library. *) +(* We define the following interfaces and structures: *) +(* finType == the packed class type of the Finite interface. *) +(* FinType m == the packed class for the Finite mixin m. *) +(* Finite.axiom e == every x : T occurs exactly once in e : seq T. *) +(* FinMixin ax_e == the Finite mixin for T, encapsulating *) +(* ax_e : Finite.axiom e for some e : seq T. *) +(* UniqFinMixin uniq_e total_e == an alternative mixin constructor that uses *) +(* uniq_e : uniq e and total_e : e =i xpredT. *) +(* PcanFinMixin fK == the Finite mixin for T, given f : T -> fT and g with fT *) +(* a finType and fK : pcancel f g. *) +(* CanFinMixin fK == the Finite mixin for T, given f : T -> fT and g with fT *) +(* a finType and fK : cancel f g. *) +(* subFinType == the join interface type for subType and finType. *) +(* [finType of T for fT] == clone for T of the finType fT. *) +(* [finType of T] == clone for T of the finType inferred for T. *) +(* [subFinType of T] == a subFinType structure for T, when T already has both *) +(* finType and subType structures. *) +(* [finMixin of T by <:] == a finType structure for T, when T has a subType *) +(* structure over an existing finType. *) +(* We define or propagate the finType structure appropriately for all basic *) +(* basic types : unit, bool, option, prod, sum, sig and sigT. We also define *) +(* a generic type constructor for finite subtypes based on an explicit *) +(* enumeration: *) +(* seq_sub s == the subType of all x \in s, where s : seq T and T has *) +(* a choiceType structure; the seq_sub s type has a *) +(* canonical finType structure. *) +(* Bounded integers are supported by the following type and operations: *) +(* 'I_n, ordinal n == the finite subType of integers i < n, whose *) +(* enumeration is {0, ..., n.-1}. 'I_n coerces to nat, *) +(* so all the integer arithmetic functions can be used *) +(* with 'I_n. *) +(* Ordinal lt_i_n == the element of 'I_n with (nat) value i, given *) +(* lt_i_n : i < n. *) +(* nat_of_ord i == the nat value of i : 'I_n (this function is a *) +(* coercion so it is not usually displayed). *) +(* ord_enum n == the explicit increasing sequence of the i : 'I_n. *) +(* cast_ord eq_n_m i == the element j : 'I_m with the same value as i : 'I_n *) +(* given eq_n_m : n = m (indeed, i : nat and j : nat *) +(* are convertible). *) +(* widen_ord le_n_m i == a j : 'I_m with the same value as i : 'I_n, given *) +(* le_n_m : n <= m. *) +(* rev_ord i == the complement to n.-1 of i : 'I_n, such that *) +(* i + rev_ord i = n.-1. *) +(* inord k == the i : 'I_n.+1 with value k (n is inferred from the *) +(* context). *) +(* sub_ord k == the i : 'I_n.+1 with value n - k (n is inferred from *) +(* the context). *) +(* ord0 == the i : 'I_n.+1 with value 0 (n is inferred from the *) +(* context). *) +(* ord_max == the i : 'I_n.+1 with value n (n is inferred from the *) +(* context). *) +(* bump h k == k.+1 if k >= h, else k (this is a nat function). *) +(* unbump h k == k.-1 if k > h, else k (this is a nat function). *) +(* lift i j == the j' : 'I_n with value bump i j, where i : 'I_n *) +(* and j : 'I_n.-1. *) +(* unlift i j == None if i = j, else Some j', where j' : 'I_n.-1 has *) +(* value unbump i j, given i, j : 'I_n. *) +(* lshift n j == the i : 'I_(m + n) with value j : 'I_m. *) +(* rshift m k == the i : 'I_(m + n) with value m + k, k : 'I_n. *) +(* unsplit u == either lshift n j or rshift m k, depending on *) +(* whether if u : 'I_m + 'I_n is inl j or inr k. *) +(* split i == the u : 'I_m + 'I_n such that i = unsplit u; the *) +(* type 'I_(m + n) of i determines the split. *) +(* Finally, every type T with a finType structure supports the following *) +(* operations: *) +(* enum A == a duplicate-free list of all the x \in A, where A is a *) +(* collective predicate over T. *) +(* #|A| == the cardinal of A, i.e., the number of x \in A. *) +(* enum_val i == the i'th item of enum A, where i : 'I_(#|A|). *) +(* enum_rank x == the i : 'I_(#|T|) such that enum_val i = x. *) +(* enum_rank_in Ax0 x == some i : 'I_(#|A|) such that enum_val i = x if *) +(* x \in A, given Ax0 : x0 \in A. *) +(* A \subset B == all x \in A satisfy x \in B. *) +(* A \proper B == all x \in A satisfy x \in B but not the converse. *) +(* [disjoint A & B] == no x \in A satisfies x \in B. *) +(* image f A == the sequence of f x for all x : T such that x \in A *) +(* (where a is an applicative predicate), of length #|P|. *) +(* The codomain of F can be any type, but image f A can *) +(* only be used as a collective predicate is it is an *) +(* eqType. *) +(* codom f == a sequence spanning the codomain of f (:= image f T). *) +(* [seq F | x : T in A] := image (fun x : T => F) A. *) +(* [seq F | x : T] := [seq F | x <- {: T}]. *) +(* [seq F | x in A], [seq F | x] == variants without casts. *) +(* iinv im_y == some x such that P x holds and f x = y, given *) +(* im_y : y \in image f P. *) +(* invF inj_f y == the x such that f x = y, for inj_j : injective f with *) +(* f : T -> T. *) +(* dinjectiveb A f == the restriction of f : T -> R to A is injective *) +(* (this is a bolean predicate, R must be an eqType). *) +(* injectiveb f == f : T -> R is injective (boolean predicate). *) +(* pred0b A == no x : T satisfies x \in A. *) +(* [forall x, P] == P (in which x can appear) is true for all values of x; *) +(* x must range over a finType. *) +(* [exists x, P] == P is true for some value of x. *) +(* [forall (x | C), P] := [forall x, C ==> P]. *) +(* [forall x in A, P] := [forall (x | x \in A), P]. *) +(* [exists (x | C), P] := [exists x, C && P]. *) +(* [exists x in A, P] := [exists (x | x \in A), P]. *) +(* and typed variants [forall x : T, P], [forall (x : T | C), P], *) +(* [exists x : T, P], [exists x : T in A, P], etc. *) +(* -> The outer brackets can be omitted when nesting finitary quantifiers, *) +(* e.g., [forall i in I, forall j in J, exists a, f i j == a]. *) +(* 'forall_pP == view for [forall x, p _], for pP : reflect .. (p _). *) +(* 'exists_pP == view for [exists x, p _], for pP : reflect .. (p _). *) +(* [pick x | P] == Some x, for an x such that P holds, or None if there *) +(* is no such x. *) +(* [pick x : T] == Some x with x : T, provided T is nonempty, else None. *) +(* [pick x in A] == Some x, with x \in A, or None if A is empty. *) +(* [pick x in A | P] == Some x, with x \in A s.t. P holds, else None. *) +(* [pick x | P & Q] := [pick x | P & Q]. *) +(* [pick x in A | P & Q] := [pick x | P & Q]. *) +(* and (un)typed variants [pick x : T | P], [pick x : T in A], [pick x], etc. *) +(* [arg min_(i < i0 | P) M] == a value of i : T minimizing M : nat, 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 of 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. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Module Finite. + +Section RawMixin. + +Variable T : eqType. + +Definition axiom e := forall x : T, count_mem x e = 1. + +Lemma uniq_enumP e : uniq e -> e =i T -> axiom e. +Proof. by move=> Ue sT x; rewrite count_uniq_mem ?sT. Qed. + +Record mixin_of := Mixin { + mixin_base : Countable.mixin_of T; + mixin_enum : seq T; + _ : axiom mixin_enum +}. + +End RawMixin. + +Section Mixins. + +Variable T : countType. + +Definition EnumMixin := + let: Countable.Pack _ (Countable.Class _ m) _ as cT := T + return forall e : seq cT, axiom e -> mixin_of cT in + @Mixin (EqType _ _) m. + +Definition UniqMixin e Ue eT := @EnumMixin e (uniq_enumP Ue eT). + +Variable n : nat. + +Definition count_enum := pmap (@pickle_inv T) (iota 0 n). + +Hypothesis ubT : forall x : T, pickle x < n. + +Lemma count_enumP : axiom count_enum. +Proof. +apply: uniq_enumP (pmap_uniq (@pickle_invK T) (iota_uniq _ _)) _ => x. +by rewrite mem_pmap -pickleK_inv map_f // mem_iota ubT. +Qed. + +Definition CountMixin := EnumMixin count_enumP. + +End Mixins. + +Section ClassDef. + +Record class_of T := Class { + base : Choice.class_of T; + mixin : mixin_of (Equality.Pack base T) +}. +Definition base2 T c := Countable.Class (@base T c) (mixin_base (mixin c)). +Local Coercion base : class_of >-> Choice.class_of. + +Structure type : Type := Pack {sort; _ : class_of sort; _ : Type}. +Local Coercion sort : type >-> Sortclass. +Variables (T : Type) (cT : type). +Definition class := let: Pack _ c _ as cT' := cT return class_of cT' in c. +Definition clone c of phant_id class c := @Pack T c T. +Let xT := let: Pack T _ _ := cT in T. +Notation xclass := (class : class_of xT). + +Definition pack b0 (m0 : mixin_of (EqType T b0)) := + fun bT b & phant_id (Choice.class bT) b => + fun m & phant_id m0 m => Pack (@Class T b m) T. + +Definition eqType := @Equality.Pack cT xclass xT. +Definition choiceType := @Choice.Pack cT xclass xT. +Definition countType := @Countable.Pack cT (base2 xclass) xT. + +End ClassDef. + +Module Import Exports. +Coercion mixin_base : mixin_of >-> Countable.mixin_of. +Coercion base : class_of >-> Choice.class_of. +Coercion mixin : class_of >-> mixin_of. +Coercion base2 : class_of >-> Countable.class_of. +Coercion sort : type >-> Sortclass. +Coercion eqType : type >-> Equality.type. +Canonical eqType. +Coercion choiceType : type >-> Choice.type. +Canonical choiceType. +Coercion countType : type >-> Countable.type. +Canonical countType. +Notation finType := type. +Notation FinType T m := (@pack T _ m _ _ id _ id). +Notation FinMixin := EnumMixin. +Notation UniqFinMixin := UniqMixin. +Notation "[ 'finType' 'of' T 'for' cT ]" := (@clone T cT _ idfun) + (at level 0, format "[ 'finType' 'of' T 'for' cT ]") : form_scope. +Notation "[ 'finType' 'of' T ]" := (@clone T _ _ id) + (at level 0, format "[ 'finType' 'of' T ]") : form_scope. +End Exports. + +Module Type EnumSig. +Parameter enum : forall cT : type, seq cT. +Axiom enumDef : enum = fun cT => mixin_enum (class cT). +End EnumSig. + +Module EnumDef : EnumSig. +Definition enum cT := mixin_enum (class cT). +Definition enumDef := erefl enum. +End EnumDef. + +Notation enum := EnumDef.enum. + +End Finite. +Export Finite.Exports. + +Canonical finEnum_unlock := Unlockable Finite.EnumDef.enumDef. + +(* Workaround for the silly syntactic uniformity restriction on coercions; *) +(* this avoids a cross-dependency between finset.v and prime.v for the *) +(* definition of the \pi(A) notation. *) +Definition fin_pred_sort (T : finType) (pT : predType T) := pred_sort pT. +Identity Coercion pred_sort_of_fin : fin_pred_sort >-> pred_sort. + +Definition enum_mem T (mA : mem_pred _) := filter mA (Finite.enum T). +Notation enum A := (enum_mem (mem A)). +Definition pick (T : finType) (P : pred T) := ohead (enum P). + +Notation "[ 'pick' x | P ]" := (pick (fun x => P%B)) + (at level 0, x ident, format "[ 'pick' x | P ]") : form_scope. +Notation "[ 'pick' x : T | P ]" := (pick (fun x : T => P%B)) + (at level 0, x ident, only parsing) : form_scope. +Definition pick_true T (x : T) := true. +Notation "[ 'pick' x : T ]" := [pick x : T | pick_true x] + (at level 0, x ident, only parsing). +Notation "[ 'pick' x ]" := [pick x : _] + (at level 0, x ident, only parsing) : form_scope. +Notation "[ 'pic' 'k' x : T ]" := [pick x : T | pick_true _] + (at level 0, x ident, format "[ 'pic' 'k' x : T ]") : form_scope. +Notation "[ 'pick' x | P & Q ]" := [pick x | P && Q ] + (at level 0, x ident, + format "[ '[hv ' 'pick' x | P '/ ' & Q ] ']'") : form_scope. +Notation "[ 'pick' x : T | P & Q ]" := [pick x : T | P && Q ] + (at level 0, x ident, only parsing) : form_scope. +Notation "[ 'pick' x 'in' A ]" := [pick x | x \in A] + (at level 0, x ident, format "[ 'pick' x 'in' A ]") : form_scope. +Notation "[ 'pick' x : T 'in' A ]" := [pick x : T | x \in A] + (at level 0, x ident, only parsing) : form_scope. +Notation "[ 'pick' x 'in' A | P ]" := [pick x | x \in A & P ] + (at level 0, x ident, + format "[ '[hv ' 'pick' x 'in' A '/ ' | P ] ']'") : form_scope. +Notation "[ 'pick' x : T 'in' A | P ]" := [pick x : T | x \in A & P ] + (at level 0, x ident, only parsing) : form_scope. +Notation "[ 'pick' x 'in' A | P & Q ]" := [pick x in A | P && Q] + (at level 0, x ident, format + "[ '[hv ' 'pick' x 'in' A '/ ' | P '/ ' & Q ] ']'") : form_scope. +Notation "[ 'pick' x : T 'in' A | P & Q ]" := [pick x : T in A | P && Q] + (at level 0, x ident, only parsing) : form_scope. + +(* We lock the definitions of card and subset to mitigate divergence of the *) +(* Coq term comparison algorithm. *) + +Local Notation card_type := (forall T : finType, mem_pred T -> nat). +Local Notation card_def := (fun T mA => size (enum_mem mA)). +Module Type CardDefSig. +Parameter card : card_type. Axiom cardEdef : card = card_def. +End CardDefSig. +Module CardDef : CardDefSig. +Definition card : card_type := card_def. Definition cardEdef := erefl card. +End CardDef. +(* Should be Include, but for a silly restriction: can't Include at toplevel! *) +Export CardDef. + +Canonical card_unlock := Unlockable cardEdef. +(* A is at level 99 to allow the notation #|G : H| in groups. *) +Notation "#| A |" := (card (mem A)) + (at level 0, A at level 99, format "#| A |") : nat_scope. + +Definition pred0b (T : finType) (P : pred T) := #|P| == 0. +Prenex Implicits pred0b. + +Module FiniteQuant. + +CoInductive quantified := Quantified of bool. + +Delimit Scope fin_quant_scope with Q. (* Bogus, only used to declare scope. *) +Bind Scope fin_quant_scope with quantified. + +Notation "F ^*" := (Quantified F) (at level 2). +Notation "F ^~" := (~~ F) (at level 2). + +Section Definitions. + +Variable T : finType. +Implicit Types (B : quantified) (x y : T). + +Definition quant0b Bp := pred0b [pred x : T | let: F^* := Bp x x in F]. +(* The first redundant argument protects the notation from Coq's K-term *) +(* display kludge; the second protects it from simpl and /=. *) +Definition ex B x y := B. +(* Binding the predicate value rather than projecting it prevents spurious *) +(* unfolding of the boolean connectives by unification. *) +Definition all B x y := let: F^* := B in F^~^*. +Definition all_in C B x y := let: F^* := B in (C ==> F)^~^*. +Definition ex_in C B x y := let: F^* := B in (C && F)^*. + +End Definitions. + +Notation "[ x | B ]" := (quant0b (fun x => B x)) (at level 0, x ident). +Notation "[ x : T | B ]" := (quant0b (fun x : T => B x)) (at level 0, x ident). + +Module Exports. + +Notation ", F" := F^* (at level 200, format ", '/ ' F") : fin_quant_scope. + +Notation "[ 'forall' x B ]" := [x | all B] + (at level 0, x at level 99, B at level 200, + format "[ '[hv' 'forall' x B ] ']'") : bool_scope. + +Notation "[ 'forall' x : T B ]" := [x : T | all B] + (at level 0, x at level 99, B at level 200, only parsing) : bool_scope. +Notation "[ 'forall' ( x | C ) B ]" := [x | all_in C B] + (at level 0, x at level 99, B at level 200, + format "[ '[hv' '[' 'forall' ( x '/ ' | C ) ']' B ] ']'") : bool_scope. +Notation "[ 'forall' ( x : T | C ) B ]" := [x : T | all_in C B] + (at level 0, x at level 99, B at level 200, only parsing) : bool_scope. +Notation "[ 'forall' x 'in' A B ]" := [x | all_in (x \in A) B] + (at level 0, x at level 99, B at level 200, + format "[ '[hv' '[' 'forall' x '/ ' 'in' A ']' B ] ']'") : bool_scope. +Notation "[ 'forall' x : T 'in' A B ]" := [x : T | all_in (x \in A) B] + (at level 0, x at level 99, B at level 200, only parsing) : bool_scope. +Notation ", 'forall' x B" := [x | all B]^* + (at level 200, x at level 99, B at level 200, + format ", '/ ' 'forall' x B") : fin_quant_scope. +Notation ", 'forall' x : T B" := [x : T | all B]^* + (at level 200, x at level 99, B at level 200, only parsing) : fin_quant_scope. +Notation ", 'forall' ( x | C ) B" := [x | all_in C B]^* + (at level 200, x at level 99, B at level 200, + format ", '/ ' '[' 'forall' ( x '/ ' | C ) ']' B") : fin_quant_scope. +Notation ", 'forall' ( x : T | C ) B" := [x : T | all_in C B]^* + (at level 200, x at level 99, B at level 200, only parsing) : fin_quant_scope. +Notation ", 'forall' x 'in' A B" := [x | all_in (x \in A) B]^* + (at level 200, x at level 99, B at level 200, + format ", '/ ' '[' 'forall' x '/ ' 'in' A ']' B") : bool_scope. +Notation ", 'forall' x : T 'in' A B" := [x : T | all_in (x \in A) B]^* + (at level 200, x at level 99, B at level 200, only parsing) : bool_scope. + +Notation "[ 'exists' x B ]" := [x | ex B]^~ + (at level 0, x at level 99, B at level 200, + format "[ '[hv' 'exists' x B ] ']'") : bool_scope. +Notation "[ 'exists' x : T B ]" := [x : T | ex B]^~ + (at level 0, x at level 99, B at level 200, only parsing) : bool_scope. +Notation "[ 'exists' ( x | C ) B ]" := [x | ex_in C B]^~ + (at level 0, x at level 99, B at level 200, + format "[ '[hv' '[' 'exists' ( x '/ ' | C ) ']' B ] ']'") : bool_scope. +Notation "[ 'exists' ( x : T | C ) B ]" := [x : T | ex_in C B]^~ + (at level 0, x at level 99, B at level 200, only parsing) : bool_scope. +Notation "[ 'exists' x 'in' A B ]" := [x | ex_in (x \in A) B]^~ + (at level 0, x at level 99, B at level 200, + format "[ '[hv' '[' 'exists' x '/ ' 'in' A ']' B ] ']'") : bool_scope. +Notation "[ 'exists' x : T 'in' A B ]" := [x : T | ex_in (x \in A) B]^~ + (at level 0, x at level 99, B at level 200, only parsing) : bool_scope. +Notation ", 'exists' x B" := [x | ex B]^~^* + (at level 200, x at level 99, B at level 200, + format ", '/ ' 'exists' x B") : fin_quant_scope. +Notation ", 'exists' x : T B" := [x : T | ex B]^~^* + (at level 200, x at level 99, B at level 200, only parsing) : fin_quant_scope. +Notation ", 'exists' ( x | C ) B" := [x | ex_in C B]^~^* + (at level 200, x at level 99, B at level 200, + format ", '/ ' '[' 'exists' ( x '/ ' | C ) ']' B") : fin_quant_scope. +Notation ", 'exists' ( x : T | C ) B" := [x : T | ex_in C B]^~^* + (at level 200, x at level 99, B at level 200, only parsing) : fin_quant_scope. +Notation ", 'exists' x 'in' A B" := [x | ex_in (x \in A) B]^~^* + (at level 200, x at level 99, B at level 200, + format ", '/ ' '[' 'exists' x '/ ' 'in' A ']' B") : bool_scope. +Notation ", 'exists' x : T 'in' A B" := [x : T | ex_in (x \in A) B]^~^* + (at level 200, x at level 99, B at level 200, only parsing) : bool_scope. + +End Exports. + +End FiniteQuant. +Export FiniteQuant.Exports. + +Definition disjoint T (A B : mem_pred _) := @pred0b T (predI A B). +Notation "[ 'disjoint' A & B ]" := (disjoint (mem A) (mem B)) + (at level 0, + format "'[hv' [ 'disjoint' '/ ' A '/' & B ] ']'") : bool_scope. + +Notation Local subset_type := (forall (T : finType) (A B : mem_pred T), bool). +Notation Local subset_def := (fun T A B => pred0b (predD A B)). +Module Type SubsetDefSig. +Parameter subset : subset_type. Axiom subsetEdef : subset = subset_def. +End SubsetDefSig. +Module Export SubsetDef : SubsetDefSig. +Definition subset : subset_type := subset_def. +Definition subsetEdef := erefl subset. +End SubsetDef. +Canonical subset_unlock := Unlockable subsetEdef. +Notation "A \subset B" := (subset (mem A) (mem B)) + (at level 70, no associativity) : bool_scope. + +Definition proper T A B := @subset T A B && ~~ subset B A. +Notation "A \proper B" := (proper (mem A) (mem B)) + (at level 70, no associativity) : bool_scope. + +(* image, xinv, inv, and ordinal operations will be defined later. *) + +Section OpsTheory. + +Variable T : finType. + +Implicit Types A B C P Q : pred T. +Implicit Types x y : T. +Implicit Type s : seq T. + +Lemma enumP : Finite.axiom (Finite.enum T). +Proof. by rewrite unlock; case T => ? [? []]. Qed. + +Section EnumPick. + +Variable P : pred T. + +Lemma enumT : enum T = Finite.enum T. +Proof. exact: filter_predT. Qed. + +Lemma mem_enum A : enum A =i A. +Proof. by move=> x; rewrite mem_filter andbC -has_pred1 has_count enumP. Qed. + +Lemma enum_uniq : uniq (enum P). +Proof. +by apply/filter_uniq/count_mem_uniq => x; rewrite enumP -enumT mem_enum. +Qed. + +Lemma enum0 : enum pred0 = Nil T. Proof. exact: filter_pred0. Qed. + +Lemma enum1 x : enum (pred1 x) = [:: x]. +Proof. +rewrite [enum _](all_pred1P x _ _); first by rewrite size_filter enumP. +by apply/allP=> y; rewrite mem_enum. +Qed. + +CoInductive pick_spec : option T -> Type := + | Pick x of P x : pick_spec (Some x) + | Nopick of P =1 xpred0 : pick_spec None. + +Lemma pickP : pick_spec (pick P). +Proof. +rewrite /pick; case: (enum _) (mem_enum P) => [|x s] Pxs /=. + by right; exact: fsym. +by left; rewrite -[P _]Pxs mem_head. +Qed. + +End EnumPick. + +Lemma eq_enum P Q : P =i Q -> enum P = enum Q. +Proof. move=> eqPQ; exact: eq_filter. Qed. + +Lemma eq_pick P Q : P =1 Q -> pick P = pick Q. +Proof. by move=> eqPQ; rewrite /pick (eq_enum eqPQ). Qed. + +Lemma cardE A : #|A| = size (enum A). +Proof. by rewrite unlock. Qed. + +Lemma eq_card A B : A =i B -> #|A| = #|B|. +Proof. by move=>eqAB; rewrite !cardE (eq_enum eqAB). Qed. + +Lemma eq_card_trans A B n : #|A| = n -> B =i A -> #|B| = n. +Proof. move <-; exact: eq_card. Qed. + +Lemma card0 : #|@pred0 T| = 0. Proof. by rewrite cardE enum0. Qed. + +Lemma cardT : #|T| = size (enum T). Proof. by rewrite cardE. Qed. + +Lemma card1 x : #|pred1 x| = 1. +Proof. by rewrite cardE enum1. Qed. + +Lemma eq_card0 A : A =i pred0 -> #|A| = 0. +Proof. exact: eq_card_trans card0. Qed. + +Lemma eq_cardT A : A =i predT -> #|A| = size (enum T). +Proof. exact: eq_card_trans cardT. Qed. + +Lemma eq_card1 x A : A =i pred1 x -> #|A| = 1. +Proof. exact: eq_card_trans (card1 x). Qed. + +Lemma cardUI A B : #|[predU A & B]| + #|[predI A & B]| = #|A| + #|B|. +Proof. by rewrite !cardE !size_filter count_predUI. Qed. + +Lemma cardID B A : #|[predI A & B]| + #|[predD A & B]| = #|A|. +Proof. +rewrite -cardUI addnC [#|predI _ _|]eq_card0 => [|x] /=. + by apply: eq_card => x; rewrite !inE andbC -andb_orl orbN. +by rewrite !inE -!andbA andbC andbA andbN. +Qed. + +Lemma cardC A : #|A| + #|[predC A]| = #|T|. +Proof. by rewrite !cardE !size_filter count_predC. Qed. + +Lemma cardU1 x A : #|[predU1 x & A]| = (x \notin A) + #|A|. +Proof. +case Ax: (x \in A). + by apply: eq_card => y; rewrite inE /=; case: eqP => // ->. +rewrite /= -(card1 x) -cardUI addnC. +rewrite [#|predI _ _|]eq_card0 => [|y /=]; first exact: eq_card. +by rewrite !inE; case: eqP => // ->. +Qed. + +Lemma card2 x y : #|pred2 x y| = (x != y).+1. +Proof. by rewrite cardU1 card1 addn1. Qed. + +Lemma cardC1 x : #|predC1 x| = #|T|.-1. +Proof. by rewrite -(cardC (pred1 x)) card1. Qed. + +Lemma cardD1 x A : #|A| = (x \in A) + #|[predD1 A & x]|. +Proof. +case Ax: (x \in A); last first. + by apply: eq_card => y; rewrite !inE /=; case: eqP => // ->. +rewrite /= -(card1 x) -cardUI addnC /=. +rewrite [#|predI _ _|]eq_card0 => [|y]; last by rewrite !inE; case: eqP. +by apply: eq_card => y; rewrite !inE; case: eqP => // ->. +Qed. + +Lemma max_card A : #|A| <= #|T|. +Proof. by rewrite -(cardC A) leq_addr. Qed. + +Lemma card_size s : #|s| <= size s. +Proof. +elim: s => [|x s IHs] /=; first by rewrite card0. +rewrite cardU1 /=; case: (~~ _) => //; exact: leqW. +Qed. + +Lemma card_uniqP s : reflect (#|s| = size s) (uniq s). +Proof. +elim: s => [|x s IHs]; first by left; exact card0. +rewrite cardU1 /= /addn; case: {+}(x \in s) => /=. + by right=> card_Ssz; have:= card_size s; rewrite card_Ssz ltnn. +by apply: (iffP IHs) => [<-| [<-]]. +Qed. + +Lemma card0_eq A : #|A| = 0 -> A =i pred0. +Proof. by move=> A0 x; apply/idP => Ax; rewrite (cardD1 x) Ax in A0. Qed. + +Lemma pred0P P : reflect (P =1 pred0) (pred0b P). +Proof. apply: (iffP eqP); [exact: card0_eq | exact: eq_card0]. Qed. + +Lemma pred0Pn P : reflect (exists x, P x) (~~ pred0b P). +Proof. +case: (pickP P) => [x Px | P0]. + by rewrite (introN (pred0P P)) => [|P0]; [left; exists x | rewrite P0 in Px]. +by rewrite -lt0n eq_card0 //; right=> [[x]]; rewrite P0. +Qed. + +Lemma card_gt0P A : reflect (exists i, i \in A) (#|A| > 0). +Proof. rewrite lt0n; exact: pred0Pn. Qed. + +Lemma subsetE A B : (A \subset B) = pred0b [predD A & B]. +Proof. by rewrite unlock. Qed. + +Lemma subsetP A B : reflect {subset A <= B} (A \subset B). +Proof. +rewrite unlock; apply: (iffP (pred0P _)) => [AB0 x | sAB x /=]. + by apply/implyP; apply/idPn; rewrite negb_imply andbC [_ && _]AB0. +by rewrite andbC -negb_imply; apply/negbF/implyP; exact: sAB. +Qed. + +Lemma subsetPn A B : + reflect (exists2 x, x \in A & x \notin B) (~~ (A \subset B)). +Proof. +rewrite unlock; apply: (iffP (pred0Pn _)) => [[x] | [x Ax nBx]]. + by case/andP; exists x. +by exists x; rewrite /= nBx. +Qed. + +Lemma subset_leq_card A B : A \subset B -> #|A| <= #|B|. +Proof. +move=> sAB. +rewrite -(cardID A B) [#|predI _ _|](@eq_card _ A) ?leq_addr //= => x. +rewrite !inE andbC; case Ax: (x \in A) => //; exact: subsetP Ax. +Qed. + +Lemma subxx_hint (mA : mem_pred T) : subset mA mA. +Proof. +by case: mA => A; have:= introT (subsetP A A); rewrite !unlock => ->. +Qed. +Hint Resolve subxx_hint. + +(* The parametrization by predType makes it easier to apply subxx. *) +Lemma subxx (pT : predType T) (pA : pT) : pA \subset pA. +Proof. by []. Qed. + +Lemma eq_subset A1 A2 : A1 =i A2 -> subset (mem A1) =1 subset (mem A2). +Proof. +move=> eqA12 [B]; rewrite !unlock; congr (_ == 0). +by apply: eq_card => x; rewrite inE /= eqA12. +Qed. + +Lemma eq_subset_r B1 B2 : B1 =i B2 -> + (@subset T)^~ (mem B1) =1 (@subset T)^~ (mem B2). +Proof. +move=> eqB12 [A]; rewrite !unlock; congr (_ == 0). +by apply: eq_card => x; rewrite !inE /= eqB12. +Qed. + +Lemma eq_subxx A B : A =i B -> A \subset B. +Proof. by move/eq_subset->. Qed. + +Lemma subset_predT A : A \subset T. +Proof. by apply/subsetP. Qed. + +Lemma predT_subset A : T \subset A -> forall x, x \in A. +Proof. move/subsetP=> allA x; exact: allA. Qed. + +Lemma subset_pred1 A x : (pred1 x \subset A) = (x \in A). +Proof. by apply/subsetP/idP=> [-> // | Ax y /eqP-> //]; exact: eqxx. Qed. + +Lemma subset_eqP A B : reflect (A =i B) ((A \subset B) && (B \subset A)). +Proof. +apply: (iffP andP) => [[sAB sBA] x| eqAB]; last by rewrite !eq_subxx. +by apply/idP/idP; apply: subsetP. +Qed. + +Lemma subset_cardP A B : #|A| = #|B| -> reflect (A =i B) (A \subset B). +Proof. +move=> eqcAB; case: (subsetP A B) (subset_eqP A B) => //= sAB. +case: (subsetP B A) => [//|[]] x Bx; apply/idPn => Ax. +case/idP: (ltnn #|A|); rewrite {2}eqcAB (cardD1 x B) Bx /=. +apply: subset_leq_card; apply/subsetP=> y Ay; rewrite inE /= andbC. +by rewrite sAB //; apply/eqP => eqyx; rewrite -eqyx Ay in Ax. +Qed. + +Lemma subset_leqif_card A B : A \subset B -> #|A| <= #|B| ?= iff (B \subset A). +Proof. +move=> sAB; split; [exact: subset_leq_card | apply/eqP/idP]. + by move/subset_cardP=> sABP; rewrite (eq_subset_r (sABP sAB)). +by move=> sBA; apply: eq_card; apply/subset_eqP; rewrite sAB. +Qed. + +Lemma subset_trans A B C : A \subset B -> B \subset C -> A \subset C. +Proof. +by move/subsetP=> sAB /subsetP=> sBC; apply/subsetP=> x /sAB; exact: sBC. +Qed. + +Lemma subset_all s A : (s \subset A) = all (mem A) s. +Proof. by exact (sameP (subsetP _ _) allP). Qed. + +Lemma properE A B : A \proper B = (A \subset B) && ~~(B \subset A). +Proof. by []. Qed. + +Lemma properP A B : + reflect (A \subset B /\ (exists2 x, x \in B & x \notin A)) (A \proper B). +Proof. +by rewrite properE; apply: (iffP andP) => [] [-> /subsetPn]. +Qed. + +Lemma proper_sub A B : A \proper B -> A \subset B. +Proof. by case/andP. Qed. + +Lemma proper_subn A B : A \proper B -> ~~ (B \subset A). +Proof. by case/andP. Qed. + +Lemma proper_trans A B C : A \proper B -> B \proper C -> A \proper C. +Proof. +case/properP=> sAB [x Bx nAx] /properP[sBC [y Cy nBy]]. +rewrite properE (subset_trans sAB) //=; apply/subsetPn; exists y => //. +by apply: contra nBy; exact: subsetP. +Qed. + +Lemma proper_sub_trans A B C : A \proper B -> B \subset C -> A \proper C. +Proof. +case/properP=> sAB [x Bx nAx] sBC; rewrite properE (subset_trans sAB) //. +by apply/subsetPn; exists x; rewrite ?(subsetP _ _ sBC). +Qed. + +Lemma sub_proper_trans A B C : A \subset B -> B \proper C -> A \proper C. +Proof. +move=> sAB /properP[sBC [x Cx nBx]]; rewrite properE (subset_trans sAB) //. +by apply/subsetPn; exists x => //; apply: contra nBx; exact: subsetP. +Qed. + +Lemma proper_card A B : A \proper B -> #|A| < #|B|. +Proof. +by case/andP=> sAB nsBA; rewrite ltn_neqAle !(subset_leqif_card sAB) andbT. +Qed. + +Lemma proper_irrefl A : ~~ (A \proper A). +Proof. by rewrite properE subxx. Qed. + +Lemma properxx A : (A \proper A) = false. +Proof. by rewrite properE subxx. Qed. + +Lemma eq_proper A B : A =i B -> proper (mem A) =1 proper (mem B). +Proof. +move=> eAB [C]; congr (_ && _); first exact: (eq_subset eAB). +by rewrite (eq_subset_r eAB). +Qed. + +Lemma eq_proper_r A B : A =i B -> + (@proper T)^~ (mem A) =1 (@proper T)^~ (mem B). +Proof. +move=> eAB [C]; congr (_ && _); first exact: (eq_subset_r eAB). +by rewrite (eq_subset eAB). +Qed. + +Lemma disjoint_sym A B : [disjoint A & B] = [disjoint B & A]. +Proof. by congr (_ == 0); apply: eq_card => x; exact: andbC. Qed. + +Lemma eq_disjoint A1 A2 : A1 =i A2 -> disjoint (mem A1) =1 disjoint (mem A2). +Proof. +by move=> eqA12 [B]; congr (_ == 0); apply: eq_card => x; rewrite !inE eqA12. +Qed. + +Lemma eq_disjoint_r B1 B2 : B1 =i B2 -> + (@disjoint T)^~ (mem B1) =1 (@disjoint T)^~ (mem B2). +Proof. +by move=> eqB12 [A]; congr (_ == 0); apply: eq_card => x; rewrite !inE eqB12. +Qed. + +Lemma subset_disjoint A B : (A \subset B) = [disjoint A & [predC B]]. +Proof. by rewrite disjoint_sym unlock. Qed. + +Lemma disjoint_subset A B : [disjoint A & B] = (A \subset [predC B]). +Proof. +by rewrite subset_disjoint; apply: eq_disjoint_r => x; rewrite !inE /= negbK. +Qed. + +Lemma disjoint_trans A B C : + A \subset B -> [disjoint B & C] -> [disjoint A & C]. +Proof. by rewrite 2!disjoint_subset; exact: subset_trans. Qed. + +Lemma disjoint0 A : [disjoint pred0 & A]. +Proof. exact/pred0P. Qed. + +Lemma eq_disjoint0 A B : A =i pred0 -> [disjoint A & B]. +Proof. by move/eq_disjoint->; exact: disjoint0. Qed. + +Lemma disjoint1 x A : [disjoint pred1 x & A] = (x \notin A). +Proof. +apply/negbRL/(sameP (pred0Pn _)). +apply: introP => [Ax | notAx [_ /andP[/eqP->]]]; last exact: negP. +by exists x; rewrite !inE eqxx. +Qed. + +Lemma eq_disjoint1 x A B : + A =i pred1 x -> [disjoint A & B] = (x \notin B). +Proof. by move/eq_disjoint->; exact: disjoint1. Qed. + +Lemma disjointU A B C : + [disjoint predU A B & C] = [disjoint A & C] && [disjoint B & C]. +Proof. +case: [disjoint A & C] / (pred0P (xpredI A C)) => [A0 | nA0] /=. + by congr (_ == 0); apply: eq_card => x; rewrite [x \in _]andb_orl A0. +apply/pred0P=> nABC; case: nA0 => x; apply/idPn=> /=; move/(_ x): nABC. +by rewrite [_ x]andb_orl; case/norP. +Qed. + +Lemma disjointU1 x A B : + [disjoint predU1 x A & B] = (x \notin B) && [disjoint A & B]. +Proof. by rewrite disjointU disjoint1. Qed. + +Lemma disjoint_cons x s B : + [disjoint x :: s & B] = (x \notin B) && [disjoint s & B]. +Proof. exact: disjointU1. Qed. + +Lemma disjoint_has s A : [disjoint s & A] = ~~ has (mem A) s. +Proof. +rewrite -(@eq_has _ (mem (enum A))) => [|x]; last exact: mem_enum. +rewrite has_sym has_filter -filter_predI -has_filter has_count -eqn0Ngt. +by rewrite -size_filter /disjoint /pred0b unlock. +Qed. + +Lemma disjoint_cat s1 s2 A : + [disjoint s1 ++ s2 & A] = [disjoint s1 & A] && [disjoint s2 & A]. +Proof. by rewrite !disjoint_has has_cat negb_or. Qed. + +End OpsTheory. + +Hint Resolve subxx_hint. + +Implicit Arguments pred0P [T P]. +Implicit Arguments pred0Pn [T P]. +Implicit Arguments subsetP [T A B]. +Implicit Arguments subsetPn [T A B]. +Implicit Arguments subset_eqP [T A B]. +Implicit Arguments card_uniqP [T s]. +Implicit Arguments properP [T A B]. +Prenex Implicits pred0P pred0Pn subsetP subsetPn subset_eqP card_uniqP. + +(**********************************************************************) +(* *) +(* Boolean quantifiers for finType *) +(* *) +(**********************************************************************) + +Section QuantifierCombinators. + +Variables (T : finType) (P : pred T) (PP : T -> Prop). +Hypothesis viewP : forall x, reflect (PP x) (P x). + +Lemma existsPP : reflect (exists x, PP x) [exists x, P x]. +Proof. by apply: (iffP pred0Pn) => -[x /viewP]; exists x. Qed. + +Lemma forallPP : reflect (forall x, PP x) [forall x, P x]. +Proof. by apply: (iffP pred0P) => /= allP x; have /viewP//=-> := allP x. Qed. + +End QuantifierCombinators. + +Notation "'exists_ view" := (existsPP (fun _ => view)) + (at level 4, right associativity, format "''exists_' view"). +Notation "'forall_ view" := (forallPP (fun _ => view)) + (at level 4, right associativity, format "''forall_' view"). + +Section Quantifiers. + +Variables (T : finType) (rT : T -> eqType). +Implicit Type (D P : pred T) (f : forall x, rT x). + +Lemma forallP P : reflect (forall x, P x) [forall x, P x]. +Proof. exact: 'forall_idP. Qed. + +Lemma eqfunP f1 f2 : reflect (forall x, f1 x = f2 x) [forall x, f1 x == f2 x]. +Proof. exact: 'forall_eqP. Qed. + +Lemma forall_inP D P : reflect (forall x, D x -> P x) [forall (x | D x), P x]. +Proof. exact: 'forall_implyP. Qed. + +Lemma eqfun_inP D f1 f2 : + reflect {in D, forall x, f1 x = f2 x} [forall (x | x \in D), f1 x == f2 x]. +Proof. by apply: (iffP 'forall_implyP) => eq_f12 x Dx; apply/eqP/eq_f12. Qed. + +Lemma existsP P : reflect (exists x, P x) [exists x, P x]. +Proof. exact: 'exists_idP. Qed. + +Lemma exists_eqP f1 f2 : + reflect (exists x, f1 x = f2 x) [exists x, f1 x == f2 x]. +Proof. exact: 'exists_eqP. Qed. + +Lemma exists_inP D P : reflect (exists2 x, D x & P x) [exists (x | D x), P x]. +Proof. by apply: (iffP 'exists_andP) => [[x []] | [x]]; exists x. Qed. + +Lemma exists_eq_inP D f1 f2 : + reflect (exists2 x, D x & f1 x = f2 x) [exists (x | D x), f1 x == f2 x]. +Proof. by apply: (iffP (exists_inP _ _)) => [] [x Dx /eqP]; exists x. Qed. + +Lemma eq_existsb P1 P2 : P1 =1 P2 -> [exists x, P1 x] = [exists x, P2 x]. +Proof. by move=> eqP12; congr (_ != 0); apply: eq_card. Qed. + +Lemma eq_existsb_in D P1 P2 : + (forall x, D x -> P1 x = P2 x) -> + [exists (x | D x), P1 x] = [exists (x | D x), P2 x]. +Proof. by move=> eqP12; apply: eq_existsb => x; apply: andb_id2l => /eqP12. Qed. + +Lemma eq_forallb P1 P2 : P1 =1 P2 -> [forall x, P1 x] = [forall x, P2 x]. +Proof. by move=> eqP12; apply/negb_inj/eq_existsb=> /= x; rewrite eqP12. Qed. + +Lemma eq_forallb_in D P1 P2 : + (forall x, D x -> P1 x = P2 x) -> + [forall (x | D x), P1 x] = [forall (x | D x), P2 x]. +Proof. +by move=> eqP12; apply: eq_forallb => i; case Di: (D i); rewrite // eqP12. +Qed. + +Lemma negb_forall P : ~~ [forall x, P x] = [exists x, ~~ P x]. +Proof. by []. Qed. + +Lemma negb_forall_in D P : + ~~ [forall (x | D x), P x] = [exists (x | D x), ~~ P x]. +Proof. by apply: eq_existsb => x; rewrite negb_imply. Qed. + +Lemma negb_exists P : ~~ [exists x, P x] = [forall x, ~~ P x]. +Proof. by apply/negbLR/esym/eq_existsb=> x; apply: negbK. Qed. + +Lemma negb_exists_in D P : + ~~ [exists (x | D x), P x] = [forall (x | D x), ~~ P x]. +Proof. by rewrite negb_exists; apply/eq_forallb => x; rewrite [~~ _]fun_if. Qed. + +End Quantifiers. + +Implicit Arguments forallP [T P]. +Implicit Arguments eqfunP [T rT f1 f2]. +Implicit Arguments forall_inP [T D P]. +Implicit Arguments eqfun_inP [T rT D f1 f2]. +Implicit Arguments existsP [T P]. +Implicit Arguments exists_eqP [T rT f1 f2]. +Implicit Arguments exists_inP [T D P]. +Implicit Arguments exists_eq_inP [T rT D f1 f2]. + +Section Extrema. + +Variables (I : finType) (i0 : I) (P : pred I) (F : I -> nat). + +Let arg_pred ord := [pred i | P i & [forall (j | P j), ord (F i) (F j)]]. + +Definition arg_min := odflt i0 (pick (arg_pred leq)). + +Definition arg_max := odflt i0 (pick (arg_pred geq)). + +CoInductive extremum_spec (ord : rel nat) : I -> Type := + ExtremumSpec i of P i & (forall j, P j -> ord (F i) (F j)) + : extremum_spec ord i. + +Hypothesis Pi0 : P i0. + +Let FP n := [exists (i | P i), F i == n]. +Let FP_F i : P i -> FP (F i). +Proof. by move=> Pi; apply/existsP; exists i; rewrite Pi /=. Qed. +Let exFP : exists n, FP n. Proof. by exists (F i0); exact: FP_F. Qed. + +Lemma arg_minP : extremum_spec leq arg_min. +Proof. +rewrite /arg_min; case: pickP => [i /andP[Pi /forallP/= min_i] | no_i]. + by split=> // j; apply/implyP. +case/ex_minnP: exFP => n ex_i min_i; case/pred0P: ex_i => i /=. +apply: contraFF (no_i i) => /andP[Pi /eqP def_n]; rewrite /= Pi. +by apply/forall_inP=> j Pj; rewrite def_n min_i ?FP_F. +Qed. + +Lemma arg_maxP : extremum_spec geq arg_max. +Proof. +rewrite /arg_max; case: pickP => [i /andP[Pi /forall_inP/= max_i] | no_i]. + by split=> // j; apply/implyP. +have (n): FP n -> n <= foldr maxn 0 (map F (enum P)). + case/existsP=> i; rewrite -[P i]mem_enum andbC /= => /andP[/eqP <-]. + elim: (enum P) => //= j e IHe; rewrite leq_max orbC !inE. + by case/predU1P=> [-> | /IHe-> //]; rewrite leqnn orbT. +case/ex_maxnP=> // n ex_i max_i; case/pred0P: ex_i => i /=. +apply: contraFF (no_i i) => /andP[Pi def_n]; rewrite /= Pi. +by apply/forall_inP=> j Pj; rewrite (eqP def_n) max_i ?FP_F. +Qed. + +End Extrema. + +Notation "[ 'arg' 'min_' ( i < i0 | P ) F ]" := + (arg_min i0 (fun i => P%B) (fun i => F)) + (at level 0, i, i0 at level 10, + format "[ 'arg' 'min_' ( i < i0 | P ) F ]") : form_scope. + +Notation "[ 'arg' 'min_' ( i < i0 'in' A ) F ]" := + [arg min_(i < i0 | i \in A) F] + (at level 0, i, i0 at level 10, + format "[ 'arg' 'min_' ( i < i0 'in' A ) F ]") : form_scope. + +Notation "[ 'arg' 'min_' ( i < i0 ) F ]" := [arg min_(i < i0 | true) F] + (at level 0, i, i0 at level 10, + format "[ 'arg' 'min_' ( i < i0 ) F ]") : form_scope. + +Notation "[ 'arg' 'max_' ( i > i0 | P ) F ]" := + (arg_max i0 (fun i => P%B) (fun i => F)) + (at level 0, i, i0 at level 10, + format "[ 'arg' 'max_' ( i > i0 | P ) F ]") : form_scope. + +Notation "[ 'arg' 'max_' ( i > i0 'in' A ) F ]" := + [arg max_(i > i0 | i \in A) F] + (at level 0, i, i0 at level 10, + format "[ 'arg' 'max_' ( i > i0 'in' A ) F ]") : form_scope. + +Notation "[ 'arg' 'max_' ( i > i0 ) F ]" := [arg max_(i > i0 | true) F] + (at level 0, i, i0 at level 10, + format "[ 'arg' 'max_' ( i > i0 ) F ]") : form_scope. + +(**********************************************************************) +(* *) +(* Boolean injectivity test for functions with a finType domain *) +(* *) +(**********************************************************************) + +Section Injectiveb. + +Variables (aT : finType) (rT : eqType) (f : aT -> rT). +Implicit Type D : pred aT. + +Definition dinjectiveb D := uniq (map f (enum D)). + +Definition injectiveb := dinjectiveb aT. + +Lemma dinjectivePn D : + reflect (exists2 x, x \in D & exists2 y, y \in [predD1 D & x] & f x = f y) + (~~ dinjectiveb D). +Proof. +apply: (iffP idP) => [injf | [x Dx [y Dxy eqfxy]]]; last first. + move: Dx; rewrite -(mem_enum D) => /rot_to[i E defE]. + rewrite /dinjectiveb -(rot_uniq i) -map_rot defE /=; apply/nandP; left. + rewrite inE /= -(mem_enum D) -(mem_rot i) defE inE in Dxy. + rewrite andb_orr andbC andbN in Dxy. + by rewrite eqfxy map_f //; case/andP: Dxy. +pose p := [pred x in D | [exists (y | y \in [predD1 D & x]), f x == f y]]. +case: (pickP p) => [x /= /andP[Dx /exists_inP[y Dxy /eqP eqfxy]] | no_p]. + by exists x; last exists y. +rewrite /dinjectiveb map_inj_in_uniq ?enum_uniq // in injf => x y Dx Dy eqfxy. +apply: contraNeq (negbT (no_p x)) => ne_xy /=; rewrite -mem_enum Dx. +by apply/existsP; exists y; rewrite /= !inE eq_sym ne_xy -mem_enum Dy eqfxy /=. +Qed. + +Lemma dinjectiveP D : reflect {in D &, injective f} (dinjectiveb D). +Proof. +rewrite -[dinjectiveb D]negbK. +case: dinjectivePn=> [noinjf | injf]; constructor. + case: noinjf => x Dx [y /andP[neqxy /= Dy] eqfxy] injf. + by case/eqP: neqxy; exact: injf. +move=> x y Dx Dy /= eqfxy; apply/eqP; apply/idPn=> nxy; case: injf. +by exists x => //; exists y => //=; rewrite inE /= eq_sym nxy. +Qed. + +Lemma injectivePn : + reflect (exists x, exists2 y, x != y & f x = f y) (~~ injectiveb). +Proof. +apply: (iffP (dinjectivePn _)) => [[x _ [y nxy eqfxy]] | [x [y nxy eqfxy]]]; + by exists x => //; exists y => //; rewrite inE /= andbT eq_sym in nxy *. +Qed. + +Lemma injectiveP : reflect (injective f) injectiveb. +Proof. apply: (iffP (dinjectiveP _)) => injf x y => [|_ _]; exact: injf. Qed. + +End Injectiveb. + +Definition image_mem T T' f mA : seq T' := map f (@enum_mem T mA). +Notation image f A := (image_mem f (mem A)). +Notation "[ 'seq' F | x 'in' A ]" := (image (fun x => F) A) + (at level 0, F at level 99, x ident, + format "'[hv' [ 'seq' F '/ ' | x 'in' A ] ']'") : seq_scope. +Notation "[ 'seq' F | x : T 'in' A ]" := (image (fun x : T => F) A) + (at level 0, F at level 99, x ident, only parsing) : seq_scope. +Notation "[ 'seq' F | x : T ]" := + [seq F | x : T in sort_of_simpl_pred (@pred_of_argType T)] + (at level 0, F at level 99, x ident, + format "'[hv' [ 'seq' F '/ ' | x : T ] ']'") : seq_scope. +Notation "[ 'seq' F , x ]" := [seq F | x : _ ] + (at level 0, F at level 99, x ident, only parsing) : seq_scope. + +Definition codom T T' f := @image_mem T T' f (mem T). + +Section Image. + +Variable T : finType. +Implicit Type A : pred T. + +Section SizeImage. + +Variables (T' : Type) (f : T -> T'). + +Lemma size_image A : size (image f A) = #|A|. +Proof. by rewrite size_map -cardE. Qed. + +Lemma size_codom : size (codom f) = #|T|. +Proof. exact: size_image. Qed. + +Lemma codomE : codom f = map f (enum T). +Proof. by []. Qed. + +End SizeImage. + +Variables (T' : eqType) (f : T -> T'). + +Lemma imageP A y : reflect (exists2 x, x \in A & y = f x) (y \in image f A). +Proof. +by apply: (iffP mapP) => [] [x Ax y_fx]; exists x; rewrite // mem_enum in Ax *. +Qed. + +Lemma codomP y : reflect (exists x, y = f x) (y \in codom f). +Proof. by apply: (iffP (imageP _ y)) => [][x]; exists x. Qed. + +Remark iinv_proof A y : y \in image f A -> {x | x \in A & f x = y}. +Proof. +move=> fy; pose b x := A x && (f x == y). +case: (pickP b) => [x /andP[Ax /eqP] | nfy]; first by exists x. +by case/negP: fy => /imageP[x Ax fx_y]; case/andP: (nfy x); rewrite fx_y. +Qed. + +Definition iinv A y fAy := s2val (@iinv_proof A y fAy). + +Lemma f_iinv A y fAy : f (@iinv A y fAy) = y. +Proof. exact: s2valP' (iinv_proof fAy). Qed. + +Lemma mem_iinv A y fAy : @iinv A y fAy \in A. +Proof. exact: s2valP (iinv_proof fAy). Qed. + +Lemma in_iinv_f A : {in A &, injective f} -> + forall x fAfx, x \in A -> @iinv A (f x) fAfx = x. +Proof. +move=> injf x fAfx Ax; apply: injf => //; [exact: mem_iinv | exact: f_iinv]. +Qed. + +Lemma preim_iinv A B y fAy : preim f B (@iinv A y fAy) = B y. +Proof. by rewrite /= f_iinv. Qed. + +Lemma image_f A x : x \in A -> f x \in image f A. +Proof. by move=> Ax; apply/imageP; exists x. Qed. + +Lemma codom_f x : f x \in codom f. +Proof. by exact: image_f. Qed. + +Lemma image_codom A : {subset image f A <= codom f}. +Proof. by move=> _ /imageP[x _ ->]; exact: codom_f. Qed. + +Lemma image_pred0 : image f pred0 =i pred0. +Proof. by move=> x; rewrite /image_mem /= enum0. Qed. + +Section Injective. + +Hypothesis injf : injective f. + +Lemma mem_image A x : (f x \in image f A) = (x \in A). +Proof. by rewrite mem_map ?mem_enum. Qed. + +Lemma pre_image A : [preim f of image f A] =i A. +Proof. by move=> x; rewrite inE /= mem_image. Qed. + +Lemma image_iinv A y (fTy : y \in codom f) : + (y \in image f A) = (iinv fTy \in A). +Proof. by rewrite -mem_image ?f_iinv. Qed. + +Lemma iinv_f x fTfx : @iinv T (f x) fTfx = x. +Proof. by apply: in_iinv_f; first exact: in2W. Qed. + +Lemma image_pre (B : pred T') : image f [preim f of B] =i [predI B & codom f]. +Proof. by move=> y; rewrite /image_mem -filter_map /= mem_filter -enumT. Qed. + +Lemma bij_on_codom (x0 : T) : {on [pred y in codom f], bijective f}. +Proof. +pose g y := iinv (valP (insigd (codom_f x0) y)). +by exists g => [x fAfx | y fAy]; first apply: injf; rewrite f_iinv insubdK. +Qed. + +Lemma bij_on_image A (x0 : T) : {on [pred y in image f A], bijective f}. +Proof. exact: subon_bij (@image_codom A) (bij_on_codom x0). Qed. + +End Injective. + +Fixpoint preim_seq s := + if s is y :: s' then + (if pick (preim f (pred1 y)) is Some x then cons x else id) (preim_seq s') + else [::]. + +Lemma map_preim (s : seq T') : {subset s <= codom f} -> map f (preim_seq s) = s. +Proof. +elim: s => //= y s IHs; case: pickP => [x /eqP fx_y | nfTy] fTs. + by rewrite /= fx_y IHs // => z s_z; apply: fTs; exact: predU1r. +by case/imageP: (fTs y (mem_head y s)) => x _ fx_y; case/eqP: (nfTy x). +Qed. + +End Image. + +Prenex Implicits codom iinv. +Implicit Arguments imageP [T T' f A y]. +Implicit Arguments codomP [T T' f y]. + +Lemma flatten_imageP (aT : finType) (rT : eqType) A (P : pred aT) (y : rT) : + reflect (exists2 x, x \in P & y \in A x) (y \in flatten [seq A x | x in P]). +Proof. +by apply: (iffP flatten_mapP) => [][x Px]; exists x; rewrite ?mem_enum in Px *. +Qed. +Implicit Arguments flatten_imageP [aT rT A P y]. + +Section CardFunImage. + +Variables (T T' : finType) (f : T -> T'). +Implicit Type A : pred T. + +Lemma leq_image_card A : #|image f A| <= #|A|. +Proof. by rewrite (cardE A) -(size_map f) card_size. Qed. + +Lemma card_in_image A : {in A &, injective f} -> #|image f A| = #|A|. +Proof. +move=> injf; rewrite (cardE A) -(size_map f); apply/card_uniqP. +rewrite map_inj_in_uniq ?enum_uniq // => x y; rewrite !mem_enum; exact: injf. +Qed. + +Lemma image_injP A : reflect {in A &, injective f} (#|image f A| == #|A|). +Proof. +apply: (iffP eqP) => [eqfA |]; last exact: card_in_image. +by apply/dinjectiveP; apply/card_uniqP; rewrite size_map -cardE. +Qed. + +Hypothesis injf : injective f. + +Lemma card_image A : #|image f A| = #|A|. +Proof. apply: card_in_image; exact: in2W. Qed. + +Lemma card_codom : #|codom f| = #|T|. +Proof. exact: card_image. Qed. + +Lemma card_preim (B : pred T') : #|[preim f of B]| = #|[predI codom f & B]|. +Proof. +rewrite -card_image /=; apply: eq_card => y. +by rewrite [y \in _]image_pre !inE andbC. +Qed. + +Hypothesis card_range : #|T| = #|T'|. + +Lemma inj_card_onto y : y \in codom f. +Proof. by move: y; apply/subset_cardP; rewrite ?card_codom ?subset_predT. Qed. + +Lemma inj_card_bij : bijective f. +Proof. +by exists (fun y => iinv (inj_card_onto y)) => y; rewrite ?iinv_f ?f_iinv. +Qed. + +End CardFunImage. + +Implicit Arguments image_injP [T T' f A]. + +Section FinCancel. + +Variables (T : finType) (f g : T -> T). + +Section Inv. + +Hypothesis injf : injective f. + +Lemma injF_onto y : y \in codom f. Proof. exact: inj_card_onto. Qed. +Definition invF y := iinv (injF_onto y). +Lemma invF_f : cancel f invF. Proof. by move=> x; exact: iinv_f. Qed. +Lemma f_invF : cancel invF f. Proof. by move=> y; exact: f_iinv. Qed. +Lemma injF_bij : bijective f. Proof. exact: inj_card_bij. Qed. + +End Inv. + +Hypothesis fK : cancel f g. + +Lemma canF_sym : cancel g f. +Proof. exact/(bij_can_sym (injF_bij (can_inj fK))). Qed. + +Lemma canF_LR x y : x = g y -> f x = y. +Proof. exact: canLR canF_sym. Qed. + +Lemma canF_RL x y : g x = y -> x = f y. +Proof. exact: canRL canF_sym. Qed. + +Lemma canF_eq x y : (f x == y) = (x == g y). +Proof. exact: (can2_eq fK canF_sym). Qed. + +Lemma canF_invF : g =1 invF (can_inj fK). +Proof. by move=> y; apply: (canLR fK); rewrite f_invF. Qed. + +End FinCancel. + +Section EqImage. + +Variables (T : finType) (T' : Type). + +Lemma eq_image (A B : pred T) (f g : T -> T') : + A =i B -> f =1 g -> image f A = image g B. +Proof. +by move=> eqAB eqfg; rewrite /image_mem (eq_enum eqAB) (eq_map eqfg). +Qed. + +Lemma eq_codom (f g : T -> T') : f =1 g -> codom f = codom g. +Proof. exact: eq_image. Qed. + +Lemma eq_invF f g injf injg : f =1 g -> @invF T f injf =1 @invF T g injg. +Proof. +by move=> eq_fg x; apply: (canLR (invF_f injf)); rewrite eq_fg f_invF. +Qed. + +End EqImage. + +(* Standard finTypes *) + +Section SeqFinType. + +Variables (T : eqType) (s : seq T). + +Record seq_sub : Type := SeqSub {ssval : T; ssvalP : in_mem ssval (@mem T _ s)}. + +Canonical seq_sub_subType := Eval hnf in [subType for ssval]. +Definition seq_sub_eqMixin := Eval hnf in [eqMixin of seq_sub by <:]. +Canonical seq_sub_eqType := Eval hnf in EqType seq_sub seq_sub_eqMixin. + +Definition seq_sub_enum : seq seq_sub := undup (pmap insub s). + +Lemma mem_seq_sub_enum x : x \in seq_sub_enum. +Proof. by rewrite mem_undup mem_pmap -valK map_f ?ssvalP. Qed. + +Lemma val_seq_sub_enum : uniq s -> map val seq_sub_enum = s. +Proof. +move=> Us; rewrite /seq_sub_enum undup_id ?pmap_sub_uniq //. +rewrite (pmap_filter (@insubK _ _ _)); apply/all_filterP. +by apply/allP => x; rewrite isSome_insub. +Qed. + +Definition seq_sub_pickle x := index x seq_sub_enum. +Definition seq_sub_unpickle n := nth None (map some seq_sub_enum) n. +Lemma seq_sub_pickleK : pcancel seq_sub_pickle seq_sub_unpickle. +Proof. +rewrite /seq_sub_unpickle => x. +by rewrite (nth_map x) ?nth_index ?index_mem ?mem_seq_sub_enum. +Qed. + +Definition seq_sub_choiceMixin := PcanChoiceMixin seq_sub_pickleK. +Canonical seq_sub_choiceType := + Eval hnf in ChoiceType seq_sub seq_sub_choiceMixin. + +Definition seq_sub_countMixin := CountMixin seq_sub_pickleK. +Canonical seq_sub_countType := Eval hnf in CountType seq_sub seq_sub_countMixin. + +Definition seq_sub_finMixin := + Eval hnf in UniqFinMixin (undup_uniq _) mem_seq_sub_enum. +Canonical seq_sub_finType := Eval hnf in FinType seq_sub seq_sub_finMixin. + +Lemma card_seq_sub : uniq s -> #|{:seq_sub}| = size s. +Proof. +by move=> Us; rewrite cardE enumT -(size_map val) unlock val_seq_sub_enum. +Qed. + +End SeqFinType. + +Canonical seq_sub_subCountType (T : choiceType) (s : seq T) := Eval hnf in [subCountType of (seq_sub s)]. + +Lemma unit_enumP : Finite.axiom [::tt]. Proof. by case. Qed. +Definition unit_finMixin := Eval hnf in FinMixin unit_enumP. +Canonical unit_finType := Eval hnf in FinType unit unit_finMixin. +Lemma card_unit : #|{: unit}| = 1. Proof. by rewrite cardT enumT unlock. Qed. + +Lemma bool_enumP : Finite.axiom [:: true; false]. Proof. by case. Qed. +Definition bool_finMixin := Eval hnf in FinMixin bool_enumP. +Canonical bool_finType := Eval hnf in FinType bool bool_finMixin. +Lemma card_bool : #|{: bool}| = 2. Proof. by rewrite cardT enumT unlock. Qed. + +Local Notation enumF T := (Finite.enum T). + +Section OptionFinType. + +Variable T : finType. + +Definition option_enum := None :: map some (enumF T). + +Lemma option_enumP : Finite.axiom option_enum. +Proof. by case=> [x|]; rewrite /= count_map (count_pred0, enumP). Qed. + +Definition option_finMixin := Eval hnf in FinMixin option_enumP. +Canonical option_finType := Eval hnf in FinType (option T) option_finMixin. + +Lemma card_option : #|{: option T}| = #|T|.+1. +Proof. by rewrite !cardT !enumT {1}unlock /= !size_map. Qed. + +End OptionFinType. + +Section TransferFinType. + +Variables (eT : countType) (fT : finType) (f : eT -> fT). + +Lemma pcan_enumP g : pcancel f g -> Finite.axiom (undup (pmap g (enumF fT))). +Proof. +move=> fK x; rewrite count_uniq_mem ?undup_uniq // mem_undup. +by rewrite mem_pmap -fK map_f // -enumT mem_enum. +Qed. + +Definition PcanFinMixin g fK := FinMixin (@pcan_enumP g fK). + +Definition CanFinMixin g (fK : cancel f g) := PcanFinMixin (can_pcan fK). + +End TransferFinType. + +Section SubFinType. + +Variables (T : choiceType) (P : pred T). +Import Finite. + +Structure subFinType := SubFinType { + subFin_sort :> subType P; + _ : mixin_of (sub_eqType subFin_sort) +}. + +Definition pack_subFinType U := + fun cT b m & phant_id (class cT) (@Class U b m) => + fun sT m' & phant_id m' m => @SubFinType sT m'. + +Implicit Type sT : subFinType. + +Definition subFin_mixin sT := + let: SubFinType _ m := sT return mixin_of (sub_eqType sT) in m. + +Coercion subFinType_subCountType sT := @SubCountType _ _ sT (subFin_mixin sT). +Canonical subFinType_subCountType. + +Coercion subFinType_finType sT := + Pack (@Class sT (sub_choiceClass sT) (subFin_mixin sT)) sT. +Canonical subFinType_finType. + +Lemma codom_val sT x : (x \in codom (val : sT -> T)) = P x. +Proof. +by apply/codomP/idP=> [[u ->]|Px]; last exists (Sub x Px); rewrite ?valP ?SubK. +Qed. + +End SubFinType. + +(* This assumes that T has both finType and subCountType structures. *) +Notation "[ 'subFinType' 'of' T ]" := (@pack_subFinType _ _ T _ _ _ id _ _ id) + (at level 0, format "[ 'subFinType' 'of' T ]") : form_scope. + +Canonical seq_sub_subFinType (T : choiceType) s := + Eval hnf in [subFinType of @seq_sub T s]. + +Section FinTypeForSub. + +Variables (T : finType) (P : pred T) (sT : subCountType P). + +Definition sub_enum : seq sT := pmap insub (enumF T). + +Lemma mem_sub_enum u : u \in sub_enum. +Proof. by rewrite mem_pmap_sub -enumT mem_enum. Qed. + +Lemma sub_enum_uniq : uniq sub_enum. +Proof. by rewrite pmap_sub_uniq // -enumT enum_uniq. Qed. + +Lemma val_sub_enum : map val sub_enum = enum P. +Proof. +rewrite pmap_filter; last exact: insubK. +by apply: eq_filter => x; exact: isSome_insub. +Qed. + +(* We can't declare a canonical structure here because we've already *) +(* stated that subType_sort and FinType.sort unify via to the *) +(* subType_finType structure. *) + +Definition SubFinMixin := UniqFinMixin sub_enum_uniq mem_sub_enum. +Definition SubFinMixin_for (eT : eqType) of phant eT := + eq_rect _ Finite.mixin_of SubFinMixin eT. + +Variable sfT : subFinType P. + +Lemma card_sub : #|sfT| = #|[pred x | P x]|. +Proof. by rewrite -(eq_card (codom_val sfT)) (card_image val_inj). Qed. + +Lemma eq_card_sub (A : pred sfT) : A =i predT -> #|A| = #|[pred x | P x]|. +Proof. exact: eq_card_trans card_sub. Qed. + +End FinTypeForSub. + +(* This assumes that T has a subCountType structure over a type that *) +(* has a finType structure. *) +Notation "[ 'finMixin' 'of' T 'by' <: ]" := + (SubFinMixin_for (Phant T) (erefl _)) + (at level 0, format "[ 'finMixin' 'of' T 'by' <: ]") : form_scope. + +(* Regression for the subFinType stack +Record myb : Type := MyB {myv : bool; _ : ~~ myv}. +Canonical myb_sub := Eval hnf in [subType for myv]. +Definition myb_eqm := Eval hnf in [eqMixin of myb by <:]. +Canonical myb_eq := Eval hnf in EqType myb myb_eqm. +Definition myb_chm := [choiceMixin of myb by <:]. +Canonical myb_ch := Eval hnf in ChoiceType myb myb_chm. +Definition myb_cntm := [countMixin of myb by <:]. +Canonical myb_cnt := Eval hnf in CountType myb myb_cntm. +Canonical myb_scnt := Eval hnf in [subCountType of myb]. +Definition myb_finm := [finMixin of myb by <:]. +Canonical myb_fin := Eval hnf in FinType myb myb_finm. +Canonical myb_sfin := Eval hnf in [subFinType of myb]. +Print Canonical Projections. +Print myb_finm. +Print myb_cntm. +*) + +Section CardSig. + +Variables (T : finType) (P : pred T). + +Definition sig_finMixin := [finMixin of {x | P x} by <:]. +Canonical sig_finType := Eval hnf in FinType {x | P x} sig_finMixin. +Canonical sig_subFinType := Eval hnf in [subFinType of {x | P x}]. + +Lemma card_sig : #|{: {x | P x}}| = #|[pred x | P x]|. +Proof. exact: card_sub. Qed. + +End CardSig. + +(**********************************************************************) +(* *) +(* Ordinal finType : {0, ... , n-1} *) +(* *) +(**********************************************************************) + +Section OrdinalSub. + +Variable n : nat. + +Inductive ordinal : predArgType := Ordinal m of m < n. + +Coercion nat_of_ord i := let: Ordinal m _ := i in m. + +Canonical ordinal_subType := [subType for nat_of_ord]. +Definition ordinal_eqMixin := Eval hnf in [eqMixin of ordinal by <:]. +Canonical ordinal_eqType := Eval hnf in EqType ordinal ordinal_eqMixin. +Definition ordinal_choiceMixin := [choiceMixin of ordinal by <:]. +Canonical ordinal_choiceType := + Eval hnf in ChoiceType ordinal ordinal_choiceMixin. +Definition ordinal_countMixin := [countMixin of ordinal by <:]. +Canonical ordinal_countType := Eval hnf in CountType ordinal ordinal_countMixin. +Canonical ordinal_subCountType := [subCountType of ordinal]. + +Lemma ltn_ord (i : ordinal) : i < n. Proof. exact: valP i. Qed. + +Lemma ord_inj : injective nat_of_ord. Proof. exact: val_inj. Qed. + +Definition ord_enum : seq ordinal := pmap insub (iota 0 n). + +Lemma val_ord_enum : map val ord_enum = iota 0 n. +Proof. +rewrite pmap_filter; last exact: insubK. +by apply/all_filterP; apply/allP=> i; rewrite mem_iota isSome_insub. +Qed. + +Lemma ord_enum_uniq : uniq ord_enum. +Proof. by rewrite pmap_sub_uniq ?iota_uniq. Qed. + +Lemma mem_ord_enum i : i \in ord_enum. +Proof. by rewrite -(mem_map ord_inj) val_ord_enum mem_iota ltn_ord. Qed. + +Definition ordinal_finMixin := + Eval hnf in UniqFinMixin ord_enum_uniq mem_ord_enum. +Canonical ordinal_finType := Eval hnf in FinType ordinal ordinal_finMixin. +Canonical ordinal_subFinType := Eval hnf in [subFinType of ordinal]. + +End OrdinalSub. + +Notation "''I_' n" := (ordinal n) + (at level 8, n at level 2, format "''I_' n"). + +Hint Resolve ltn_ord. + +Section OrdinalEnum. + +Variable n : nat. + +Lemma val_enum_ord : map val (enum 'I_n) = iota 0 n. +Proof. by rewrite enumT unlock val_ord_enum. Qed. + +Lemma size_enum_ord : size (enum 'I_n) = n. +Proof. by rewrite -(size_map val) val_enum_ord size_iota. Qed. + +Lemma card_ord : #|'I_n| = n. +Proof. by rewrite cardE size_enum_ord. Qed. + +Lemma nth_enum_ord i0 m : m < n -> nth i0 (enum 'I_n) m = m :> nat. +Proof. +by move=> ?; rewrite -(nth_map _ 0) (size_enum_ord, val_enum_ord) // nth_iota. +Qed. + +Lemma nth_ord_enum (i0 i : 'I_n) : nth i0 (enum 'I_n) i = i. +Proof. apply: val_inj; exact: nth_enum_ord. Qed. + +Lemma index_enum_ord (i : 'I_n) : index i (enum 'I_n) = i. +Proof. +by rewrite -{1}(nth_ord_enum i i) index_uniq ?(enum_uniq, size_enum_ord). +Qed. + +End OrdinalEnum. + +Lemma widen_ord_proof n m (i : 'I_n) : n <= m -> i < m. +Proof. exact: leq_trans. Qed. +Definition widen_ord n m le_n_m i := Ordinal (@widen_ord_proof n m i le_n_m). + +Lemma cast_ord_proof n m (i : 'I_n) : n = m -> i < m. +Proof. by move <-. Qed. +Definition cast_ord n m eq_n_m i := Ordinal (@cast_ord_proof n m i eq_n_m). + +Lemma cast_ord_id n eq_n i : cast_ord eq_n i = i :> 'I_n. +Proof. exact: val_inj. Qed. + +Lemma cast_ord_comp n1 n2 n3 eq_n2 eq_n3 i : + @cast_ord n2 n3 eq_n3 (@cast_ord n1 n2 eq_n2 i) = + cast_ord (etrans eq_n2 eq_n3) i. +Proof. exact: val_inj. Qed. + +Lemma cast_ordK n1 n2 eq_n : + cancel (@cast_ord n1 n2 eq_n) (cast_ord (esym eq_n)). +Proof. by move=> i; exact: val_inj. Qed. + +Lemma cast_ordKV n1 n2 eq_n : + cancel (cast_ord (esym eq_n)) (@cast_ord n1 n2 eq_n). +Proof. by move=> i; exact: val_inj. Qed. + +Lemma cast_ord_inj n1 n2 eq_n : injective (@cast_ord n1 n2 eq_n). +Proof. exact: can_inj (cast_ordK eq_n). Qed. + +Lemma rev_ord_proof n (i : 'I_n) : n - i.+1 < n. +Proof. by case: n i => [|n] [i lt_i_n] //; rewrite ltnS subSS leq_subr. Qed. +Definition rev_ord n i := Ordinal (@rev_ord_proof n i). + +Lemma rev_ordK n : involutive (@rev_ord n). +Proof. +by case: n => [|n] [i lti] //; apply: val_inj; rewrite /= !subSS subKn. +Qed. + +Lemma rev_ord_inj {n} : injective (@rev_ord n). +Proof. exact: inv_inj (@rev_ordK n). Qed. + +(* bijection between any finType T and the Ordinal finType of its cardinal *) +Section EnumRank. + +Variable T : finType. +Implicit Type A : pred T. + +Lemma enum_rank_subproof x0 A : x0 \in A -> 0 < #|A|. +Proof. by move=> Ax0; rewrite (cardD1 x0) Ax0. Qed. + +Definition enum_rank_in x0 A (Ax0 : x0 \in A) x := + insubd (Ordinal (@enum_rank_subproof x0 [eta A] Ax0)) (index x (enum A)). + +Definition enum_rank x := @enum_rank_in x T (erefl true) x. + +Lemma enum_default A : 'I_(#|A|) -> T. +Proof. by rewrite cardE; case: (enum A) => [|//] []. Qed. + +Definition enum_val A i := nth (@enum_default [eta A] i) (enum A) i. +Prenex Implicits enum_val. + +Lemma enum_valP A i : @enum_val A i \in A. +Proof. by rewrite -mem_enum mem_nth -?cardE. Qed. + +Lemma enum_val_nth A x i : @enum_val A i = nth x (enum A) i. +Proof. by apply: set_nth_default; rewrite cardE in i *; apply: ltn_ord. Qed. + +Lemma nth_image T' y0 (f : T -> T') A (i : 'I_#|A|) : + nth y0 (image f A) i = f (enum_val i). +Proof. by rewrite -(nth_map _ y0) // -cardE. Qed. + +Lemma nth_codom T' y0 (f : T -> T') (i : 'I_#|T|) : + nth y0 (codom f) i = f (enum_val i). +Proof. exact: nth_image. Qed. + +Lemma nth_enum_rank_in x00 x0 A Ax0 : + {in A, cancel (@enum_rank_in x0 A Ax0) (nth x00 (enum A))}. +Proof. +move=> x Ax; rewrite /= insubdK ?nth_index ?mem_enum //. +by rewrite cardE [_ \in _]index_mem mem_enum. +Qed. + +Lemma nth_enum_rank x0 : cancel enum_rank (nth x0 (enum T)). +Proof. by move=> x; apply: nth_enum_rank_in. Qed. + +Lemma enum_rankK_in x0 A Ax0 : + {in A, cancel (@enum_rank_in x0 A Ax0) enum_val}. +Proof. by move=> x; apply: nth_enum_rank_in. Qed. + +Lemma enum_rankK : cancel enum_rank enum_val. +Proof. by move=> x; apply: enum_rankK_in. Qed. + +Lemma enum_valK_in x0 A Ax0 : cancel enum_val (@enum_rank_in x0 A Ax0). +Proof. +move=> x; apply: ord_inj; rewrite insubdK; last first. + by rewrite cardE [_ \in _]index_mem mem_nth // -cardE. +by rewrite index_uniq ?enum_uniq // -cardE. +Qed. + +Lemma enum_valK : cancel enum_val enum_rank. +Proof. by move=> x; apply: enum_valK_in. Qed. + +Lemma enum_rank_inj : injective enum_rank. +Proof. exact: can_inj enum_rankK. Qed. + +Lemma enum_val_inj A : injective (@enum_val A). +Proof. by move=> i; apply: can_inj (enum_valK_in (enum_valP i)) (i). Qed. + +Lemma enum_val_bij_in x0 A : x0 \in A -> {on A, bijective (@enum_val A)}. +Proof. +move=> Ax0; exists (enum_rank_in Ax0) => [i _|]; last exact: enum_rankK_in. +exact: enum_valK_in. +Qed. + +Lemma enum_rank_bij : bijective enum_rank. +Proof. by move: enum_rankK enum_valK; exists (@enum_val T). Qed. + +Lemma enum_val_bij : bijective (@enum_val T). +Proof. by move: enum_rankK enum_valK; exists enum_rank. Qed. + +(* Due to the limitations of the Coq unification patterns, P can only be *) +(* inferred from the premise of this lemma, not its conclusion. As a result *) +(* this lemma will only be usable in forward chaining style. *) +Lemma fin_all_exists U (P : forall x : T, U x -> Prop) : + (forall x, exists u, P x u) -> (exists u, forall x, P x (u x)). +Proof. +move=> ex_u; pose Q m x := enum_rank x < m -> {ux | P x ux}. +suffices: forall m, m <= #|T| -> exists w : forall x, Q m x, True. + case/(_ #|T|)=> // w _; pose u x := sval (w x (ltn_ord _)). + by exists u => x; rewrite {}/u; case: (w x _). +elim=> [|m IHm] ltmX; first by have w x: Q 0 x by []; exists w. +have{IHm} [w _] := IHm (ltnW ltmX); pose i := Ordinal ltmX. +have [u Pu] := ex_u (enum_val i); suffices w' x: Q m.+1 x by exists w'. +rewrite /Q ltnS leq_eqVlt (val_eqE _ i); case: eqP => [def_i _ | _ /w //]. +by rewrite -def_i enum_rankK in u Pu; exists u. +Qed. + +Lemma fin_all_exists2 U (P Q : forall x : T, U x -> Prop) : + (forall x, exists2 u, P x u & Q x u) -> + (exists2 u, forall x, P x (u x) & forall x, Q x (u x)). +Proof. +move=> ex_u; have (x): exists u, P x u /\ Q x u by have [u] := ex_u x; exists u. +by case/fin_all_exists=> u /all_and2[]; exists u. +Qed. + +End EnumRank. + +Implicit Arguments enum_val_inj [[T] [A] x1 x2]. +Implicit Arguments enum_rank_inj [[T] x1 x2]. +Prenex Implicits enum_val enum_rank. + +Lemma enum_rank_ord n i : enum_rank i = cast_ord (esym (card_ord n)) i. +Proof. +by apply: val_inj; rewrite insubdK ?index_enum_ord // card_ord [_ \in _]ltn_ord. +Qed. + +Lemma enum_val_ord n i : enum_val i = cast_ord (card_ord n) i. +Proof. +by apply: canLR (@enum_rankK _) _; apply: val_inj; rewrite enum_rank_ord. +Qed. + +(* The integer bump / unbump operations. *) + +Definition bump h i := (h <= i) + i. +Definition unbump h i := i - (h < i). + +Lemma bumpK h : cancel (bump h) (unbump h). +Proof. +rewrite /bump /unbump => i. +have [le_hi | lt_ih] := leqP h i; first by rewrite ltnS le_hi subn1. +by rewrite ltnNge ltnW ?subn0. +Qed. + +Lemma neq_bump h i : h != bump h i. +Proof. +rewrite /bump eqn_leq; have [le_hi | lt_ih] := leqP h i. + by rewrite ltnNge le_hi andbF. +by rewrite leqNgt lt_ih. +Qed. + +Lemma unbumpKcond h i : bump h (unbump h i) = (i == h) + i. +Proof. +rewrite /bump /unbump leqNgt -subSKn. +case: (ltngtP i h) => /= [-> | ltih | ->] //; last by rewrite ltnn. +by rewrite subn1 /= leqNgt !(ltn_predK ltih, ltih, add1n). +Qed. + +Lemma unbumpK h : {in predC1 h, cancel (unbump h) (bump h)}. +Proof. by move=> i; move/negbTE=> neq_h_i; rewrite unbumpKcond neq_h_i. Qed. + +Lemma bump_addl h i k : bump (k + h) (k + i) = k + bump h i. +Proof. by rewrite /bump leq_add2l addnCA. Qed. + +Lemma bumpS h i : bump h.+1 i.+1 = (bump h i).+1. +Proof. exact: addnS. Qed. + +Lemma unbump_addl h i k : unbump (k + h) (k + i) = k + unbump h i. +Proof. +apply: (can_inj (bumpK (k + h))). +by rewrite bump_addl !unbumpKcond eqn_add2l addnCA. +Qed. + +Lemma unbumpS h i : unbump h.+1 i.+1 = (unbump h i).+1. +Proof. exact: unbump_addl 1. Qed. + +Lemma leq_bump h i j : (i <= bump h j) = (unbump h i <= j). +Proof. +rewrite /bump leq_subLR. +case: (leqP i h) (leqP h j) => [le_i_h | lt_h_i] [le_h_j | lt_j_h] //. + by rewrite leqW (leq_trans le_i_h). +by rewrite !(leqNgt i) ltnW (leq_trans _ lt_h_i). +Qed. + +Lemma leq_bump2 h i j : (bump h i <= bump h j) = (i <= j). +Proof. by rewrite leq_bump bumpK. Qed. + +Lemma bumpC h1 h2 i : + bump h1 (bump h2 i) = bump (bump h1 h2) (bump (unbump h2 h1) i). +Proof. +rewrite {1 5}/bump -leq_bump addnCA; congr (_ + (_ + _)). +rewrite 2!leq_bump /unbump /bump; case: (leqP h1 h2) => [le_h12 | lt_h21]. + by rewrite subn0 ltnS le_h12 subn1. +by rewrite subn1 (ltn_predK lt_h21) (leqNgt h1) lt_h21 subn0. +Qed. + +(* The lift operations on ordinals; to avoid a messy dependent type, *) +(* unlift is a partial operation (returns an option). *) + +Lemma lift_subproof n h (i : 'I_n.-1) : bump h i < n. +Proof. by case: n i => [[]|n] //= i; rewrite -addnS (leq_add (leq_b1 _)). Qed. + +Definition lift n (h : 'I_n) (i : 'I_n.-1) := Ordinal (lift_subproof h i). + +Lemma unlift_subproof n (h : 'I_n) (u : {j | j != h}) : unbump h (val u) < n.-1. +Proof. +case: n h u => [|n h] [] //= j ne_jh. +rewrite -(leq_bump2 h.+1) bumpS unbumpK // /bump. +case: (ltngtP n h) => [|_|eq_nh]; rewrite ?(leqNgt _ h) ?ltn_ord //. +by rewrite ltn_neqAle [j <= _](valP j) {2}eq_nh andbT. +Qed. + +Definition unlift n (h i : 'I_n) := + omap (fun u : {j | j != h} => Ordinal (unlift_subproof u)) (insub i). + +CoInductive unlift_spec n h i : option 'I_n.-1 -> Type := + | UnliftSome j of i = lift h j : unlift_spec h i (Some j) + | UnliftNone of i = h : unlift_spec h i None. + +Lemma unliftP n (h i : 'I_n) : unlift_spec h i (unlift h i). +Proof. +rewrite /unlift; case: insubP => [u nhi | ] def_i /=; constructor. + by apply: val_inj; rewrite /= def_i unbumpK. +by rewrite negbK in def_i; exact/eqP. +Qed. + +Lemma neq_lift n (h : 'I_n) i : h != lift h i. +Proof. exact: neq_bump. Qed. + +Lemma unlift_none n (h : 'I_n) : unlift h h = None. +Proof. by case: unliftP => // j Dh; case/eqP: (neq_lift h j). Qed. + +Lemma unlift_some n (h i : 'I_n) : + h != i -> {j | i = lift h j & unlift h i = Some j}. +Proof. +rewrite eq_sym => /eqP neq_ih. +by case Dui: (unlift h i) / (unliftP h i) => [j Dh|//]; exists j. +Qed. + +Lemma lift_inj n (h : 'I_n) : injective (lift h). +Proof. +move=> i1 i2; move/eqP; rewrite [_ == _](can_eq (@bumpK _)) => eq_i12. +exact/eqP. +Qed. + +Lemma liftK n (h : 'I_n) : pcancel (lift h) (unlift h). +Proof. +by move=> i; case: (unlift_some (neq_lift h i)) => j; move/lift_inj->. +Qed. + +(* Shifting and splitting indices, for cutting and pasting arrays *) + +Lemma lshift_subproof m n (i : 'I_m) : i < m + n. +Proof. by apply: leq_trans (valP i) _; exact: leq_addr. Qed. + +Lemma rshift_subproof m n (i : 'I_n) : m + i < m + n. +Proof. by rewrite ltn_add2l. Qed. + +Definition lshift m n (i : 'I_m) := Ordinal (lshift_subproof n i). +Definition rshift m n (i : 'I_n) := Ordinal (rshift_subproof m i). + +Lemma split_subproof m n (i : 'I_(m + n)) : i >= m -> i - m < n. +Proof. by move/subSn <-; rewrite leq_subLR. Qed. + +Definition split m n (i : 'I_(m + n)) : 'I_m + 'I_n := + match ltnP (i) m with + | LtnNotGeq lt_i_m => inl _ (Ordinal lt_i_m) + | GeqNotLtn ge_i_m => inr _ (Ordinal (split_subproof ge_i_m)) + end. + +CoInductive split_spec m n (i : 'I_(m + n)) : 'I_m + 'I_n -> bool -> Type := + | SplitLo (j : 'I_m) of i = j :> nat : split_spec i (inl _ j) true + | SplitHi (k : 'I_n) of i = m + k :> nat : split_spec i (inr _ k) false. + +Lemma splitP m n (i : 'I_(m + n)) : split_spec i (split i) (i < m). +Proof. +rewrite /split {-3}/leq. +by case: (@ltnP i m) => cmp_i_m //=; constructor; rewrite ?subnKC. +Qed. + +Definition unsplit m n (jk : 'I_m + 'I_n) := + match jk with inl j => lshift n j | inr k => rshift m k end. + +Lemma ltn_unsplit m n (jk : 'I_m + 'I_n) : (unsplit jk < m) = jk. +Proof. by case: jk => [j|k]; rewrite /= ?ltn_ord // ltnNge leq_addr. Qed. + +Lemma splitK m n : cancel (@split m n) (@unsplit m n). +Proof. by move=> i; apply: val_inj; case: splitP. Qed. + +Lemma unsplitK m n : cancel (@unsplit m n) (@split m n). +Proof. +move=> jk; have:= ltn_unsplit jk. +by do [case: splitP; case: jk => //= i j] => [|/addnI] => /ord_inj->. +Qed. + +Section OrdinalPos. + +Variable n' : nat. +Local Notation n := n'.+1. + +Definition ord0 := Ordinal (ltn0Sn n'). +Definition ord_max := Ordinal (ltnSn n'). + +Lemma leq_ord (i : 'I_n) : i <= n'. Proof. exact: valP i. Qed. + +Lemma sub_ord_proof m : n' - m < n. +Proof. by rewrite ltnS leq_subr. Qed. +Definition sub_ord m := Ordinal (sub_ord_proof m). + +Lemma sub_ordK (i : 'I_n) : n' - (n' - i) = i. +Proof. by rewrite subKn ?leq_ord. Qed. + +Definition inord m : 'I_n := insubd ord0 m. + +Lemma inordK m : m < n -> inord m = m :> nat. +Proof. by move=> lt_m; rewrite val_insubd lt_m. Qed. + +Lemma inord_val (i : 'I_n) : inord i = i. +Proof. by rewrite /inord /insubd valK. Qed. + +Lemma enum_ordS : enum 'I_n = ord0 :: map (lift ord0) (enum 'I_n'). +Proof. +apply: (inj_map val_inj); rewrite val_enum_ord /= -map_comp. +by rewrite (map_comp (addn 1)) val_enum_ord -iota_addl. +Qed. + +Lemma lift_max (i : 'I_n') : lift ord_max i = i :> nat. +Proof. by rewrite /= /bump leqNgt ltn_ord. Qed. + +Lemma lift0 (i : 'I_n') : lift ord0 i = i.+1 :> nat. Proof. by []. Qed. + +End OrdinalPos. + +Implicit Arguments ord0 [[n']]. +Implicit Arguments ord_max [[n']]. +Implicit Arguments inord [[n']]. +Implicit Arguments sub_ord [[n']]. + +(* Product of two fintypes which is a fintype *) +Section ProdFinType. + +Variable T1 T2 : finType. + +Definition prod_enum := [seq (x1, x2) | x1 <- enum T1, x2 <- enum T2]. + +Lemma predX_prod_enum (A1 : pred T1) (A2 : pred T2) : + count [predX A1 & A2] prod_enum = #|A1| * #|A2|. +Proof. +rewrite !cardE !size_filter -!enumT /prod_enum. +elim: (enum T1) => //= x1 s1 IHs; rewrite count_cat {}IHs count_map /preim /=. +by case: (x1 \in A1); rewrite ?count_pred0. +Qed. + +Lemma prod_enumP : Finite.axiom prod_enum. +Proof. +by case=> x1 x2; rewrite (predX_prod_enum (pred1 x1) (pred1 x2)) !card1. +Qed. + +Definition prod_finMixin := Eval hnf in FinMixin prod_enumP. +Canonical prod_finType := Eval hnf in FinType (T1 * T2) prod_finMixin. + +Lemma cardX (A1 : pred T1) (A2 : pred T2) : #|[predX A1 & A2]| = #|A1| * #|A2|. +Proof. by rewrite -predX_prod_enum unlock size_filter unlock. Qed. + +Lemma card_prod : #|{: T1 * T2}| = #|T1| * #|T2|. +Proof. by rewrite -cardX; apply: eq_card; case. Qed. + +Lemma eq_card_prod (A : pred (T1 * T2)) : A =i predT -> #|A| = #|T1| * #|T2|. +Proof. exact: eq_card_trans card_prod. Qed. + +End ProdFinType. + +Section TagFinType. + +Variables (I : finType) (T_ : I -> finType). + +Definition tag_enum := + flatten [seq [seq Tagged T_ x | x <- enumF (T_ i)] | i <- enumF I]. + +Lemma tag_enumP : Finite.axiom tag_enum. +Proof. +case=> i x; rewrite -(enumP i) /tag_enum -enumT. +elim: (enum I) => //= j e IHe. +rewrite count_cat count_map {}IHe; congr (_ + _). +rewrite -size_filter -cardE /=; case: eqP => [-> | ne_j_i]. + by apply: (@eq_card1 _ x) => y; rewrite -topredE /= tagged_asE ?eqxx. +by apply: eq_card0 => y. +Qed. + +Definition tag_finMixin := Eval hnf in FinMixin tag_enumP. +Canonical tag_finType := Eval hnf in FinType {i : I & T_ i} tag_finMixin. + +Lemma card_tagged : + #|{: {i : I & T_ i}}| = sumn (map (fun i => #|T_ i|) (enum I)). +Proof. +rewrite cardE !enumT {1}unlock size_flatten /shape -map_comp. +by congr (sumn _); apply: eq_map => i; rewrite /= size_map -enumT -cardE. +Qed. + +End TagFinType. + +Section SumFinType. + +Variables T1 T2 : finType. + +Definition sum_enum := + [seq inl _ x | x <- enumF T1] ++ [seq inr _ y | y <- enumF T2]. + +Lemma sum_enum_uniq : uniq sum_enum. +Proof. +rewrite cat_uniq -!enumT !(enum_uniq, map_inj_uniq); try by move=> ? ? []. +by rewrite andbT; apply/hasP=> [[_ /mapP[x _ ->] /mapP[]]]. +Qed. + +Lemma mem_sum_enum u : u \in sum_enum. +Proof. by case: u => x; rewrite mem_cat -!enumT map_f ?mem_enum ?orbT. Qed. + +Definition sum_finMixin := Eval hnf in UniqFinMixin sum_enum_uniq mem_sum_enum. +Canonical sum_finType := Eval hnf in FinType (T1 + T2) sum_finMixin. + +Lemma card_sum : #|{: T1 + T2}| = #|T1| + #|T2|. +Proof. by rewrite !cardT !enumT {1}unlock size_cat !size_map. Qed. + + +End SumFinType. diff --git a/mathcomp/discrete/generic_quotient.v b/mathcomp/discrete/generic_quotient.v new file mode 100644 index 0000000..1d9bd56 --- /dev/null +++ b/mathcomp/discrete/generic_quotient.v @@ -0,0 +1,727 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +(* -*- coding : utf-8 -*- *) + +Require Import ssreflect ssrfun ssrbool eqtype ssrnat choice seq fintype. + +(*****************************************************************************) +(* Provided a base type T, this files defines an interface for quotients Q *) +(* of the type T with explicit functions for canonical surjection (\pi *) +(* : T -> Q) and for choosing a representative (repr : Q -> T). It then *) +(* provide a helper to quotient T by a decidable equivalence relation (e *) +(* : rel T) if T is a choiceType (or encodable as a choiceType modulo e). *) +(* *) +(* See "Pragamatic Quotient Types in Coq", proceedings of ITP2013, *) +(* by Cyril Cohen. *) +(* *) +(* *** Generic Quotienting *** *) +(* QuotClass (reprK : cancel repr pi) == builds the quotient which *) +(* canonical surjection function is pi and which *) +(* representative selection function is repr. *) +(* QuotType Q class == packs the quotClass class to build a quotType *) +(* You may declare such elements as Canonical *) +(* \pi_Q x == the class in Q of the element x of T *) +(* \pi x == the class of x where Q is inferred from the context *) +(* repr c == canonical representative in T of the class c *) +(* [quotType of Q] == clone of the canonical quotType structure of Q on T *) +(* x = y %[mod Q] := \pi_Q x = \pi_Q y *) +(* <-> x and y are equal modulo Q *) +(* x <> y %[mod Q] := \pi_Q x <> \pi_Q y *) +(* x == y %[mod Q] := \pi_Q x == \pi_Q y *) +(* x != y %[mod Q] := \pi_Q x != \pi_Q y *) +(* *) +(* The quotient_scope is delimited by %qT *) +(* The most useful lemmas are piE and reprK *) +(* *) +(* *** Morphisms *** *) +(* One may declare existing functions and predicates as liftings of some *) +(* morphisms for a quotient. *) +(* PiMorph1 pi_f == where pi_f : {morph \pi : x / f x >-> fq x} *) +(* declares fq : Q -> Q as the lifting of f : T -> T *) +(* PiMorph2 pi_g == idem with pi_g : {morph \pi : x y / g x y >-> gq x y} *) +(* PiMono1 pi_p == idem with pi_p : {mono \pi : x / p x >-> pq x} *) +(* PiMono2 pi_r == idem with pi_r : {morph \pi : x y / r x y >-> rq x y} *) +(* PiMorph11 pi_f == idem with pi_f : {morph \pi : x / f x >-> fq x} *) +(* where fq : Q -> Q' and f : T -> T'. *) +(* PiMorph eq == Most general declaration of compatibility, *) +(* /!\ use with caution /!\ *) +(* One can use the following helpers to build the liftings which may or *) +(* may not satisfy the above properties (but if they do not, it is *) +(* probably not a good idea to define them): *) +(* lift_op1 Q f := lifts f : T -> T *) +(* lift_op2 Q g := lifts g : T -> T -> T *) +(* lift_fun1 Q p := lifts p : T -> R *) +(* lift_fun2 Q r := lifts r : T -> T -> R *) +(* lift_op11 Q Q' f := lifts f : T -> T' *) +(* There is also the special case of constants and embedding functions *) +(* that one may define and declare as compatible with Q using: *) +(* lift_cst Q x := lifts x : T to Q *) +(* PiConst c := declare the result c of the previous construction as *) +(* compatible with Q *) +(* lift_embed Q e := lifts e : R -> T to R -> Q *) +(* PiEmbed f := declare the result f of the previous construction as *) +(* compatible with Q *) +(* *) +(* *** Quotients that have an eqType structure *** *) +(* Having a canonical (eqQuotType e) structure enables piE to replace terms *) +(* of the form (x == y) by terms of the form (e x' y') if x and y are *) +(* canonical surjections of some x' and y'. *) +(* EqQuotType e Q m == builds an (eqQuotType e) structure on Q from the *) +(* morphism property m *) +(* where m : {mono \pi : x y / e x y >-> x == y} *) +(* [eqQuotType of Q] == clones the canonical eqQuotType structure of Q *) +(* *) +(* *** Equivalence and quotient by an equivalence *** *) +(* EquivRel r er es et == builds an equiv_rel structure based on the *) +(* reflexivity, symmetry and transitivity property *) +(* of a boolean relation. *) +(* {eq_quot e} == builds the quotType of T by equiv *) +(* where e : rel T is an equiv_rel *) +(* and T is a choiceType or a (choiceTypeMod e) *) +(* it is canonically an eqType, a choiceType, *) +(* a quotType and an eqQuotType. *) +(* x = y %[mod_eq e] := x = y %[mod {eq_quot e}] *) +(* <-> x and y are equal modulo e *) +(* ... *) +(*****************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Reserved Notation "\pi_ Q" (at level 0, format "\pi_ Q"). +Reserved Notation "\pi" (at level 0, format "\pi"). +Reserved Notation "{pi_ Q a }" + (at level 0, Q at next level, format "{pi_ Q a }"). +Reserved Notation "{pi a }" (at level 0, format "{pi a }"). +Reserved Notation "x == y %[mod_eq e ]" (at level 70, y at next level, + no associativity, format "'[hv ' x '/' == y '/' %[mod_eq e ] ']'"). +Reserved Notation "x = y %[mod_eq e ]" (at level 70, y at next level, + no associativity, format "'[hv ' x '/' = y '/' %[mod_eq e ] ']'"). +Reserved Notation "x != y %[mod_eq e ]" (at level 70, y at next level, + no associativity, format "'[hv ' x '/' != y '/' %[mod_eq e ] ']'"). +Reserved Notation "x <> y %[mod_eq e ]" (at level 70, y at next level, + no associativity, format "'[hv ' x '/' <> y '/' %[mod_eq e ] ']'"). +Reserved Notation "{eq_quot e }" (at level 0, e at level 0, + format "{eq_quot e }", only parsing). + +Delimit Scope quotient_scope with qT. +Local Open Scope quotient_scope. + +(*****************************************) +(* Definition of the quotient interface. *) +(*****************************************) + +Section QuotientDef. + +Variable T : Type. + +Record quot_mixin_of qT := QuotClass { + quot_repr : qT -> T; + quot_pi : T -> qT; + _ : cancel quot_repr quot_pi +}. + +Notation quot_class_of := quot_mixin_of. + +Record quotType := QuotTypePack { + quot_sort :> Type; + quot_class : quot_class_of quot_sort; + _ : Type +}. + +Definition QuotType_pack qT m := @QuotTypePack qT m qT. + +Variable qT : quotType. +Definition pi_phant of phant qT := quot_pi (quot_class qT). +Local Notation "\pi" := (pi_phant (Phant qT)). +Definition repr_of := quot_repr (quot_class qT). + +Lemma repr_ofK : cancel repr_of \pi. +Proof. by rewrite /pi_phant /repr_of /=; case:qT=> [? []]. Qed. + +Definition QuotType_clone (Q : Type) qT cT + of phant_id (quot_class qT) cT := @QuotTypePack Q cT Q. + +End QuotientDef. + +(****************************) +(* Protecting some symbols. *) +(****************************) + +Module Type PiSig. +Parameter f : forall (T : Type) (qT : quotType T), phant qT -> T -> qT. +Axiom E : f = pi_phant. +End PiSig. + +Module Pi : PiSig. +Definition f := pi_phant. +Definition E := erefl f. +End Pi. + +Module MPi : PiSig. +Definition f := pi_phant. +Definition E := erefl f. +End MPi. + +Module Type ReprSig. +Parameter f : forall (T : Type) (qT : quotType T), qT -> T. +Axiom E : f = repr_of. +End ReprSig. + +Module Repr : ReprSig. +Definition f := repr_of. +Definition E := erefl f. +End Repr. + +(*******************) +(* Fancy Notations *) +(*******************) + +Notation repr := Repr.f. +Notation "\pi_ Q" := (@Pi.f _ _ (Phant Q)) : quotient_scope. +Notation "\pi" := (@Pi.f _ _ (Phant _)) (only parsing) : quotient_scope. +Notation "x == y %[mod Q ]" := (\pi_Q x == \pi_Q y) : quotient_scope. +Notation "x = y %[mod Q ]" := (\pi_Q x = \pi_Q y) : quotient_scope. +Notation "x != y %[mod Q ]" := (\pi_Q x != \pi_Q y) : quotient_scope. +Notation "x <> y %[mod Q ]" := (\pi_Q x <> \pi_Q y) : quotient_scope. + +Local Notation "\mpi" := (@MPi.f _ _ (Phant _)). +Canonical mpi_unlock := Unlockable MPi.E. +Canonical pi_unlock := Unlockable Pi.E. +Canonical repr_unlock := Unlockable Repr.E. + +Notation quot_class_of := quot_mixin_of. +Notation QuotType Q m := (@QuotType_pack _ Q m). +Notation "[ 'quotType' 'of' Q ]" := (@QuotType_clone _ Q _ _ id) + (at level 0, format "[ 'quotType' 'of' Q ]") : form_scope. + +Implicit Arguments repr [T qT]. +Prenex Implicits repr. + +(************************) +(* Exporting the theory *) +(************************) + +Section QuotTypeTheory. + +Variable T : Type. +Variable qT : quotType T. + +Lemma reprK : cancel repr \pi_qT. +Proof. by move=> x; rewrite !unlock repr_ofK. Qed. + +CoInductive pi_spec (x : T) : T -> Type := + PiSpec y of x = y %[mod qT] : pi_spec x y. + +Lemma piP (x : T) : pi_spec x (repr (\pi_qT x)). +Proof. by constructor; rewrite reprK. Qed. + +Lemma mpiE : \mpi =1 \pi_qT. +Proof. by move=> x; rewrite !unlock. Qed. + +Lemma quotW P : (forall y : T, P (\pi_qT y)) -> forall x : qT, P x. +Proof. by move=> Py x; rewrite -[x]reprK; apply: Py. Qed. + +Lemma quotP P : (forall y : T, repr (\pi_qT y) = y -> P (\pi_qT y)) + -> forall x : qT, P x. +Proof. by move=> Py x; rewrite -[x]reprK; apply: Py; rewrite reprK. Qed. + +End QuotTypeTheory. + +(*******************) +(* About morphisms *) +(*******************) + +(* This was pi_morph T (x : T) := PiMorph { pi_op : T; _ : x = pi_op }. *) +Structure equal_to T (x : T) := EqualTo { + equal_val : T; + _ : x = equal_val +}. +Lemma equal_toE (T : Type) (x : T) (m : equal_to x) : equal_val m = x. +Proof. by case: m. Qed. + +Notation piE := (@equal_toE _ _). + +Canonical equal_to_pi T (qT : quotType T) (x : T) := + @EqualTo _ (\pi_qT x) (\pi x) (erefl _). + +Implicit Arguments EqualTo [T x equal_val]. +Prenex Implicits EqualTo. + +Section Morphism. + +Variables T U : Type. +Variable (qT : quotType T). +Variable (qU : quotType U). + +Variable (f : T -> T) (g : T -> T -> T) (p : T -> U) (r : T -> T -> U). +Variable (fq : qT -> qT) (gq : qT -> qT -> qT) (pq : qT -> U) (rq : qT -> qT -> U). +Variable (h : T -> U) (hq : qT -> qU). +Hypothesis pi_f : {morph \pi : x / f x >-> fq x}. +Hypothesis pi_g : {morph \pi : x y / g x y >-> gq x y}. +Hypothesis pi_p : {mono \pi : x / p x >-> pq x}. +Hypothesis pi_r : {mono \pi : x y / r x y >-> rq x y}. +Hypothesis pi_h : forall (x : T), \pi_qU (h x) = hq (\pi_qT x). +Variables (a b : T) (x : equal_to (\pi_qT a)) (y : equal_to (\pi_qT b)). + +(* Internal Lemmmas : do not use directly *) +Lemma pi_morph1 : \pi (f a) = fq (equal_val x). Proof. by rewrite !piE. Qed. +Lemma pi_morph2 : \pi (g a b) = gq (equal_val x) (equal_val y). Proof. by rewrite !piE. Qed. +Lemma pi_mono1 : p a = pq (equal_val x). Proof. by rewrite !piE. Qed. +Lemma pi_mono2 : r a b = rq (equal_val x) (equal_val y). Proof. by rewrite !piE. Qed. +Lemma pi_morph11 : \pi (h a) = hq (equal_val x). Proof. by rewrite !piE. Qed. + +End Morphism. + +Implicit Arguments pi_morph1 [T qT f fq]. +Implicit Arguments pi_morph2 [T qT g gq]. +Implicit Arguments pi_mono1 [T U qT p pq]. +Implicit Arguments pi_mono2 [T U qT r rq]. +Implicit Arguments pi_morph11 [T U qT qU h hq]. +Prenex Implicits pi_morph1 pi_morph2 pi_mono1 pi_mono2 pi_morph11. + +Notation "{pi_ Q a }" := (equal_to (\pi_Q a)) : quotient_scope. +Notation "{pi a }" := (equal_to (\pi a)) : quotient_scope. + +(* Declaration of morphisms *) +Notation PiMorph pi_x := (EqualTo pi_x). +Notation PiMorph1 pi_f := + (fun a (x : {pi a}) => EqualTo (pi_morph1 pi_f a x)). +Notation PiMorph2 pi_g := + (fun a b (x : {pi a}) (y : {pi b}) => EqualTo (pi_morph2 pi_g a b x y)). +Notation PiMono1 pi_p := + (fun a (x : {pi a}) => EqualTo (pi_mono1 pi_p a x)). +Notation PiMono2 pi_r := + (fun a b (x : {pi a}) (y : {pi b}) => EqualTo (pi_mono2 pi_r a b x y)). +Notation PiMorph11 pi_f := + (fun a (x : {pi a}) => EqualTo (pi_morph11 pi_f a x)). + +(* lifiting helpers *) +Notation lift_op1 Q f := (locked (fun x : Q => \pi_Q (f (repr x)) : Q)). +Notation lift_op2 Q g := + (locked (fun x y : Q => \pi_Q (g (repr x) (repr y)) : Q)). +Notation lift_fun1 Q f := (locked (fun x : Q => f (repr x))). +Notation lift_fun2 Q g := (locked (fun x y : Q => g (repr x) (repr y))). +Notation lift_op11 Q Q' f := (locked (fun x : Q => \pi_Q' (f (repr x)) : Q')). + +(* constant declaration *) +Notation lift_cst Q x := (locked (\pi_Q x : Q)). +Notation PiConst a := (@EqualTo _ _ a (lock _)). + +(* embedding declaration, please don't redefine \pi *) +Notation lift_embed qT e := (locked (fun x => \pi_qT (e x) : qT)). + +Lemma eq_lock T T' e : e =1 (@locked (T -> T') (fun x : T => e x)). +Proof. by rewrite -lock. Qed. +Prenex Implicits eq_lock. + +Notation PiEmbed e := + (fun x => @EqualTo _ _ (e x) (eq_lock (fun _ => \pi _) _)). + +(********************) +(* About eqQuotType *) +(********************) + +Section EqQuotTypeStructure. + +Variable T : Type. +Variable eq_quot_op : rel T. + +Definition eq_quot_mixin_of (Q : Type) (qc : quot_class_of T Q) + (ec : Equality.class_of Q) := + {mono \pi_(QuotTypePack qc Q) : x y / + eq_quot_op x y >-> @eq_op (Equality.Pack ec Q) x y}. + +Record eq_quot_class_of (Q : Type) : Type := EqQuotClass { + eq_quot_quot_class :> quot_class_of T Q; + eq_quot_eq_mixin :> Equality.class_of Q; + pi_eq_quot_mixin :> eq_quot_mixin_of eq_quot_quot_class eq_quot_eq_mixin +}. + +Record eqQuotType : Type := EqQuotTypePack { + eq_quot_sort :> Type; + _ : eq_quot_class_of eq_quot_sort; + _ : Type +}. + +Implicit Type eqT : eqQuotType. + +Definition eq_quot_class eqT : eq_quot_class_of eqT := + let: EqQuotTypePack _ cT _ as qT' := eqT return eq_quot_class_of qT' in cT. + +Canonical eqQuotType_eqType eqT := EqType eqT (eq_quot_class eqT). +Canonical eqQuotType_quotType eqT := QuotType eqT (eq_quot_class eqT). + +Coercion eqQuotType_eqType : eqQuotType >-> eqType. +Coercion eqQuotType_quotType : eqQuotType >-> quotType. + +Definition EqQuotType_pack Q := + fun (qT : quotType T) (eT : eqType) qc ec + of phant_id (quot_class qT) qc & phant_id (Equality.class eT) ec => + fun m => EqQuotTypePack (@EqQuotClass Q qc ec m) Q. + +Definition EqQuotType_clone (Q : Type) eqT cT + of phant_id (eq_quot_class eqT) cT := @EqQuotTypePack Q cT Q. + +Lemma pi_eq_quot eqT : {mono \pi_eqT : x y / eq_quot_op x y >-> x == y}. +Proof. by case: eqT => [] ? []. Qed. + +Canonical pi_eq_quot_mono eqT := PiMono2 (pi_eq_quot eqT). + +End EqQuotTypeStructure. + +Notation EqQuotType e Q m := (@EqQuotType_pack _ e Q _ _ _ _ id id m). +Notation "[ 'eqQuotType' e 'of' Q ]" := (@EqQuotType_clone _ e Q _ _ id) + (at level 0, format "[ 'eqQuotType' e 'of' Q ]") : form_scope. + +(**************************************************************************) +(* Even if a quotType is a natural subType, we do not make this subType *) +(* canonical, to allow the user to define the subtyping he wants. However *) +(* one can: *) +(* - get the eqMixin and the choiceMixin by subtyping *) +(* - get the subType structure and maybe declare it Canonical. *) +(**************************************************************************) + +Module QuotSubType. +Section SubTypeMixin. + +Variable T : eqType. +Variable qT : quotType T. + +Definition Sub x (px : repr (\pi_qT x) == x) := \pi_qT x. + +Lemma qreprK x Px : repr (@Sub x Px) = x. +Proof. by rewrite /Sub (eqP Px). Qed. + +Lemma sortPx (x : qT) : repr (\pi_qT (repr x)) == repr x. +Proof. by rewrite !reprK eqxx. Qed. + +Lemma sort_Sub (x : qT) : x = Sub (sortPx x). +Proof. by rewrite /Sub reprK. Qed. + +Lemma reprP K (PK : forall x Px, K (@Sub x Px)) u : K u. +Proof. by rewrite (sort_Sub u); apply: PK. Qed. + +Canonical subType := SubType _ _ _ reprP qreprK. +Definition eqMixin := Eval hnf in [eqMixin of qT by <:]. + +Canonical eqType := EqType qT eqMixin. + +End SubTypeMixin. + +Definition choiceMixin (T : choiceType) (qT : quotType T) := + Eval hnf in [choiceMixin of qT by <:]. +Canonical choiceType (T : choiceType) (qT : quotType T) := + ChoiceType qT (@choiceMixin T qT). + +Definition countMixin (T : countType) (qT : quotType T) := + Eval hnf in [countMixin of qT by <:]. +Canonical countType (T : countType) (qT : quotType T) := + CountType qT (@countMixin T qT). + +Section finType. +Variables (T : finType) (qT : quotType T). +Canonical subCountType := [subCountType of qT]. +Definition finMixin := Eval hnf in [finMixin of qT by <:]. +End finType. + +End QuotSubType. + +Notation "[ 'subType' Q 'of' T 'by' %/ ]" := +(@SubType T _ Q _ _ (@QuotSubType.reprP _ _) (@QuotSubType.qreprK _ _)) +(at level 0, format "[ 'subType' Q 'of' T 'by' %/ ]") : form_scope. + +Notation "[ 'eqMixin' 'of' Q 'by' <:%/ ]" := + (@QuotSubType.eqMixin _ _: Equality.class_of Q) + (at level 0, format "[ 'eqMixin' 'of' Q 'by' <:%/ ]") : form_scope. + +Notation "[ 'choiceMixin' 'of' Q 'by' <:%/ ]" := + (@QuotSubType.choiceMixin _ _: Choice.mixin_of Q) + (at level 0, format "[ 'choiceMixin' 'of' Q 'by' <:%/ ]") : form_scope. + +Notation "[ 'countMixin' 'of' Q 'by' <:%/ ]" := + (@QuotSubType.countMixin _ _: Countable.mixin_of Q) + (at level 0, format "[ 'countMixin' 'of' Q 'by' <:%/ ]") : form_scope. + +Notation "[ 'finMixin' 'of' Q 'by' <:%/ ]" := + (@QuotSubType.finMixin _ _: Finite.mixin_of Q) + (at level 0, format "[ 'finMixin' 'of' Q 'by' <:%/ ]") : form_scope. + +(****************************************************) +(* Definition of a (decidable) equivalence relation *) +(****************************************************) + +Section EquivRel. + +Variable T : Type. + +Lemma left_trans (e : rel T) : + symmetric e -> transitive e -> left_transitive e. +Proof. by move=> s t ? * ?; apply/idP/idP; apply: t; rewrite // s. Qed. + +Lemma right_trans (e : rel T) : + symmetric e -> transitive e -> right_transitive e. +Proof. by move=> s t ? * x; rewrite ![e x _]s; apply: left_trans. Qed. + +CoInductive equiv_class_of (equiv : rel T) := + EquivClass of reflexive equiv & symmetric equiv & transitive equiv. + +Record equiv_rel := EquivRelPack { + equiv :> rel T; + _ : equiv_class_of equiv +}. + +Variable e : equiv_rel. + +Definition equiv_class := + let: EquivRelPack _ ce as e' := e return equiv_class_of e' in ce. + +Definition equiv_pack (r : rel T) ce of phant_id ce equiv_class := + @EquivRelPack r ce. + +Lemma equiv_refl x : e x x. Proof. by case: e => [] ? []. Qed. +Lemma equiv_sym : symmetric e. Proof. by case: e => [] ? []. Qed. +Lemma equiv_trans : transitive e. Proof. by case: e => [] ? []. Qed. + +Lemma eq_op_trans (T' : eqType) : transitive (@eq_op T'). +Proof. by move=> x y z; move/eqP->; move/eqP->. Qed. + +Lemma equiv_ltrans: left_transitive e. +Proof. by apply: left_trans; [apply: equiv_sym|apply: equiv_trans]. Qed. + +Lemma equiv_rtrans: right_transitive e. +Proof. by apply: right_trans; [apply: equiv_sym|apply: equiv_trans]. Qed. + +End EquivRel. + +Hint Resolve equiv_refl. + +Notation EquivRel r er es et := (@EquivRelPack _ r (EquivClass er es et)). +Notation "[ 'equiv_rel' 'of' e ]" := (@equiv_pack _ _ e _ id) + (at level 0, format "[ 'equiv_rel' 'of' e ]") : form_scope. + +(**************************************************) +(* Encoding to another type modulo an equivalence *) +(**************************************************) + +Section EncodingModuloRel. + +Variables (D E : Type) (ED : E -> D) (DE : D -> E) (e : rel D). + +CoInductive encModRel_class_of (r : rel D) := + EncModRelClassPack of (forall x, r x x -> r (ED (DE x)) x) & (r =2 e). + +Record encModRel := EncModRelPack { + enc_mod_rel :> rel D; + _ : encModRel_class_of enc_mod_rel +}. + +Variable r : encModRel. + +Definition encModRelClass := + let: EncModRelPack _ c as r' := r return encModRel_class_of r' in c. + +Definition encModRelP (x : D) : r x x -> r (ED (DE x)) x. +Proof. by case: r => [] ? [] /= he _ /he. Qed. + +Definition encModRelE : r =2 e. Proof. by case: r => [] ? []. Qed. + +Definition encoded_equiv : rel E := [rel x y | r (ED x) (ED y)]. + +End EncodingModuloRel. + +Notation EncModRelClass m := + (EncModRelClassPack (fun x _ => m x) (fun _ _ => erefl _)). +Notation EncModRel r m := (@EncModRelPack _ _ _ _ _ r (EncModRelClass m)). + +Section EncodingModuloEquiv. + +Variables (D E : Type) (ED : E -> D) (DE : D -> E) (e : equiv_rel D). +Variable (r : encModRel ED DE e). + +Lemma enc_mod_rel_is_equiv : equiv_class_of (enc_mod_rel r). +Proof. +split => [x|x y|y x z]; rewrite !encModRelE //; first by rewrite equiv_sym. +by move=> exy /(equiv_trans exy). +Qed. + +Definition enc_mod_rel_equiv_rel := EquivRelPack enc_mod_rel_is_equiv. + +Definition encModEquivP (x : D) : r (ED (DE x)) x. +Proof. by rewrite encModRelP ?encModRelE. Qed. + +Local Notation e' := (encoded_equiv r). + +Lemma encoded_equivE : e' =2 [rel x y | e (ED x) (ED y)]. +Proof. by move=> x y; rewrite /encoded_equiv /= encModRelE. Qed. +Local Notation e'E := encoded_equivE. + +Lemma encoded_equiv_is_equiv : equiv_class_of e'. +Proof. +split => [x|x y|y x z]; rewrite !e'E //=; first by rewrite equiv_sym. +by move=> exy /(equiv_trans exy). +Qed. + +Canonical encoded_equiv_equiv_rel := EquivRelPack encoded_equiv_is_equiv. + +Lemma encoded_equivP x : e' (DE (ED x)) x. +Proof. by rewrite /encoded_equiv /= encModEquivP. Qed. + +End EncodingModuloEquiv. + +(**************************************) +(* Quotient by a equivalence relation *) +(**************************************) + +Module EquivQuot. +Section EquivQuot. + +Variables (D : Type) (C : choiceType) (CD : C -> D) (DC : D -> C). +Variables (eD : equiv_rel D) (encD : encModRel CD DC eD). +Notation eC := (encoded_equiv encD). + +Definition canon x := choose (eC x) (x). + +Record equivQuotient := EquivQuotient { + erepr : C; + _ : (frel canon) erepr erepr +}. + +Definition type_of of (phantom (rel _) encD) := equivQuotient. + +Lemma canon_id : forall x, (invariant canon canon) x. +Proof. +move=> x /=; rewrite /canon (@eq_choose _ _ (eC x)). + by rewrite (@choose_id _ (eC x) _ x) ?chooseP ?equiv_refl. +by move=> y; apply: equiv_ltrans; rewrite equiv_sym /= chooseP. +Qed. + +Definition pi := locked (fun x => EquivQuotient (canon_id x)). + +Lemma ereprK : cancel erepr pi. +Proof. +unlock pi; case=> x hx; move/eqP:(hx)=> hx'. +exact: (@val_inj _ _ [subType for erepr]). +Qed. + +Local Notation encDE := (encModRelE encD). +Local Notation encDP := (encModEquivP encD). +Canonical encD_equiv_rel := EquivRelPack (enc_mod_rel_is_equiv encD). + +Lemma pi_CD (x y : C) : reflect (pi x = pi y) (eC x y). +Proof. +apply: (iffP idP) => hxy. + apply: (can_inj ereprK); unlock pi canon => /=. + rewrite -(@eq_choose _ (eC x) (eC y)); last first. + by move=> z; rewrite /eC /=; apply: equiv_ltrans. + by apply: choose_id; rewrite ?equiv_refl //. +rewrite (equiv_trans (chooseP (equiv_refl _ _))) //=. +move: hxy => /(f_equal erepr) /=; unlock pi canon => /= ->. +by rewrite equiv_sym /= chooseP. +Qed. + +Lemma pi_DC (x y : D) : + reflect (pi (DC x) = pi (DC y)) (eD x y). +Proof. +apply: (iffP idP)=> hxy. + apply/pi_CD; rewrite /eC /=. + by rewrite (equiv_ltrans (encDP _)) (equiv_rtrans (encDP _)) /= encDE. +rewrite -encDE -(equiv_ltrans (encDP _)) -(equiv_rtrans (encDP _)) /=. +exact/pi_CD. +Qed. + +Lemma equivQTP : cancel (CD \o erepr) (pi \o DC). +Proof. +by move=> x; rewrite /= (pi_CD _ (erepr x) _) ?ereprK /eC /= ?encDP. +Qed. + +Local Notation qT := (type_of (Phantom (rel D) encD)). +Definition quotClass := QuotClass equivQTP. +Canonical quotType := QuotType qT quotClass. + +Lemma eqmodP x y : reflect (x = y %[mod qT]) (eD x y). +Proof. by apply: (iffP (pi_DC _ _)); rewrite !unlock. Qed. + +Fact eqMixin : Equality.mixin_of qT. Proof. exact: CanEqMixin ereprK. Qed. +Canonical eqType := EqType qT eqMixin. +Definition choiceMixin := CanChoiceMixin ereprK. +Canonical choiceType := ChoiceType qT choiceMixin. + +Lemma eqmodE x y : x == y %[mod qT] = eD x y. +Proof. exact: sameP eqP (@eqmodP _ _). Qed. + +Canonical eqQuotType := EqQuotType eD qT eqmodE. + +End EquivQuot. +End EquivQuot. + +Canonical EquivQuot.quotType. +Canonical EquivQuot.eqType. +Canonical EquivQuot.choiceType. +Canonical EquivQuot.eqQuotType. + +Notation "{eq_quot e }" := +(@EquivQuot.type_of _ _ _ _ _ _ (Phantom (rel _) e)) : quotient_scope. +Notation "x == y %[mod_eq r ]" := (x == y %[mod {eq_quot r}]) : quotient_scope. +Notation "x = y %[mod_eq r ]" := (x = y %[mod {eq_quot r}]) : quotient_scope. +Notation "x != y %[mod_eq r ]" := (x != y %[mod {eq_quot r}]) : quotient_scope. +Notation "x <> y %[mod_eq r ]" := (x <> y %[mod {eq_quot r}]) : quotient_scope. + +(***********************************************************) +(* If the type is directly a choiceType, no need to encode *) +(***********************************************************) + +Section DefaultEncodingModuloRel. + +Variables (D : choiceType) (r : rel D). + +Definition defaultEncModRelClass := + @EncModRelClassPack D D id id r r (fun _ rxx => rxx) (fun _ _ => erefl _). + +Canonical defaultEncModRel := EncModRelPack defaultEncModRelClass. + +End DefaultEncodingModuloRel. + +(***************************************************) +(* Recovering a potential countable type structure *) +(***************************************************) + +Section CountEncodingModuloRel. + +Variables (D : Type) (C : countType) (CD : C -> D) (DC : D -> C). +Variables (eD : equiv_rel D) (encD : encModRel CD DC eD). +Notation eC := (encoded_equiv encD). + +Fact eq_quot_countMixin : Countable.mixin_of {eq_quot encD}. +Proof. exact: CanCountMixin (@EquivQuot.ereprK _ _ _ _ _ _). Qed. +Canonical eq_quot_countType := CountType {eq_quot encD} eq_quot_countMixin. + +End CountEncodingModuloRel. + +Section EquivQuotTheory. + +Variables (T : choiceType) (e : equiv_rel T) (Q : eqQuotType e). + +Lemma eqmodE x y : x == y %[mod_eq e] = e x y. +Proof. by rewrite pi_eq_quot. Qed. + +Lemma eqmodP x y : reflect (x = y %[mod_eq e]) (e x y). +Proof. by rewrite -eqmodE; apply/eqP. Qed. + +End EquivQuotTheory. + +Prenex Implicits eqmodE eqmodP. + +Section EqQuotTheory. + +Variables (T : Type) (e : rel T) (Q : eqQuotType e). + +Lemma eqquotE x y : x == y %[mod Q] = e x y. +Proof. by rewrite pi_eq_quot. Qed. + +Lemma eqquotP x y : reflect (x = y %[mod Q]) (e x y). +Proof. by rewrite -eqquotE; apply/eqP. Qed. + +End EqQuotTheory. + +Prenex Implicits eqquotE eqquotP. diff --git a/mathcomp/discrete/path.v b/mathcomp/discrete/path.v new file mode 100644 index 0000000..804e673 --- /dev/null +++ b/mathcomp/discrete/path.v @@ -0,0 +1,890 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. + +(******************************************************************************) +(* The basic theory of paths over an eqType; this file is essentially a *) +(* complement to seq.v. Paths are non-empty sequences that obey a progression *) +(* relation. They are passed around in three parts: the head and tail of the *) +(* sequence, and a proof of (boolean) predicate asserting the progression. *) +(* This "exploded" view is rarely embarrassing, as the first two parameters *) +(* are usually inferred from the type of the third; on the contrary, it saves *) +(* the hassle of constantly constructing and destructing a dependent record. *) +(* We define similarly cycles, for which we allow the empty sequence, *) +(* which represents a non-rooted empty cycle; by contrast, the "empty" path *) +(* from a point x is the one-item sequence containing only x. *) +(* We allow duplicates; uniqueness, if desired (as is the case for several *) +(* geometric constructions), must be asserted separately. We do provide *) +(* shorthand, but only for cycles, because the equational properties of *) +(* "path" and "uniq" are unfortunately incompatible (esp. wrt "cat"). *) +(* We define notations for the common cases of function paths, where the *) +(* progress relation is actually a function. In detail: *) +(* path e x p == x :: p is an e-path [:: x_0; x_1; ... ; x_n], i.e., we *) +(* e x_i x_{i+1} for all i < n. The path x :: p starts at x *) +(* and ends at last x p. *) +(* fpath f x p == x :: p is an f-path, where f is a function, i.e., p is of *) +(* the form [:: f x; f (f x); ...]. This is just a notation *) +(* for path (frel f) x p. *) +(* sorted e s == s is an e-sorted sequence: either s = [::], or s = x :: p *) +(* is an e-path (this is oten used with e = leq or ltn). *) +(* cycle e c == c is an e-cycle: either c = [::], or c = x :: p with *) +(* x :: (rcons p x) an e-path. *) +(* fcycle f c == c is an f-cycle, for a function f. *) +(* traject f x n == the f-path of size n starting at x *) +(* := [:: x; f x; ...; iter n.-1 f x] *) +(* looping f x n == the f-paths of size greater than n starting at x loop *) +(* back, or, equivalently, traject f x n contains all *) +(* iterates of f at x. *) +(* merge e s1 s2 == the e-sorted merge of sequences s1 and s2: this is always *) +(* a permutation of s1 ++ s2, and is e-sorted when s1 and s2 *) +(* are and e is total. *) +(* sort e s == a permutation of the sequence s, that is e-sorted when e *) +(* is total (computed by a merge sort with the merge function *) +(* above). *) +(* mem2 s x y == x, then y occur in the sequence (path) s; this is *) +(* non-strict: mem2 s x x = (x \in s). *) +(* next c x == the successor of the first occurrence of x in the sequence *) +(* c (viewed as a cycle), or x if x \notin c. *) +(* prev c x == the predecessor of the first occurrence of x in the *) +(* sequence c (viewed as a cycle), or x if x \notin c. *) +(* arc c x y == the sub-arc of the sequece c (viewed as a cycle) starting *) +(* at the first occurrence of x in c, and ending just before *) +(* the next ocurrence of y (in cycle order); arc c x y *) +(* returns an unspecified sub-arc of c if x and y do not both *) +(* occur in c. *) +(* ucycle e c <-> ucycleb e c (ucycle e c is a Coercion target of type Prop) *) +(* ufcycle f c <-> c is a simple f-cycle, for a function f. *) +(* shorten x p == the tail a duplicate-free subpath of x :: p with the same *) +(* endpoints (x and last x p), obtained by removing all loops *) +(* from x :: p. *) +(* rel_base e e' h b <-> the function h is a functor from relation e to *) +(* relation e', EXCEPT at points whose image under h satisfy *) +(* the "base" predicate b: *) +(* e' (h x) (h y) = e x y UNLESS b (h x) holds *) +(* This is the statement of the side condition of the path *) +(* functorial mapping lemma map_path. *) +(* fun_base f f' h b <-> the function h is a functor from function f to f', *) +(* except at the preimage of predicate b under h. *) +(* We also provide three segmenting dependently-typed lemmas (splitP, splitPl *) +(* and splitPr) whose elimination split a path x0 :: p at an internal point x *) +(* as follows: *) +(* - splitP applies when x \in p; it replaces p with (rcons p1 x ++ p2), so *) +(* that x appears explicitly at the end of the left part. The elimination *) +(* of splitP will also simultaneously replace take (index x p) with p1 and *) +(* drop (index x p).+1 p with p2. *) +(* - splitPl applies when x \in x0 :: p; it replaces p with p1 ++ p2 and *) +(* simulaneously generates an equation x = last x0 p. *) +(* - splitPr applies when x \in p; it replaces p with (p1 ++ x :: p2), so x *) +(* appears explicitly at the start of the right part. *) +(* The parts p1 and p2 are computed using index/take/drop in all cases, but *) +(* only splitP attemps to subsitute the explicit values. The substitution of *) +(* p can be deferred using the dependent equation generation feature of *) +(* ssreflect, e.g.: case/splitPr def_p: {1}p / x_in_p => [p1 p2] generates *) +(* the equation p = p1 ++ p2 instead of performing the substitution outright. *) +(* Similarly, eliminating the loop removal lemma shortenP simultaneously *) +(* replaces shorten e x p with a fresh constant p', and last x p with *) +(* last x p'. *) +(* Note that although all "path" functions actually operate on the *) +(* underlying sequence, we provide a series of lemmas that define their *) +(* interaction with thepath and cycle predicates, e.g., the cat_path equation *) +(* can be used to split the path predicate after splitting the underlying *) +(* sequence. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Section Paths. + +Variables (n0 : nat) (T : Type). + +Section Path. + +Variables (x0_cycle : T) (e : rel T). + +Fixpoint path x (p : seq T) := + if p is y :: p' then e x y && path y p' else true. + +Lemma cat_path x p1 p2 : path x (p1 ++ p2) = path x p1 && path (last x p1) p2. +Proof. by elim: p1 x => [|y p1 Hrec] x //=; rewrite Hrec -!andbA. Qed. + +Lemma rcons_path x p y : path x (rcons p y) = path x p && e (last x p) y. +Proof. by rewrite -cats1 cat_path /= andbT. Qed. + +Lemma pathP x p x0 : + reflect (forall i, i < size p -> e (nth x0 (x :: p) i) (nth x0 p i)) + (path x p). +Proof. +elim: p x => [|y p IHp] x /=; first by left. +apply: (iffP andP) => [[e_xy /IHp e_p [] //] | e_p]. +by split; [exact: (e_p 0) | apply/(IHp y) => i; exact: e_p i.+1]. +Qed. + +Definition cycle p := if p is x :: p' then path x (rcons p' x) else true. + +Lemma cycle_path p : cycle p = path (last x0_cycle p) p. +Proof. by case: p => //= x p; rewrite rcons_path andbC. Qed. + +Lemma rot_cycle p : cycle (rot n0 p) = cycle p. +Proof. +case: n0 p => [|n] [|y0 p] //=; first by rewrite /rot /= cats0. +rewrite /rot /= -{3}(cat_take_drop n p) -cats1 -catA cat_path. +case: (drop n p) => [|z0 q]; rewrite /= -cats1 !cat_path /= !andbT andbC //. +by rewrite last_cat; repeat bool_congr. +Qed. + +Lemma rotr_cycle p : cycle (rotr n0 p) = cycle p. +Proof. by rewrite -rot_cycle rotrK. Qed. + +End Path. + +Lemma eq_path e e' : e =2 e' -> path e =2 path e'. +Proof. by move=> ee' x p; elim: p x => //= y p IHp x; rewrite ee' IHp. Qed. + +Lemma eq_cycle e e' : e =2 e' -> cycle e =1 cycle e'. +Proof. by move=> ee' [|x p] //=; exact: eq_path. Qed. + +Lemma sub_path e e' : subrel e e' -> forall x p, path e x p -> path e' x p. +Proof. by move=> ee' x p; elim: p x => //= y p IHp x /andP[/ee'-> /IHp]. Qed. + +Lemma rev_path e x p : + path e (last x p) (rev (belast x p)) = path (fun z => e^~ z) x p. +Proof. +elim: p x => //= y p IHp x; rewrite rev_cons rcons_path -{}IHp andbC. +by rewrite -(last_cons x) -rev_rcons -lastI rev_cons last_rcons. +Qed. + +End Paths. + +Implicit Arguments pathP [T e x p]. +Prenex Implicits pathP. + +Section EqPath. + +Variables (n0 : nat) (T : eqType) (x0_cycle : T) (e : rel T). +Implicit Type p : seq T. + +CoInductive split x : seq T -> seq T -> seq T -> Type := + Split p1 p2 : split x (rcons p1 x ++ p2) p1 p2. + +Lemma splitP p x (i := index x p) : + x \in p -> split x p (take i p) (drop i.+1 p). +Proof. +move=> p_x; have lt_ip: i < size p by rewrite index_mem. +by rewrite -{1}(cat_take_drop i p) (drop_nth x lt_ip) -cat_rcons nth_index. +Qed. + +CoInductive splitl x1 x : seq T -> Type := + Splitl p1 p2 of last x1 p1 = x : splitl x1 x (p1 ++ p2). + +Lemma splitPl x1 p x : x \in x1 :: p -> splitl x1 x p. +Proof. +rewrite inE; case: eqP => [->| _ /splitP[]]; first by rewrite -(cat0s p). +by split; exact: last_rcons. +Qed. + +CoInductive splitr x : seq T -> Type := + Splitr p1 p2 : splitr x (p1 ++ x :: p2). + +Lemma splitPr p x : x \in p -> splitr x p. +Proof. by case/splitP=> p1 p2; rewrite cat_rcons. Qed. + +Fixpoint next_at x y0 y p := + match p with + | [::] => if x == y then y0 else x + | y' :: p' => if x == y then y' else next_at x y0 y' p' + end. + +Definition next p x := if p is y :: p' then next_at x y y p' else x. + +Fixpoint prev_at x y0 y p := + match p with + | [::] => if x == y0 then y else x + | y' :: p' => if x == y' then y else prev_at x y0 y' p' + end. + +Definition prev p x := if p is y :: p' then prev_at x y y p' else x. + +Lemma next_nth p x : + next p x = if x \in p then + if p is y :: p' then nth y p' (index x p) else x + else x. +Proof. +case: p => //= y0 p. +elim: p {2 3 5}y0 => [|y' p IHp] y /=; rewrite (eq_sym y) inE; + by case: ifP => // _; exact: IHp. +Qed. + +Lemma prev_nth p x : + prev p x = if x \in p then + if p is y :: p' then nth y p (index x p') else x + else x. +Proof. +case: p => //= y0 p; rewrite inE orbC. +elim: p {2 5}y0 => [|y' p IHp] y; rewrite /= ?inE // (eq_sym y'). +by case: ifP => // _; exact: IHp. +Qed. + +Lemma mem_next p x : (next p x \in p) = (x \in p). +Proof. +rewrite next_nth; case p_x: (x \in p) => //. +case: p (index x p) p_x => [|y0 p'] //= i _; rewrite inE. +have [lt_ip | ge_ip] := ltnP i (size p'); first by rewrite orbC mem_nth. +by rewrite nth_default ?eqxx. +Qed. + +Lemma mem_prev p x : (prev p x \in p) = (x \in p). +Proof. +rewrite prev_nth; case p_x: (x \in p) => //; case: p => [|y0 p] // in p_x *. +by apply mem_nth; rewrite /= ltnS index_size. +Qed. + +(* ucycleb is the boolean predicate, but ucycle is defined as a Prop *) +(* so that it can be used as a coercion target. *) +Definition ucycleb p := cycle e p && uniq p. +Definition ucycle p : Prop := cycle e p && uniq p. + +(* Projections, used for creating local lemmas. *) +Lemma ucycle_cycle p : ucycle p -> cycle e p. +Proof. by case/andP. Qed. + +Lemma ucycle_uniq p : ucycle p -> uniq p. +Proof. by case/andP. Qed. + +Lemma next_cycle p x : cycle e p -> x \in p -> e x (next p x). +Proof. +case: p => //= y0 p; elim: p {1 3 5}y0 => [|z p IHp] y /=; rewrite inE. + by rewrite andbT; case: (x =P y) => // ->. +by case/andP=> eyz /IHp; case: (x =P y) => // ->. +Qed. + +Lemma prev_cycle p x : cycle e p -> x \in p -> e (prev p x) x. +Proof. +case: p => //= y0 p; rewrite inE orbC. +elim: p {1 5}y0 => [|z p IHp] y /=; rewrite ?inE. + by rewrite andbT; case: (x =P y0) => // ->. +by case/andP=> eyz /IHp; case: (x =P z) => // ->. +Qed. + +Lemma rot_ucycle p : ucycle (rot n0 p) = ucycle p. +Proof. by rewrite /ucycle rot_uniq rot_cycle. Qed. + +Lemma rotr_ucycle p : ucycle (rotr n0 p) = ucycle p. +Proof. by rewrite /ucycle rotr_uniq rotr_cycle. Qed. + +(* The "appears no later" partial preorder defined by a path. *) + +Definition mem2 p x y := y \in drop (index x p) p. + +Lemma mem2l p x y : mem2 p x y -> x \in p. +Proof. +by rewrite /mem2 -!index_mem size_drop ltn_subRL; apply/leq_ltn_trans/leq_addr. +Qed. + +Lemma mem2lf {p x y} : x \notin p -> mem2 p x y = false. +Proof. exact/contraNF/mem2l. Qed. + +Lemma mem2r p x y : mem2 p x y -> y \in p. +Proof. +by rewrite -[in y \in p](cat_take_drop (index x p) p) mem_cat orbC /mem2 => ->. +Qed. + +Lemma mem2rf {p x y} : y \notin p -> mem2 p x y = false. +Proof. exact/contraNF/mem2r. Qed. + +Lemma mem2_cat p1 p2 x y : + mem2 (p1 ++ p2) x y = mem2 p1 x y || mem2 p2 x y || (x \in p1) && (y \in p2). +Proof. +rewrite [LHS]/mem2 index_cat fun_if if_arg !drop_cat addKn. +case: ifPn => [p1x | /mem2lf->]; last by rewrite ltnNge leq_addr orbF. +by rewrite index_mem p1x mem_cat -orbA (orb_idl (@mem2r _ _ _)). +Qed. + +Lemma mem2_splice p1 p3 x y p2 : + mem2 (p1 ++ p3) x y -> mem2 (p1 ++ p2 ++ p3) x y. +Proof. +by rewrite !mem2_cat mem_cat andb_orr orbC => /or3P[]->; rewrite ?orbT. +Qed. + +Lemma mem2_splice1 p1 p3 x y z : + mem2 (p1 ++ p3) x y -> mem2 (p1 ++ z :: p3) x y. +Proof. exact: mem2_splice [::z]. Qed. + +Lemma mem2_cons x p y z : + mem2 (x :: p) y z = (if x == y then z \in x :: p else mem2 p y z). +Proof. by rewrite [LHS]/mem2 /=; case: ifP. Qed. + +Lemma mem2_seq1 x y z : mem2 [:: x] y z = (y == x) && (z == x). +Proof. by rewrite mem2_cons eq_sym inE. Qed. + +Lemma mem2_last y0 p x : mem2 p x (last y0 p) = (x \in p). +Proof. +apply/idP/idP; first exact: mem2l; rewrite -index_mem /mem2 => p_x. +by rewrite -nth_last -(subnKC p_x) -nth_drop mem_nth // size_drop subnSK. +Qed. + +Lemma mem2l_cat {p1 p2 x} : x \notin p1 -> mem2 (p1 ++ p2) x =1 mem2 p2 x. +Proof. by move=> p1'x y; rewrite mem2_cat (negPf p1'x) mem2lf ?orbF. Qed. + +Lemma mem2r_cat {p1 p2 x y} : y \notin p2 -> mem2 (p1 ++ p2) x y = mem2 p1 x y. +Proof. +by move=> p2'y; rewrite mem2_cat (negPf p2'y) -orbA orbC andbF mem2rf. +Qed. + +Lemma mem2lr_splice {p1 p2 p3 x y} : + x \notin p2 -> y \notin p2 -> mem2 (p1 ++ p2 ++ p3) x y = mem2 (p1 ++ p3) x y. +Proof. +move=> p2'x p2'y; rewrite catA !mem2_cat !mem_cat. +by rewrite (negPf p2'x) (negPf p2'y) (mem2lf p2'x) andbF !orbF. +Qed. + +CoInductive split2r x y : seq T -> Type := + Split2r p1 p2 of y \in x :: p2 : split2r x y (p1 ++ x :: p2). + +Lemma splitP2r p x y : mem2 p x y -> split2r x y p. +Proof. +move=> pxy; have px := mem2l pxy. +have:= pxy; rewrite /mem2 (drop_nth x) ?index_mem ?nth_index //. +by case/splitP: px => p1 p2; rewrite cat_rcons. +Qed. + +Fixpoint shorten x p := + if p is y :: p' then + if x \in p then shorten x p' else y :: shorten y p' + else [::]. + +CoInductive shorten_spec x p : T -> seq T -> Type := + ShortenSpec p' of path e x p' & uniq (x :: p') & subpred (mem p') (mem p) : + shorten_spec x p (last x p') p'. + +Lemma shortenP x p : path e x p -> shorten_spec x p (last x p) (shorten x p). +Proof. +move=> e_p; have: x \in x :: p by exact: mem_head. +elim: p x {1 3 5}x e_p => [|y2 p IHp] x y1. + by rewrite mem_seq1 => _ /eqP->. +rewrite inE orbC /= => /andP[ey12 /IHp {IHp}IHp]. +case: ifPn => [y2p_x _ | not_y2p_x /eqP def_x]. + have [p' e_p' Up' p'p] := IHp _ y2p_x. + by split=> // y /p'p; exact: predU1r. +have [p' e_p' Up' p'p] := IHp y2 (mem_head y2 p). +have{p'p} p'p z: z \in y2 :: p' -> z \in y2 :: p. + by rewrite !inE; case: (z == y2) => // /p'p. +rewrite -(last_cons y1) def_x; split=> //=; first by rewrite ey12. +by rewrite (contra (p'p y1)) -?def_x. +Qed. + +End EqPath. + + +(* Ordered paths and sorting. *) + +Section SortSeq. + +Variable T : eqType. +Variable leT : rel T. + +Definition sorted s := if s is x :: s' then path leT x s' else true. + +Lemma path_sorted x s : path leT x s -> sorted s. +Proof. by case: s => //= y s /andP[]. Qed. + +Lemma path_min_sorted x s : + {in s, forall y, leT x y} -> path leT x s = sorted s. +Proof. by case: s => //= y s -> //; exact: mem_head. Qed. + +Section Transitive. + +Hypothesis leT_tr : transitive leT. + +Lemma subseq_order_path x s1 s2 : + subseq s1 s2 -> path leT x s2 -> path leT x s1. +Proof. +elim: s2 x s1 => [|y s2 IHs] x [|z s1] //= {IHs}/(IHs y). +case: eqP => [-> | _] IHs /andP[] => [-> // | leTxy /IHs /=]. +by case/andP=> /(leT_tr leTxy)->. +Qed. + +Lemma order_path_min x s : path leT x s -> all (leT x) s. +Proof. +move/subseq_order_path=> le_x_s; apply/allP=> y. +by rewrite -sub1seq => /le_x_s/andP[]. +Qed. + +Lemma subseq_sorted s1 s2 : subseq s1 s2 -> sorted s2 -> sorted s1. +Proof. +case: s1 s2 => [|x1 s1] [|x2 s2] //= sub_s12 /(subseq_order_path sub_s12). +by case: eqP => [-> | _ /andP[]]. +Qed. + +Lemma sorted_filter a s : sorted s -> sorted (filter a s). +Proof. exact: subseq_sorted (filter_subseq a s). Qed. + +Lemma sorted_uniq : irreflexive leT -> forall s, sorted s -> uniq s. +Proof. +move=> leT_irr; elim=> //= x s IHs s_ord. +rewrite (IHs (path_sorted s_ord)) andbT; apply/negP=> s_x. +by case/allPn: (order_path_min s_ord); exists x; rewrite // leT_irr. +Qed. + +Lemma eq_sorted : antisymmetric leT -> + forall s1 s2, sorted s1 -> sorted s2 -> perm_eq s1 s2 -> s1 = s2. +Proof. +move=> leT_asym; elim=> [|x1 s1 IHs1] s2 //= ord_s1 ord_s2 eq_s12. + by case: {+}s2 (perm_eq_size eq_s12). +have s2_x1: x1 \in s2 by rewrite -(perm_eq_mem eq_s12) mem_head. +case: s2 s2_x1 eq_s12 ord_s2 => //= x2 s2; rewrite in_cons. +case: eqP => [<- _| ne_x12 /= s2_x1] eq_s12 ord_s2. + by rewrite {IHs1}(IHs1 s2) ?(@path_sorted x1) // -(perm_cons x1). +case: (ne_x12); apply: leT_asym; rewrite (allP (order_path_min ord_s2)) //. +have: x2 \in x1 :: s1 by rewrite (perm_eq_mem eq_s12) mem_head. +case/predU1P=> [eq_x12 | s1_x2]; first by case ne_x12. +by rewrite (allP (order_path_min ord_s1)). +Qed. + +Lemma eq_sorted_irr : irreflexive leT -> + forall s1 s2, sorted s1 -> sorted s2 -> s1 =i s2 -> s1 = s2. +Proof. +move=> leT_irr s1 s2 s1_sort s2_sort eq_s12. +have: antisymmetric leT. + by move=> m n /andP[? ltnm]; case/idP: (leT_irr m); exact: leT_tr ltnm. +by move/eq_sorted; apply=> //; apply: uniq_perm_eq => //; exact: sorted_uniq. +Qed. + +End Transitive. + +Hypothesis leT_total : total leT. + +Fixpoint merge s1 := + if s1 is x1 :: s1' then + let fix merge_s1 s2 := + if s2 is x2 :: s2' then + if leT x2 x1 then x2 :: merge_s1 s2' else x1 :: merge s1' s2 + else s1 in + merge_s1 + else id. + +Lemma merge_path x s1 s2 : + path leT x s1 -> path leT x s2 -> path leT x (merge s1 s2). +Proof. +elim: s1 s2 x => //= x1 s1 IHs1. +elim=> //= x2 s2 IHs2 x /andP[le_x_x1 ord_s1] /andP[le_x_x2 ord_s2]. +case: ifP => le_x21 /=; first by rewrite le_x_x2 {}IHs2 // le_x21. +by rewrite le_x_x1 IHs1 //=; have:= leT_total x2 x1; rewrite le_x21 /= => ->. +Qed. + +Lemma merge_sorted s1 s2 : sorted s1 -> sorted s2 -> sorted (merge s1 s2). +Proof. +case: s1 s2 => [|x1 s1] [|x2 s2] //= ord_s1 ord_s2. +case: ifP => le_x21 /=. + by apply: (@merge_path x2 (x1 :: s1)) => //=; rewrite le_x21. +by apply: merge_path => //=; have:= leT_total x2 x1; rewrite le_x21 /= => ->. +Qed. + +Lemma perm_merge s1 s2 : perm_eql (merge s1 s2) (s1 ++ s2). +Proof. +apply/perm_eqlP; rewrite perm_eq_sym; elim: s1 s2 => //= x1 s1 IHs1. +elim=> [|x2 s2 IHs2]; rewrite /= ?cats0 //. +case: ifP => _ /=; last by rewrite perm_cons. +by rewrite (perm_catCA (_ :: _) [::x2]) perm_cons. +Qed. + +Lemma mem_merge s1 s2 : merge s1 s2 =i s1 ++ s2. +Proof. by apply: perm_eq_mem; rewrite perm_merge. Qed. + +Lemma size_merge s1 s2 : size (merge s1 s2) = size (s1 ++ s2). +Proof. by apply: perm_eq_size; rewrite perm_merge. Qed. + +Lemma merge_uniq s1 s2 : uniq (merge s1 s2) = uniq (s1 ++ s2). +Proof. by apply: perm_eq_uniq; rewrite perm_merge. Qed. + +Fixpoint merge_sort_push s1 ss := + match ss with + | [::] :: ss' | [::] as ss' => s1 :: ss' + | s2 :: ss' => [::] :: merge_sort_push (merge s1 s2) ss' + end. + +Fixpoint merge_sort_pop s1 ss := + if ss is s2 :: ss' then merge_sort_pop (merge s1 s2) ss' else s1. + +Fixpoint merge_sort_rec ss s := + if s is [:: x1, x2 & s'] then + let s1 := if leT x1 x2 then [:: x1; x2] else [:: x2; x1] in + merge_sort_rec (merge_sort_push s1 ss) s' + else merge_sort_pop s ss. + +Definition sort := merge_sort_rec [::]. + +Lemma sort_sorted s : sorted (sort s). +Proof. +rewrite /sort; have allss: all sorted [::] by []. +elim: {s}_.+1 {-2}s [::] allss (ltnSn (size s)) => // n IHn s ss allss. +have: sorted s -> sorted (merge_sort_pop s ss). + elim: ss allss s => //= s2 ss IHss /andP[ord_s2 ord_ss] s ord_s. + exact: IHss ord_ss _ (merge_sorted ord_s ord_s2). +case: s => [|x1 [|x2 s _]]; try by auto. +move/ltnW/IHn; apply=> {n IHn s}; set s1 := if _ then _ else _. +have: sorted s1 by exact: (@merge_sorted [::x2] [::x1]). +elim: ss {x1 x2}s1 allss => /= [|s2 ss IHss] s1; first by rewrite andbT. +case/andP=> ord_s2 ord_ss ord_s1. +by case: {1}s2=> /= [|_ _]; [rewrite ord_s1 | exact: IHss (merge_sorted _ _)]. +Qed. + +Lemma perm_sort s : perm_eql (sort s) s. +Proof. +rewrite /sort; apply/perm_eqlP; pose catss := foldr (@cat T) [::]. +rewrite perm_eq_sym -{1}[s]/(catss [::] ++ s). +elim: {s}_.+1 {-2}s [::] (ltnSn (size s)) => // n IHn s ss. +have: perm_eq (catss ss ++ s) (merge_sort_pop s ss). + elim: ss s => //= s2 ss IHss s1; rewrite -{IHss}(perm_eqrP (IHss _)). + by rewrite perm_catC catA perm_catC perm_cat2l -perm_merge. +case: s => // x1 [//|x2 s _]; move/ltnW; move/IHn=> {n IHn}IHs. +rewrite -{IHs}(perm_eqrP (IHs _)) ifE; set s1 := if_expr _ _ _. +rewrite (catA _ [::_;_] s) {s}perm_cat2r. +apply: (@perm_eq_trans _ (catss ss ++ s1)). + by rewrite perm_cat2l /s1 -ifE; case: ifP; rewrite // (perm_catC [::_]). +elim: ss {x1 x2}s1 => /= [|s2 ss IHss] s1; first by rewrite cats0. +rewrite perm_catC; case def_s2: {2}s2=> /= [|y s2']; first by rewrite def_s2. +by rewrite catA -{IHss}(perm_eqrP (IHss _)) perm_catC perm_cat2l -perm_merge. +Qed. + +Lemma mem_sort s : sort s =i s. +Proof. by apply: perm_eq_mem; rewrite perm_sort. Qed. + +Lemma size_sort s : size (sort s) = size s. +Proof. by apply: perm_eq_size; rewrite perm_sort. Qed. + +Lemma sort_uniq s : uniq (sort s) = uniq s. +Proof. by apply: perm_eq_uniq; rewrite perm_sort. Qed. + +Lemma perm_sortP : transitive leT -> antisymmetric leT -> + forall s1 s2, reflect (sort s1 = sort s2) (perm_eq s1 s2). +Proof. +move=> leT_tr leT_asym s1 s2. +apply: (iffP idP) => eq12; last by rewrite -perm_sort eq12 perm_sort. +apply: eq_sorted; rewrite ?sort_sorted //. +by rewrite perm_sort (perm_eqlP eq12) -perm_sort. +Qed. + +End SortSeq. + +Lemma rev_sorted (T : eqType) (leT : rel T) s : + sorted leT (rev s) = sorted (fun y x => leT x y) s. +Proof. by case: s => //= x p; rewrite -rev_path lastI rev_rcons. Qed. + +Lemma ltn_sorted_uniq_leq s : sorted ltn s = uniq s && sorted leq s. +Proof. +case: s => //= n s; elim: s n => //= m s IHs n. +rewrite inE ltn_neqAle negb_or IHs -!andbA. +case sn: (n \in s); last do !bool_congr. +rewrite andbF; apply/and5P=> [[ne_nm lenm _ _ le_ms]]; case/negP: ne_nm. +rewrite eqn_leq lenm; exact: (allP (order_path_min leq_trans le_ms)). +Qed. + +Lemma iota_sorted i n : sorted leq (iota i n). +Proof. by elim: n i => // [[|n] //= IHn] i; rewrite IHn leqW. Qed. + +Lemma iota_ltn_sorted i n : sorted ltn (iota i n). +Proof. by rewrite ltn_sorted_uniq_leq iota_sorted iota_uniq. Qed. + +(* Function trajectories. *) + +Notation fpath f := (path (coerced_frel f)). +Notation fcycle f := (cycle (coerced_frel f)). +Notation ufcycle f := (ucycle (coerced_frel f)). + +Prenex Implicits path next prev cycle ucycle mem2. + +Section Trajectory. + +Variables (T : Type) (f : T -> T). + +Fixpoint traject x n := if n is n'.+1 then x :: traject (f x) n' else [::]. + +Lemma trajectS x n : traject x n.+1 = x :: traject (f x) n. +Proof. by []. Qed. + +Lemma trajectSr x n : traject x n.+1 = rcons (traject x n) (iter n f x). +Proof. by elim: n x => //= n IHn x; rewrite IHn -iterSr. Qed. + +Lemma last_traject x n : last x (traject (f x) n) = iter n f x. +Proof. by case: n => // n; rewrite iterSr trajectSr last_rcons. Qed. + +Lemma traject_iteri x n : + traject x n = iteri n (fun i => rcons^~ (iter i f x)) [::]. +Proof. by elim: n => //= n <-; rewrite -trajectSr. Qed. + +Lemma size_traject x n : size (traject x n) = n. +Proof. by elim: n x => //= n IHn x //=; rewrite IHn. Qed. + +Lemma nth_traject i n : i < n -> forall x, nth x (traject x n) i = iter i f x. +Proof. +elim: n => // n IHn; rewrite ltnS leq_eqVlt => le_i_n x. +rewrite trajectSr nth_rcons size_traject. +case: ltngtP le_i_n => [? _||->] //; exact: IHn. +Qed. + +End Trajectory. + +Section EqTrajectory. + +Variables (T : eqType) (f : T -> T). + +Lemma eq_fpath f' : f =1 f' -> fpath f =2 fpath f'. +Proof. by move/eq_frel/eq_path. Qed. + +Lemma eq_fcycle f' : f =1 f' -> fcycle f =1 fcycle f'. +Proof. by move/eq_frel/eq_cycle. Qed. + +Lemma fpathP x p : reflect (exists n, p = traject f (f x) n) (fpath f x p). +Proof. +elim: p x => [|y p IHp] x; first by left; exists 0. +rewrite /= andbC; case: IHp => [fn_p | not_fn_p]; last first. + by right=> [] [[//|n]] [<- fn_p]; case: not_fn_p; exists n. +apply: (iffP eqP) => [-> | [[] // _ []//]]. +by have [n ->] := fn_p; exists n.+1. +Qed. + +Lemma fpath_traject x n : fpath f x (traject f (f x) n). +Proof. by apply/(fpathP x); exists n. Qed. + +Definition looping x n := iter n f x \in traject f x n. + +Lemma loopingP x n : + reflect (forall m, iter m f x \in traject f x n) (looping x n). +Proof. +apply: (iffP idP) => loop_n; last exact: loop_n. +case: n => // n in loop_n *; elim=> [|m /= IHm]; first exact: mem_head. +move: (fpath_traject x n) loop_n; rewrite /looping !iterS -last_traject /=. +move: (iter m f x) IHm => y /splitPl[p1 p2 def_y]. +rewrite cat_path last_cat def_y; case: p2 => // z p2 /and3P[_ /eqP-> _] _. +by rewrite inE mem_cat mem_head !orbT. +Qed. + +Lemma trajectP x n y : + reflect (exists2 i, i < n & y = iter i f x) (y \in traject f x n). +Proof. +elim: n x => [|n IHn] x /=; first by right; case. +rewrite inE; have [-> | /= neq_xy] := eqP; first by left; exists 0. +apply: {IHn}(iffP (IHn _)) => [[i] | [[|i]]] // lt_i_n ->. + by exists i.+1; rewrite ?iterSr. +by exists i; rewrite ?iterSr. +Qed. + +Lemma looping_uniq x n : uniq (traject f x n.+1) = ~~ looping x n. +Proof. +rewrite /looping; elim: n x => [|n IHn] x //. +rewrite {-3}[n.+1]lock /= -lock {}IHn -iterSr -negb_or inE; congr (~~ _). +apply: orb_id2r => /trajectP no_loop. +apply/idP/eqP => [/trajectP[m le_m_n def_x] | {1}<-]; last first. + by rewrite iterSr -last_traject mem_last. +have loop_m: looping x m.+1 by rewrite /looping iterSr -def_x mem_head. +have/trajectP[[|i] // le_i_m def_fn1x] := loopingP _ _ loop_m n.+1. +by case: no_loop; exists i; rewrite -?iterSr // -ltnS (leq_trans le_i_m). +Qed. + +End EqTrajectory. + +Implicit Arguments fpathP [T f x p]. +Implicit Arguments loopingP [T f x n]. +Implicit Arguments trajectP [T f x n y]. +Prenex Implicits traject fpathP loopingP trajectP. + +Section UniqCycle. + +Variables (n0 : nat) (T : eqType) (e : rel T) (p : seq T). + +Hypothesis Up : uniq p. + +Lemma prev_next : cancel (next p) (prev p). +Proof. +move=> x; rewrite prev_nth mem_next next_nth; case p_x: (x \in p) => //. +case def_p: p Up p_x => // [y q]; rewrite -{-1}def_p => /= /andP[not_qy Uq] p_x. +rewrite -{2}(nth_index y p_x); congr (nth y _ _); set i := index x p. +have: ~~ (size q < i) by rewrite -index_mem -/i def_p leqNgt in p_x. +case: ltngtP => // [lt_i_q | ->] _; first by rewrite index_uniq. +by apply/eqP; rewrite nth_default // eqn_leq index_size leqNgt index_mem. +Qed. + +Lemma next_prev : cancel (prev p) (next p). +Proof. +move=> x; rewrite next_nth mem_prev prev_nth; case p_x: (x \in p) => //. +case def_p: p p_x => // [y q]; rewrite -def_p => p_x. +rewrite index_uniq //; last by rewrite def_p ltnS index_size. +case q_x: (x \in q); first exact: nth_index. +rewrite nth_default; last by rewrite leqNgt index_mem q_x. +by apply/eqP; rewrite def_p inE q_x orbF eq_sym in p_x. +Qed. + +Lemma cycle_next : fcycle (next p) p. +Proof. +case def_p: {-2}p Up => [|x q] Uq //. +apply/(pathP x)=> i; rewrite size_rcons => le_i_q. +rewrite -cats1 -cat_cons nth_cat le_i_q /= next_nth {}def_p mem_nth //. +rewrite index_uniq // nth_cat /= ltn_neqAle andbC -ltnS le_i_q. +by case: (i =P _) => //= ->; rewrite subnn nth_default. +Qed. + +Lemma cycle_prev : cycle (fun x y => x == prev p y) p. +Proof. +apply: etrans cycle_next; symmetry; case def_p: p => [|x q] //. +apply: eq_path; rewrite -def_p; exact (can2_eq prev_next next_prev). +Qed. + +Lemma cycle_from_next : (forall x, x \in p -> e x (next p x)) -> cycle e p. +Proof. +case: p (next p) cycle_next => //= [x q] n; rewrite -(belast_rcons x q x). +move: {q}(rcons q x) => q n_q; move/allP. +by elim: q x n_q => //= _ q IHq x /andP[/eqP <- n_q] /andP[-> /IHq->]. +Qed. + +Lemma cycle_from_prev : (forall x, x \in p -> e (prev p x) x) -> cycle e p. +Proof. +move=> e_p; apply: cycle_from_next => x p_x. +by rewrite -{1}[x]prev_next e_p ?mem_next. +Qed. + +Lemma next_rot : next (rot n0 p) =1 next p. +Proof. +move=> x; have n_p := cycle_next; rewrite -(rot_cycle n0) in n_p. +case p_x: (x \in p); last by rewrite !next_nth mem_rot p_x. +by rewrite (eqP (next_cycle n_p _)) ?mem_rot. +Qed. + +Lemma prev_rot : prev (rot n0 p) =1 prev p. +Proof. +move=> x; have p_p := cycle_prev; rewrite -(rot_cycle n0) in p_p. +case p_x: (x \in p); last by rewrite !prev_nth mem_rot p_x. +by rewrite (eqP (prev_cycle p_p _)) ?mem_rot. +Qed. + +End UniqCycle. + +Section UniqRotrCycle. + +Variables (n0 : nat) (T : eqType) (p : seq T). + +Hypothesis Up : uniq p. + +Lemma next_rotr : next (rotr n0 p) =1 next p. Proof. exact: next_rot. Qed. + +Lemma prev_rotr : prev (rotr n0 p) =1 prev p. Proof. exact: prev_rot. Qed. + +End UniqRotrCycle. + +Section UniqCycleRev. + +Variable T : eqType. +Implicit Type p : seq T. + +Lemma prev_rev p : uniq p -> prev (rev p) =1 next p. +Proof. +move=> Up x; case p_x: (x \in p); last first. + by rewrite next_nth prev_nth mem_rev p_x. +case/rot_to: p_x (Up) => [i q def_p] Urp; rewrite -rev_uniq in Urp. +rewrite -(prev_rotr i Urp); do 2 rewrite -(prev_rotr 1) ?rotr_uniq //. +rewrite -rev_rot -(next_rot i Up) {i p Up Urp}def_p. +by case: q => // y q; rewrite !rev_cons !(=^~ rcons_cons, rotr1_rcons) /= eqxx. +Qed. + +Lemma next_rev p : uniq p -> next (rev p) =1 prev p. +Proof. by move=> Up x; rewrite -{2}[p]revK prev_rev // rev_uniq. Qed. + +End UniqCycleRev. + +Section MapPath. + +Variables (T T' : Type) (h : T' -> T) (e : rel T) (e' : rel T'). + +Definition rel_base (b : pred T) := + forall x' y', ~~ b (h x') -> e (h x') (h y') = e' x' y'. + +Lemma map_path b x' p' (Bb : rel_base b) : + ~~ has (preim h b) (belast x' p') -> + path e (h x') (map h p') = path e' x' p'. +Proof. by elim: p' x' => [|y' p' IHp'] x' //= /norP[/Bb-> /IHp'->]. Qed. + +End MapPath. + +Section MapEqPath. + +Variables (T T' : eqType) (h : T' -> T) (e : rel T) (e' : rel T'). + +Hypothesis Ih : injective h. + +Lemma mem2_map x' y' p' : mem2 (map h p') (h x') (h y') = mem2 p' x' y'. +Proof. by rewrite {1}/mem2 (index_map Ih) -map_drop mem_map. Qed. + +Lemma next_map p : uniq p -> forall x, next (map h p) (h x) = h (next p x). +Proof. +move=> Up x; case p_x: (x \in p); last by rewrite !next_nth (mem_map Ih) p_x. +case/rot_to: p_x => i p' def_p. +rewrite -(next_rot i Up); rewrite -(map_inj_uniq Ih) in Up. +rewrite -(next_rot i Up) -map_rot {i p Up}def_p /=. +by case: p' => [|y p''] //=; rewrite !eqxx. +Qed. + +Lemma prev_map p : uniq p -> forall x, prev (map h p) (h x) = h (prev p x). +Proof. +move=> Up x; rewrite -{1}[x](next_prev Up) -(next_map Up). +by rewrite prev_next ?map_inj_uniq. +Qed. + +End MapEqPath. + +Definition fun_base (T T' : eqType) (h : T' -> T) f f' := + rel_base h (frel f) (frel f'). + +Section CycleArc. + +Variable T : eqType. +Implicit Type p : seq T. + +Definition arc p x y := let px := rot (index x p) p in take (index y px) px. + +Lemma arc_rot i p : uniq p -> {in p, arc (rot i p) =2 arc p}. +Proof. +move=> Up x p_x y; congr (fun q => take (index y q) q); move: Up p_x {y}. +rewrite -{1 2 5 6}(cat_take_drop i p) /rot cat_uniq => /and3P[_ Up12 _]. +rewrite !drop_cat !take_cat !index_cat mem_cat orbC. +case p2x: (x \in drop i p) => /= => [_ | p1x]. + rewrite index_mem p2x [x \in _](negbTE (hasPn Up12 _ p2x)) /= addKn. + by rewrite ltnNge leq_addr catA. +by rewrite p1x index_mem p1x addKn ltnNge leq_addr /= catA. +Qed. + +Lemma left_arc x y p1 p2 (p := x :: p1 ++ y :: p2) : + uniq p -> arc p x y = x :: p1. +Proof. +rewrite /arc /p [index x _]/= eqxx rot0 -cat_cons cat_uniq index_cat. +move: (x :: p1) => xp1 /and3P[_ /norP[/= /negbTE-> _] _]. +by rewrite eqxx addn0 take_size_cat. +Qed. + +Lemma right_arc x y p1 p2 (p := x :: p1 ++ y :: p2) : + uniq p -> arc p y x = y :: p2. +Proof. +rewrite -[p]cat_cons -rot_size_cat rot_uniq => Up. +by rewrite arc_rot ?left_arc ?mem_head. +Qed. + +CoInductive rot_to_arc_spec p x y := + RotToArcSpec i p1 p2 of x :: p1 = arc p x y + & y :: p2 = arc p y x + & rot i p = x :: p1 ++ y :: p2 : + rot_to_arc_spec p x y. + +Lemma rot_to_arc p x y : + uniq p -> x \in p -> y \in p -> x != y -> rot_to_arc_spec p x y. +Proof. +move=> Up p_x p_y ne_xy; case: (rot_to p_x) (p_y) (Up) => [i q def_p] q_y. +rewrite -(mem_rot i) def_p inE eq_sym (negbTE ne_xy) in q_y. +rewrite -(rot_uniq i) def_p. +case/splitPr: q / q_y def_p => q1 q2 def_p Uq12; exists i q1 q2 => //. + by rewrite -(arc_rot i Up p_x) def_p left_arc. +by rewrite -(arc_rot i Up p_y) def_p right_arc. +Qed. + +End CycleArc. + +Prenex Implicits arc. + diff --git a/mathcomp/discrete/prime.v b/mathcomp/discrete/prime.v new file mode 100644 index 0000000..fa39012 --- /dev/null +++ b/mathcomp/discrete/prime.v @@ -0,0 +1,1404 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq path fintype. +Require Import div bigop. + +(******************************************************************************) +(* This file contains the definitions of: *) +(* prime p <=> p is a prime. *) +(* primes m == the sorted list of prime divisors of m > 1, else [::]. *) +(* pfactor == the type of prime factors, syntax (p ^ e)%pfactor. *) +(* prime_decomp m == the list of prime factors of m > 1, sorted by primes. *) +(* logn p m == the e such that (p ^ e) \in prime_decomp n, else 0. *) +(* trunc_log p m == the largest e such that p ^ e <= m, or 0 if p or m is 0. *) +(* pdiv n == the smallest prime divisor of n > 1, else 1. *) +(* max_pdiv n == the largest prime divisor of n > 1, else 1. *) +(* divisors m == the sorted list of divisors of m > 0, else [::]. *) +(* totient n == the Euler totient (#|{i < n | i and n coprime}|). *) +(* nat_pred == the type of explicit collective nat predicates. *) +(* := simpl_pred nat. *) +(* -> We allow the coercion nat >-> nat_pred, interpreting p as pred1 p. *) +(* -> We define a predType for nat_pred, enabling the notation p \in pi. *) +(* -> We don't have nat_pred >-> pred, which would imply nat >-> Funclass. *) +(* pi^' == the complement of pi : nat_pred, i.e., the nat_pred such *) +(* that (p \in pi^') = (p \notin pi). *) +(* \pi(n) == the set of prime divisors of n, i.e., the nat_pred such *) +(* that (p \in \pi(n)) = (p \in primes n). *) +(* \pi(A) == the set of primes of #|A|, with A a collective predicate *) +(* over a finite Type. *) +(* -> The notation \pi(A) is implemented with a collapsible Coercion, so *) +(* the type of A must coerce to finpred_class (e.g., by coercing to *) +(* {set T}), not merely implement the predType interface (as seq T *) +(* does). *) +(* -> The expression #|A| will only appear in \pi(A) after simplification *) +(* collapses the coercion stack, so it is advisable to do so early on. *) +(* pi.-nat n <=> n > 0 and all prime divisors of n are in pi. *) +(* n`_pi == the pi-part of n -- the largest pi.-nat divisor of n. *) +(* := \prod_(0 <= p < n.+1 | p \in pi) p ^ logn p n. *) +(* -> The nat >-> nat_pred coercion lets us write p.-nat n and n`_p. *) +(* In addition to the lemmas relevant to these definitions, this file also *) +(* contains the dvdn_sum lemma, so that bigop.v doesn't depend on div.v. *) +(******************************************************************************) + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(* The complexity of any arithmetic operation with the Peano representation *) +(* is pretty dreadful, so using algorithms for "harder" problems such as *) +(* factoring, that are geared for efficient artihmetic leads to dismal *) +(* performance -- it takes a significant time, for instance, to compute the *) +(* divisors of just a two-digit number. On the other hand, for Peano *) +(* integers, prime factoring (and testing) is linear-time with a small *) +(* constant factor -- indeed, the same as converting in and out of a binary *) +(* representation. This is implemented by the code below, which is then *) +(* used to give the "standard" definitions of prime, primes, and divisors, *) +(* which can then be used casually in proofs with moderately-sized numeric *) +(* values (indeed, the code here performs well for up to 6-digit numbers). *) + +(* We start with faster mod-2 functions. *) + +Fixpoint edivn2 q r := if r is r'.+2 then edivn2 q.+1 r' else (q, r). + +Lemma edivn2P n : edivn_spec n 2 (edivn2 0 n). +Proof. +rewrite -[n]odd_double_half addnC -{1}[n./2]addn0 -{1}mul2n mulnC. +elim: n./2 {1 4}0 => [|r IHr] q; first by case (odd n) => /=. +rewrite addSnnS; exact: IHr. +Qed. + +Fixpoint elogn2 e q r {struct q} := + match q, r with + | 0, _ | _, 0 => (e, q) + | q'.+1, 1 => elogn2 e.+1 q' q' + | q'.+1, r'.+2 => elogn2 e q' r' + end. + +CoInductive elogn2_spec n : nat * nat -> Type := + Elogn2Spec e m of n = 2 ^ e * m.*2.+1 : elogn2_spec n (e, m). + +Lemma elogn2P n : elogn2_spec n.+1 (elogn2 0 n n). +Proof. +rewrite -{1}[n.+1]mul1n -[1]/(2 ^ 0) -{1}(addKn n n) addnn. +elim: n {1 4 6}n {2 3}0 (leqnn n) => [|q IHq] [|[|r]] e //=; last first. + by move/ltnW; exact: IHq. +clear 1; rewrite subn1 -[_.-1.+1]doubleS -mul2n mulnA -expnSr. +rewrite -{1}(addKn q q) addnn; exact: IHq. +Qed. + +Definition ifnz T n (x y : T) := if n is 0 then y else x. + +CoInductive ifnz_spec T n (x y : T) : T -> Type := + | IfnzPos of n > 0 : ifnz_spec n x y x + | IfnzZero of n = 0 : ifnz_spec n x y y. + +Lemma ifnzP T n (x y : T) : ifnz_spec n x y (ifnz n x y). +Proof. by case: n => [|n]; [right | left]. Qed. + +(* For pretty-printing. *) +Definition NumFactor (f : nat * nat) := ([Num of f.1], f.2). + +Definition pfactor p e := p ^ e. + +Definition cons_pfactor (p e : nat) pd := ifnz e ((p, e) :: pd) pd. + +Notation Local "p ^? e :: pd" := (cons_pfactor p e pd) + (at level 30, e at level 30, pd at level 60) : nat_scope. + +Section prime_decomp. + +Import NatTrec. + +Fixpoint prime_decomp_rec m k a b c e := + let p := k.*2.+1 in + if a is a'.+1 then + if b - (ifnz e 1 k - c) is b'.+1 then + [rec m, k, a', b', ifnz c c.-1 (ifnz e p.-2 1), e] else + if (b == 0) && (c == 0) then + let b' := k + a' in [rec b'.*2.+3, k, a', b', k.-1, e.+1] else + let bc' := ifnz e (ifnz b (k, 0) (edivn2 0 c)) (b, c) in + p ^? e :: ifnz a' [rec m, k.+1, a'.-1, bc'.1 + a', bc'.2, 0] [:: (m, 1)] + else if (b == 0) && (c == 0) then [:: (p, e.+2)] else p ^? e :: [:: (m, 1)] +where "[ 'rec' m , k , a , b , c , e ]" := (prime_decomp_rec m k a b c e). + +Definition prime_decomp n := + let: (e2, m2) := elogn2 0 n.-1 n.-1 in + if m2 < 2 then 2 ^? e2 :: 3 ^? m2 :: [::] else + let: (a, bc) := edivn m2.-2 3 in + let: (b, c) := edivn (2 - bc) 2 in + 2 ^? e2 :: [rec m2.*2.+1, 1, a, b, c, 0]. + +(* The list of divisors and the Euler function are computed directly from *) +(* the decomposition, using a merge_sort variant sort the divisor list. *) + +Definition add_divisors f divs := + let: (p, e) := f in + let add1 divs' := merge leq (map (NatTrec.mul p) divs') divs in + iter e add1 divs. + +Definition add_totient_factor f m := let: (p, e) := f in p.-1 * p ^ e.-1 * m. + +End prime_decomp. + +Definition primes n := unzip1 (prime_decomp n). + +Definition prime p := if prime_decomp p is [:: (_ , 1)] then true else false. + +Definition nat_pred := simpl_pred nat. + +Definition pi_unwrapped_arg := nat. +Definition pi_wrapped_arg := wrapped nat. +Coercion unwrap_pi_arg (wa : pi_wrapped_arg) : pi_unwrapped_arg := unwrap wa. +Coercion pi_arg_of_nat (n : nat) := Wrap n : pi_wrapped_arg. +Coercion pi_arg_of_fin_pred T pT (A : @fin_pred_sort T pT) : pi_wrapped_arg := + Wrap #|A|. + +Definition pi_of (n : pi_unwrapped_arg) : nat_pred := [pred p in primes n]. + +Notation "\pi ( n )" := (pi_of n) + (at level 2, format "\pi ( n )") : nat_scope. +Notation "\p 'i' ( A )" := \pi(#|A|) + (at level 2, format "\p 'i' ( A )") : nat_scope. + +Definition pdiv n := head 1 (primes n). + +Definition max_pdiv n := last 1 (primes n). + +Definition divisors n := foldr add_divisors [:: 1] (prime_decomp n). + +Definition totient n := foldr add_totient_factor (n > 0) (prime_decomp n). + +(* Correctness of the decomposition algorithm. *) + +Lemma prime_decomp_correct : + let pd_val pd := \prod_(f <- pd) pfactor f.1 f.2 in + let lb_dvd q m := ~~ has [pred d | d %| m] (index_iota 2 q) in + let pf_ok f := lb_dvd f.1 f.1 && (0 < f.2) in + let pd_ord q pd := path ltn q (unzip1 pd) in + let pd_ok q n pd := [/\ n = pd_val pd, all pf_ok pd & pd_ord q pd] in + forall n, n > 0 -> pd_ok 1 n (prime_decomp n). +Proof. +rewrite unlock => pd_val lb_dvd pf_ok pd_ord pd_ok. +have leq_pd_ok m p q pd: q <= p -> pd_ok p m pd -> pd_ok q m pd. + rewrite /pd_ok /pd_ord; case: pd => [|[r _] pd] //= leqp [<- ->]. + by case/andP=> /(leq_trans _)->. +have apd_ok m e q p pd: lb_dvd p p || (e == 0) -> q < p -> + pd_ok p m pd -> pd_ok q (p ^ e * m) (p ^? e :: pd). +- case: e => [|e]; rewrite orbC /= => pr_p ltqp. + rewrite mul1n; apply: leq_pd_ok; exact: ltnW. + by rewrite /pd_ok /pd_ord /pf_ok /= pr_p ltqp => [[<- -> ->]]. +case=> // n _; rewrite /prime_decomp. +case: elogn2P => e2 m2 -> {n}; case: m2 => [|[|abc]]; try exact: apd_ok. +rewrite [_.-2]/= !ltnS ltn0 natTrecE; case: edivnP => a bc ->{abc}. +case: edivnP => b c def_bc /= ltc2 ltbc3; apply: (apd_ok) => //. +move def_m: _.*2.+1 => m; set k := {2}1; rewrite -[2]/k.*2; set e := 0. +pose p := k.*2.+1; rewrite -{1}[m]mul1n -[1]/(p ^ e)%N. +have{def_m bc def_bc ltc2 ltbc3}: + let kb := (ifnz e k 1).*2 in + [&& k > 0, p < m, lb_dvd p m, c < kb & lb_dvd p p || (e == 0)] + /\ m + (b * kb + c).*2 = p ^ 2 + (a * p).*2. +- rewrite -{-2}def_m; split=> //=; last first. + by rewrite -def_bc addSn -doubleD 2!addSn -addnA subnKC // addnC. + rewrite ltc2 /lb_dvd /index_iota /= dvdn2 -def_m. + by rewrite [_.+2]lock /= odd_double. +move: {2}a.+1 (ltnSn a) => n; clearbody k e. +elim: n => // n IHn in a k p m b c e *; rewrite ltnS => le_a_n []. +set kb := _.*2; set d := _ + c => /and5P[lt0k ltpm leppm ltc pr_p def_m]. +have def_k1: k.-1.+1 = k := ltn_predK lt0k. +have def_kb1: kb.-1.+1 = kb by rewrite /kb -def_k1; case e. +have eq_bc_0: (b == 0) && (c == 0) = (d == 0). + by rewrite addn_eq0 muln_eq0 orbC -def_kb1. +have lt1p: 1 < p by rewrite ltnS double_gt0. +have co_p_2: coprime p 2 by rewrite /coprime gcdnC gcdnE modn2 /= odd_double. +have if_d0: d = 0 -> [/\ m = (p + a.*2) * p, lb_dvd p p & lb_dvd p (p + a.*2)]. + move=> d0; have{d0 def_m} def_m: m = (p + a.*2) * p. + by rewrite d0 addn0 -mulnn -!mul2n mulnA -mulnDl in def_m *. + split=> //; apply/hasPn=> r /(hasPn leppm); apply: contra => /= dv_r. + by rewrite def_m dvdn_mull. + by rewrite def_m dvdn_mulr. +case def_a: a => [|a'] /= in le_a_n *; rewrite !natTrecE -/p {}eq_bc_0. + case: d if_d0 def_m => [[//| def_m {pr_p}pr_p pr_m'] _ | d _ def_m] /=. + rewrite def_m def_a addn0 mulnA -2!expnSr. + by split; rewrite /pd_ord /pf_ok /= ?muln1 ?pr_p ?leqnn. + apply: apd_ok; rewrite // /pd_ok /= /pfactor expn1 muln1 /pd_ord /= ltpm. + rewrite /pf_ok !andbT /=; split=> //; apply: contra leppm. + case/hasP=> r /=; rewrite mem_index_iota => /andP[lt1r ltrm] dvrm; apply/hasP. + have [ltrp | lepr] := ltnP r p. + by exists r; rewrite // mem_index_iota lt1r. + case/dvdnP: dvrm => q def_q; exists q; last by rewrite def_q /= dvdn_mulr. + rewrite mem_index_iota -(ltn_pmul2r (ltnW lt1r)) -def_q mul1n ltrm. + move: def_m; rewrite def_a addn0 -(@ltn_pmul2r p) // mulnn => <-. + apply: (@leq_ltn_trans m); first by rewrite def_q leq_mul. + by rewrite -addn1 leq_add2l. +have def_k2: k.*2 = ifnz e 1 k * kb. + by rewrite /kb; case: (e) => [|e']; rewrite (mul1n, muln2). +case def_b': (b - _) => [|b']; last first. + have ->: ifnz e k.*2.-1 1 = kb.-1 by rewrite /kb; case e. + apply: IHn => {n le_a_n}//; rewrite -/p -/kb; split=> //. + rewrite lt0k ltpm leppm pr_p andbT /=. + by case: ifnzP; [move/ltn_predK->; exact: ltnW | rewrite def_kb1]. + apply: (@addIn p.*2). + rewrite -2!addnA -!doubleD -addnA -mulSnr -def_a -def_m /d. + have ->: b * kb = b' * kb + (k.*2 - c * kb + kb). + rewrite addnCA addnC -mulSnr -def_b' def_k2 -mulnBl -mulnDl subnK //. + by rewrite ltnW // -subn_gt0 def_b'. + rewrite -addnA; congr (_ + (_ + _).*2). + case: (c) ltc; first by rewrite -addSnnS def_kb1 subn0 addn0 addnC. + rewrite /kb; case e => [[] // _ | e' c' _] /=; last first. + by rewrite subnDA subnn addnC addSnnS. + by rewrite mul1n -doubleB -doubleD subn1 !addn1 def_k1. +have ltdp: d < p. + move/eqP: def_b'; rewrite subn_eq0 -(@leq_pmul2r kb); last first. + by rewrite -def_kb1. + rewrite mulnBl -def_k2 ltnS -(leq_add2r c); move/leq_trans; apply. + have{ltc} ltc: c < k.*2. + by apply: (leq_trans ltc); rewrite leq_double /kb; case e. + rewrite -{2}(subnK (ltnW ltc)) leq_add2r leq_sub2l //. + by rewrite -def_kb1 mulnS leq_addr. +case def_d: d if_d0 => [|d'] => [[//|{def_m ltdp pr_p} def_m pr_p pr_m'] | _]. + rewrite eqxx -doubleS -addnS -def_a doubleD -addSn -/p def_m. + rewrite mulnCA mulnC -expnSr. + apply: IHn => {n le_a_n}//; rewrite -/p -/kb; split. + rewrite lt0k -addn1 leq_add2l {1}def_a pr_m' pr_p /= def_k1 -addnn. + by rewrite leq_addr. + rewrite -addnA -doubleD addnCA def_a addSnnS def_k1 -(addnC k) -mulnSr. + rewrite -[_.*2.+1]/p mulnDl doubleD addnA -mul2n mulnA mul2n -mulSn. + by rewrite -/p mulnn. +have next_pm: lb_dvd p.+2 m. + rewrite /lb_dvd /index_iota 2!subSS subn0 -(subnK lt1p) iota_add. + rewrite has_cat; apply/norP; split=> //=; rewrite orbF subnKC // orbC. + apply/norP; split; apply/dvdnP=> [[q def_q]]. + case/hasP: leppm; exists 2; first by rewrite /p -(subnKC lt0k). + by rewrite /= def_q dvdn_mull // dvdn2 /= odd_double. + move/(congr1 (dvdn p)): def_m; rewrite -mulnn -!mul2n mulnA -mulnDl. + rewrite dvdn_mull // dvdn_addr; last by rewrite def_q dvdn_mull. + case/dvdnP=> r; rewrite mul2n => def_r; move: ltdp (congr1 odd def_r). + rewrite odd_double -ltn_double {1}def_r -mul2n ltn_pmul2r //. + by case: r def_r => [|[|[]]] //; rewrite def_d // mul1n /= odd_double. +apply: apd_ok => //; case: a' def_a le_a_n => [|a'] def_a => [_ | lta] /=. + rewrite /pd_ok /= /pfactor expn1 muln1 /pd_ord /= ltpm /pf_ok !andbT /=. + split=> //; apply: contra next_pm. + case/hasP=> q; rewrite mem_index_iota => /andP[lt1q ltqm] dvqm; apply/hasP. + have [ltqp | lepq] := ltnP q p.+2. + by exists q; rewrite // mem_index_iota lt1q. + case/dvdnP: dvqm => r def_r; exists r; last by rewrite def_r /= dvdn_mulr. + rewrite mem_index_iota -(ltn_pmul2r (ltnW lt1q)) -def_r mul1n ltqm /=. + rewrite -(@ltn_pmul2l p.+2) //; apply: (@leq_ltn_trans m). + by rewrite def_r mulnC leq_mul. + rewrite -addn2 mulnn sqrnD mul2n muln2 -addnn addnCA -addnA addnCA addnA. + by rewrite def_a mul1n in def_m; rewrite -def_m addnS -addnA ltnS leq_addr. +set bc := ifnz _ _ _; apply: leq_pd_ok (leqnSn _) _. +rewrite -doubleS -{1}[m]mul1n -[1]/(k.+1.*2.+1 ^ 0)%N. +apply: IHn; first exact: ltnW. +rewrite doubleS -/p [ifnz 0 _ _]/=; do 2?split => //. + rewrite orbT next_pm /= -(leq_add2r d.*2) def_m 2!addSnnS -doubleS leq_add. + - move: ltc; rewrite /kb {}/bc andbT; case e => //= e' _; case: ifnzP => //. + by case: edivn2P. + - by rewrite -{1}[p]muln1 -mulnn ltn_pmul2l. + by rewrite leq_double def_a mulSn (leq_trans ltdp) ?leq_addr. +rewrite mulnDl !muln2 -addnA addnCA doubleD addnCA. +rewrite (_ : _ + bc.2 = d); last first. + rewrite /d {}/bc /kb -muln2. + case: (e) (b) def_b' => //= _ []; first by case: edivn2P. + by case c; do 2?case; rewrite // mul1n /= muln2. +rewrite def_m 3!doubleS addnC -(addn2 p) sqrnD mul2n muln2 -3!addnA. +congr (_ + _); rewrite 4!addnS -!doubleD; congr _.*2.+2.+2. +by rewrite def_a -add2n mulnDl -addnA -muln2 -mulnDr mul2n. +Qed. + +Lemma primePn n : + reflect (n < 2 \/ exists2 d, 1 < d < n & d %| n) (~~ prime n). +Proof. +rewrite /prime; case: n => [|[|p2]]; try by do 2!left. +case: (@prime_decomp_correct p2.+2) => //; rewrite unlock. +case: prime_decomp => [|[q [|[|e]]] pd] //=; last first; last by rewrite andbF. + rewrite {1}/pfactor 2!expnS -!mulnA /=. + case: (_ ^ _ * _) => [|u -> _ /andP[lt1q _]]; first by rewrite !muln0. + left; right; exists q; last by rewrite dvdn_mulr. + have lt0q := ltnW lt1q; rewrite lt1q -{1}[q]muln1 ltn_pmul2l //. + by rewrite -[2]muln1 leq_mul. +rewrite {1}/pfactor expn1; case: pd => [|[r e] pd] /=; last first. + case: e => [|e] /=; first by rewrite !andbF. + rewrite {1}/pfactor expnS -mulnA. + case: (_ ^ _ * _) => [|u -> _ /and3P[lt1q ltqr _]]; first by rewrite !muln0. + left; right; exists q; last by rewrite dvdn_mulr. + by rewrite lt1q -{1}[q]mul1n ltn_mul // -[q.+1]muln1 leq_mul. +rewrite muln1 !andbT => def_q pr_q lt1q; right=> [[]] // [d]. +by rewrite def_q -mem_index_iota => in_d_2q dv_d_q; case/hasP: pr_q; exists d. +Qed. + +Lemma primeP p : + reflect (p > 1 /\ forall d, d %| p -> xpred2 1 p d) (prime p). +Proof. +rewrite -[prime p]negbK; have [npr_p | pr_p] := primePn p. + right=> [[lt1p pr_p]]; case: npr_p => [|[d n1pd]]. + by rewrite ltnNge lt1p. + by move/pr_p=> /orP[] /eqP def_d; rewrite def_d ltnn ?andbF in n1pd. +have [lep1 | lt1p] := leqP; first by case: pr_p; left. +left; split=> // d dv_d_p; apply/norP=> [[nd1 ndp]]; case: pr_p; right. +exists d; rewrite // andbC 2!ltn_neqAle ndp eq_sym nd1. +by have lt0p := ltnW lt1p; rewrite dvdn_leq // (dvdn_gt0 lt0p). +Qed. + +Lemma prime_nt_dvdP d p : prime p -> d != 1 -> reflect (d = p) (d %| p). +Proof. +case/primeP=> _ min_p d_neq1; apply: (iffP idP) => [/min_p|-> //]. +by rewrite (negPf d_neq1) /= => /eqP. +Qed. + +Implicit Arguments primeP [p]. +Implicit Arguments primePn [n]. +Prenex Implicits primePn primeP. + +Lemma prime_gt1 p : prime p -> 1 < p. +Proof. by case/primeP. Qed. + +Lemma prime_gt0 p : prime p -> 0 < p. +Proof. by move/prime_gt1; exact: ltnW. Qed. + +Hint Resolve prime_gt1 prime_gt0. + +Lemma prod_prime_decomp n : + n > 0 -> n = \prod_(f <- prime_decomp n) f.1 ^ f.2. +Proof. by case/prime_decomp_correct. Qed. + +Lemma even_prime p : prime p -> p = 2 \/ odd p. +Proof. +move=> pr_p; case odd_p: (odd p); [by right | left]. +have: 2 %| p by rewrite dvdn2 odd_p. +by case/primeP: pr_p => _ dv_p /dv_p/(2 =P p). +Qed. + +Lemma prime_oddPn p : prime p -> reflect (p = 2) (~~ odd p). +Proof. +by move=> p_pr; apply: (iffP idP) => [|-> //]; case/even_prime: p_pr => ->. +Qed. + +Lemma odd_prime_gt2 p : odd p -> prime p -> p > 2. +Proof. by move=> odd_p /prime_gt1; apply: odd_gt2. Qed. + +Lemma mem_prime_decomp n p e : + (p, e) \in prime_decomp n -> [/\ prime p, e > 0 & p ^ e %| n]. +Proof. +case: (posnP n) => [-> //| /prime_decomp_correct[def_n mem_pd ord_pd pd_pe]]. +have /andP[pr_p ->] := allP mem_pd _ pd_pe; split=> //; last first. + case/splitPr: pd_pe def_n => pd1 pd2 ->. + by rewrite big_cat big_cons /= mulnCA dvdn_mulr. +have lt1p: 1 < p. + apply: (allP (order_path_min ltn_trans ord_pd)). + by apply/mapP; exists (p, e). +apply/primeP; split=> // d dv_d_p; apply/norP=> [[nd1 ndp]]. +case/hasP: pr_p; exists d => //. +rewrite mem_index_iota andbC 2!ltn_neqAle ndp eq_sym nd1. +by have lt0p := ltnW lt1p; rewrite dvdn_leq // (dvdn_gt0 lt0p). +Qed. + +Lemma prime_coprime p m : prime p -> coprime p m = ~~ (p %| m). +Proof. +case/primeP=> p_gt1 p_pr; apply/eqP/negP=> [d1 | ndv_pm]. + case/dvdnP=> k def_m; rewrite -(addn0 m) def_m gcdnMDl gcdn0 in d1. + by rewrite d1 in p_gt1. +by apply: gcdn_def => // d /p_pr /orP[] /eqP->. +Qed. + +Lemma dvdn_prime2 p q : prime p -> prime q -> (p %| q) = (p == q). +Proof. +move=> pr_p pr_q; apply: negb_inj. +by rewrite eqn_dvd negb_and -!prime_coprime // coprime_sym orbb. +Qed. + +Lemma Euclid_dvdM m n p : prime p -> (p %| m * n) = (p %| m) || (p %| n). +Proof. +move=> pr_p; case dv_pm: (p %| m); first exact: dvdn_mulr. +by rewrite Gauss_dvdr // prime_coprime // dv_pm. +Qed. + +Lemma Euclid_dvd1 p : prime p -> (p %| 1) = false. +Proof. by rewrite dvdn1; case: eqP => // ->. Qed. + +Lemma Euclid_dvdX m n p : prime p -> (p %| m ^ n) = (p %| m) && (n > 0). +Proof. +case: n => [|n] pr_p; first by rewrite andbF Euclid_dvd1. +by apply: (inv_inj negbK); rewrite !andbT -!prime_coprime // coprime_pexpr. +Qed. + +Lemma mem_primes p n : (p \in primes n) = [&& prime p, n > 0 & p %| n]. +Proof. +rewrite andbCA; case: posnP => [-> // | /= n_gt0]. +apply/mapP/andP=> [[[q e]]|[pr_p]] /=. + case/mem_prime_decomp=> pr_q e_gt0; case/dvdnP=> u -> -> {p}. + by rewrite -(prednK e_gt0) expnS mulnCA dvdn_mulr. +rewrite {1}(prod_prime_decomp n_gt0) big_seq. +apply big_ind => [| u v IHu IHv | [q e] /= mem_qe dv_p_qe]. +- by rewrite Euclid_dvd1. +- by rewrite Euclid_dvdM // => /orP[]. +exists (q, e) => //=; case/mem_prime_decomp: mem_qe => pr_q _ _. +by rewrite Euclid_dvdX // dvdn_prime2 // in dv_p_qe; case: eqP dv_p_qe. +Qed. + +Lemma sorted_primes n : sorted ltn (primes n). +Proof. +by case: (posnP n) => [-> // | /prime_decomp_correct[_ _]]; exact: path_sorted. +Qed. + +Lemma eq_primes m n : (primes m =i primes n) <-> (primes m = primes n). +Proof. +split=> [eqpr| -> //]. +by apply: (eq_sorted_irr ltn_trans ltnn); rewrite ?sorted_primes. +Qed. + +Lemma primes_uniq n : uniq (primes n). +Proof. exact: (sorted_uniq ltn_trans ltnn (sorted_primes n)). Qed. + +(* The smallest prime divisor *) + +Lemma pi_pdiv n : (pdiv n \in \pi(n)) = (n > 1). +Proof. +case: n => [|[|n]] //; rewrite /pdiv !inE /primes. +have:= prod_prime_decomp (ltn0Sn n.+1); rewrite unlock. +by case: prime_decomp => //= pf pd _; rewrite mem_head. +Qed. + +Lemma pdiv_prime n : 1 < n -> prime (pdiv n). +Proof. by rewrite -pi_pdiv mem_primes; case/and3P. Qed. + +Lemma pdiv_dvd n : pdiv n %| n. +Proof. +by case: n (pi_pdiv n) => [|[|n]] //; rewrite mem_primes=> /and3P[]. +Qed. + +Lemma pi_max_pdiv n : (max_pdiv n \in \pi(n)) = (n > 1). +Proof. +rewrite !inE -pi_pdiv /max_pdiv /pdiv !inE. +by case: (primes n) => //= p ps; rewrite mem_head mem_last. +Qed. + +Lemma max_pdiv_prime n : n > 1 -> prime (max_pdiv n). +Proof. by rewrite -pi_max_pdiv mem_primes => /andP[]. Qed. + +Lemma max_pdiv_dvd n : max_pdiv n %| n. +Proof. +by case: n (pi_max_pdiv n) => [|[|n]] //; rewrite mem_primes => /andP[]. +Qed. + +Lemma pdiv_leq n : 0 < n -> pdiv n <= n. +Proof. by move=> n_gt0; rewrite dvdn_leq // pdiv_dvd. Qed. + +Lemma max_pdiv_leq n : 0 < n -> max_pdiv n <= n. +Proof. by move=> n_gt0; rewrite dvdn_leq // max_pdiv_dvd. Qed. + +Lemma pdiv_gt0 n : 0 < pdiv n. +Proof. by case: n => [|[|n]] //; rewrite prime_gt0 ?pdiv_prime. Qed. + +Lemma max_pdiv_gt0 n : 0 < max_pdiv n. +Proof. by case: n => [|[|n]] //; rewrite prime_gt0 ?max_pdiv_prime. Qed. +Hint Resolve pdiv_gt0 max_pdiv_gt0. + +Lemma pdiv_min_dvd m d : 1 < d -> d %| m -> pdiv m <= d. +Proof. +move=> lt1d dv_d_m; case: (posnP m) => [->|mpos]; first exact: ltnW. +rewrite /pdiv; apply: leq_trans (pdiv_leq (ltnW lt1d)). +have: pdiv d \in primes m. + by rewrite mem_primes mpos pdiv_prime // (dvdn_trans (pdiv_dvd d)). +case: (primes m) (sorted_primes m) => //= p pm ord_pm. +rewrite inE => /predU1P[-> //|]. +move/(allP (order_path_min ltn_trans ord_pm)); exact: ltnW. +Qed. + +Lemma max_pdiv_max n p : p \in \pi(n) -> p <= max_pdiv n. +Proof. +rewrite /max_pdiv !inE => n_p. +case/splitPr: n_p (sorted_primes n) => p1 p2; rewrite last_cat -cat_rcons /=. +rewrite headI /= cat_path -(last_cons 0) -headI last_rcons; case/andP=> _. +move/(order_path_min ltn_trans); case/lastP: p2 => //= p2 q. +by rewrite all_rcons last_rcons ltn_neqAle -andbA => /and3P[]. +Qed. + +Lemma ltn_pdiv2_prime n : 0 < n -> n < pdiv n ^ 2 -> prime n. +Proof. +case def_n: n => [|[|n']] // _; rewrite -def_n => lt_n_p2. +suffices ->: n = pdiv n by rewrite pdiv_prime ?def_n. +apply/eqP; rewrite eqn_leq leqNgt andbC pdiv_leq; last by rewrite def_n. +move: lt_n_p2; rewrite ltnNge; apply: contra => lt_pm_m. +case/dvdnP: (pdiv_dvd n) => q def_q. +rewrite {2}def_q -mulnn leq_pmul2r // pdiv_min_dvd //. + by rewrite -[pdiv n]mul1n {2}def_q ltn_pmul2r in lt_pm_m. +by rewrite def_q dvdn_mulr. +Qed. + +Lemma primePns n : + reflect (n < 2 \/ exists p, [/\ prime p, p ^ 2 <= n & p %| n]) (~~ prime n). +Proof. +apply: (iffP idP) => [npr_p|]; last first. + case=> [|[p [pr_p le_p2_n dv_p_n]]]; first by case: n => [|[]]. + apply/negP=> pr_n; move: dv_p_n le_p2_n; rewrite dvdn_prime2 //; move/eqP->. + by rewrite leqNgt -{1}[n]muln1 -mulnn ltn_pmul2l ?prime_gt1 ?prime_gt0. +case: leqP => [lt1p|]; [right | by left]. +exists (pdiv n); rewrite pdiv_dvd pdiv_prime //; split=> //. +by case: leqP npr_p => //; move/ltn_pdiv2_prime->; auto. +Qed. + +Implicit Arguments primePns [n]. +Prenex Implicits primePns. + +Lemma pdivP n : n > 1 -> {p | prime p & p %| n}. +Proof. by move=> lt1n; exists (pdiv n); rewrite ?pdiv_dvd ?pdiv_prime. Qed. + +Lemma primes_mul m n p : m > 0 -> n > 0 -> + (p \in primes (m * n)) = (p \in primes m) || (p \in primes n). +Proof. +move=> m_gt0 n_gt0; rewrite !mem_primes muln_gt0 m_gt0 n_gt0. +by case pr_p: (prime p); rewrite // Euclid_dvdM. +Qed. + +Lemma primes_exp m n : n > 0 -> primes (m ^ n) = primes m. +Proof. +case: n => // n _; rewrite expnS; case: (posnP m) => [-> //| m_gt0]. +apply/eq_primes => /= p; elim: n => [|n IHn]; first by rewrite muln1. +by rewrite primes_mul ?(expn_gt0, expnS, IHn, orbb, m_gt0). +Qed. + +Lemma primes_prime p : prime p -> primes p = [::p]. +Proof. +move=> pr_p; apply: (eq_sorted_irr ltn_trans ltnn) => // [|q]. + exact: sorted_primes. +rewrite mem_seq1 mem_primes prime_gt0 //=. +by apply/andP/idP=> [[pr_q q_p] | /eqP-> //]; rewrite -dvdn_prime2. +Qed. + +Lemma coprime_has_primes m n : m > 0 -> n > 0 -> + coprime m n = ~~ has (mem (primes m)) (primes n). +Proof. +move=> m_gt0 n_gt0; apply/eqnP/hasPn=> [mn1 p | no_p_mn]. + rewrite /= !mem_primes m_gt0 n_gt0 /= => /andP[pr_p p_n]. + have:= prime_gt1 pr_p; rewrite pr_p ltnNge -mn1 /=; apply: contra => p_m. + by rewrite dvdn_leq ?gcdn_gt0 ?m_gt0 // dvdn_gcd ?p_m. +case: (ltngtP (gcdn m n) 1) => //; first by rewrite ltnNge gcdn_gt0 ?m_gt0. +move/pdiv_prime; set p := pdiv _ => pr_p. +move/implyP: (no_p_mn p); rewrite /= !mem_primes m_gt0 n_gt0 pr_p /=. +by rewrite !(dvdn_trans (pdiv_dvd _)) // (dvdn_gcdl, dvdn_gcdr). +Qed. + +Lemma pdiv_id p : prime p -> pdiv p = p. +Proof. by move=> p_pr; rewrite /pdiv primes_prime. Qed. + +Lemma pdiv_pfactor p k : prime p -> pdiv (p ^ k.+1) = p. +Proof. by move=> p_pr; rewrite /pdiv primes_exp ?primes_prime. Qed. + +(* Primes are unbounded. *) + +Lemma dvdn_fact m n : 0 < m <= n -> m %| n`!. +Proof. +case: m => //= m; elim: n => //= n IHn; rewrite ltnS leq_eqVlt. +by case/predU1P=> [-> | /IHn]; [apply: dvdn_mulr | apply: dvdn_mull]. +Qed. + +Lemma prime_above m : {p | m < p & prime p}. +Proof. +have /pdivP[p pr_p p_dv_m1]: 1 < m`! + 1 by rewrite addn1 ltnS fact_gt0. +exists p => //; rewrite ltnNge; apply: contraL p_dv_m1 => p_le_m. +by rewrite dvdn_addr ?dvdn_fact ?prime_gt0 // gtnNdvd ?prime_gt1. +Qed. + +(* "prime" logarithms and p-parts. *) + +Fixpoint logn_rec d m r := + match r, edivn m d with + | r'.+1, (_.+1 as m', 0) => (logn_rec d m' r').+1 + | _, _ => 0 + end. + +Definition logn p m := if prime p then logn_rec p m m else 0. + +Lemma lognE p m : + logn p m = if [&& prime p, 0 < m & p %| m] then (logn p (m %/ p)).+1 else 0. +Proof. +rewrite /logn /dvdn; case p_pr: (prime p) => //. +rewrite /divn modn_def; case def_m: {2 3}m => [|m'] //=. +case: edivnP def_m => [[|q] [|r] -> _] // def_m; congr _.+1; rewrite [_.1]/=. +have{m def_m}: q < m'. + by rewrite -ltnS -def_m addn0 mulnC -{1}[q.+1]mul1n ltn_pmul2r // prime_gt1. +elim: {m' q}_.+1 {-2}m' q.+1 (ltnSn m') (ltn0Sn q) => // s IHs. +case=> [[]|r] //= m; rewrite ltnS => lt_rs m_gt0 le_mr. +rewrite -{3}[m]prednK //=; case: edivnP => [[|q] [|_] def_q _] //. +have{def_q} lt_qm': q < m.-1. + by rewrite -[q.+1]muln1 -ltnS prednK // def_q addn0 ltn_pmul2l // prime_gt1. +have{le_mr} le_m'r: m.-1 <= r by rewrite -ltnS prednK. +by rewrite (IHs r) ?(IHs m.-1) // ?(leq_trans lt_qm', leq_trans _ lt_rs). +Qed. + +Lemma logn_gt0 p n : (0 < logn p n) = (p \in primes n). +Proof. by rewrite lognE -mem_primes; case: {+}(p \in _). Qed. + +Lemma ltn_log0 p n : n < p -> logn p n = 0. +Proof. by case: n => [|n] ltnp; rewrite lognE ?andbF // gtnNdvd ?andbF. Qed. + +Lemma logn0 p : logn p 0 = 0. +Proof. by rewrite /logn if_same. Qed. + +Lemma logn1 p : logn p 1 = 0. +Proof. by rewrite lognE dvdn1 /= andbC; case: eqP => // ->. Qed. + +Lemma pfactor_gt0 p n : 0 < p ^ logn p n. +Proof. by rewrite expn_gt0 lognE; case: (posnP p) => // ->. Qed. +Hint Resolve pfactor_gt0. + +Lemma pfactor_dvdn p n m : prime p -> m > 0 -> (p ^ n %| m) = (n <= logn p m). +Proof. +move=> p_pr; elim: n m => [|n IHn] m m_gt0; first exact: dvd1n. +rewrite lognE p_pr m_gt0 /=; case dv_pm: (p %| m); last first. + apply/dvdnP=> [] [/= q def_m]. + by rewrite def_m expnS mulnCA dvdn_mulr in dv_pm. +case/dvdnP: dv_pm m_gt0 => q ->{m}; rewrite muln_gt0 => /andP[p_gt0 q_gt0]. +by rewrite expnSr dvdn_pmul2r // mulnK // IHn. +Qed. + +Lemma pfactor_dvdnn p n : p ^ logn p n %| n. +Proof. +case: n => // n; case pr_p: (prime p); first by rewrite pfactor_dvdn. +by rewrite lognE pr_p dvd1n. +Qed. + +Lemma logn_prime p q : prime q -> logn p q = (p == q). +Proof. +move=> pr_q; have q_gt0 := prime_gt0 pr_q; rewrite lognE q_gt0 /=. +case pr_p: (prime p); last by case: eqP pr_p pr_q => // -> ->. +by rewrite dvdn_prime2 //; case: eqP => // ->; rewrite divnn q_gt0 logn1. +Qed. + +Lemma pfactor_coprime p n : + prime p -> n > 0 -> {m | coprime p m & n = m * p ^ logn p n}. +Proof. +move=> p_pr n_gt0; set k := logn p n. +have dv_pk_n: p ^ k %| n by rewrite pfactor_dvdn. +exists (n %/ p ^ k); last by rewrite divnK. +rewrite prime_coprime // -(@dvdn_pmul2r (p ^ k)) ?expn_gt0 ?prime_gt0 //. +by rewrite -expnS divnK // pfactor_dvdn // ltnn. +Qed. + +Lemma pfactorK p n : prime p -> logn p (p ^ n) = n. +Proof. +move=> p_pr; have pn_gt0: p ^ n > 0 by rewrite expn_gt0 prime_gt0. +apply/eqP; rewrite eqn_leq -pfactor_dvdn // dvdnn andbT. +by rewrite -(leq_exp2l _ _ (prime_gt1 p_pr)) dvdn_leq // pfactor_dvdn. +Qed. + +Lemma pfactorKpdiv p n : prime p -> logn (pdiv (p ^ n)) (p ^ n) = n. +Proof. by case: n => // n p_pr; rewrite pdiv_pfactor ?pfactorK. Qed. + +Lemma dvdn_leq_log p m n : 0 < n -> m %| n -> logn p m <= logn p n. +Proof. +move=> n_gt0 dv_m_n; have m_gt0 := dvdn_gt0 n_gt0 dv_m_n. +case p_pr: (prime p); last by do 2!rewrite lognE p_pr /=. +by rewrite -pfactor_dvdn //; apply: dvdn_trans dv_m_n; rewrite pfactor_dvdn. +Qed. + +Lemma ltn_logl p n : 0 < n -> logn p n < n. +Proof. +move=> n_gt0; have [p_gt1 | p_le1] := boolP (1 < p). + by rewrite (leq_trans (ltn_expl _ p_gt1)) // dvdn_leq ?pfactor_dvdnn. +by rewrite lognE (contraNF (@prime_gt1 _)). +Qed. + +Lemma logn_Gauss p m n : coprime p m -> logn p (m * n) = logn p n. +Proof. +move=> co_pm; case p_pr: (prime p); last by rewrite /logn p_pr. +have [-> | n_gt0] := posnP n; first by rewrite muln0. +have [m0 | m_gt0] := posnP m; first by rewrite m0 prime_coprime ?dvdn0 in co_pm. +have mn_gt0: m * n > 0 by rewrite muln_gt0 m_gt0. +apply/eqP; rewrite eqn_leq andbC dvdn_leq_log ?dvdn_mull //. +set k := logn p _; have: p ^ k %| m * n by rewrite pfactor_dvdn. +by rewrite Gauss_dvdr ?coprime_expl // -pfactor_dvdn. +Qed. + +Lemma lognM p m n : 0 < m -> 0 < n -> logn p (m * n) = logn p m + logn p n. +Proof. +case p_pr: (prime p); last by rewrite /logn p_pr. +have xlp := pfactor_coprime p_pr. +case/xlp=> m' co_m' def_m /xlp[n' co_n' def_n] {xlp}. +by rewrite {1}def_m {1}def_n mulnCA -mulnA -expnD !logn_Gauss // pfactorK. +Qed. + +Lemma lognX p m n : logn p (m ^ n) = n * logn p m. +Proof. +case p_pr: (prime p); last by rewrite /logn p_pr muln0. +elim: n => [|n IHn]; first by rewrite logn1. +have [->|m_gt0] := posnP m; first by rewrite exp0n // lognE andbF muln0. +by rewrite expnS lognM ?IHn // expn_gt0 m_gt0. +Qed. + +Lemma logn_div p m n : m %| n -> logn p (n %/ m) = logn p n - logn p m. +Proof. +rewrite dvdn_eq => /eqP def_n. +case: (posnP n) => [-> |]; first by rewrite div0n logn0. +by rewrite -{1 3}def_n muln_gt0 => /andP[q_gt0 m_gt0]; rewrite lognM ?addnK. +Qed. + +Lemma dvdn_pfactor p d n : prime p -> + reflect (exists2 m, m <= n & d = p ^ m) (d %| p ^ n). +Proof. +move=> p_pr; have pn_gt0: p ^ n > 0 by rewrite expn_gt0 prime_gt0. +apply: (iffP idP) => [dv_d_pn|[m le_m_n ->]]; last first. + by rewrite -(subnK le_m_n) expnD dvdn_mull. +exists (logn p d); first by rewrite -(pfactorK n p_pr) dvdn_leq_log. +have d_gt0: d > 0 by exact: dvdn_gt0 dv_d_pn. +case: (pfactor_coprime p_pr d_gt0) => q co_p_q def_d. +rewrite {1}def_d ((q =P 1) _) ?mul1n // -dvdn1. +suff: q %| p ^ n * 1 by rewrite Gauss_dvdr // coprime_sym coprime_expl. +by rewrite muln1 (dvdn_trans _ dv_d_pn) // def_d dvdn_mulr. +Qed. + +Lemma prime_decompE n : prime_decomp n = [seq (p, logn p n) | p <- primes n]. +Proof. +case: n => // n; pose f0 := (0, 0); rewrite -map_comp. +apply: (@eq_from_nth _ f0) => [|i lt_i_n]; first by rewrite size_map. +rewrite (nth_map f0) //; case def_f: (nth _ _ i) => [p e] /=. +congr (_, _); rewrite [n.+1]prod_prime_decomp //. +have: (p, e) \in prime_decomp n.+1 by rewrite -def_f mem_nth. +case/mem_prime_decomp=> pr_p _ _. +rewrite (big_nth f0) big_mkord (bigD1 (Ordinal lt_i_n)) //=. +rewrite def_f mulnC logn_Gauss ?pfactorK //. +apply big_ind => [|m1 m2 com1 com2| [j ltj] /=]; first exact: coprimen1. + by rewrite coprime_mulr com1. +rewrite -val_eqE /= => nji; case def_j: (nth _ _ j) => [q e1] /=. +have: (q, e1) \in prime_decomp n.+1 by rewrite -def_j mem_nth. +case/mem_prime_decomp=> pr_q e1_gt0 _; rewrite coprime_pexpr //. +rewrite prime_coprime // dvdn_prime2 //; apply: contra nji => eq_pq. +rewrite -(nth_uniq 0 _ _ (primes_uniq n.+1)) ?size_map //=. +by rewrite !(nth_map f0) // def_f def_j /= eq_sym. +Qed. + +(* Some combinatorial formulae. *) + +Lemma divn_count_dvd d n : n %/ d = \sum_(1 <= i < n.+1) (d %| i). +Proof. +have [-> | d_gt0] := posnP d; first by rewrite big_add1 divn0 big1. +apply: (@addnI (d %| 0)); rewrite -(@big_ltn _ 0 _ 0 _ (dvdn d)) // big_mkord. +rewrite (partition_big (fun i : 'I_n.+1 => inord (i %/ d)) 'I_(n %/ d).+1) //=. +rewrite dvdn0 add1n -{1}[_.+1]card_ord -sum1_card; apply: eq_bigr => [[q ?] _]. +rewrite (bigD1 (inord (q * d))) /eq_op /= !inordK ?ltnS -?leq_divRL ?mulnK //. +rewrite dvdn_mull ?big1 // => [[i /= ?] /andP[/eqP <- /negPf]]. +by rewrite eq_sym dvdn_eq inordK ?ltnS ?leq_div2r // => ->. +Qed. + +Lemma logn_count_dvd p n : prime p -> logn p n = \sum_(1 <= k < n) (p ^ k %| n). +Proof. +rewrite big_add1 => p_prime; case: n => [|n]; first by rewrite logn0 big_geq. +rewrite big_mkord -big_mkcond (eq_bigl _ _ (fun _ => pfactor_dvdn _ _ _)) //=. +by rewrite big_ord_narrow ?sum1_card ?card_ord // -ltnS ltn_logl. +Qed. + +(* Truncated real log. *) + +Definition trunc_log p n := + let fix loop n k := + if k is k'.+1 then if p <= n then (loop (n %/ p) k').+1 else 0 else 0 + in loop n n. + +Lemma trunc_log_bounds p n : + 1 < p -> 0 < n -> let k := trunc_log p n in p ^ k <= n < p ^ k.+1. +Proof. +rewrite {+}/trunc_log => p_gt1; have p_gt0 := ltnW p_gt1. +elim: n {-2 5}n (leqnn n) => [|m IHm] [|n] //=; rewrite ltnS => le_n_m _. +have [le_p_n | // ] := leqP p _; rewrite 2!expnSr -leq_divRL -?ltn_divLR //. +by apply: IHm; rewrite ?divn_gt0 // -ltnS (leq_trans (ltn_Pdiv _ _)). +Qed. + +Lemma trunc_log_ltn p n : 1 < p -> n < p ^ (trunc_log p n).+1. +Proof. +have [-> | n_gt0] := posnP n; first by move=> /ltnW; rewrite expn_gt0. +by case/trunc_log_bounds/(_ n_gt0)/andP. +Qed. + +Lemma trunc_logP p n : 1 < p -> 0 < n -> p ^ trunc_log p n <= n. +Proof. by move=> p_gt1 /(trunc_log_bounds p_gt1)/andP[]. Qed. + +Lemma trunc_log_max p k j : 1 < p -> p ^ j <= k -> j <= trunc_log p k. +Proof. +move=> p_gt1 le_pj_k; rewrite -ltnS -(@ltn_exp2l p) //. +exact: leq_ltn_trans (trunc_log_ltn _ _). +Qed. + +(* pi- parts *) + +(* Testing for membership in set of prime factors. *) + +Canonical nat_pred_pred := Eval hnf in [predType of nat_pred]. + +Coercion nat_pred_of_nat (p : nat) : nat_pred := pred1 p. + +Section NatPreds. + +Variables (n : nat) (pi : nat_pred). + +Definition negn : nat_pred := [predC pi]. + +Definition pnat : pred nat := fun m => (m > 0) && all (mem pi) (primes m). + +Definition partn := \prod_(0 <= p < n.+1 | p \in pi) p ^ logn p n. + +End NatPreds. + +Notation "pi ^'" := (negn pi) (at level 2, format "pi ^'") : nat_scope. + +Notation "pi .-nat" := (pnat pi) (at level 2, format "pi .-nat") : nat_scope. + +Notation "n `_ pi" := (partn n pi) : nat_scope. + +Section PnatTheory. + +Implicit Types (n p : nat) (pi rho : nat_pred). + +Lemma negnK pi : pi^'^' =i pi. +Proof. move=> p; exact: negbK. Qed. + +Lemma eq_negn pi1 pi2 : pi1 =i pi2 -> pi1^' =i pi2^'. +Proof. by move=> eq_pi n; rewrite 3!inE /= eq_pi. Qed. + +Lemma eq_piP m n : \pi(m) =i \pi(n) <-> \pi(m) = \pi(n). +Proof. +rewrite /pi_of; have eqs := eq_sorted_irr ltn_trans ltnn. +by split=> [|-> //]; move/(eqs _ _ (sorted_primes m) (sorted_primes n)) ->. +Qed. + +Lemma part_gt0 pi n : 0 < n`_pi. +Proof. exact: prodn_gt0. Qed. +Hint Resolve part_gt0. + +Lemma sub_in_partn pi1 pi2 n : + {in \pi(n), {subset pi1 <= pi2}} -> n`_pi1 %| n`_pi2. +Proof. +move=> pi12; rewrite ![n`__]big_mkcond /=. +apply (big_ind2 (fun m1 m2 => m1 %| m2)) => // [*|p _]; first exact: dvdn_mul. +rewrite lognE -mem_primes; case: ifP => pi1p; last exact: dvd1n. +by case: ifP => pr_p; [rewrite pi12 | rewrite if_same]. +Qed. + +Lemma eq_in_partn pi1 pi2 n : {in \pi(n), pi1 =i pi2} -> n`_pi1 = n`_pi2. +Proof. +by move=> pi12; apply/eqP; rewrite eqn_dvd ?sub_in_partn // => p /pi12->. +Qed. + +Lemma eq_partn pi1 pi2 n : pi1 =i pi2 -> n`_pi1 = n`_pi2. +Proof. by move=> pi12; apply: eq_in_partn => p _. Qed. + +Lemma partnNK pi n : n`_pi^'^' = n`_pi. +Proof. by apply: eq_partn; exact: negnK. Qed. + +Lemma widen_partn m pi n : + n <= m -> n`_pi = \prod_(0 <= p < m.+1 | p \in pi) p ^ logn p n. +Proof. +move=> le_n_m; rewrite big_mkcond /=. +rewrite [n`_pi](big_nat_widen _ _ m.+1) // big_mkcond /=. +apply: eq_bigr => p _; rewrite ltnS lognE. +by case: and3P => [[_ n_gt0 p_dv_n]|]; rewrite ?if_same // andbC dvdn_leq. +Qed. + +Lemma partn0 pi : 0`_pi = 1. +Proof. by apply: big1_seq => [] [|n]; rewrite andbC. Qed. + +Lemma partn1 pi : 1`_pi = 1. +Proof. by apply: big1_seq => [] [|[|n]]; rewrite andbC. Qed. + +Lemma partnM pi m n : m > 0 -> n > 0 -> (m * n)`_pi = m`_pi * n`_pi. +Proof. +have le_pmul m' n': m' > 0 -> n' <= m' * n' by move/prednK <-; exact: leq_addr. +move=> mpos npos; rewrite !(@widen_partn (n * m)) 3?(le_pmul, mulnC) //. +rewrite !big_mkord -big_split; apply: eq_bigr => p _ /=. +by rewrite lognM // expnD. +Qed. + +Lemma partnX pi m n : (m ^ n)`_pi = m`_pi ^ n. +Proof. +elim: n => [|n IHn]; first exact: partn1. +rewrite expnS; case: (posnP m) => [->|m_gt0]; first by rewrite partn0 exp1n. +by rewrite expnS partnM ?IHn // expn_gt0 m_gt0. +Qed. + +Lemma partn_dvd pi m n : n > 0 -> m %| n -> m`_pi %| n`_pi. +Proof. +move=> n_gt0 dvmn; case/dvdnP: dvmn n_gt0 => q ->{n}. +by rewrite muln_gt0 => /andP[q_gt0 m_gt0]; rewrite partnM ?dvdn_mull. +Qed. + +Lemma p_part p n : n`_p = p ^ logn p n. +Proof. +case (posnP (logn p n)) => [log0 |]. + by rewrite log0 [n`_p]big1_seq // => q; case/andP; move/eqnP->; rewrite log0. +rewrite logn_gt0 mem_primes; case/and3P=> _ n_gt0 dv_p_n. +have le_p_n: p < n.+1 by rewrite ltnS dvdn_leq. +by rewrite [n`_p]big_mkord (big_pred1 (Ordinal le_p_n)). +Qed. + +Lemma p_part_eq1 p n : (n`_p == 1) = (p \notin \pi(n)). +Proof. +rewrite mem_primes p_part lognE; case: and3P => // [[p_pr _ _]]. +by rewrite -dvdn1 pfactor_dvdn // logn1. +Qed. + +Lemma p_part_gt1 p n : (n`_p > 1) = (p \in \pi(n)). +Proof. by rewrite ltn_neqAle part_gt0 andbT eq_sym p_part_eq1 negbK. Qed. + +Lemma primes_part pi n : primes n`_pi = filter (mem pi) (primes n). +Proof. +have ltnT := ltn_trans. +case: (posnP n) => [-> | n_gt0]; first by rewrite partn0. +apply: (eq_sorted_irr ltnT ltnn); rewrite ?(sorted_primes, sorted_filter) //. +move=> p; rewrite mem_filter /= !mem_primes n_gt0 part_gt0 /=. +apply/andP/and3P=> [[p_pr] | [pi_p p_pr dv_p_n]]. + rewrite /partn; apply big_ind => [|n1 n2 IHn1 IHn2|q pi_q]. + - by rewrite dvdn1; case: eqP p_pr => // ->. + - by rewrite Euclid_dvdM //; case/orP. + rewrite -{1}(expn1 p) pfactor_dvdn // lognX muln_gt0. + rewrite logn_gt0 mem_primes n_gt0 - andbA /=; case/and3P=> pr_q dv_q_n. + by rewrite logn_prime //; case: eqP => // ->. +have le_p_n: p < n.+1 by rewrite ltnS dvdn_leq. +rewrite [n`_pi]big_mkord (bigD1 (Ordinal le_p_n)) //= dvdn_mulr //. +by rewrite lognE p_pr n_gt0 dv_p_n expnS dvdn_mulr. +Qed. + +Lemma filter_pi_of n m : n < m -> filter \pi(n) (index_iota 0 m) = primes n. +Proof. +move=> lt_n_m; have ltnT := ltn_trans; apply: (eq_sorted_irr ltnT ltnn). +- by rewrite sorted_filter // iota_ltn_sorted. +- exact: sorted_primes. +move=> p; rewrite mem_filter mem_index_iota /= mem_primes; case: and3P => //. +case=> _ n_gt0 dv_p_n; apply: leq_ltn_trans lt_n_m; exact: dvdn_leq. +Qed. + +Lemma partn_pi n : n > 0 -> n`_\pi(n) = n. +Proof. +move=> n_gt0; rewrite {3}(prod_prime_decomp n_gt0) prime_decompE big_map. +by rewrite -[n`__]big_filter filter_pi_of. +Qed. + +Lemma partnT n : n > 0 -> n`_predT = n. +Proof. +move=> n_gt0; rewrite -{2}(partn_pi n_gt0) {2}/partn big_mkcond /=. +by apply: eq_bigr => p _; rewrite -logn_gt0; case: (logn p _). +Qed. + +Lemma partnC pi n : n > 0 -> n`_pi * n`_pi^' = n. +Proof. +move=> n_gt0; rewrite -{3}(partnT n_gt0) /partn. +do 2!rewrite mulnC big_mkcond /=; rewrite -big_split; apply: eq_bigr => p _ /=. +by rewrite mulnC inE /=; case: (p \in pi); rewrite /= (muln1, mul1n). +Qed. + +Lemma dvdn_part pi n : n`_pi %| n. +Proof. by case: n => // n; rewrite -{2}[n.+1](@partnC pi) // dvdn_mulr. Qed. + +Lemma logn_part p m : logn p m`_p = logn p m. +Proof. +case p_pr: (prime p); first by rewrite p_part pfactorK. +by rewrite lognE (lognE p m) p_pr. +Qed. + +Lemma partn_lcm pi m n : m > 0 -> n > 0 -> (lcmn m n)`_pi = lcmn m`_pi n`_pi. +Proof. +move=> m_gt0 n_gt0; have p_gt0: lcmn m n > 0 by rewrite lcmn_gt0 m_gt0. +apply/eqP; rewrite eqn_dvd dvdn_lcm !partn_dvd ?dvdn_lcml ?dvdn_lcmr //. +rewrite -(dvdn_pmul2r (part_gt0 pi^' (lcmn m n))) partnC // dvdn_lcm !andbT. +rewrite -{1}(partnC pi m_gt0) andbC -{1}(partnC pi n_gt0). +by rewrite !dvdn_mul ?partn_dvd ?dvdn_lcml ?dvdn_lcmr. +Qed. + +Lemma partn_gcd pi m n : m > 0 -> n > 0 -> (gcdn m n)`_pi = gcdn m`_pi n`_pi. +Proof. +move=> m_gt0 n_gt0; have p_gt0: gcdn m n > 0 by rewrite gcdn_gt0 m_gt0. +apply/eqP; rewrite eqn_dvd dvdn_gcd !partn_dvd ?dvdn_gcdl ?dvdn_gcdr //=. +rewrite -(dvdn_pmul2r (part_gt0 pi^' (gcdn m n))) partnC // dvdn_gcd. +rewrite -{3}(partnC pi m_gt0) andbC -{3}(partnC pi n_gt0). +by rewrite !dvdn_mul ?partn_dvd ?dvdn_gcdl ?dvdn_gcdr. +Qed. + +Lemma partn_biglcm (I : finType) (P : pred I) F pi : + (forall i, P i -> F i > 0) -> + (\big[lcmn/1%N]_(i | P i) F i)`_pi = \big[lcmn/1%N]_(i | P i) (F i)`_pi. +Proof. +move=> F_gt0; set m := \big[lcmn/1%N]_(i | P i) F i. +have m_gt0: 0 < m by elim/big_ind: m => // p q p_gt0; rewrite lcmn_gt0 p_gt0. +apply/eqP; rewrite eqn_dvd andbC; apply/andP; split. + by apply/dvdn_biglcmP=> i Pi; rewrite partn_dvd // (@biglcmn_sup _ i). +rewrite -(dvdn_pmul2r (part_gt0 pi^' m)) partnC //. +apply/dvdn_biglcmP=> i Pi; rewrite -(partnC pi (F_gt0 i Pi)) dvdn_mul //. + by rewrite (@biglcmn_sup _ i). +by rewrite partn_dvd // (@biglcmn_sup _ i). +Qed. + +Lemma partn_biggcd (I : finType) (P : pred I) F pi : + #|SimplPred P| > 0 -> (forall i, P i -> F i > 0) -> + (\big[gcdn/0]_(i | P i) F i)`_pi = \big[gcdn/0]_(i | P i) (F i)`_pi. +Proof. +move=> ntP F_gt0; set d := \big[gcdn/0]_(i | P i) F i. +have d_gt0: 0 < d. + case/card_gt0P: ntP => i /= Pi; have:= F_gt0 i Pi. + rewrite !lt0n -!dvd0n; apply: contra => dv0d. + by rewrite (dvdn_trans dv0d) // (@biggcdn_inf _ i). +apply/eqP; rewrite eqn_dvd; apply/andP; split. + by apply/dvdn_biggcdP=> i Pi; rewrite partn_dvd ?F_gt0 // (@biggcdn_inf _ i). +rewrite -(dvdn_pmul2r (part_gt0 pi^' d)) partnC //. +apply/dvdn_biggcdP=> i Pi; rewrite -(partnC pi (F_gt0 i Pi)) dvdn_mul //. + by rewrite (@biggcdn_inf _ i). +by rewrite partn_dvd ?F_gt0 // (@biggcdn_inf _ i). +Qed. + +Lemma sub_in_pnat pi rho n : + {in \pi(n), {subset pi <= rho}} -> pi.-nat n -> rho.-nat n. +Proof. +rewrite /pnat => subpi /andP[-> pi_n]. +apply/allP=> p pr_p; apply: subpi => //; exact: (allP pi_n). +Qed. + +Lemma eq_in_pnat pi rho n : {in \pi(n), pi =i rho} -> pi.-nat n = rho.-nat n. +Proof. by move=> eqpi; apply/idP/idP; apply: sub_in_pnat => p /eqpi->. Qed. + +Lemma eq_pnat pi rho n : pi =i rho -> pi.-nat n = rho.-nat n. +Proof. by move=> eqpi; apply: eq_in_pnat => p _. Qed. + +Lemma pnatNK pi n : pi^'^'.-nat n = pi.-nat n. +Proof. exact: eq_pnat (negnK pi). Qed. + +Lemma pnatI pi rho n : [predI pi & rho].-nat n = pi.-nat n && rho.-nat n. +Proof. by rewrite /pnat andbCA all_predI !andbA andbb. Qed. + +Lemma pnat_mul pi m n : pi.-nat (m * n) = pi.-nat m && pi.-nat n. +Proof. +rewrite /pnat muln_gt0 andbCA -andbA andbCA. +case: posnP => // n_gt0; case: posnP => //= m_gt0. +apply/allP/andP=> [pi_mn | [pi_m pi_n] p]. + by split; apply/allP=> p m_p; apply: pi_mn; rewrite primes_mul // m_p ?orbT. +rewrite primes_mul // => /orP[]; [exact: (allP pi_m) | exact: (allP pi_n)]. +Qed. + +Lemma pnat_exp pi m n : pi.-nat (m ^ n) = pi.-nat m || (n == 0). +Proof. by case: n => [|n]; rewrite orbC // /pnat expn_gt0 orbC primes_exp. Qed. + +Lemma part_pnat pi n : pi.-nat n`_pi. +Proof. +rewrite /pnat primes_part part_gt0. +by apply/allP=> p; rewrite mem_filter => /andP[]. +Qed. + +Lemma pnatE pi p : prime p -> pi.-nat p = (p \in pi). +Proof. by move=> pr_p; rewrite /pnat prime_gt0 ?primes_prime //= andbT. Qed. + +Lemma pnat_id p : prime p -> p.-nat p. +Proof. by move=> pr_p; rewrite pnatE ?inE /=. Qed. + +Lemma coprime_pi' m n : m > 0 -> n > 0 -> coprime m n = \pi(m)^'.-nat n. +Proof. +by move=> m_gt0 n_gt0; rewrite /pnat n_gt0 all_predC coprime_has_primes. +Qed. + +Lemma pnat_pi n : n > 0 -> \pi(n).-nat n. +Proof. rewrite /pnat => ->; exact/allP. Qed. + +Lemma pi_of_dvd m n : m %| n -> n > 0 -> {subset \pi(m) <= \pi(n)}. +Proof. +move=> m_dv_n n_gt0 p; rewrite !mem_primes n_gt0 => /and3P[-> _ p_dv_m]. +exact: dvdn_trans p_dv_m m_dv_n. +Qed. + +Lemma pi_ofM m n : m > 0 -> n > 0 -> \pi(m * n) =i [predU \pi(m) & \pi(n)]. +Proof. move=> m_gt0 n_gt0 p; exact: primes_mul. Qed. + +Lemma pi_of_part pi n : n > 0 -> \pi(n`_pi) =i [predI \pi(n) & pi]. +Proof. by move=> n_gt0 p; rewrite /pi_of primes_part mem_filter andbC. Qed. + +Lemma pi_of_exp p n : n > 0 -> \pi(p ^ n) = \pi(p). +Proof. by move=> n_gt0; rewrite /pi_of primes_exp. Qed. + +Lemma pi_of_prime p : prime p -> \pi(p) =i (p : nat_pred). +Proof. by move=> pr_p q; rewrite /pi_of primes_prime // mem_seq1. Qed. + +Lemma p'natEpi p n : n > 0 -> p^'.-nat n = (p \notin \pi(n)). +Proof. by case: n => // n _; rewrite /pnat all_predC has_pred1. Qed. + +Lemma p'natE p n : prime p -> p^'.-nat n = ~~ (p %| n). +Proof. +case: n => [|n] p_pr; first by case: p p_pr. +by rewrite p'natEpi // mem_primes p_pr. +Qed. + +Lemma pnatPpi pi n p : pi.-nat n -> p \in \pi(n) -> p \in pi. +Proof. by case/andP=> _ /allP; exact. Qed. + +Lemma pnat_dvd m n pi : m %| n -> pi.-nat n -> pi.-nat m. +Proof. by case/dvdnP=> q ->; rewrite pnat_mul; case/andP. Qed. + +Lemma pnat_div m n pi : m %| n -> pi.-nat n -> pi.-nat (n %/ m). +Proof. +case/dvdnP=> q ->; rewrite pnat_mul andbC => /andP[]. +by case: m => // m _; rewrite mulnK. +Qed. + +Lemma pnat_coprime pi m n : pi.-nat m -> pi^'.-nat n -> coprime m n. +Proof. +case/andP=> m_gt0 pi_m /andP[n_gt0 pi'_n]. +rewrite coprime_has_primes //; apply/hasPn=> p /(allP pi'_n). +apply: contra; exact: allP. +Qed. + +Lemma p'nat_coprime pi m n : pi^'.-nat m -> pi.-nat n -> coprime m n. +Proof. by move=> pi'm pi_n; rewrite (pnat_coprime pi'm) ?pnatNK. Qed. + +Lemma sub_pnat_coprime pi rho m n : + {subset rho <= pi^'} -> pi.-nat m -> rho.-nat n -> coprime m n. +Proof. +by move=> pi'rho pi_m; move/(sub_in_pnat (in1W pi'rho)); exact: pnat_coprime. +Qed. + +Lemma coprime_partC pi m n : coprime m`_pi n`_pi^'. +Proof. by apply: (@pnat_coprime pi); exact: part_pnat. Qed. + +Lemma pnat_1 pi n : pi.-nat n -> pi^'.-nat n -> n = 1. +Proof. +by move=> pi_n pi'_n; rewrite -(eqnP (pnat_coprime pi_n pi'_n)) gcdnn. +Qed. + +Lemma part_pnat_id pi n : pi.-nat n -> n`_pi = n. +Proof. +case/andP=> n_gt0 pi_n. +rewrite -{2}(partnT n_gt0) /partn big_mkcond; apply: eq_bigr=> p _. +case: (posnP (logn p n)) => [-> |]; first by rewrite if_same. +by rewrite logn_gt0 => /(allP pi_n)/= ->. +Qed. + +Lemma part_p'nat pi n : pi^'.-nat n -> n`_pi = 1. +Proof. +case/andP=> n_gt0 pi'_n; apply: big1_seq => p /andP[pi_p _]. +case: (posnP (logn p n)) => [-> //|]. +by rewrite logn_gt0; move/(allP pi'_n); case/negP. +Qed. + +Lemma partn_eq1 pi n : n > 0 -> (n`_pi == 1) = pi^'.-nat n. +Proof. +move=> n_gt0; apply/eqP/idP=> [pi_n_1|]; last exact: part_p'nat. +by rewrite -(partnC pi n_gt0) pi_n_1 mul1n part_pnat. +Qed. + +Lemma pnatP pi n : + n > 0 -> reflect (forall p, prime p -> p %| n -> p \in pi) (pi.-nat n). +Proof. +move=> n_gt0; rewrite /pnat n_gt0. +apply: (iffP allP) => /= pi_n p => [pr_p p_n|]. + by rewrite pi_n // mem_primes pr_p n_gt0. +by rewrite mem_primes n_gt0 /=; case/andP; move: p. +Qed. + +Lemma pi_pnat pi p n : p.-nat n -> p \in pi -> pi.-nat n. +Proof. +move=> p_n pi_p; have [n_gt0 _] := andP p_n. +by apply/pnatP=> // q q_pr /(pnatP _ n_gt0 p_n _ q_pr)/eqnP->. +Qed. + +Lemma p_natP p n : p.-nat n -> {k | n = p ^ k}. +Proof. by move=> p_n; exists (logn p n); rewrite -p_part part_pnat_id. Qed. + +Lemma pi'_p'nat pi p n : pi^'.-nat n -> p \in pi -> p^'.-nat n. +Proof. +move=> pi'n pi_p; apply: sub_in_pnat pi'n => q _. +by apply: contraNneq => ->. +Qed. + +Lemma pi_p'nat p pi n : pi.-nat n -> p \in pi^' -> p^'.-nat n. +Proof. by move=> pi_n; apply: pi'_p'nat; rewrite pnatNK. Qed. + +Lemma partn_part pi rho n : {subset pi <= rho} -> n`_rho`_pi = n`_pi. +Proof. +move=> pi_sub_rho; have [->|n_gt0] := posnP n; first by rewrite !partn0 partn1. +rewrite -{2}(partnC rho n_gt0) partnM //. +suffices: pi^'.-nat n`_rho^' by move/part_p'nat->; rewrite muln1. +apply: sub_in_pnat (part_pnat _ _) => q _; apply: contra; exact: pi_sub_rho. +Qed. + +Lemma partnI pi rho n : n`_[predI pi & rho] = n`_pi`_rho. +Proof. +rewrite -(@partnC [predI pi & rho] _`_rho) //. +symmetry; rewrite 2?partn_part; try by move=> p /andP []. +rewrite mulnC part_p'nat ?mul1n // pnatNK pnatI part_pnat andbT. +exact: pnat_dvd (dvdn_part _ _) (part_pnat _ _). +Qed. + +Lemma odd_2'nat n : odd n = 2^'.-nat n. +Proof. by case: n => // n; rewrite p'natE // dvdn2 negbK. Qed. + +End PnatTheory. +Hint Resolve part_gt0. + +(************************************) +(* Properties of the divisors list. *) +(************************************) + +Lemma divisors_correct n : n > 0 -> + [/\ uniq (divisors n), sorted leq (divisors n) + & forall d, (d \in divisors n) = (d %| n)]. +Proof. +move/prod_prime_decomp=> def_n; rewrite {4}def_n {def_n}. +have: all prime (primes n) by apply/allP=> p; rewrite mem_primes; case/andP. +have:= primes_uniq n; rewrite /primes /divisors; move/prime_decomp: n. +elim=> [|[p e] pd] /=; first by split=> // d; rewrite big_nil dvdn1 mem_seq1. +rewrite big_cons /=; move: (foldr _ _ pd) => divs. +move=> IHpd /andP[npd_p Upd] /andP[pr_p pr_pd]. +have lt0p: 0 < p by exact: prime_gt0. +have {IHpd Upd}[Udivs Odivs mem_divs] := IHpd Upd pr_pd. +have ndivs_p m: p * m \notin divs. + suffices: p \notin divs; rewrite !mem_divs. + by apply: contra => /dvdnP[n ->]; rewrite mulnCA dvdn_mulr. + have ndv_p_1: ~~(p %| 1) by rewrite dvdn1 neq_ltn orbC prime_gt1. + rewrite big_seq; elim/big_ind: _ => [//|u v npu npv|[q f] /= pd_qf]. + by rewrite Euclid_dvdM //; apply/norP. + elim: (f) => // f'; rewrite expnS Euclid_dvdM // orbC negb_or => -> {f'}/=. + have pd_q: q \in unzip1 pd by apply/mapP; exists (q, f). + by apply: contra npd_p; rewrite dvdn_prime2 // ?(allP pr_pd) // => /eqP->. +elim: e => [|e] /=; first by split=> // d; rewrite mul1n. +have Tmulp_inj: injective (NatTrec.mul p). + by move=> u v /eqP; rewrite !natTrecE eqn_pmul2l // => /eqP. +move: (iter e _ _) => divs' [Udivs' Odivs' mem_divs']; split=> [||d]. +- rewrite merge_uniq cat_uniq map_inj_uniq // Udivs Udivs' andbT /=. + apply/hasP=> [[d dv_d /mapP[d' _ def_d]]]. + by case/idPn: dv_d; rewrite def_d natTrecE. +- rewrite (merge_sorted leq_total) //; case: (divs') Odivs' => //= d ds. + rewrite (@map_path _ _ _ _ leq xpred0) ?has_pred0 // => u v _. + by rewrite !natTrecE leq_pmul2l. +rewrite mem_merge mem_cat; case dv_d_p: (p %| d). + case/dvdnP: dv_d_p => d' ->{d}; rewrite mulnC (negbTE (ndivs_p d')) orbF. + rewrite expnS -mulnA dvdn_pmul2l // -mem_divs'. + by rewrite -(mem_map Tmulp_inj divs') natTrecE. +case pdiv_d: (_ \in _). + by case/mapP: pdiv_d dv_d_p => d' _ ->; rewrite natTrecE dvdn_mulr. +rewrite mem_divs Gauss_dvdr // coprime_sym. +by rewrite coprime_expl ?prime_coprime ?dv_d_p. +Qed. + +Lemma sorted_divisors n : sorted leq (divisors n). +Proof. by case: (posnP n) => [-> | /divisors_correct[]]. Qed. + +Lemma divisors_uniq n : uniq (divisors n). +Proof. by case: (posnP n) => [-> | /divisors_correct[]]. Qed. + +Lemma sorted_divisors_ltn n : sorted ltn (divisors n). +Proof. by rewrite ltn_sorted_uniq_leq divisors_uniq sorted_divisors. Qed. + +Lemma dvdn_divisors d m : 0 < m -> (d %| m) = (d \in divisors m). +Proof. by case/divisors_correct. Qed. + +Lemma divisor1 n : 1 \in divisors n. +Proof. by case: n => // n; rewrite -dvdn_divisors // dvd1n. Qed. + +Lemma divisors_id n : 0 < n -> n \in divisors n. +Proof. by move/dvdn_divisors <-. Qed. + +(* Big sum / product lemmas*) + +Lemma dvdn_sum d I r (K : pred I) F : + (forall i, K i -> d %| F i) -> d %| \sum_(i <- r | K i) F i. +Proof. move=> dF; elim/big_ind: _ => //; exact: dvdn_add. Qed. + +Lemma dvdn_partP n m : 0 < n -> + reflect (forall p, p \in \pi(n) -> n`_p %| m) (n %| m). +Proof. +move=> n_gt0; apply: (iffP idP) => n_dvd_m => [p _|]. + apply: dvdn_trans n_dvd_m; exact: dvdn_part. +have [-> // | m_gt0] := posnP m. +rewrite -(partnT n_gt0) -(partnT m_gt0). +rewrite !(@widen_partn (m + n)) ?leq_addl ?leq_addr // /in_mem /=. +elim/big_ind2: _ => // [* | q _]; first exact: dvdn_mul. +have [-> // | ] := posnP (logn q n); rewrite logn_gt0 => q_n. +have pr_q: prime q by move: q_n; rewrite mem_primes; case/andP. +by have:= n_dvd_m q q_n; rewrite p_part !pfactor_dvdn // pfactorK. +Qed. + +Lemma modn_partP n a b : 0 < n -> + reflect (forall p : nat, p \in \pi(n) -> a = b %[mod n`_p]) (a == b %[mod n]). +Proof. +move=> n_gt0; wlog le_b_a: a b / b <= a. + move=> IH; case: (leqP b a) => [|/ltnW] /IH {IH}// IH. + by rewrite eq_sym; apply: (iffP IH) => eqab p; move/eqab. +rewrite eqn_mod_dvd //; apply: (iffP (dvdn_partP _ n_gt0)) => eqab p /eqab; + by rewrite -eqn_mod_dvd // => /eqP. +Qed. + +(* The Euler totient function *) + +Lemma totientE n : + n > 0 -> totient n = \prod_(p <- primes n) (p.-1 * p ^ (logn p n).-1). +Proof. +move=> n_gt0; rewrite /totient n_gt0 prime_decompE unlock. +by elim: (primes n) => //= [p pr ->]; rewrite !natTrecE. +Qed. + +Lemma totient_gt0 n : (0 < totient n) = (0 < n). +Proof. +case: n => // n; rewrite totientE // big_seq_cond prodn_cond_gt0 // => p. +by rewrite mem_primes muln_gt0 expn_gt0; case: p => [|[|]]. +Qed. + +Lemma totient_pfactor p e : + prime p -> e > 0 -> totient (p ^ e) = p.-1 * p ^ e.-1. +Proof. +move=> p_pr e_gt0; rewrite totientE ?expn_gt0 ?prime_gt0 //. +by rewrite primes_exp // primes_prime // unlock /= muln1 pfactorK. +Qed. + +Lemma totient_coprime m n : + coprime m n -> totient (m * n) = totient m * totient n. +Proof. +move=> co_mn; have [-> //| m_gt0] := posnP m. +have [->|n_gt0] := posnP n; first by rewrite !muln0. +rewrite !totientE ?muln_gt0 ?m_gt0 //. +have /(eq_big_perm _)->: perm_eq (primes (m * n)) (primes m ++ primes n). + apply: uniq_perm_eq => [||p]; first exact: primes_uniq. + by rewrite cat_uniq !primes_uniq -coprime_has_primes // co_mn. + by rewrite mem_cat primes_mul. +rewrite big_cat /= !big_seq. +congr (_ * _); apply: eq_bigr => p; rewrite mem_primes => /and3P[_ _ dvp]. + rewrite (mulnC m) logn_Gauss //; move: co_mn. + by rewrite -(divnK dvp) coprime_mull => /andP[]. +rewrite logn_Gauss //; move: co_mn. +by rewrite coprime_sym -(divnK dvp) coprime_mull => /andP[]. +Qed. + +Lemma totient_count_coprime n : totient n = \sum_(0 <= d < n) coprime n d. +Proof. +elim: {n}_.+1 {-2}n (ltnSn n) => // m IHm n; rewrite ltnS => le_n_m. +case: (leqP n 1) => [|lt1n]; first by rewrite unlock; case: (n) => [|[]]. +pose p := pdiv n; have p_pr: prime p by exact: pdiv_prime. +have p1 := prime_gt1 p_pr; have p0 := ltnW p1. +pose np := n`_p; pose np' := n`_p^'. +have co_npp': coprime np np' by rewrite coprime_partC. +have [n0 np0 np'0]: [/\ n > 0, np > 0 & np' > 0] by rewrite ltnW ?part_gt0. +have def_n: n = np * np' by rewrite partnC. +have lnp0: 0 < logn p n by rewrite lognE p_pr n0 pdiv_dvd. +pose in_mod k (k0 : k > 0) d := Ordinal (ltn_pmod d k0). +rewrite {1}def_n totient_coprime // {IHm}(IHm np') ?big_mkord; last first. + apply: leq_trans le_n_m; rewrite def_n ltn_Pmull //. + by rewrite /np p_part -(expn0 p) ltn_exp2l. +have ->: totient np = #|[pred d : 'I_np | coprime np d]|. + rewrite {1}[np]p_part totient_pfactor //=; set q := p ^ _. + apply: (@addnI (1 * q)); rewrite -mulnDl [1 + _]prednK // mul1n. + have def_np: np = p * q by rewrite -expnS prednK // -p_part. + pose mulp := [fun d : 'I_q => in_mod _ np0 (p * d)]. + rewrite -def_np -{1}[np]card_ord -(cardC (mem (codom mulp))). + rewrite card_in_image => [|[d1 ltd1] [d2 ltd2] /= _ _ []]; last first. + move/eqP; rewrite def_np -!muln_modr ?modn_small //. + by rewrite eqn_pmul2l // => eq_op12; exact/eqP. + rewrite card_ord; congr (q + _); apply: eq_card => d /=. + rewrite !inE [np in coprime np _]p_part coprime_pexpl ?prime_coprime //. + congr (~~ _); apply/codomP/idP=> [[d' -> /=] | /dvdnP[r def_d]]. + by rewrite def_np -muln_modr // dvdn_mulr. + do [rewrite mulnC; case: d => d ltd /=] in def_d *. + have ltr: r < q by rewrite -(ltn_pmul2l p0) -def_np -def_d. + by exists (Ordinal ltr); apply: val_inj; rewrite /= -def_d modn_small. +pose h (d : 'I_n) := (in_mod _ np0 d, in_mod _ np'0 d). +pose h' (d : 'I_np * 'I_np') := in_mod _ n0 (chinese np np' d.1 d.2). +rewrite -!big_mkcond -sum_nat_const pair_big (reindex_onto h h') => [|[d d'] _]. + apply: eq_bigl => [[d ltd] /=]; rewrite !inE /= -val_eqE /= andbC. + rewrite !coprime_modr def_n -chinese_mod // -coprime_mull -def_n. + by rewrite modn_small ?eqxx. +apply/eqP; rewrite /eq_op /= /eq_op /= !modn_dvdm ?dvdn_part //. +by rewrite chinese_modl // chinese_modr // !modn_small ?eqxx ?ltn_ord. +Qed. + + + diff --git a/mathcomp/discrete/tuple.v b/mathcomp/discrete/tuple.v new file mode 100644 index 0000000..ba64bd0 --- /dev/null +++ b/mathcomp/discrete/tuple.v @@ -0,0 +1,412 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +(******************************************************************************) +(* Tuples, i.e., sequences with a fixed (known) length. We define: *) +(* n.-tuple T == the type of n-tuples of elements of type T. *) +(* [tuple of s] == the tuple whose underlying sequence (value) is s. *) +(* The size of s must be known: specifically, Coq must *) +(* be able to infer a Canonical tuple projecting on s. *) +(* in_tuple s == the (size s)-tuple with value s. *) +(* [tuple] == the empty tuple, and *) +(* [tuple x1; ..; xn] == the explicit n.-tuple <x1; ..; xn>. *) +(* [tuple E | i < n] == the n.-tuple with general term E (i : 'I_n is bound *) +(* in E). *) +(* tcast Emn t == the m-tuple t cast as an n-tuple using Emn : m = n. *) +(* As n.-tuple T coerces to seq t, all seq operations (size, nth, ...) can be *) +(* applied to t : n.-tuple T; we provide a few specialized instances when *) +(* avoids the need for a default value. *) +(* tsize t == the size of t (the n in n.-tuple T) *) +(* tnth t i == the i'th component of t, where i : 'I_n. *) +(* [tnth t i] == the i'th component of t, where i : nat and i < n *) +(* is convertible to true. *) +(* thead t == the first element of t, when n is m.+1 for some m. *) +(* Most seq constructors (cons, behead, cat, rcons, belast, take, drop, rot, *) +(* map, ...) can be used to build tuples via the [tuple of s] construct. *) +(* Tuples are actually a subType of seq, and inherit all combinatorial *) +(* structures, including the finType structure. *) +(* Some useful lemmas and definitions: *) +(* tuple0 : [tuple] is the only 0.-tuple *) +(* tupleP : elimination view for n.+1.-tuple *) +(* ord_tuple n : the n.-tuple of all i : 'I_n *) +(******************************************************************************) + +Section Def. + +Variables (n : nat) (T : Type). + +Structure tuple_of : Type := Tuple {tval :> seq T; _ : size tval == n}. + +Canonical tuple_subType := Eval hnf in [subType for tval]. + +Implicit Type t : tuple_of. + +Definition tsize of tuple_of := n. + +Lemma size_tuple t : size t = n. +Proof. exact: (eqP (valP t)). Qed. + +Lemma tnth_default t : 'I_n -> T. +Proof. by rewrite -(size_tuple t); case: (tval t) => [|//] []. Qed. + +Definition tnth t i := nth (tnth_default t i) t i. + +Lemma tnth_nth x t i : tnth t i = nth x t i. +Proof. by apply: set_nth_default; rewrite size_tuple. Qed. + +Lemma map_tnth_enum t : map (tnth t) (enum 'I_n) = t. +Proof. +case def_t: {-}(val t) => [|x0 t']. + by rewrite [enum _]size0nil // -cardE card_ord -(size_tuple t) def_t. +apply: (@eq_from_nth _ x0) => [|i]; rewrite size_map. + by rewrite -cardE size_tuple card_ord. +move=> lt_i_e; have lt_i_n: i < n by rewrite -cardE card_ord in lt_i_e. +by rewrite (nth_map (Ordinal lt_i_n)) // (tnth_nth x0) nth_enum_ord. +Qed. + +Lemma eq_from_tnth t1 t2 : tnth t1 =1 tnth t2 -> t1 = t2. +Proof. +by move/eq_map=> eq_t; apply: val_inj; rewrite /= -!map_tnth_enum eq_t. +Qed. + +Definition tuple t mkT : tuple_of := + mkT (let: Tuple _ tP := t return size t == n in tP). + +Lemma tupleE t : tuple (fun sP => @Tuple t sP) = t. +Proof. by case: t. Qed. + +End Def. + +Notation "n .-tuple" := (tuple_of n) + (at level 2, format "n .-tuple") : type_scope. + +Notation "{ 'tuple' n 'of' T }" := (n.-tuple T : predArgType) + (at level 0, only parsing) : form_scope. + +Notation "[ 'tuple' 'of' s ]" := (tuple (fun sP => @Tuple _ _ s sP)) + (at level 0, format "[ 'tuple' 'of' s ]") : form_scope. + +Notation "[ 'tnth' t i ]" := (tnth t (@Ordinal (tsize t) i (erefl true))) + (at level 0, t, i at level 8, format "[ 'tnth' t i ]") : form_scope. + +Canonical nil_tuple T := Tuple (isT : @size T [::] == 0). +Canonical cons_tuple n T x (t : n.-tuple T) := + Tuple (valP t : size (x :: t) == n.+1). + +Notation "[ 'tuple' x1 ; .. ; xn ]" := [tuple of x1 :: .. [:: xn] ..] + (at level 0, format "[ 'tuple' '[' x1 ; '/' .. ; '/' xn ']' ]") + : form_scope. + +Notation "[ 'tuple' ]" := [tuple of [::]] + (at level 0, format "[ 'tuple' ]") : form_scope. + +Section CastTuple. + +Variable T : Type. + +Definition in_tuple (s : seq T) := Tuple (eqxx (size s)). + +Definition tcast m n (eq_mn : m = n) t := + let: erefl in _ = n := eq_mn return n.-tuple T in t. + +Lemma tcastE m n (eq_mn : m = n) t i : + tnth (tcast eq_mn t) i = tnth t (cast_ord (esym eq_mn) i). +Proof. by case: n / eq_mn in i *; rewrite cast_ord_id. Qed. + +Lemma tcast_id n (eq_nn : n = n) t : tcast eq_nn t = t. +Proof. by rewrite (eq_axiomK eq_nn). Qed. + +Lemma tcastK m n (eq_mn : m = n) : cancel (tcast eq_mn) (tcast (esym eq_mn)). +Proof. by case: n / eq_mn. Qed. + +Lemma tcastKV m n (eq_mn : m = n) : cancel (tcast (esym eq_mn)) (tcast eq_mn). +Proof. by case: n / eq_mn. Qed. + +Lemma tcast_trans m n p (eq_mn : m = n) (eq_np : n = p) t: + tcast (etrans eq_mn eq_np) t = tcast eq_np (tcast eq_mn t). +Proof. by case: n / eq_mn eq_np; case: p /. Qed. + +Lemma tvalK n (t : n.-tuple T) : in_tuple t = tcast (esym (size_tuple t)) t. +Proof. by apply: val_inj => /=; case: _ / (esym _). Qed. + +Lemma in_tupleE s : in_tuple s = s :> seq T. Proof. by []. Qed. + +End CastTuple. + +Section SeqTuple. + +Variables (n m : nat) (T U rT : Type). +Implicit Type t : n.-tuple T. + +Lemma rcons_tupleP t x : size (rcons t x) == n.+1. +Proof. by rewrite size_rcons size_tuple. Qed. +Canonical rcons_tuple t x := Tuple (rcons_tupleP t x). + +Lemma nseq_tupleP x : @size T (nseq n x) == n. +Proof. by rewrite size_nseq. Qed. +Canonical nseq_tuple x := Tuple (nseq_tupleP x). + +Lemma iota_tupleP : size (iota m n) == n. +Proof. by rewrite size_iota. Qed. +Canonical iota_tuple := Tuple iota_tupleP. + +Lemma behead_tupleP t : size (behead t) == n.-1. +Proof. by rewrite size_behead size_tuple. Qed. +Canonical behead_tuple t := Tuple (behead_tupleP t). + +Lemma belast_tupleP x t : size (belast x t) == n. +Proof. by rewrite size_belast size_tuple. Qed. +Canonical belast_tuple x t := Tuple (belast_tupleP x t). + +Lemma cat_tupleP t (u : m.-tuple T) : size (t ++ u) == n + m. +Proof. by rewrite size_cat !size_tuple. Qed. +Canonical cat_tuple t u := Tuple (cat_tupleP t u). + +Lemma take_tupleP t : size (take m t) == minn m n. +Proof. by rewrite size_take size_tuple eqxx. Qed. +Canonical take_tuple t := Tuple (take_tupleP t). + +Lemma drop_tupleP t : size (drop m t) == n - m. +Proof. by rewrite size_drop size_tuple. Qed. +Canonical drop_tuple t := Tuple (drop_tupleP t). + +Lemma rev_tupleP t : size (rev t) == n. +Proof. by rewrite size_rev size_tuple. Qed. +Canonical rev_tuple t := Tuple (rev_tupleP t). + +Lemma rot_tupleP t : size (rot m t) == n. +Proof. by rewrite size_rot size_tuple. Qed. +Canonical rot_tuple t := Tuple (rot_tupleP t). + +Lemma rotr_tupleP t : size (rotr m t) == n. +Proof. by rewrite size_rotr size_tuple. Qed. +Canonical rotr_tuple t := Tuple (rotr_tupleP t). + +Lemma map_tupleP f t : @size rT (map f t) == n. +Proof. by rewrite size_map size_tuple. Qed. +Canonical map_tuple f t := Tuple (map_tupleP f t). + +Lemma scanl_tupleP f x t : @size rT (scanl f x t) == n. +Proof. by rewrite size_scanl size_tuple. Qed. +Canonical scanl_tuple f x t := Tuple (scanl_tupleP f x t). + +Lemma pairmap_tupleP f x t : @size rT (pairmap f x t) == n. +Proof. by rewrite size_pairmap size_tuple. Qed. +Canonical pairmap_tuple f x t := Tuple (pairmap_tupleP f x t). + +Lemma zip_tupleP t (u : n.-tuple U) : size (zip t u) == n. +Proof. by rewrite size1_zip !size_tuple. Qed. +Canonical zip_tuple t u := Tuple (zip_tupleP t u). + +Lemma allpairs_tupleP f t (u : m.-tuple U) : @size rT (allpairs f t u) == n * m. +Proof. by rewrite size_allpairs !size_tuple. Qed. +Canonical allpairs_tuple f t u := Tuple (allpairs_tupleP f t u). + +Definition thead (u : n.+1.-tuple T) := tnth u ord0. + +Lemma tnth0 x t : tnth [tuple of x :: t] ord0 = x. +Proof. by []. Qed. + +Lemma theadE x t : thead [tuple of x :: t] = x. +Proof. by []. Qed. + +Lemma tuple0 : all_equal_to ([tuple] : 0.-tuple T). +Proof. by move=> t; apply: val_inj; case: t => [[]]. Qed. + +CoInductive tuple1_spec : n.+1.-tuple T -> Type := + Tuple1spec x t : tuple1_spec [tuple of x :: t]. + +Lemma tupleP u : tuple1_spec u. +Proof. +case: u => [[|x s] //= sz_s]; pose t := @Tuple n _ s sz_s. +rewrite (_ : Tuple _ = [tuple of x :: t]) //; exact: val_inj. +Qed. + +Lemma tnth_map f t i : tnth [tuple of map f t] i = f (tnth t i) :> rT. +Proof. by apply: nth_map; rewrite size_tuple. Qed. + +End SeqTuple. + +Lemma tnth_behead n T (t : n.+1.-tuple T) i : + tnth [tuple of behead t] i = tnth t (inord i.+1). +Proof. by case/tupleP: t => x t; rewrite !(tnth_nth x) inordK ?ltnS. Qed. + +Lemma tuple_eta n T (t : n.+1.-tuple T) : t = [tuple of thead t :: behead t]. +Proof. by case/tupleP: t => x t; exact: val_inj. Qed. + +Section TupleQuantifiers. + +Variables (n : nat) (T : Type). +Implicit Types (a : pred T) (t : n.-tuple T). + +Lemma forallb_tnth a t : [forall i, a (tnth t i)] = all a t. +Proof. +apply: negb_inj; rewrite -has_predC -has_map negb_forall. +apply/existsP/(has_nthP true) => [[i a_t_i] | [i lt_i_n a_t_i]]. + by exists i; rewrite ?size_tuple // -tnth_nth tnth_map. +rewrite size_tuple in lt_i_n; exists (Ordinal lt_i_n). +by rewrite -tnth_map (tnth_nth true). +Qed. + +Lemma existsb_tnth a t : [exists i, a (tnth t i)] = has a t. +Proof. by apply: negb_inj; rewrite negb_exists -all_predC -forallb_tnth. Qed. + +Lemma all_tnthP a t : reflect (forall i, a (tnth t i)) (all a t). +Proof. by rewrite -forallb_tnth; apply: forallP. Qed. + +Lemma has_tnthP a t : reflect (exists i, a (tnth t i)) (has a t). +Proof. by rewrite -existsb_tnth; apply: existsP. Qed. + +End TupleQuantifiers. + +Implicit Arguments all_tnthP [n T a t]. +Implicit Arguments has_tnthP [n T a t]. + +Section EqTuple. + +Variables (n : nat) (T : eqType). + +Definition tuple_eqMixin := Eval hnf in [eqMixin of n.-tuple T by <:]. +Canonical tuple_eqType := Eval hnf in EqType (n.-tuple T) tuple_eqMixin. + +Canonical tuple_predType := + Eval hnf in mkPredType (fun t : n.-tuple T => mem_seq t). + +Lemma memtE (t : n.-tuple T) : mem t = mem (tval t). +Proof. by []. Qed. + +Lemma mem_tnth i (t : n.-tuple T) : tnth t i \in t. +Proof. by rewrite mem_nth ?size_tuple. Qed. + +Lemma memt_nth x0 (t : n.-tuple T) i : i < n -> nth x0 t i \in t. +Proof. by move=> i_lt_n; rewrite mem_nth ?size_tuple. Qed. + +Lemma tnthP (t : n.-tuple T) x : reflect (exists i, x = tnth t i) (x \in t). +Proof. +apply: (iffP idP) => [/(nthP x)[i ltin <-] | [i ->]]; last exact: mem_tnth. +by rewrite size_tuple in ltin; exists (Ordinal ltin); rewrite (tnth_nth x). +Qed. + +Lemma seq_tnthP (s : seq T) x : x \in s -> {i | x = tnth (in_tuple s) i}. +Proof. +move=> s_x; pose i := index x s; have lt_i: i < size s by rewrite index_mem. +by exists (Ordinal lt_i); rewrite (tnth_nth x) nth_index. +Qed. + +End EqTuple. + +Definition tuple_choiceMixin n (T : choiceType) := + [choiceMixin of n.-tuple T by <:]. + +Canonical tuple_choiceType n (T : choiceType) := + Eval hnf in ChoiceType (n.-tuple T) (tuple_choiceMixin n T). + +Definition tuple_countMixin n (T : countType) := + [countMixin of n.-tuple T by <:]. + +Canonical tuple_countType n (T : countType) := + Eval hnf in CountType (n.-tuple T) (tuple_countMixin n T). + +Canonical tuple_subCountType n (T : countType) := + Eval hnf in [subCountType of n.-tuple T]. + +Module Type FinTupleSig. +Section FinTupleSig. +Variables (n : nat) (T : finType). +Parameter enum : seq (n.-tuple T). +Axiom enumP : Finite.axiom enum. +Axiom size_enum : size enum = #|T| ^ n. +End FinTupleSig. +End FinTupleSig. + +Module FinTuple : FinTupleSig. +Section FinTuple. +Variables (n : nat) (T : finType). + +Definition enum : seq (n.-tuple T) := + let extend e := flatten (codom (fun x => map (cons x) e)) in + pmap insub (iter n extend [::[::]]). + +Lemma enumP : Finite.axiom enum. +Proof. +case=> /= t t_n; rewrite -(count_map _ (pred1 t)) (pmap_filter (@insubK _ _ _)). +rewrite count_filter -(@eq_count _ (pred1 t)) => [|s /=]; last first. + by rewrite isSome_insub; case: eqP=> // ->. +elim: n t t_n => [|m IHm] [|x t] //= {IHm}/IHm; move: (iter m _ _) => em IHm. +transitivity (x \in T : nat); rewrite // -mem_enum codomE. +elim: (fintype.enum T) (enum_uniq T) => //= y e IHe /andP[/negPf ney]. +rewrite count_cat count_map inE /preim /= {1}/eq_op /= eq_sym => /IHe->. +by case: eqP => [->|_]; rewrite ?(ney, count_pred0, IHm). +Qed. + +Lemma size_enum : size enum = #|T| ^ n. +Proof. +rewrite /= cardE size_pmap_sub; elim: n => //= m IHm. +rewrite expnS /codom /image_mem; elim: {2 3}(fintype.enum T) => //= x e IHe. +by rewrite count_cat {}IHe count_map IHm. +Qed. + +End FinTuple. +End FinTuple. + +Section UseFinTuple. + +Variables (n : nat) (T : finType). + +Canonical tuple_finMixin := Eval hnf in FinMixin (@FinTuple.enumP n T). +Canonical tuple_finType := Eval hnf in FinType (n.-tuple T) tuple_finMixin. +Canonical tuple_subFinType := Eval hnf in [subFinType of n.-tuple T]. + +Lemma card_tuple : #|{:n.-tuple T}| = #|T| ^ n. +Proof. by rewrite [#|_|]cardT enumT unlock FinTuple.size_enum. Qed. + +Lemma enum_tupleP (A : pred T) : size (enum A) == #|A|. +Proof. by rewrite -cardE. Qed. +Canonical enum_tuple A := Tuple (enum_tupleP A). + +Definition ord_tuple : n.-tuple 'I_n := Tuple (introT eqP (size_enum_ord n)). +Lemma val_ord_tuple : val ord_tuple = enum 'I_n. Proof. by []. Qed. + +Lemma tuple_map_ord U (t : n.-tuple U) : t = [tuple of map (tnth t) ord_tuple]. +Proof. by apply: val_inj => /=; rewrite map_tnth_enum. Qed. + +Lemma tnth_ord_tuple i : tnth ord_tuple i = i. +Proof. +apply: val_inj; rewrite (tnth_nth i) -(nth_map _ 0) ?size_tuple //. +by rewrite /= enumT unlock val_ord_enum nth_iota. +Qed. + +Section ImageTuple. + +Variables (T' : Type) (f : T -> T') (A : pred T). + +Canonical image_tuple : #|A|.-tuple T' := [tuple of image f A]. +Canonical codom_tuple : #|T|.-tuple T' := [tuple of codom f]. + +End ImageTuple. + +Section MkTuple. + +Variables (T' : Type) (f : 'I_n -> T'). + +Definition mktuple := map_tuple f ord_tuple. + +Lemma tnth_mktuple i : tnth mktuple i = f i. +Proof. by rewrite tnth_map tnth_ord_tuple. Qed. + +Lemma nth_mktuple x0 (i : 'I_n) : nth x0 mktuple i = f i. +Proof. by rewrite -tnth_nth tnth_mktuple. Qed. + +End MkTuple. + +End UseFinTuple. + +Notation "[ 'tuple' F | i < n ]" := (mktuple (fun i : 'I_n => F)) + (at level 0, i at level 0, + format "[ '[hv' 'tuple' F '/' | i < n ] ']'") : form_scope. + + |
