diff options
| author | Cyril Cohen | 2015-07-17 18:03:31 +0200 |
|---|---|---|
| committer | Cyril Cohen | 2015-07-17 18:03:31 +0200 |
| commit | 532de9b68384a114c6534a0736ed024c900447f9 (patch) | |
| tree | e100a6a7839bf7548ab8a9e053033f8eef3c7492 /mathcomp | |
| parent | f180c539a00fd83d8b3b5fd2d5710eb16e971e2e (diff) | |
Updating files + reorganizing everything
Diffstat (limited to 'mathcomp')
211 files changed, 14025 insertions, 5875 deletions
diff --git a/mathcomp/Make b/mathcomp/Make new file mode 100644 index 0000000..d8fb252 --- /dev/null +++ b/mathcomp/Make @@ -0,0 +1,182 @@ +algebra/all_algebra.v +algebra/finalg.v +algebra/fraction.v +algebra/intdiv.v +algebra/interval.v +algebra/matrix.v +algebra/mxalgebra.v +algebra/mxpoly.v +algebra/polydiv.v +algebra/poly.v +algebra/polyXY.v +algebra/rat.v +algebra/ring_quotient.v +algebra/ssralg.v +algebra/ssrint.v +algebra/ssrnum.v +algebra/vector.v +algebra/zmodp.v +all/all.v +basic/all_basic.v +basic/bigop.v +basic/binomial.v +basic/choice.v +basic/div.v +basic/finfun.v +basic/fingraph.v +basic/finset.v +basic/fintype.v +basic/generic_quotient.v +basic/path.v +basic/prime.v +basic/tuple.v +character/all_character.v +character/character.v +character/classfun.v +character/finfield.v +character/inertia.v +character/integral_char.v +character/mxabelem.v +character/mxrepresentation.v +character/vcharacter.v +field/algC.v +field/algebraics_fundamentals.v +field/algnum.v +field/all_field.v +field/closed_field.v +field/countalg.v +field/cyclotomic.v +field/falgebra.v +field/fieldext.v +field/galois.v +field/separable.v +fingroup/action.v +fingroup/all_fingroup.v +fingroup/automorphism.v +fingroup/fingroup.v +fingroup/gproduct.v +fingroup/morphism.v +fingroup/perm.v +fingroup/presentation.v +fingroup/quotient.v +odd_order/BGappendixAB.v +odd_order/BGappendixC.v +odd_order/BGsection10.v +odd_order/BGsection11.v +odd_order/BGsection12.v +odd_order/BGsection13.v +odd_order/BGsection14.v +odd_order/BGsection15.v +odd_order/BGsection16.v +odd_order/BGsection1.v +odd_order/BGsection2.v +odd_order/BGsection3.v +odd_order/BGsection4.v +odd_order/BGsection5.v +odd_order/BGsection6.v +odd_order/BGsection7.v +odd_order/BGsection8.v +odd_order/BGsection9.v +odd_order/PFsection10.v +odd_order/PFsection11.v +odd_order/PFsection12.v +odd_order/PFsection13.v +odd_order/PFsection14.v +odd_order/PFsection1.v +odd_order/PFsection2.v +odd_order/PFsection3.v +odd_order/PFsection4.v +odd_order/PFsection5.v +odd_order/PFsection6.v +odd_order/PFsection7.v +odd_order/PFsection8.v +odd_order/PFsection9.v +odd_order/stripped_odd_order_theorem.v +odd_order/wielandt_fixpoint.v +real_closed/all_real_closed.v +real_closed/bigenough.v +real_closed/cauchyreals.v +real_closed/complex.v +real_closed/mxtens.v +real_closed/ordered_qelim.v +real_closed/polyorder.v +real_closed/polyrcf.v +real_closed/qe_rcf_th.v +real_closed/qe_rcf.v +real_closed/realalg.v +solvable/abelian.v +solvable/all_solvable.v +solvable/alt.v +solvable/burnside_app.v +solvable/center.v +solvable/commutator.v +solvable/cyclic.v +solvable/extraspecial.v +solvable/extremal.v +solvable/finmodule.v +solvable/frobenius.v +solvable/gfunctor.v +solvable/gseries.v +solvable/hall.v +solvable/jordanholder.v +solvable/maximal.v +solvable/nilpotent.v +solvable/pgroup.v +solvable/primitive_action.v +solvable/sylow.v +ssreflect/all_ssreflect.v +ssreflect/eqtype.v +ssreflect/seq.v +ssreflect/ssrbool.v +ssreflect/ssreflect.v +ssreflect/ssrfun.v +ssreflect/ssrmatching.v +ssreflect/ssrnat.v +ssrtest/absevarprop.v +ssrtest/binders_of.v +ssrtest/binders.v +ssrtest/caseview.v +ssrtest/congr.v +ssrtest/deferclear.v +ssrtest/dependent_type_err.v +ssrtest/elim2.v +ssrtest/elim_pattern.v +ssrtest/elim.v +ssrtest/first_n.v +ssrtest/gen_have.v +ssrtest/gen_pattern.v +ssrtest/havesuff.v +ssrtest/have_TC.v +ssrtest/have_transp.v +ssrtest/have_view_idiom.v +ssrtest/if_isnt.v +ssrtest/indetLHS.v +ssrtest/intro_beta.v +ssrtest/intro_noop.v +ssrtest/ipatalternation.v +ssrtest/ltac_have.v +ssrtest/ltac_in.v +ssrtest/move_after.v +ssrtest/multiview.v +ssrtest/occarrow.v +ssrtest/patnoX.v +ssrtest/rewpatterns.v +ssrtest/set_lamda.v +ssrtest/set_pattern.v +ssrtest/ssrsyntax1.v +ssrtest/ssrsyntax2.v +ssrtest/tc.v +ssrtest/testmx.v +ssrtest/typeof.v +ssrtest/unkeyed.v +ssrtest/view_case.v +ssrtest/wlogletin.v +ssrtest/wlog_suff.v +ssrtest/wlong_intro.v +ssreflect.ml4 +ssreflect.mllib +ssrmatching.ml4 +ssrmatching.mli + +-I . +-R . mathcomp diff --git a/mathcomp/Makefile b/mathcomp/Makefile index fcd8dfe..11419d3 100644 --- a/mathcomp/Makefile +++ b/mathcomp/Makefile @@ -1,37 +1,44 @@ +H=@ -.PHONY: ssreflect algebra fingroup all odd_order solvable field character discrete real_closed ssrtest +ifeq "$(COQBIN)" "" +COQBIN=$(dir $(shell which coqtop))/ +endif -export COQPATH=$(PWD)/.. +BRANCH_coq = $(shell $(COQBIN)/coqtop -v | head -1 | sed 's/.*version \([0-9]\.[0-9]\)[^ ]* .*/v\1/') -all: ssreflect algebra fingroup odd_order solvable field character discrete real_closed ssrtest - $(MAKE) -C $@ +HASH_coq = $(shell echo Quit. | $(COQBIN)/coqtop 2>&1 | head -1 | sed 's/^.*(\([a-f0-9]*\)).*/\1/' ) -ssreflect: - $(MAKE) -C $@ +HASH_coq_v85beta1 = eaa3d0b15adf4eb11ffb00ab087746a5b15c4d5d -algebra: fingroup - $(MAKE) -C $@ -fingroup: discrete - $(MAKE) -C $@ +ifeq "$(HASH_coq)" "$(HASH_coq_v85beta1)" +V=v8.5beta1 +else +V=$(BRANCH_coq) +endif -odd_order: field - $(MAKE) -C $@ +OLD_MAKEFLAGS:=$(MAKEFLAGS) +MAKEFLAGS+=-B -solvable: algebra - $(MAKE) -C $@ +.DEFAULT_GOAL := all -field: solvable - $(MAKE) -C $@ +%: + $(H)[ -e Makefile.coq ] || $(call coqmakefile) + # Override COQDEP to find only the "right" copy of .ml files + $(H)MAKEFLAGS=$(OLD_MAKEFLAGS) $(MAKE) --no-print-directory \ + -f Makefile.coq $* \ + COQDEP='$(COQBIN)/coqdep -exclude-dir "plugin" -c' -character: field - $(MAKE) -C $@ +define coqmakefile + (echo "Generating Makefile.coq for Coq $(V) with COQBIN=$(COQBIN)";\ + ln -sf ssreflect/plugin/$(V)/ssreflect.mllib .;\ + ln -sf ssreflect/plugin/$(V)/ssrmatching.mli .;\ + ln -sf ssreflect/plugin/$(V)/ssrmatching.ml4 .;\ + ln -sf ssreflect/plugin/$(V)/ssreflect.ml4 .;\ + $(COQBIN)/coq_makefile -f Make -o Makefile.coq) +endef -discrete: ssreflect - $(MAKE) -C $@ - -real_closed: algebra - $(MAKE) -C $@ - -ssrtest: algebra - $(MAKE) -C $@ +clean: + $(H)MAKEFLAGS=$(OLD_MAKEFLAGS) $(MAKE) --no-print-directory \ + -f Makefile.coq clean + $(H)rm -f Makefile.coq diff --git a/mathcomp/algebra/Make b/mathcomp/algebra/Make index d9f4c63..7d12cb4 100644 --- a/mathcomp/algebra/Make +++ b/mathcomp/algebra/Make @@ -1,4 +1,4 @@ -all.v +all_algebra.v finalg.v fraction.v intdiv.v @@ -16,6 +16,5 @@ ssrint.v ssrnum.v vector.v zmodp.v -cyclic.v -R . mathcomp.algebra diff --git a/mathcomp/algebra/all.v b/mathcomp/algebra/all_algebra.v index 8a93ca9..8a93ca9 100644 --- a/mathcomp/algebra/all.v +++ b/mathcomp/algebra/all_algebra.v diff --git a/mathcomp/algebra/finalg.v b/mathcomp/algebra/finalg.v index 13daa31..0155a1b 100644 --- a/mathcomp/algebra/finalg.v +++ b/mathcomp/algebra/finalg.v @@ -1,12 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import choice fintype finset. -From mathcomp.fingroup -Require Import fingroup morphism perm action. -Require Import ssralg. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp +Require Import ssralg finset fingroup morphism perm action. (*****************************************************************************) (* This file clones the entire ssralg hierachy for finite types; this allows *) @@ -155,7 +152,7 @@ Lemma zmodMgE x y : (x * y)%g = x + y. Proof. by []. Qed. Lemma zmodXgE n x : (x ^+ n)%g = x *+ n. Proof. by []. Qed. Lemma zmod_mulgC x y : commute x y. Proof. exact: GRing.addrC. Qed. Lemma zmod_abelian (A : {set U}) : abelian A. -Proof. by apply/centsP=> x _ y _; exact: zmod_mulgC. Qed. +Proof. by apply/centsP=> x _ y _; apply: zmod_mulgC. Qed. End AdditiveGroup. @@ -447,11 +444,11 @@ Lemma unit_mul_proof u v : val u * val v \is a GRing.unit. Proof. by rewrite (GRing.unitrMr _ (valP u)) ?(valP v). Qed. Definition unit_mul u v := Unit phR (unit_mul_proof u v). Lemma unit_muluA : associative unit_mul. -Proof. move=> u v w; apply: val_inj; exact: GRing.mulrA. Qed. +Proof. by move=> u v w; apply: val_inj; apply: GRing.mulrA. Qed. Lemma unit_mul1u : left_id unit1 unit_mul. -Proof. move=> u; apply: val_inj; exact: GRing.mul1r. Qed. +Proof. by move=> u; apply: val_inj; apply: GRing.mul1r. Qed. Lemma unit_mulVu : left_inverse unit1 unit_inv unit_mul. -Proof. move=> u; apply: val_inj; exact: GRing.mulVr (valP u). Qed. +Proof. by move=> u; apply: val_inj; apply: GRing.mulVr (valP u). Qed. Definition unit_GroupMixin := FinGroup.Mixin unit_muluA unit_mul1u unit_mulVu. Canonical unit_baseFinGroupType := @@ -829,12 +826,12 @@ Lemma decidable : GRing.DecidableField.axiom sat. Proof. move=> e f; elim: f e; try by move=> f1 IH1 f2 IH2 e /=; case IH1; case IH2; constructor; tauto. -- by move=> b e; exact: idP. -- by move=> t1 t2 e; exact: eqP. -- by move=> t e; exact: idP. +- by move=> b e; apply: idP. +- by move=> t1 t2 e; apply: eqP. +- by move=> t e; apply: idP. - by move=> f IH e /=; case: IH; constructor. -- by move=> i f IH e; apply: (iffP existsP) => [] [x fx]; exists x; exact/IH. -by move=> i f IH e; apply: (iffP forallP) => f_ x; exact/IH. +- by move=> i f IH e; apply: (iffP existsP) => [] [x fx]; exists x; apply/IH. +by move=> i f IH e; apply: (iffP forallP) => f_ x; apply/IH. Qed. Definition DecidableFieldMixin := DecFieldMixin decidable. diff --git a/mathcomp/algebra/fraction.v b/mathcomp/algebra/fraction.v index 9ce2718..896a287 100644 --- a/mathcomp/algebra/fraction.v +++ b/mathcomp/algebra/fraction.v @@ -1,10 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice tuple bigop generic_quotient. -Require Import ssralg poly polydiv. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat div seq choice tuple. +From mathcomp +Require Import bigop ssralg poly polydiv generic_quotient. (* This file builds the field of fraction of any integral domain. *) (* The main result of this file is the existence of the field *) diff --git a/mathcomp/algebra/intdiv.v b/mathcomp/algebra/intdiv.v index cdfde3c..4e4148a 100644 --- a/mathcomp/algebra/intdiv.v +++ b/mathcomp/algebra/intdiv.v @@ -1,13 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype tuple finfun bigop prime. -From mathcomp.fingroup -Require Import perm. -Require Import ssralg poly ssrnum ssrint rat. -Require Import polydiv finalg zmodp matrix mxalgebra vector. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp +Require Import fintype tuple finfun bigop prime ssralg poly ssrnum ssrint rat. +From mathcomp +Require Import polydiv finalg perm zmodp matrix mxalgebra vector. (******************************************************************************) (* This file provides various results on divisibility of integers. *) @@ -489,7 +487,7 @@ have [-> | m_gt0] := posnP m; first by rewrite modz0. case: n => n; first by rewrite modz_nat gcdn_modr. rewrite modNz_nat // NegzE abszN {2}(divn_eq n m) -addnS gcdnMDl. rewrite -addrA -opprD -intS /=; set m1 := _.+1. -have le_m1m: (m1 <= m)%N by exact: ltn_pmod. +have le_m1m: (m1 <= m)%N by apply: ltn_pmod. by rewrite subzn // !(gcdnC m) -{2 3}(subnK le_m1m) gcdnDl gcdnDr gcdnC. Qed. @@ -851,13 +849,13 @@ Qed. Local Notation pZtoQ := (map_poly (intr : int -> rat)). Lemma size_rat_int_poly p : size (pZtoQ p) = size p. -Proof. by apply: size_map_inj_poly; first exact: intr_inj. Qed. +Proof. by apply: size_map_inj_poly; first apply: intr_inj. Qed. Lemma rat_poly_scale (p : {poly rat}) : {q : {poly int} & {a | a != 0 & p = a%:~R^-1 *: pZtoQ q}}. Proof. pose a := \prod_(i < size p) denq p`_i. -have nz_a: a != 0 by apply/prodf_neq0=> i _; exact: denq_neq0. +have nz_a: a != 0 by apply/prodf_neq0=> i _; apply: denq_neq0. exists (map_poly numq (a%:~R *: p)), a => //. apply: canRL (scalerK _) _; rewrite ?intr_eq0 //. apply/polyP=> i; rewrite !(coefZ, coef_map_id0) // numqK // Qint_def mulrC. @@ -927,7 +925,7 @@ wlog [j a'Mij]: m n M i Da le_mn / {j | ~~ (a %| M i j)%Z}; last first. have{leA} ltA: (`|b| < A)%N. rewrite -ltnS (leq_trans _ leA) // ltnS ltn_neqAle andbC. rewrite dvdn_leq ?absz_gt0 ? dvdn_gcdl //=. - by rewrite (contraNneq _ a'Mij) ?dvdzE // => <-; exact: dvdn_gcdr. + by rewrite (contraNneq _ a'Mij) ?dvdzE // => <-; apply: dvdn_gcdr. pose t2 := [fun j : 'I_2 => [tuple _; _]`_j : int]; pose a1 := M i 1. pose Uul := \matrix_(k, j) t2 (t2 u (- (a1 %/ b)%Z) j) (t2 v (a %/ b)%Z j) k. pose U : 'M_(2 + n) := block_mx Uul 0 0 1%:M; pose M1 := M *m U. @@ -995,7 +993,7 @@ exists (block_mx 1 0 Ml L). exists (block_mx 1 Mu 0 R). by rewrite unitmxE det_ublock det_scalar1 mul1r. exists (1 :: d); set D1 := \matrix_(i, j) _ in dM1. - by rewrite /= path_min_sorted // => g _; exact: dvd1n. + by rewrite /= path_min_sorted // => g _; apply: dvd1n. rewrite [D in _ *m D *m _](_ : _ = block_mx 1 0 0 D1); last first. by apply/matrixP=> i j; do 3?[rewrite ?mxE ?ord1 //=; case: splitP => ? ->]. rewrite !mulmx_block !(mul0mx, mulmx0, addr0) !mulmx1 add0r mul1mx -Da -dM1. @@ -1023,7 +1021,7 @@ have [K kerK]: {K : 'M_(k, m) | map_mx intr K == kermx S}%MS. apply/matrixP=> i j; rewrite 3!mxE mulrC [d](bigD1 (i, j)) // rmorphM mulrA. by rewrite -numqE -rmorphM numq_int. suffices nz_d: d%:Q != 0 by rewrite !eqmx_scale // !eq_row_base andbb. - by rewrite intr_eq0; apply/prodf_neq0 => i _; exact: denq_neq0. + by rewrite intr_eq0; apply/prodf_neq0 => i _; apply: denq_neq0. have [L _ [G uG [D _ defK]]] := int_Smith_normal_form K. pose Gud := castmx (Dm, Em) G; pose G'lr := castmx (Em, Dm) (invmx G). have{K L D defK kerK} kerGu: map_mx intr (usubmx Gud) *m S = 0. @@ -1046,7 +1044,7 @@ have{K L D defK kerK} kerGu: map_mx intr (usubmx Gud) *m S = 0. by rewrite -{2}[Gud]vsubmxK map_col_mx mul_row_col mul0mx addr0. pose T := map_mx intr (dsubmx Gud) *m S. have{kerGu} defS: map_mx intr (rsubmx G'lr) *m T = S. - have: G'lr *m Gud = 1%:M by rewrite /G'lr /Gud; case: _ / (Dm); exact: mulVmx. + have: G'lr *m Gud = 1%:M by rewrite /G'lr /Gud; case: _ / (Dm); apply: mulVmx. rewrite -{1}[G'lr]hsubmxK -[Gud]vsubmxK mulmxA mul_row_col -map_mxM. move/(canRL (addKr _))->; rewrite -mulNmx raddfD /= map_mx1 map_mxM /=. by rewrite mulmxDl -mulmxA kerGu mulmx0 add0r mul1mx. diff --git a/mathcomp/algebra/interval.v b/mathcomp/algebra/interval.v index bed19fa..fbc6f16 100644 --- a/mathcomp/algebra/interval.v +++ b/mathcomp/algebra/interval.v @@ -1,12 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype bigop finset. -From mathcomp.fingroup -Require Import fingroup. -Require Import ssralg zmodp ssrint ssrnum. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq div choice fintype. +From mathcomp +Require Import bigop ssralg finset fingroup zmodp ssrint ssrnum. (*****************************************************************************) (* This file provide support for intervals in numerical and real domains. *) @@ -31,8 +28,8 @@ Require Import ssralg zmodp ssrint ssrnum. (* when R is a real domain: we could provide a specific theory for this *) (* important case. *) (* *) -(* See also "Formal proofs in real algebraic geometry: from ordered fields *) -(* to quantifier elimination", LMCS journal, 2012 *) +(* See also ``Formal proofs in real algebraic geometry: from ordered fields *) +(* to quantifier elimination'', LMCS journal, 2012 *) (* by Cyril Cohen and Assia Mahboubi *) (* *) (* And "Formalized algebraic numbers: construction and first-order theory" *) @@ -148,10 +145,10 @@ Lemma lersifxx x b: (x <= x ?< if b) = ~~ b. Proof. by case: b; rewrite /= lterr. Qed. Lemma lersif_trans x y z b1 b2 : - x <= y ?< if b1 -> y <= z ?< if b2 -> x <= z ?< if b1 || b2. + x <= y ?< if b1 -> y <= z ?< if b2 -> x <= z ?< if b1 || b2. Proof. -move: b1 b2 => [] [] //=; -by [exact: ler_trans|exact: ler_lt_trans|exact: ltr_le_trans|exact: ltr_trans]. +case: b1; first by case: b2; [apply: ltr_trans | apply: ltr_le_trans]. +by case: b2; [apply: ler_lt_trans | apply: ler_trans]. Qed. Lemma lersifW b x y : x <= y ?< if b -> x <= y. diff --git a/mathcomp/algebra/matrix.v b/mathcomp/algebra/matrix.v index 9339e67..b842c67 100644 --- a/mathcomp/algebra/matrix.v +++ b/mathcomp/algebra/matrix.v @@ -1,12 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype finfun bigop prime binomial finset. -From mathcomp.fingroup -Require Import fingroup perm. -Require Import ssralg finalg zmodp. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div choice fintype. +From mathcomp +Require Import finfun bigop prime binomial ssralg finset fingroup finalg. +From mathcomp +Require Import perm zmodp. (******************************************************************************) (* Basic concrete linear algebra : definition of type for matrices, and all *) @@ -229,7 +228,7 @@ Proof. by move=> i j; rewrite unlock /fun_of_matrix /= ffunE. Qed. Lemma matrixP (A B : matrix) : A =2 B <-> A = B. Proof. rewrite /fun_of_matrix; split=> [/= eqAB | -> //]. -by apply/val_inj/ffunP=> [[i j]]; exact: eqAB. +by apply/val_inj/ffunP=> [[i j]]; apply: eqAB. Qed. End MatrixDef. @@ -573,7 +572,7 @@ Proof. by apply/matrixP=> i j; rewrite mxE row_mxEr. Qed. Lemma hsubmxK A : row_mx (lsubmx A) (rsubmx A) = A. Proof. apply/matrixP=> i j; rewrite !mxE. -case: splitP => k Dk //=; rewrite !mxE //=; congr (A _ _); exact: val_inj. +by case: splitP => k Dk //=; rewrite !mxE //=; congr (A _ _); apply: val_inj. Qed. Lemma col_mxEu A1 A2 i j : col_mx A1 A2 (lshift m2 i) j = A1 i j. @@ -1960,7 +1959,7 @@ Lemma mulmxE : mulmx = *%R. Proof. by []. Qed. Lemma idmxE : 1%:M = 1 :> 'M_n. Proof. by []. Qed. Lemma scalar_mx_is_multiplicative : multiplicative (@scalar_mx n). -Proof. by split=> //; exact: scalar_mxM. Qed. +Proof. by split=> //; apply: scalar_mxM. Qed. Canonical scalar_mx_rmorphism := AddRMorphism scalar_mx_is_multiplicative. End MatrixRing. @@ -2120,7 +2119,7 @@ Proof. by rewrite map_mx_sub map_mx1 map_pid_mx. Qed. Lemma map_mx_is_multiplicative n' (n := n'.+1) : multiplicative ((map_mx f) n n). -Proof. by split; [exact: map_mxM | exact: map_mx1]. Qed. +Proof. by split; [apply: map_mxM | apply: map_mx1]. Qed. Canonical map_mx_rmorphism n' := AddRMorphism (map_mx_is_multiplicative n'). @@ -2321,7 +2320,7 @@ transitivity (\sum_(f : F) \sum_(s : 'S_n) (-1) ^+ s * \prod_i AB s i (f i)). rewrite (bigID (fun f : F => injectiveb f)) /= addrC big1 ?add0r => [|f Uf]. rewrite (reindex (@pval _)) /=; last first. pose in_Sn := insubd (1%g : 'S_n). - by exists in_Sn => /= f Uf; first apply: val_inj; exact: insubdK. + by exists in_Sn => /= f Uf; first apply: val_inj; apply: insubdK. apply: eq_big => /= [s | s _]; rewrite ?(valP s) // big_distrr /=. rewrite (reindex_inj (mulgI s)); apply: eq_bigr => t _ /=. rewrite big_split /= mulrA mulrCA mulrA mulrCA mulrA. @@ -2766,13 +2765,13 @@ Local Notation "A ^f" := (map_mx f A) : ring_scope. Lemma map_mx_inj m n : injective ((map_mx f) m n). Proof. move=> A B eq_AB; apply/matrixP=> i j. -by move/matrixP/(_ i j): eq_AB; rewrite !mxE; exact: fmorph_inj. +by move/matrixP/(_ i j): eq_AB; rewrite !mxE; apply: fmorph_inj. Qed. Lemma map_mx_is_scalar n (A : 'M_n) : is_scalar_mx A^f = is_scalar_mx A. Proof. rewrite /is_scalar_mx; case: (insub _) => // i. -by rewrite mxE -map_scalar_mx inj_eq //; exact: map_mx_inj. +by rewrite mxE -map_scalar_mx inj_eq //; apply: map_mx_inj. Qed. Lemma map_unitmx n (A : 'M_n) : (A^f \in unitmx) = (A \in unitmx). @@ -2829,7 +2828,7 @@ Proof. elim: n => [|n IHn] /= in A *; first exact: is_perm_mx1. set A' := _ - _; move/(_ A'): IHn; case: cormen_lup => [[P L U]] {A'}/=. rewrite (is_perm_mxMr _ (perm_mx_is_perm _ _)). -case/is_perm_mxP => s ->; exact: lift0_mx_is_perm. +by case/is_perm_mxP => s ->; apply: lift0_mx_is_perm. Qed. Lemma cormen_lup_correct n (A : 'M_n.+1) : @@ -2861,7 +2860,7 @@ Proof. elim: n => [|n IHn] /= in A i j *; first by rewrite [i]ord1 [j]ord1 mxE. set A' := _ - _; move/(_ A'): IHn; case: cormen_lup => [[P L U]] {A'}/= Ll. rewrite !mxE split1; case: unliftP => [i'|] -> /=; rewrite !mxE split1. - by case: unliftP => [j'|] -> //; exact: Ll. + by case: unliftP => [j'|] -> //; apply: Ll. by case: unliftP => [j'|] ->; rewrite /= mxE. Qed. @@ -2871,7 +2870,7 @@ Proof. elim: n => [|n IHn] /= in A i j *; first by rewrite [i]ord1. set A' := _ - _; move/(_ A'): IHn; case: cormen_lup => [[P L U]] {A'}/= Uu. rewrite !mxE split1; case: unliftP => [i'|] -> //=; rewrite !mxE split1. -by case: unliftP => [j'|] ->; [exact: Uu | rewrite /= mxE]. +by case: unliftP => [j'|] ->; [apply: Uu | rewrite /= mxE]. Qed. End CormenLUP. diff --git a/mathcomp/algebra/mxalgebra.v b/mathcomp/algebra/mxalgebra.v index cb1717c..bb5caab 100644 --- a/mathcomp/algebra/mxalgebra.v +++ b/mathcomp/algebra/mxalgebra.v @@ -1,12 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype finfun bigop prime binomial finset. -From mathcomp.fingroup -Require Import fingroup perm. -Require Import ssralg finalg zmodp matrix. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div choice fintype. +From mathcomp +Require Import finfun bigop prime binomial ssralg finset fingroup finalg. +From mathcomp +Require Import perm zmodp matrix. (*****************************************************************************) (* In this file we develop the rank and row space theory of matrices, based *) @@ -440,7 +439,7 @@ Proof. set MN := M *m N; set rMN := \rank _. pose L : 'M_(rMN, m) := pid_mx rMN *m invmx (col_ebase MN). pose U : 'M_(n, rMN) := invmx (row_ebase MN) *m pid_mx rMN. -suffices: L *m M *m (N *m U) = 1%:M by exact: mulmx1_min. +suffices: L *m M *m (N *m U) = 1%:M by apply: mulmx1_min. rewrite mulmxA -(mulmxA L) -[M *m N]mulmx_ebase -/MN. by rewrite !mulmxA mulmxKV // mulmxK // !pid_mx_id /rMN ?pid_mx_1. Qed. @@ -531,14 +530,14 @@ Proof. by case/submxP=> D ->; rewrite mulmxA submxMl. Qed. Lemma submx_trans m1 m2 m3 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) : (A <= B -> B <= C -> A <= C)%MS. -Proof. by case/submxP=> D ->{A}; exact: mulmx_sub. Qed. +Proof. by case/submxP=> D ->{A}; apply: mulmx_sub. Qed. Lemma ltmx_sub_trans m1 m2 m3 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) : (A < B)%MS -> (B <= C)%MS -> (A < C)%MS. Proof. case/andP=> sAB ltAB sBC; rewrite ltmxE (submx_trans sAB) //. -by apply: contra ltAB; exact: submx_trans. +by apply: contra ltAB; apply: submx_trans. Qed. Lemma sub_ltmx_trans m1 m2 m3 n @@ -546,11 +545,11 @@ Lemma sub_ltmx_trans m1 m2 m3 n (A <= B)%MS -> (B < C)%MS -> (A < C)%MS. Proof. move=> sAB /andP[sBC ltBC]; rewrite ltmxE (submx_trans sAB) //. -by apply: contra ltBC => sCA; exact: submx_trans sAB. +by apply: contra ltBC => sCA; apply: submx_trans sAB. Qed. Lemma ltmx_trans m n : transitive (@ltmx m m n). -Proof. by move=> A B C; move/ltmxW; exact: sub_ltmx_trans. Qed. +Proof. by move=> A B C; move/ltmxW; apply: sub_ltmx_trans. Qed. Lemma ltmx_irrefl m n : irreflexive (@ltmx m m n). Proof. by move=> A; rewrite /ltmx submx_refl andbF. Qed. @@ -563,7 +562,7 @@ Lemma submx0null m1 m2 n (A : 'M[F]_(m1, n)) : Proof. by case/submxP=> D; rewrite mulmx0. Qed. Lemma submx0 m n (A : 'M_(m, n)) : (A <= (0 : 'M_n))%MS = (A == 0). -Proof. by apply/idP/eqP=> [|->]; [exact: submx0null | exact: sub0mx]. Qed. +Proof. by apply/idP/eqP=> [|->]; [apply: submx0null | apply: sub0mx]. Qed. Lemma lt0mx m n (A : 'M_(m, n)) : ((0 : 'M_n) < A)%MS = (A != 0). Proof. by rewrite /ltmx sub0mx submx0. Qed. @@ -572,7 +571,7 @@ Lemma ltmx0 m n (A : 'M[F]_(m, n)) : (A < (0 : 'M_n))%MS = false. Proof. by rewrite /ltmx sub0mx andbF. Qed. Lemma eqmx0P m n (A : 'M_(m, n)) : reflect (A = 0) (A == (0 : 'M_n))%MS. -Proof. by rewrite submx0 sub0mx andbT; exact: eqP. Qed. +Proof. by rewrite submx0 sub0mx andbT; apply: eqP. Qed. Lemma eqmx_eq0 m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : (A :=: B)%MS -> (A == 0) = (B == 0). @@ -588,7 +587,7 @@ Lemma summx_sub m1 m2 n (B : 'M_(m2, n)) I (r : seq I) (P : pred I) (A_ : I -> 'M_(m1, n)) : (forall i, P i -> A_ i <= B)%MS -> ((\sum_(i <- r | P i) A_ i)%R <= B)%MS. Proof. -move=> leAB; elim/big_ind: _ => // [|A1 A2]; [exact: sub0mx | exact: addmx_sub]. +by move=> leAB; elim/big_ind: _ => // [|C D]; [apply/sub0mx | apply/addmx_sub]. Qed. Lemma scalemx_sub m1 m2 n a (A : 'M_(m1, n)) (B : 'M_(m2, n)) : @@ -608,7 +607,7 @@ Lemma row_subP m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : reflect (forall i, row i A <= B)%MS (A <= B)%MS. Proof. apply: (iffP idP) => [sAB i|sAB]. - by apply: submx_trans sAB; exact: row_sub. + by apply: submx_trans sAB; apply: row_sub. rewrite submxE; apply/eqP/row_matrixP=> i; apply/eqP. by rewrite row_mul row0 -submxE. Qed. @@ -624,7 +623,7 @@ Implicit Arguments rV_subP [m1 m2 n A B]. Lemma row_subPn m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : reflect (exists i, ~~ (row i A <= B)%MS) (~~ (A <= B)%MS). -Proof. by rewrite (sameP row_subP forallP) negb_forall; exact: existsP. Qed. +Proof. by rewrite (sameP row_subP forallP) negb_forall; apply: existsP. Qed. Lemma sub_rVP n (u v : 'rV_n) : reflect (exists a, u = a *: v) (u <= v)%MS. Proof. @@ -636,14 +635,14 @@ Qed. Lemma rank_rV n (v : 'rV_n) : \rank v = (v != 0). Proof. case: eqP => [-> | nz_v]; first by rewrite mxrank0. -by apply/eqP; rewrite eqn_leq rank_leq_row lt0n mxrank_eq0; exact/eqP. +by apply/eqP; rewrite eqn_leq rank_leq_row lt0n mxrank_eq0; apply/eqP. Qed. Lemma rowV0Pn m n (A : 'M_(m, n)) : reflect (exists2 v : 'rV_n, v <= A & v != 0)%MS (A != 0). Proof. rewrite -submx0; apply: (iffP idP) => [| [v svA]]; last first. - by rewrite -submx0; exact: contra (submx_trans _). + by rewrite -submx0; apply: contra (submx_trans _). by case/row_subPn=> i; rewrite submx0; exists (row i A); rewrite ?row_sub. Qed. @@ -651,7 +650,7 @@ Lemma rowV0P m n (A : 'M_(m, n)) : reflect (forall v : 'rV_n, v <= A -> v = 0)%MS (A == 0). Proof. rewrite -[A == 0]negbK; case: rowV0Pn => IH. - by right; case: IH => v svA nzv IH; case/eqP: nzv; exact: IH. + by right; case: IH => v svA nzv IH; case/eqP: nzv; apply: IH. by left=> v svA; apply/eqP; apply/idPn=> nzv; case: IH; exists v. Qed. @@ -703,7 +702,7 @@ Lemma mxrank_unit n (A : 'M_n) : A \in unitmx -> \rank A = n. Proof. by rewrite -row_full_unit =>/eqnP. Qed. Lemma mxrank1 n : \rank (1%:M : 'M_n) = n. -Proof. by apply: mxrank_unit; exact: unitmx1. Qed. +Proof. by apply: mxrank_unit; apply: unitmx1. Qed. Lemma mxrank_delta m n i j : \rank (delta_mx i j : 'M_(m, n)) = 1%N. Proof. @@ -736,8 +735,8 @@ Lemma eqmxP m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : Proof. apply: (iffP andP) => [[sAB sBA] | eqAB]; last by rewrite !eqAB. split=> [|m3 C]; first by apply/eqP; rewrite eqn_leq !mxrankS. -split; first by apply/idP/idP; exact: submx_trans. -by apply/idP/idP=> sC; exact: submx_trans sC _. +split; first by apply/idP/idP; apply: submx_trans. +by apply/idP/idP=> sC; apply: submx_trans sC _. Qed. Implicit Arguments eqmxP [m1 m2 n A B]. @@ -872,7 +871,7 @@ by rewrite qidmx_eq1 row_full_unit unitmx1 => /eqP. Qed. Lemma genmx_id m n (A : 'M_(m, n)) : (<<<<A>>>> = <<A>>)%MS. -Proof. by apply: eq_genmx; exact: genmxE. Qed. +Proof. by apply: eq_genmx; apply: genmxE. Qed. Lemma row_base_free m n (A : 'M_(m, n)) : row_free (row_base A). Proof. by apply/eqnP; rewrite eq_row_base. Qed. @@ -899,7 +898,7 @@ Qed. Lemma mxrank_leqif_eq m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : (A <= B)%MS -> \rank A <= \rank B ?= iff (A == B)%MS. -Proof. by move=> sAB; rewrite sAB; exact: mxrank_leqif_sup. Qed. +Proof. by move=> sAB; rewrite sAB; apply: mxrank_leqif_sup. Qed. Lemma ltmxErank m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : (A < B)%MS = (A <= B)%MS && (\rank A < \rank B). @@ -1044,7 +1043,7 @@ Qed. Lemma addsmx_addKr n m1 m2 (A B : 'M_(m1, n)) (C : 'M_(m2, n)) : (B <= C)%MS -> ((A + B)%R + C :=: A + C)%MS. -Proof. by rewrite -!(addsmxC C) addrC; exact: addsmx_addKl. Qed. +Proof. by rewrite -!(addsmxC C) addrC; apply: addsmx_addKl. Qed. Lemma adds_eqmx m1 m2 m3 m4 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) (D : 'M_(m4, n)) : @@ -1114,7 +1113,7 @@ Lemma sub_sumsmxP P m n (A : 'M_(m, n)) (B_ : I -> 'M_n) : (A <= \sum_(i | P i) B_ i)%MS. Proof. apply: (iffP idP) => [| [u_ ->]]; last first. - by apply: summx_sub_sums => i _; exact: submxMl. + by apply: summx_sub_sums => i _; apply: submxMl. elim: {P}_.+1 {-2}P A (ltnSn #|P|) => // b IHb P A. case: (pickP P) => [i Pi | P0 _]; last first. rewrite big_pred0 //; move/submx0null->. @@ -1137,7 +1136,7 @@ Qed. Lemma sumsmxMr P n (A_ : I -> 'M[F]_n) (B : 'M_n) : ((\sum_(i | P i) A_ i)%MS *m B :=: \sum_(i | P i) (A_ i *m B))%MS. Proof. -by apply: eqmx_trans (sumsmxMr_gen _ _ _) (eqmx_sums _) => i _; exact: genmxE. +by apply: eqmx_trans (sumsmxMr_gen _ _ _) (eqmx_sums _) => i _; apply: genmxE. Qed. Lemma rank_pid_mx m n r : r <= m -> r <= n -> \rank (pid_mx r : 'M_(m, n)) = r. @@ -1194,7 +1193,7 @@ Lemma mulmx0_rank_max m n p (A : 'M_(m, n)) (B : 'M_(n, p)) : A *m B = 0 -> \rank A + \rank B <= n. Proof. move=> AB0; rewrite -{3}(subnK (rank_leq_row B)) leq_add2r. -rewrite -mxrank_ker mxrankS //; exact/sub_kermxP. +by rewrite -mxrank_ker mxrankS //; apply/sub_kermxP. Qed. Lemma mxrank_Frobenius m n p q (A : 'M_(m, n)) B (C : 'M_(p, q)) : @@ -1261,14 +1260,14 @@ move=> eqABid /eqmxP eqAB. have{eqABid eqAB} eqAB: equivmx A (qidmx A) =1 equivmx B (qidmx B). by move=> C; rewrite /equivmx eqABid !eqAB. rewrite {1}/capmx_norm (eq_choose eqAB). -by apply: choose_id; first rewrite -eqAB; exact: capmx_witnessP. +by apply: choose_id; first rewrite -eqAB; apply: capmx_witnessP. Qed. Let capmx_nopP m n (A : 'M_(m, n)) : equivmx_spec A (qidmx A) (capmx_nop A). Proof. rewrite /capmx_nop; case: (eqVneq m n) => [-> | ne_mn] in A *. by rewrite conform_mx_id. -rewrite nonconform_mx ?ne_mn //; exact: capmx_normP. +by rewrite nonconform_mx ?ne_mn //; apply: capmx_normP. Qed. Let sub_qidmx m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : @@ -1340,7 +1339,7 @@ Qed. Lemma capmx_idPl n m1 m2 (A : 'M_(m1, n)) (B : 'M_(m2, n)) : reflect (A :&: B :=: A)%MS (A <= B)%MS. -Proof. by rewrite capmxC; exact: capmx_idPr. Qed. +Proof. by rewrite capmxC; apply: capmx_idPr. Qed. Lemma capmxS m1 m2 m3 m4 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) (D : 'M_(m4, n)) : @@ -1449,7 +1448,7 @@ Qed. Lemma matrix_modr m1 m2 m3 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) (C : 'M_(m3, n)) : (C <= A -> (A :&: B) + C :=: A :&: (B + C))%MS. -Proof. by rewrite !(capmxC A) -!(addsmxC C); exact: matrix_modl. Qed. +Proof. by rewrite !(capmxC A) -!(addsmxC C); apply: matrix_modl. Qed. Lemma capmx_compl m n (A : 'M_(m, n)) : (A :&: A^C)%MS = 0. Proof. @@ -1477,7 +1476,7 @@ Lemma mxrank_injP m n p (A : 'M_(m, n)) (f : 'M_(n, p)) : reflect (\rank (A *m f) = \rank A) ((A :&: kermx f)%MS == 0). Proof. rewrite -mxrank_eq0 -(eqn_add2l (\rank (A *m f))). -by rewrite mxrank_mul_ker addn0 eq_sym; exact: eqP. +by rewrite mxrank_mul_ker addn0 eq_sym; apply: eqP. Qed. Lemma mxrank_disjoint_sum m1 m2 n (A : 'M_(m1, n)) (B : 'M_(m2, n)) : @@ -1828,7 +1827,7 @@ Qed. Lemma mxdirect_addsP (A : 'M_(m1, n)) (B : 'M_(m2, n)) : reflect (A :&: B = 0)%MS (mxdirect (A + B)). -Proof. by rewrite mxdirect_addsE !mxdirect_trivial; exact: eqP. Qed. +Proof. by rewrite mxdirect_addsE !mxdirect_trivial; apply: eqP. Qed. End BinaryDirect. @@ -1870,7 +1869,7 @@ Lemma mxdirect_sumsE (S_ : I -> mxsum_expr n n) (xunwrap := unwrap) : Proof. apply: (iffP (mxdirect_sums_recP _)) => [dxS | [dxS_ dxS] i Pi]. by do [split; last apply/mxdirect_sumsP] => i; case/dxS. -by split; [exact: dxS_ | exact: mxdirect_sumsP Pi]. +by split; [apply: dxS_ | apply: mxdirect_sumsP Pi]. Qed. End NaryDirect. @@ -1937,7 +1936,7 @@ Lemma eigenspaceP a m (W : 'M_(m, n)) : reflect (W *m g = a *: W) (W <= eigenspace a)%MS. Proof. rewrite (sameP (sub_kermxP _ _) eqP). -by rewrite mulmxBr subr_eq0 mul_mx_scalar; exact: eqP. +by rewrite mulmxBr subr_eq0 mul_mx_scalar; apply: eqP. Qed. Lemma eigenvalueP a : @@ -1961,7 +1960,7 @@ have nz_aij: a_ i - a_ j != 0. case: (sub_dsumsmx dxVi' (sub0mx 1 _)) => C _ _ uniqC. rewrite -(eqmx_eq0 (eqmx_scale _ nz_aij)). rewrite (uniqC (fun k => (a_ i - a_ k) *: u k)) => // [|k Pi'k|]. -- by rewrite -(uniqC (fun _ => 0)) ?big1 // => k Pi'k; exact: sub0mx. +- by rewrite -(uniqC (fun _ => 0)) ?big1 // => k Pi'k; apply: sub0mx. - by rewrite scalemx_sub ?Vi'u. rewrite -{1}(subrr (v *m g)) {1}def_vg def_v scaler_sumr mulmx_suml -sumrB. by apply: eq_bigr => k /Vi'u/eigenspaceP->; rewrite scalerBl. @@ -2220,7 +2219,7 @@ Lemma memmx_subP m1 m2 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) : reflect (forall A, A \in R1 -> A \in R2) (R1 <= R2)%MS. Proof. apply: (iffP idP) => [sR12 A R1_A | sR12]; first exact: submx_trans sR12. -by apply/rV_subP=> vA; rewrite -(vec_mxK vA); exact: sR12. +by apply/rV_subP=> vA; rewrite -(vec_mxK vA); apply: sR12. Qed. Implicit Arguments memmx_subP [m1 m2 n R1 R2]. @@ -2270,7 +2269,7 @@ rewrite -{2}rankS (ltn_leqif (mxrank_leqif_sup sSR)). apply: (iffP idP) => [/row_subPn[i] | [A sAR]]. rewrite -[row i R]vec_mxK memmx1; set A := vec_mx _ => nsA. by exists A; rewrite // vec_mxK row_sub. -by rewrite -memmx1; apply: contra; exact: submx_trans. +by rewrite -memmx1; apply/contra/submx_trans. Qed. Definition mulsmx m1 m2 n (R1 : 'A[F]_(m1, n)) (R2 : 'A_(m2, n)) := @@ -2313,7 +2312,7 @@ Lemma mulsmxS m1 m2 m3 m4 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) (R1 <= R3 -> R2 <= R4 -> R1 * R2 <= R3 * R4)%MS. Proof. move=> sR13 sR24; apply/mulsmx_subP=> A1 A2 R_A1 R_A2. -by apply: mem_mulsmx; [exact: submx_trans sR13 | exact: submx_trans sR24]. +by apply: mem_mulsmx; [apply: submx_trans sR13 | apply: submx_trans sR24]. Qed. Lemma muls_eqmx m1 m2 m3 m4 n (R1 : 'A_(m1, n)) (R2 : 'A_(m2, n)) @@ -2481,7 +2480,7 @@ Qed. Implicit Arguments cent_mxP [m n B R]. Lemma scalar_mx_cent m n a (R : 'A_(m, n)) : (a%:M \in 'C(R))%MS. -Proof. by apply/cent_mxP=> A _; exact: scalar_mxC. Qed. +Proof. by apply/cent_mxP=> A _; apply: scalar_mxC. Qed. Lemma center_mx_sub m n (R : 'A_(m, n)) : ('Z(R) <= R)%MS. Proof. exact: capmxSl. Qed. @@ -2762,7 +2761,7 @@ by rewrite map_mx_sub ? map_mxM. Qed. Lemma map_center_mx m n (E : 'A_(m, n)) : (('Z(E))^f :=: 'Z(E^f))%MS. -Proof. by rewrite /center_mx -map_cent_mx; exact: map_capmx. Qed. +Proof. by rewrite /center_mx -map_cent_mx; apply: map_capmx. Qed. End MapMatrixSpaces. diff --git a/mathcomp/algebra/mxpoly.v b/mathcomp/algebra/mxpoly.v index bc5a998..7071687 100644 --- a/mathcomp/algebra/mxpoly.v +++ b/mathcomp/algebra/mxpoly.v @@ -1,12 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import div fintype tuple finfun bigop. -From mathcomp.fingroup -Require Import fingroup perm. -Require Import ssralg zmodp matrix mxalgebra poly polydiv. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq div fintype tuple. +From mathcomp +Require Import finfun bigop fingroup perm ssralg zmodp matrix mxalgebra. +From mathcomp +Require Import poly polydiv. (******************************************************************************) (* This file provides basic support for formal computation with matrices, *) @@ -384,7 +383,7 @@ pose Msize (A : M_RX) := \max_i \max_j size (A i j). pose phi (A : M_RX) := \poly_(k < Msize A) \matrix_(i, j) (A i j)`_k. have coef_phi A i j k: (phi A)`_k i j = (A i j)`_k. rewrite coef_poly; case: (ltnP k _) => le_m_k; rewrite mxE // nth_default //. - apply: leq_trans (leq_trans (leq_bigmax i) le_m_k); exact: (leq_bigmax j). + by apply: leq_trans (leq_trans (leq_bigmax i) le_m_k); apply: (leq_bigmax j). have phi_is_rmorphism : rmorphism phi. do 2?[split=> [A B|]]; apply/polyP=> k; apply/matrixP=> i j; last 1 first. - rewrite coef_phi mxE coefMn !coefC. @@ -496,7 +495,7 @@ Qed. Lemma minpoly_mx_ring : mxring Ad. Proof. -apply/andP; split; first by apply/mulsmx_subP; exact: minpoly_mxM. +apply/andP; split; first by apply/mulsmx_subP; apply: minpoly_mxM. apply/mxring_idP; exists 1%:M; split=> *; rewrite ?mulmx1 ?mul1mx //. by rewrite -mxrank_eq0 mxrank1. exact: minpoly_mx1. @@ -562,7 +561,7 @@ by rewrite -rmorphM horner_mx_C -rmorphD /= scalar_mx_is_scalar. Qed. Lemma mxminpoly_dvd_char : p_A %| char_poly A. -Proof. by apply: mxminpoly_min; exact: Cayley_Hamilton. Qed. +Proof. by apply: mxminpoly_min; apply: Cayley_Hamilton. Qed. Lemma eigenvalue_root_min a : eigenvalue A a = root p_A a. Proof. @@ -784,7 +783,7 @@ have{mon_p pw0 intRp intRq}: genI S. split; set S1 := _ ++ _; first exists p. split=> // i; rewrite -[p]coefK coef_poly; case: ifP => // lt_i_p. by apply: genS; rewrite mem_cat orbC mem_nth. - have: all (mem S1) S1 by exact/allP. + have: all (mem S1) S1 by apply/allP. elim: {-1}S1 => //= y S2 IH /andP[S1y S12]; split; last exact: IH. have{q S S1 IH S1y S12 intRp intRq} [q mon_q qx0]: integralOver RtoK y. by move: S1y; rewrite mem_cat => /orP[]; [apply: intRq | apply: intRp]. diff --git a/mathcomp/algebra/poly.v b/mathcomp/algebra/poly.v index 2757b30..9b2366e 100644 --- a/mathcomp/algebra/poly.v +++ b/mathcomp/algebra/poly.v @@ -1,11 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype bigop binomial. -Require Import ssralg. - +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div choice fintype. +From mathcomp +Require Import bigop ssralg binomial. (******************************************************************************) (* This file provides a library for univariate polynomials over ring *) @@ -195,7 +193,7 @@ Proof. split=> [eq_pq | -> //]; apply: poly_inj. without loss lt_pq: p q eq_pq / size p < size q. move=> IH; case: (ltngtP (size p) (size q)); try by move/IH->. - move/(@eq_from_nth _ 0); exact. + by move/(@eq_from_nth _ 0); apply. case: q => q nz_q /= in lt_pq eq_pq *; case/eqP: nz_q. by rewrite (last_nth 0) -(subnKC lt_pq) /= -eq_pq nth_default ?leq_addr. Qed. @@ -236,7 +234,7 @@ by rewrite !polyseqC !eqxx nz_c. Qed. Lemma polyseqK p : Poly p = p. -Proof. by apply: poly_inj; exact: PolyK (valP p). Qed. +Proof. by apply: poly_inj; apply: PolyK (valP p). Qed. Lemma size_Poly s : size (Poly s) <= size s. Proof. @@ -364,7 +362,7 @@ Lemma nil_poly p : nilp p = (p == 0). Proof. exact: size_poly_eq0. Qed. Lemma poly0Vpos p : {p = 0} + {size p > 0}. -Proof. by rewrite lt0n size_poly_eq0; exact: eqVneq. Qed. +Proof. by rewrite lt0n size_poly_eq0; apply: eqVneq. Qed. Lemma polySpred p : p != 0 -> size p = (size p).-1.+1. Proof. by rewrite -size_poly_eq0 -lt0n => /prednK. Qed. @@ -612,7 +610,7 @@ Lemma polyC_mul : {morph polyC : a b / a * b}. Proof. by move=> a b; apply/polyP=> [[|i]]; rewrite coefCM !coefC ?simp. Qed. Fact polyC_multiplicative : multiplicative polyC. -Proof. by split; first exact: polyC_mul. Qed. +Proof. by split; first apply: polyC_mul. Qed. Canonical polyC_rmorphism := AddRMorphism polyC_multiplicative. Lemma polyC_exp n : {morph polyC : c / c ^+ n}. @@ -744,7 +742,7 @@ Lemma poly_ind (K : {poly R} -> Type) : K 0 -> (forall p c, K p -> K (p * 'X + c%:P)) -> (forall p, K p). Proof. move=> K0 Kcons p; rewrite -[p]polyseqK. -elim: {p}(p : seq R) => //= p c IHp; rewrite cons_poly_def; exact: Kcons. +by elim: {p}(p : seq R) => //= p c IHp; rewrite cons_poly_def; apply: Kcons. Qed. Lemma polyseqXsubC a : 'X - a%:P = [:: - a; 1] :> seq R. @@ -804,7 +802,7 @@ Lemma size_polyXn n : size 'X^n = n.+1. Proof. by rewrite polyseqXn size_rcons size_nseq. Qed. Lemma commr_polyXn p n : GRing.comm p 'X^n. -Proof. by apply: commrX; exact: commr_polyX. Qed. +Proof. by apply: commrX; apply: commr_polyX. Qed. Lemma lead_coefXn n : lead_coef 'X^n = 1. Proof. by rewrite /lead_coef nth_last polyseqXn last_rcons. Qed. @@ -848,7 +846,7 @@ Lemma monicX : 'X \is monic. Proof. exact/eqP/lead_coefX. Qed. Lemma monicXn n : 'X^n \is monic. Proof. exact/eqP/lead_coefXn. Qed. Lemma monic_neq0 p : p \is monic -> p != 0. -Proof. by rewrite -lead_coef_eq0 => /eqP->; exact: oner_neq0. Qed. +Proof. by rewrite -lead_coef_eq0 => /eqP->; apply: oner_neq0. Qed. Lemma lead_coef_monicM p q : p \is monic -> lead_coef (p * q) = lead_coef q. Proof. @@ -898,7 +896,7 @@ Proof. exact/eqP/lead_coefXsubC. Qed. Lemma monic_prod_XsubC I rI (P : pred I) (F : I -> R) : \prod_(i <- rI | P i) ('X - (F i)%:P) \is monic. -Proof. by apply: monic_prod => i _; exact: monicXsubC. Qed. +Proof. by apply: monic_prod => i _; apply: monicXsubC. Qed. Lemma size_prod_XsubC I rI (F : I -> R) : size (\prod_(i <- rI) ('X - (F i)%:P)) = (size rI).+1. @@ -968,7 +966,7 @@ Qed. Lemma monic_comreg p : p \is monic -> GRing.comm p (lead_coef p)%:P /\ GRing.rreg (lead_coef p). -Proof. by move/monicP->; split; [exact: commr1 | exact: rreg1]. Qed. +Proof. by move/monicP->; split; [apply: commr1 | apply: rreg1]. Qed. (* Horner evaluation of polynomials *) Implicit Types s rs : seq R. @@ -1191,7 +1189,7 @@ by rewrite /root_of_unity rootE hornerD hornerN hornerXn hornerC subr_eq0. Qed. Lemma unity_rootP n z : reflect (z ^+ n = 1) (n.-unity_root z). -Proof. by rewrite unity_rootE; exact: eqP. Qed. +Proof. by rewrite unity_rootE; apply: eqP. Qed. Definition primitive_root_of_unity n z := (n > 0) && [forall i : 'I_n, i.+1.-unity_root z == (i.+1 == n)]. @@ -1257,7 +1255,7 @@ set q := (n %/ d)%N; rewrite /q.-primitive_root ltn_divRL // n_gt0. apply/forallP=> i; rewrite unity_rootE -exprM -prim_order_dvd. rewrite -(divnK d_dv_n) -/q -(divnK d_dv_k) mulnAC dvdn_pmul2r //. apply/eqP; apply/idP/idP=> [|/eqP->]; last by rewrite dvdn_mull. -rewrite Gauss_dvdr; first by rewrite eqn_leq ltn_ord; exact: dvdn_leq. +rewrite Gauss_dvdr; first by rewrite eqn_leq ltn_ord; apply: dvdn_leq. by rewrite /coprime gcdnC -(eqn_pmul2r d_gt0) mul1n muln_gcdl !divnK. Qed. @@ -1326,7 +1324,7 @@ by rewrite qualifE polyseqC; case: eqP => [->|] /=; rewrite ?andbT ?rpred0. Qed. Fact polyOver_addr_closed : addr_closed (polyOver kS). -Proof. +Proof. split=> [|p q Sp Sq]; first exact: polyOver0. by apply/polyOverP=> i; rewrite coefD rpredD ?(polyOverP _). Qed. @@ -1356,7 +1354,7 @@ Canonical polyOver_semiringPred := SemiringPred polyOver_mulr_closed. Lemma polyOverZ : {in kS & polyOver kS, forall c p, c *: p \is a polyOver kS}. Proof. -by move=> c p Sc /polyOverP Sp; apply/polyOverP=> i; rewrite coefZ rpredM ?Sp. +by move=> c p Sc /polyOverP Sp; apply/polyOverP=> i; rewrite coefZ rpredM ?Sp. Qed. Lemma polyOverX : 'X \in polyOver kS. @@ -1461,7 +1459,7 @@ by rewrite deriv_mulC IHp !mulrDl -!mulrA !commr_polyX !addrA. Qed. Definition derivE := Eval lazy beta delta [morphism_2 morphism_1] in - (derivZ, deriv_mulC, derivC, derivX, derivMXaddC, derivXsubC, derivM, derivB, + (derivZ, deriv_mulC, derivC, derivX, derivMXaddC, derivXsubC, derivM, derivB, derivD, derivN, derivXn, derivM, derivMn). (* Iterated derivative. *) @@ -1544,7 +1542,7 @@ by apply: leq_trans le_p_n _; apply leq_addr. Qed. Lemma lt_size_deriv (p : {poly R}) : p != 0 -> size p^`() < size p. -Proof. by move=> /polySpred->; exact: size_poly. Qed. +Proof. by move=> /polySpred->; apply: size_poly. Qed. (* A normalising version of derivation to get the division by n! in Taylor *) @@ -2035,10 +2033,10 @@ Canonical polynomial_algType := Eval hnf in [algType R of polynomial R for poly_algType]. Lemma hornerM p q x : (p * q).[x] = p.[x] * q.[x]. -Proof. by rewrite hornerM_comm //; exact: mulrC. Qed. +Proof. by rewrite hornerM_comm //; apply: mulrC. Qed. Lemma horner_exp p x n : (p ^+ n).[x] = p.[x] ^+ n. -Proof. by rewrite horner_exp_comm //; exact: mulrC. Qed. +Proof. by rewrite horner_exp_comm //; apply: mulrC. Qed. Lemma horner_prod I r (P : pred I) (F : I -> {poly R}) x : (\prod_(i <- r | P i) F i).[x] = \prod_(i <- r | P i) (F i).[x]. @@ -2053,7 +2051,7 @@ Lemma horner_evalE x p : horner_eval x p = p.[x]. Proof. by []. Qed. Fact horner_eval_is_lrmorphism x : lrmorphism_for *%R (horner_eval x). Proof. -have cxid: commr_rmorph idfun x by exact: mulrC. +have cxid: commr_rmorph idfun x by apply: mulrC. have evalE : horner_eval x =1 horner_morph cxid. by move=> p; congr _.[x]; rewrite map_poly_id. split=> [|c p]; last by rewrite !evalE /= -linearZ. @@ -2067,7 +2065,7 @@ Canonical horner_eval_lrmorphism x := [lrmorphism of horner_eval x]. Fact comp_poly_multiplicative q : multiplicative (comp_poly q). Proof. split=> [p1 p2|]; last by rewrite comp_polyC. -by rewrite /comp_poly rmorphM hornerM_comm //; exact: mulrC. +by rewrite /comp_poly rmorphM hornerM_comm //; apply: mulrC. Qed. Canonical comp_poly_rmorphism q := AddRMorphism (comp_poly_multiplicative q). Canonical comp_poly_lrmorphism q := [lrmorphism of comp_poly q]. @@ -2281,6 +2279,19 @@ rewrite coefZ (nth_default 0 (leq_trans _ le_qp_i)) ?mulr0 //=. by rewrite polySpred ?expf_neq0 // !size_exp -(subnKC q_gt1) ltn_pmul2l. Qed. +Theorem max_poly_roots p rs : + p != 0 -> all (root p) rs -> uniq rs -> size rs < size p. +Proof. +elim: rs p => [p pn0 _ _ | r rs ihrs p pn0] /=; first by rewrite size_poly_gt0. +case/andP => rpr arrs /andP [rnrs urs]; case/factor_theorem: rpr => q epq. +case: (altP (q =P 0)) => [q0 | ?]; first by move: pn0; rewrite epq q0 mul0r eqxx. +have -> : size p = (size q).+1. + by rewrite epq size_Mmonic ?monicXsubC // size_XsubC addnC. +suff /eq_in_all h : {in rs, root q =1 root p} by apply: ihrs => //; rewrite h. +move=> x xrs; rewrite epq rootM root_XsubC orbC; case: (altP (x =P r)) => // exr. +by move: rnrs; rewrite -exr xrs. +Qed. + End PolynomialIdomain. Section MapFieldPoly. @@ -2400,10 +2411,6 @@ elim: rs => //= r rs ->; congr (_ && _); rewrite -has_pred1 -all_predC. by apply: eq_all => t; rewrite /diff_roots mulrC eqxx unitfE subr_eq0. Qed. -Theorem max_poly_roots p rs : - p != 0 -> all (root p) rs -> uniq rs -> size rs < size p. -Proof. by rewrite -uniq_rootsE; exact: max_ring_poly_roots. Qed. - Section UnityRoots. Variable n : nat. diff --git a/mathcomp/algebra/polyXY.v b/mathcomp/algebra/polyXY.v index d79b2bf..770b4cc 100644 --- a/mathcomp/algebra/polyXY.v +++ b/mathcomp/algebra/polyXY.v @@ -1,12 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import choice div fintype tuple finfun bigop binomial. -From mathcomp.fingroup -Require Import fingroup perm. -Require Import ssralg zmodp matrix mxalgebra poly polydiv mxpoly. +From mathcomp +Require Import ssrfun ssrbool choice eqtype ssrnat seq div fintype. +From mathcomp +Require Import tuple finfun bigop fingroup perm ssralg zmodp matrix mxalgebra. +From mathcomp +Require Import poly polydiv mxpoly binomial. (******************************************************************************) (* This file provides additional primitives and theory for bivariate *) diff --git a/mathcomp/algebra/polydiv.v b/mathcomp/algebra/polydiv.v index 5910586..3c27a2b 100644 --- a/mathcomp/algebra/polydiv.v +++ b/mathcomp/algebra/polydiv.v @@ -1,10 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import bigop choice fintype. -Require Import ssralg poly. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp +Require Import bigop ssralg poly. (******************************************************************************) (* This file provides a library for the basic theory of Euclidean and pseudo- *) @@ -132,7 +131,7 @@ Definition redivp_expanded_def p q := if q == 0 then (0%N, 0, p) else redivp_rec q 0 0 p (size p). Fact redivp_key : unit. Proof. by []. Qed. Definition redivp : {poly R} -> {poly R} -> nat * {poly R} * {poly R} := - locked_with redivp_key redivp_expanded_def. + locked_with redivp_key redivp_expanded_def. Canonical redivp_unlockable := [unlockable fun redivp]. Definition rdivp p q := ((redivp p q).1).2. @@ -154,22 +153,22 @@ Proof. by rewrite /rdivp unlock eqxx. Qed. Lemma rdivp_small p q : size p < size q -> rdivp p q = 0. Proof. -rewrite /rdivp unlock; case: eqP => Eq; first by rewrite Eq size_poly0. -by case sp: (size p) => [| s] hs /=; rewrite sp hs. +rewrite /rdivp unlock; have [-> | _ ltpq] := eqP; first by rewrite size_poly0. +by case: (size p) => [|s]; rewrite /= ltpq. Qed. -Lemma leq_rdivp p q : (size (rdivp p q) <= size p). +Lemma leq_rdivp p q : size (rdivp p q) <= size p. Proof. -case: (ltnP (size p) (size q)); first by move/rdivp_small->; rewrite size_poly0. -rewrite /rdivp /rmodp /rscalp unlock; case q0 : (q == 0) => /=. - by rewrite size_poly0. -have : (size (0 : {poly R})) <= size p by rewrite size_poly0. -move: (leqnn (size p)); move: {2 3 4 6}(size p)=> A. -elim: (size p) 0%N (0 : {poly R}) {1 3 4}p (leqnn (size p)) => [| n ihn] k q1 r. - by move/size_poly_leq0P->; rewrite /= size_poly0 lt0n size_poly_eq0 q0. -move=> /= hrn hr hq1 hq; case: ltnP=> //= hqr. -have sr: 0 < size r by apply: leq_trans hqr; rewrite size_poly_gt0 q0. +have [/rdivp_small->|] := ltnP (size p) (size q); first by rewrite size_poly0. +rewrite /rdivp /rmodp /rscalp unlock. +case q0: (q == 0) => /=; first by rewrite size_poly0. +have: size (0 : {poly R}) <= size p by rewrite size_poly0. +move: (leqnn (size p)); move: {2 3 4 6}(size p) => A. +elim: (size p) 0%N (0 : {poly R}) {1 3 4}p (leqnn (size p)) => [|n ihn] k q1 r. + by move/size_poly_leq0P->; rewrite /= size_poly0 lt0n size_poly_eq0 q0. +move=> /= hrn hr hq1 hq; case: ltnP => //= hqr. have sq: 0 < size q by rewrite size_poly_gt0 q0. +have sr: 0 < size r by apply: leq_trans sq hqr. apply: ihn => //. - apply/leq_sizeP => j hnj. rewrite coefB -scalerAl coefZ coefXnM ltn_subRL ltnNge. @@ -487,7 +486,7 @@ Qed. Lemma rdivp_eq p : p * (lead_coef d ^+ (rscalp p d))%:P = (rdivp p d) * d + (rmodp p d). Proof. -rewrite /rdivp /rmodp /rscalp; case: comm_redivpP=> k q1 r1 Hc _; exact: Hc. +by rewrite /rdivp /rmodp /rscalp; case: comm_redivpP=> k q1 r1 Hc _; apply: Hc. Qed. (* section variables impose an inconvenient order on parameters *) @@ -543,7 +542,7 @@ by rewrite polyC_exp; apply: commrX. Qed. Lemma rdvdpp : rdvdp d d. -Proof. apply/eqP; exact: rmodpp. Qed. +Proof. by apply/eqP; apply: rmodpp. Qed. Lemma rdivpK p : rdvdp d p -> (rdivp p d) * d = p * (lead_coef d ^+ rscalp p d)%:P. @@ -617,12 +616,12 @@ Proof. by rewrite -[p * d]addr0 rdivp_addl_mul rdiv0p addr0. Qed. Lemma rmodp_mull p : rmodp (p * d) d = 0. Proof. -apply: rmodp_mull; rewrite (eqP mond); [exact: commr1 | exact: rreg1]. +by apply: rmodp_mull; rewrite (eqP mond); [apply: commr1 | apply: rreg1]. Qed. Lemma rmodpp : rmodp d d = 0. Proof. -apply: rmodpp; rewrite (eqP mond); [exact: commr1 | exact: rreg1]. +by apply: rmodpp; rewrite (eqP mond); [apply: commr1 | apply: rreg1]. Qed. Lemma rmodp_addl_mul_small q r : @@ -644,12 +643,12 @@ Proof. have -> : rmodp q d = q - (rdivp q d) * d. by rewrite {2}(rdivp_eq q) addrAC subrr add0r. rewrite mulrDr rmodp_add -mulNr mulrA. -rewrite -{2}[rmodp _ _]addr0; congr (_ + _); exact: rmodp_mull. +by rewrite -{2}[rmodp _ _]addr0; congr (_ + _); apply: rmodp_mull. Qed. Lemma rdvdpp : rdvdp d d. Proof. -apply: rdvdpp; rewrite (eqP mond); [exact: commr1 | exact: rreg1]. +by apply: rdvdpp; rewrite (eqP mond); [apply: commr1 | apply: rreg1]. Qed. (* section variables impose an inconvenient order on parameters *) @@ -664,7 +663,7 @@ Qed. Lemma rdvdp_mull p : rdvdp d (p * d). Proof. -apply: rdvdp_mull; rewrite (eqP mond) //; [exact: commr1 | exact: rreg1]. +by apply: rdvdp_mull; rewrite (eqP mond) //; [apply: commr1 | apply: rreg1]. Qed. Lemma rdvdpP p : reflect (exists qq, p = qq * d) (rdvdp d p). @@ -699,7 +698,7 @@ Lemma rdvdp_XsubCl p x : rdvdp ('X - x%:P) p = root p x. Proof. have [HcX Hr] := (monic_comreg (monicXsubC x)). apply/rmodp_eq0P/factor_theorem; last first. - case=> p1 ->; apply: rmodp_mull; exact: monicXsubC. + by case=> p1 ->; apply: rmodp_mull; apply: monicXsubC. move=> e0; exists (rdivp p ('X - x%:P)). by rewrite {1}(rdivp_eq (monicXsubC x) p) e0 addr0. Qed. @@ -742,7 +741,7 @@ Qed. Lemma rdivp_eq d p : (lead_coef d ^+ (rscalp p d)) *: p = (rdivp p d) * d + (rmodp p d). Proof. -rewrite /rdivp /rmodp /rscalp; case: redivpP=> k q1 r1 Hc _; exact: Hc. +by rewrite /rdivp /rmodp /rscalp; case: redivpP=> k q1 r1 Hc _; apply: Hc. Qed. Lemma rdvdp_eqP d p : rdvdp_spec p d (rmodp p d) (rdvdp d p). @@ -775,8 +774,8 @@ Lemma uniq_roots_rdvdp p rs : all (root p) rs -> uniq_roots rs -> rdvdp (\prod_(z <- rs) ('X - z%:P)) p. Proof. -move=> rrs; case/(uniq_roots_prod_XsubC rrs)=> q ->; apply: RingMonic.rdvdp_mull. -exact: monic_prod_XsubC. +move=> rrs; case/(uniq_roots_prod_XsubC rrs)=> q ->. +exact/RingMonic.rdvdp_mull/monic_prod_XsubC. Qed. End UnitRingPseudoDivision. @@ -900,7 +899,7 @@ Qed. Lemma edivp_eq d q r : size r < size d -> lead_coef d \in GRing.unit -> edivp (q * d + r) d = (0%N, q, r). Proof. -have hC : GRing.comm d (lead_coef d)%:P by exact: mulrC. +have hC : GRing.comm d (lead_coef d)%:P by apply: mulrC. move=> hsrd hu; rewrite unlock hu; case et: (redivp _ _) => [[s qq] rr]. have cdn0 : lead_coef d != 0. by move: hu; case d0: (lead_coef d == 0) => //; rewrite (eqP d0) unitr0. @@ -973,9 +972,9 @@ rewrite modpE RingComRreg.rmodp_mull ?scaler0 ?if_same //. by apply/rregP; rewrite lead_coef_eq0. Qed. -Lemma mulKp p q : q != 0 -> +Lemma mulKp p q : q != 0 -> q * p %/ q = lead_coef q ^+ scalp (p * q) q *: p. -Proof. move=> ?; rewrite mulrC; exact: mulpK. Qed. +Proof. by move=> nzq; rewrite mulrC; apply: mulpK. Qed. Lemma divpp p : p != 0 -> p %/ p = (lead_coef p ^+ scalp p p)%:P. Proof. @@ -1089,7 +1088,7 @@ Lemma ltn_divpl d q p : d != 0 -> (size (q %/ d) < size p) = (size q < size (p * d)). Proof. move=> dn0; have sd : size d > 0 by rewrite size_poly_gt0 dn0. -have: (lead_coef d) ^+ (scalp q d) != 0 by exact: lc_expn_scalp_neq0. +have: (lead_coef d) ^+ (scalp q d) != 0 by apply: lc_expn_scalp_neq0. move/size_scale; move/(_ q)<-; rewrite divp_eq; case quo0 : (q %/ d == 0). rewrite (eqP quo0) mul0r add0r size_poly0. case p0 : (p == 0); first by rewrite (eqP p0) mul0r size_poly0 ltnn ltn0. @@ -1170,7 +1169,7 @@ Proof. by move=> pq hq; apply: contraL pq=> /eqP ->; rewrite dvd0p. Qed. Lemma dvdp1 d : (d %| 1) = ((size d) == 1%N). Proof. rewrite /dvdp modpE; case ud: (lead_coef d \in GRing.unit); last exact: rdvdp1. -rewrite -size_poly_eq0 size_scale; first by rewrite size_poly_eq0; exact: rdvdp1. +rewrite -size_poly_eq0 size_scale; first by rewrite size_poly_eq0 -rdvdp1. by rewrite invr_eq0 expf_neq0 //; apply: contraTneq ud => ->; rewrite unitr0. Qed. @@ -1284,7 +1283,7 @@ by rewrite mulrDl -!scalerAl -Eq1 -Eq2 !scalerA mulrC addrC scalerDr. Qed. Lemma dvdp_addl n d m : d %| n -> (d %| m + n) = (d %| m). -Proof. by rewrite addrC; exact: dvdp_addr. Qed. +Proof. by rewrite addrC; apply: dvdp_addr. Qed. Lemma dvdp_add d m n : d %| m -> d %| n -> d %| m + n. Proof. by move/dvdp_addr->. Qed. @@ -1334,10 +1333,10 @@ by apply: (@eq_dvdp _ (q2 * q1) _ _ sn0); rewrite -scalerA Hq2 scalerAr Hq1 mulr Qed. Lemma dvdp_mulIl p q : p %| p * q. -Proof. by apply: dvdp_mulr; exact: dvdpp. Qed. +Proof. by apply: dvdp_mulr; apply: dvdpp. Qed. Lemma dvdp_mulIr p q : q %| p * q. -Proof. by apply: dvdp_mull; exact: dvdpp. Qed. +Proof. by apply: dvdp_mull; apply: dvdpp. Qed. Lemma dvdp_mul2r r p q : r != 0 -> (p * r %| q * r) = (p %| q). Proof. @@ -1401,10 +1400,10 @@ by rewrite exprS mulf_neq0. Qed. Lemma dvdp_XsubCl p x : ('X - x%:P) %| p = root p x. -Proof. rewrite dvdpE; exact: Ring.rdvdp_XsubCl. Qed. +Proof. by rewrite dvdpE; apply: Ring.rdvdp_XsubCl. Qed. Lemma polyXsubCP p x : reflect (p.[x] = 0) (('X - x%:P) %| p). -Proof. rewrite dvdpE; exact: Ring.polyXsubCP. Qed. +Proof. by rewrite dvdpE; apply: Ring.polyXsubCP. Qed. Lemma eqp_div_XsubC p c : (p == (p %/ ('X - c%:P)) * ('X - c%:P)) = ('X - c%:P %| p). @@ -1549,15 +1548,13 @@ by rewrite /eqp; case/andP=> dd' d'd dp; apply: (dvdp_trans d'd). Qed. Lemma dvdp_scaler c m n : c != 0 -> m %| c *: n = (m %| n). -Proof. move=> cn0; apply: eqp_dvdr; exact: eqp_scale. Qed. +Proof. by move=> cn0; apply: eqp_dvdr; apply: eqp_scale. Qed. Lemma dvdp_scalel c m n : c != 0 -> (c *: m %| n) = (m %| n). -Proof. move=> cn0; apply: eqp_dvdl; exact: eqp_scale. Qed. +Proof. by move=> cn0; apply: eqp_dvdl; apply: eqp_scale. Qed. Lemma dvdp_opp d p : d %| (- p) = (d %| p). -Proof. -by apply: eqp_dvdr; rewrite -scaleN1r eqp_scale // oppr_eq0 oner_eq0. -Qed. +Proof. by apply: eqp_dvdr; rewrite -scaleN1r eqp_scale ?oppr_eq0 ?oner_eq0. Qed. Lemma eqp_mul2r r p q : r != 0 -> (p * r %= q * r) = (p %= q). Proof. by move => nz_r; rewrite /eqp !dvdp_mul2r. Qed. @@ -1602,7 +1599,7 @@ case (p =P 0)=> [->|]; [|move/eqP => Hp]. move: pq; rewrite dvdp_eq; set c := _ ^+ _; set x := _ %/ _; move/eqP=> eqpq. move:(eqpq); move/(congr1 (size \o (@polyseq R)))=> /=. have cn0 : c != 0 by rewrite expf_neq0 // lead_coef_eq0. -rewrite (@eqp_size _ q); last by exact: eqp_scale. +rewrite (@eqp_size _ q); last by apply: eqp_scale. rewrite size_mul ?p0 // => [-> HH|]; last first. apply/eqP=> HH; move: eqpq; rewrite HH mul0r. by move/eqP; rewrite scale_poly_eq0 (negPf Hq) (negPf cn0). @@ -1814,7 +1811,7 @@ rewrite gcdpE modp_mull gcd0p size_mul //; case: ifP; first by rewrite eqpxx. rewrite (polySpred mn0) addSn /= -{1}[size n]add0n ltn_add2r; move/negbT. rewrite -ltnNge prednK ?size_poly_gt0 // leq_eqVlt ltnS leqn0 size_poly_eq0. rewrite (negPf mn0) orbF; case/size_poly1P=> c cn0 -> {mn0 m}; rewrite mul_polyC. -suff -> : n %% (c *: n) = 0 by rewrite gcd0p; exact: eqp_scale. +suff -> : n %% (c *: n) = 0 by rewrite gcd0p; apply: eqp_scale. by apply/modp_eq0P; rewrite dvdp_scalel. Qed. @@ -1832,7 +1829,7 @@ Qed. Lemma gcdp_scaler c m n : c != 0 -> gcdp m (c *: n) %= gcdp m n. Proof. move=> cn0; apply: eqp_trans (gcdpC _ _) _. -apply: eqp_trans (gcdp_scalel _ _ _) _ => //; exact: gcdpC. +by apply: eqp_trans (gcdp_scalel _ _ _) _ => //; apply: gcdpC. Qed. Lemma dvdp_gcd_idl m n : m %| n -> gcdp m n %= m. @@ -1845,15 +1842,15 @@ by rewrite expf_neq0 // lead_coef_eq0. Qed. Lemma dvdp_gcd_idr m n : n %| m -> gcdp m n %= n. -Proof. move/dvdp_gcd_idl => h; apply: eqp_trans h; exact: gcdpC. Qed. +Proof. by move/dvdp_gcd_idl => h; apply: eqp_trans h; apply: gcdpC. Qed. Lemma gcdp_exp p k l : gcdp (p ^+ k) (p ^+ l) %= p ^+ minn k l. Proof. wlog leqmn: k l / k <= l. move=> hwlog; case: (leqP k l); first exact: hwlog. - move/ltnW; rewrite minnC; move/hwlog=> h; apply: eqp_trans h; exact: gcdpC. + by move/ltnW; rewrite minnC; move/hwlog=> h; apply: eqp_trans h; apply: gcdpC. rewrite (minn_idPl leqmn); move/subnK: leqmn<-; rewrite exprD. -apply: eqp_trans (gcdp_mull _ _) _; exact: eqpxx. +by apply: eqp_trans (gcdp_mull _ _) _; apply: eqpxx. Qed. Lemma gcdp_eq0 p q : gcdp p q == 0 = (p == 0) && (q == 0). @@ -1918,7 +1915,7 @@ Lemma gcdp_def d m n : gcdp m n %= d. Proof. move=> dm dn h; rewrite /eqp dvdp_gcd dm dn !andbT. -apply: h; [exact: dvdp_gcdl | exact: dvdp_gcdr]. +by apply: h; [apply: dvdp_gcdl | apply: dvdp_gcdr]. Qed. Definition coprimep p q := size (gcdp p q) == 1%N. @@ -2019,7 +2016,7 @@ by rewrite -!gcdp_eqp1; move/(eqp_gcdr p) => h1; apply: (eqp_ltrans h1). Qed. Lemma eqp_coprimepl p q r : q %= r -> coprimep q p = coprimep r p. -Proof. rewrite !(coprimep_sym _ p); exact: eqp_coprimepr. Qed. +Proof. by rewrite !(coprimep_sym _ p); apply: eqp_coprimepr. Qed. (* This should be implemented with an extended remainder sequence *) Fixpoint egcdp_rec p q k {struct k} : {poly R} * {poly R} := @@ -2072,7 +2069,7 @@ rewrite gcdpE ltnNge qsp //= (eqp_ltrans (gcdpC _ _)); split; last first. have : size v + size (p %/ q) > 0 by rewrite addn_gt0 size_poly_gt0 vn0. have : size q + size (p %/ q) > 0 by rewrite addn_gt0 size_poly_gt0 qn0. do 2! move/prednK=> {1}<-; rewrite ltnS => h; apply: leq_trans h _. - rewrite size_divp // addnBA; last by apply: leq_trans qsp; exact: leq_pred. + rewrite size_divp // addnBA; last by apply: leq_trans qsp; apply: leq_pred. rewrite addnC -addnBA ?leq_pred //; move: qn0; rewrite -size_poly_eq0 -lt0n. by move/prednK=> {1}<-; rewrite subSnn addn1. - by rewrite size_scale // lc_expn_scalp_neq0. @@ -2110,7 +2107,7 @@ Lemma Bezout_coprimepP : forall p q, reflect (exists u, u.1 * p + u.2 * q %= 1) (coprimep p q). Proof. move=> p q; rewrite -gcdp_eqp1; apply:(iffP idP)=> [g1|]. - case: (Bezoutp p q) => [[u v] Puv]; exists (u, v); exact: eqp_trans g1. + by case: (Bezoutp p q) => [[u v] Puv]; exists (u, v); apply: eqp_trans g1. case=>[[u v]]; rewrite eqp_sym=> Puv; rewrite /eqp (eqp_dvdr _ Puv). by rewrite dvdp_addr dvdp_mull ?dvdp_gcdl ?dvdp_gcdr //= dvd1p. Qed. @@ -2134,7 +2131,7 @@ by rewrite mulrA dvdp_mull ?dvdpp. Qed. Lemma Gauss_dvdpr p q d: coprimep d q -> (d %| q * p) = (d %| p). -Proof. rewrite mulrC; exact: Gauss_dvdpl. Qed. +Proof. by rewrite mulrC; apply: Gauss_dvdpl. Qed. (* This could be simplified with the introduction of lcmp *) Lemma Gauss_dvdp m n p : coprimep m n -> (m * n %| p) = (m %| p) && (n %| p). @@ -2171,12 +2168,12 @@ Proof. by move=> co_pn; rewrite mulrC Gauss_gcdpr. Qed. Lemma coprimep_mulr p q r : coprimep p (q * r) = (coprimep p q && coprimep p r). Proof. -apply/coprimepP/andP=> [hp|[/coprimepP hq hr]]. - split; apply/coprimepP=> d dp dq; rewrite hp //; - [exact: dvdp_mulr|exact: dvdp_mull]. +apply/coprimepP/andP=> [hp | [/coprimepP-hq hr]]. + by split; apply/coprimepP=> d dp dq; rewrite hp //; + [apply/dvdp_mulr | apply/dvdp_mull]. move=> d dp dqr; move/(_ _ dp) in hq. rewrite Gauss_dvdpl in dqr; first exact: hq. -by move/coprimep_dvdr:hr; apply. +by move/coprimep_dvdr: hr; apply. Qed. Lemma coprimep_mull p q r: coprimep (q * r) p = (coprimep q p && coprimep r p). @@ -2203,7 +2200,7 @@ Lemma coprimep_expl k m n : coprimep m n -> coprimep (m ^+ k) n. Proof. by case: k => [|k] co_pm; rewrite ?coprime1p // coprimep_pexpl. Qed. Lemma coprimep_expr k m n : coprimep m n -> coprimep m (n ^+ k). -Proof. by rewrite !(coprimep_sym m); exact: coprimep_expl. Qed. +Proof. by rewrite !(coprimep_sym m); apply: coprimep_expl. Qed. Lemma gcdp_mul2l p q r : gcdp (p * q) (p * r) %= (p * gcdp q r). Proof. @@ -2383,7 +2380,7 @@ have sp' : size (p %/ (gcdp p q)) <= k. by rewrite [_ == _]cop ltnS leqn0 size_poly_eq0 (negPf gn0). case (ihk _ sp')=> r' dr'p'; first rewrite p'n0 orbF=> cr'q maxr'. constructor=> //=; rewrite ?(negPf p0) ?orbF //. - apply: (dvdp_trans dr'p'); apply: divp_dvd; exact: dvdp_gcdl. + exact/(dvdp_trans dr'p')/divp_dvd/dvdp_gcdl. move=> d dp cdq; apply: maxr'; last by rewrite cdq. case dpq: (d %| gcdp p q). move: (dpq); rewrite dvdp_gcd dp /= => dq; apply: dvdUp; move: cdq. @@ -2589,7 +2586,7 @@ Lemma mulpK p : p * q %/ q = p. Proof. by rewrite mulpK ?monic_neq0 // (eqP monq) expr1n scale1r. Qed. Lemma mulKp p : q * p %/ q = p. -Proof. by rewrite mulrC; exact: mulpK. Qed. +Proof. by rewrite mulrC; apply: mulpK. Qed. End MonicDivisor. @@ -2741,7 +2738,7 @@ by rewrite -lead_coef_eq0; apply: contraTneq ulcd => ->; rewrite unitr0. Qed. Lemma mulKp q : (d * q) %/ d = q. -Proof. rewrite mulrC; exact: mulpK. Qed. +Proof. by rewrite mulrC; apply: mulpK. Qed. Lemma divp_addl_mul_small q r : size r < size d -> (q * d + r) %/ d = q. @@ -2797,7 +2794,7 @@ by move/divpK: hdm<-; rewrite mulrA dvdp_mull // dvdpp. Qed. Lemma divp_mulAC m n : d %| m -> m %/ d * n = m * n %/ d. -Proof. by move=> hdm; rewrite mulrC (mulrC m); exact: divp_mulA. Qed. +Proof. by move=> hdm; rewrite mulrC (mulrC m); apply: divp_mulA. Qed. Lemma divp_mulCA p q : d %| p -> d %| q -> p * (q %/ d) = q * (p %/ d). Proof. by move=> hdp hdq; rewrite mulrC divp_mulAC // divp_mulA. Qed. @@ -2806,7 +2803,7 @@ Lemma modp_mul p q : (p * (q %% d)) %% d = (p * q) %% d. Proof. have -> : q %% d = q - q %/ d * d by rewrite {2}(divp_eq q) -addrA addrC subrK. rewrite mulrDr modp_add // -mulNr mulrA -{2}[_ %% _]addr0; congr (_ + _). -by apply/eqP; apply: dvdp_mull; exact: dvdpp. +by apply/eqP; apply: dvdp_mull; apply: dvdpp. Qed. End UnitDivisor. @@ -2853,7 +2850,7 @@ Lemma divp_divl r p q : Proof. move=> ulcr ulcp. have e : q = (q %/ p %/ r) * (p * r) + ((q %/ p) %% r * p + q %% p). - rewrite addrA (mulrC p) mulrA -mulrDl; rewrite -divp_eq //; exact: divp_eq. + by rewrite addrA (mulrC p) mulrA -mulrDl; rewrite -divp_eq //; apply: divp_eq. have pn0 : p != 0. by rewrite -lead_coef_eq0; apply: contraTneq ulcp => ->; rewrite unitr0. have rn0 : r != 0. @@ -2966,7 +2963,7 @@ Lemma mulpK p q : q != 0 -> p * q %/ q = p. Proof. by move=> qn0; rewrite IdomainUnit.mulpK // unitfE lead_coef_eq0. Qed. Lemma mulKp p q : q != 0 -> q * p %/ q = p. -Proof. by rewrite mulrC; exact: mulpK. Qed. +Proof. by rewrite mulrC; apply: mulpK. Qed. Lemma divp_scalel c p q : (c *: p) %/ q = c *: (p %/ q). Proof. @@ -3015,7 +3012,7 @@ Qed. Lemma eqp_mod p1 p2 q1 q2 : p1 %= p2 -> q1 %= q2 -> p1 %% q1 %= p2 %% q2. Proof. move=> e1 e2; apply: eqp_trans (eqp_modpr _ e2). -apply: eqp_trans (eqp_modpl _ e1); exact: eqpxx. +by apply: eqp_trans (eqp_modpl _ e1); apply: eqpxx. Qed. Lemma eqp_divr (d m n : {poly F}) : m %= n -> (d %/ m) %= (d %/ n). @@ -3028,7 +3025,7 @@ Qed. Lemma eqp_div p1 p2 q1 q2 : p1 %= p2 -> q1 %= q2 -> p1 %/ q1 %= p2 %/ q2. Proof. move=> e1 e2; apply: eqp_trans (eqp_divr _ e2). -apply: eqp_trans (eqp_divl _ e1); exact: eqpxx. +by apply: eqp_trans (eqp_divl _ e1); apply: eqpxx. Qed. Lemma eqp_gdcor p q r : q %= r -> gdcop p q %= gdcop p r. @@ -3036,7 +3033,7 @@ Proof. move=> eqr; rewrite /gdcop (eqp_size eqr). move: (size r)=> n; elim: n p q r eqr => [|n ihn] p q r; first by rewrite eqpxx. move=> eqr /=; rewrite (eqp_coprimepl p eqr); case: ifP => _ //; apply: ihn. -apply: eqp_div => //; exact: eqp_gcdl. +by apply: eqp_div => //; apply: eqp_gcdl. Qed. Lemma eqp_gdcol p q r : q %= r -> gdcop q p %= gdcop r p. @@ -3048,7 +3045,7 @@ elim: n p q r eqr {1 3}p (eqpxx p) => [|n ihn] p q r eqr s esp /=. suff rn0 : r != 0 by rewrite (negPf nq0) (negPf rn0) eqpxx. by apply: contraTneq eqr => ->; rewrite eqp0. rewrite (eqp_coprimepr _ eqr) (eqp_coprimepl _ esp); case: ifP=> _ //. -apply: ihn => //; apply: eqp_div => //; exact: eqp_gcd. +by apply: ihn => //; apply: eqp_div => //; apply: eqp_gcd. Qed. Lemma eqp_rgdco_gdco q p : rgdcop q p %= gdcop q p. @@ -3140,7 +3137,7 @@ by apply: IdomainUnit.divp_mulA; rewrite unitfE lead_coef_eq0. Qed. Lemma divp_mulAC d m n : d %| m -> m %/ d * n = m * n %/ d. -Proof. by move=> hdm; rewrite mulrC (mulrC m); exact: divp_mulA. Qed. +Proof. by move=> hdm; rewrite mulrC (mulrC m); apply: divp_mulA. Qed. Lemma divp_mulCA d p q : d %| p -> d %| q -> p * (q %/ d) = q * (p %/ d). Proof. by move=> hdp hdq; rewrite mulrC divp_mulAC // divp_mulA. Qed. @@ -3209,14 +3206,14 @@ Qed. Lemma edivp_eq d q r : size r < size d -> edivp (q * d + r) d = (0%N, q, r). Proof. move=> srd; apply: Idomain.edivp_eq ; rewrite // unitfE lead_coef_eq0. -rewrite -size_poly_gt0; exact: leq_trans srd. +by rewrite -size_poly_gt0; apply: leq_trans srd. Qed. Lemma modp_mul p q m : (p * (q %% m)) %% m = (p * q) %% m. Proof. have ->: q %% m = q - q %/ m * m by rewrite {2}(divp_eq q m) -addrA addrC subrK. rewrite mulrDr modp_add // -mulNr mulrA -{2}[_ %% _]addr0; congr (_ + _). -by apply/eqP; apply: dvdp_mull; exact: dvdpp. +by apply/eqP; apply: dvdp_mull; apply: dvdpp. Qed. Lemma dvdpP p q : reflect (exists qq, p = qq * q) (q %| p). @@ -3362,10 +3359,10 @@ Proof. by rewrite /eqp !dvdp_map. Qed. Lemma gcdp_map p q : (gcdp p q)^f = gcdp p^f q^f. Proof. wlog lt_p_q: p q / size p < size q. - move=> IH; case: (ltnP (size p) (size q)) => [|le_q_p]; first exact: IH. + move=> IHpq; case: (ltnP (size p) (size q)) => [|le_q_p]; first exact: IHpq. rewrite gcdpE (gcdpE p^f) !size_map_poly ltnNge le_q_p /= -map_modp. case: (eqVneq q 0) => [-> | q_nz]; first by rewrite rmorph0 !gcdp0. - by rewrite IH ?ltn_modp. + by rewrite IHpq ?ltn_modp. elim: {q}_.+1 p {-2}q (ltnSn (size q)) lt_p_q => // m IHm p q le_q_m lt_p_q. rewrite gcdpE (gcdpE p^f) !size_map_poly lt_p_q -map_modp. case: (eqVneq p 0) => [-> | q_nz]; first by rewrite rmorph0 !gcdp0. diff --git a/mathcomp/algebra/rat.v b/mathcomp/algebra/rat.v index 398d8de..d82fab0 100644 --- a/mathcomp/algebra/rat.v +++ b/mathcomp/algebra/rat.v @@ -1,10 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype bigop. -Require Import ssralg ssrnum ssrint. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp +Require Import bigop ssralg div ssrnum ssrint. (******************************************************************************) (* This file defines a datatype for rational numbers and equips it with a *) @@ -727,7 +726,7 @@ Variable F : numFieldType. Fact ratr_is_rmorphism : rmorphism (@ratr F). Proof. -have injZtoQ: @injective rat int intr by exact: intr_inj. +have injZtoQ: @injective rat int intr by apply: intr_inj. have nz_den x: (denq x)%:~R != 0 :> F by rewrite intr_eq0 denq_eq0. do 2?split; rewrite /ratr ?divr1 // => x y; last first. rewrite mulrC mulrAC; apply: canLR (mulKf (nz_den _)) _; rewrite !mulrA. diff --git a/mathcomp/algebra/ring_quotient.v b/mathcomp/algebra/ring_quotient.v index ea41744..c1427f3 100644 --- a/mathcomp/algebra/ring_quotient.v +++ b/mathcomp/algebra/ring_quotient.v @@ -1,10 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool ssrnat eqtype seq. -From mathcomp.discrete -Require Import choice generic_quotient. -Require Import ssralg. +From mathcomp +Require Import eqtype choice ssreflect ssrbool ssrnat ssrfun seq. +From mathcomp +Require Import ssralg generic_quotient. + (******************************************************************************) (* This file describes quotients of algebraic structures. *) @@ -47,15 +46,15 @@ Require Import ssralg. (* right, prime ideal of the ring R. *) (* *) (* The formalization of ideals features the following constructions: *) -(* nontrivial_ideal S == the collective predicate (S : pred R) on the *) +(* proper_ideal S == the collective predicate (S : pred R) on the *) (* ring R is stable by the ring product and does *) (* contain R's one. *) (* prime_idealr_closed S := u * v \in S -> (u \in S) || (v \in S) *) (* idealr_closed S == the collective predicate (S : pred R) on the *) (* ring R represents a (right) ideal. This *) -(* implies its being a nontrivial_ideal. *) +(* implies its being a proper_ideal. *) (* *) -(* MkIdeal idealS == packs idealS : nontrivial_ideal S into an *) +(* MkIdeal idealS == packs idealS : proper_ideal S into an *) (* idealr S interface structure associating the *) (* idealr_closed property to the canonical *) (* pred_key S (see ssrbool), which must already *) @@ -428,7 +427,7 @@ Notation UnitRingQuotMixin Q mU mV := Section IdealDef. -Definition nontrivial_ideal (R : ringType) (S : predPredType R) : Prop := +Definition proper_ideal (R : ringType) (S : predPredType R) : Prop := 1 \notin S /\ forall a, {in S, forall u, a * u \in S}. Definition prime_idealr_closed (R : ringType) (S : predPredType R) : Prop := @@ -437,18 +436,18 @@ Definition prime_idealr_closed (R : ringType) (S : predPredType R) : Prop := Definition idealr_closed (R : ringType) (S : predPredType R) := [/\ 0 \in S, 1 \notin S & forall a, {in S &, forall u v, a * u + v \in S}]. -Lemma idealr_closed_nontrivial R S : @idealr_closed R S -> nontrivial_ideal S. +Lemma idealr_closed_nontrivial R S : @idealr_closed R S -> proper_ideal S. Proof. by case=> S0 S1 hS; split => // a x xS; rewrite -[_ * _]addr0 hS. Qed. Lemma idealr_closedB R S : @idealr_closed R S -> zmod_closed S. Proof. by case=> S0 _ hS; split=> // x y xS yS; rewrite -mulN1r addrC hS. Qed. Coercion idealr_closedB : idealr_closed >-> zmod_closed. -Coercion idealr_closed_nontrivial : idealr_closed >-> nontrivial_ideal. +Coercion idealr_closed_nontrivial : idealr_closed >-> proper_ideal. Structure idealr (R : ringType) (S : predPredType R) := MkIdeal { idealr_zmod :> zmodPred S; - _ : nontrivial_ideal S + _ : proper_ideal S }. Structure prime_idealr (R : ringType) (S : predPredType R) := MkPrimeIdeal { @@ -457,7 +456,7 @@ Structure prime_idealr (R : ringType) (S : predPredType R) := MkPrimeIdeal { }. Definition Idealr (R : ringType) (I : predPredType R) (zmodI : zmodPred I) - (kI : keyed_pred zmodI) : nontrivial_ideal I -> idealr I. + (kI : keyed_pred zmodI) : proper_ideal I -> idealr I. Proof. by move=> kI1; split => //. Qed. Section IdealTheory. diff --git a/mathcomp/algebra/ssralg.v b/mathcomp/algebra/ssralg.v index 820a4b9..cadf4c2 100644 --- a/mathcomp/algebra/ssralg.v +++ b/mathcomp/algebra/ssralg.v @@ -1,9 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype finfun bigop prime binomial. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat div seq choice fintype. +From mathcomp +Require Import finfun bigop prime binomial. (******************************************************************************) (* The algebraic part of the Algebraic Hierarchy, as described in *) @@ -703,12 +703,18 @@ Proof. by move=> x y; rewrite -addrA addrN addr0. Qed. Lemma addrNK : @rev_right_loop V V -%R +%R. Proof. by move=> x y; rewrite -addrA addNr addr0. Qed. Definition subrK := addrNK. +Lemma subKr x : involutive (fun y => x - y). +Proof. by move=> y; apply: (canLR (addrK _)); rewrite addrC subrK. Qed. Lemma addrI : @right_injective V V V +%R. -Proof. move=> x; exact: can_inj (addKr x). Qed. +Proof. by move=> x; apply: can_inj (addKr x). Qed. Lemma addIr : @left_injective V V V +%R. -Proof. move=> y; exact: can_inj (addrK y). Qed. +Proof. by move=> y; apply: can_inj (addrK y). Qed. +Lemma subrI : right_injective (fun x y => x - y). +Proof. by move=> x; apply: can_inj (subKr x). Qed. +Lemma subIr : left_injective (fun x y => x - y). +Proof. by move=> y; apply: addIr. Qed. Lemma opprK : @involutive V -%R. -Proof. by move=> x; apply: (@addIr (- x)); rewrite addNr addrN. Qed. +Proof. by move=> x; apply: (@subIr x); rewrite addNr addrN. Qed. Lemma oppr_inj : @injective V V -%R. Proof. exact: inv_inj opprK. Qed. Lemma oppr0 : -0 = 0 :> V. @@ -719,13 +725,14 @@ Proof. by rewrite (inv_eq opprK) oppr0. Qed. Lemma subr0 x : x - 0 = x. Proof. by rewrite oppr0 addr0. Qed. Lemma sub0r x : 0 - x = - x. Proof. by rewrite add0r. Qed. +Lemma opprB x y : - (x - y) = y - x. +Proof. by apply: (canRL (addrK x)); rewrite addrC subKr. Qed. + Lemma opprD : {morph -%R: x y / x + y : V}. -Proof. -by move=> x y; apply: (@addrI (x + y)); rewrite addrA subrr addrAC addrK subrr. -Qed. +Proof. by move=> x y; rewrite -[y in LHS]opprK opprB addrC. Qed. -Lemma opprB x y : - (x - y) = y - x. -Proof. by rewrite opprD addrC opprK. Qed. +Lemma subr0_eq x y : x - y = 0 -> x = y. +Proof. by rewrite -(subrr y) => /addIr. Qed. Lemma subr_eq x y z : (x - z == y) = (x == y + z). Proof. exact: can2_eq (subrK z) (addrK z) x y. Qed. @@ -734,7 +741,7 @@ Lemma subr_eq0 x y : (x - y == 0) = (x == y). Proof. by rewrite subr_eq add0r. Qed. Lemma addr_eq0 x y : (x + y == 0) = (x == - y). -Proof. by rewrite -[x == _]subr_eq0 opprK. Qed. +Proof. by rewrite -[y in LHS]opprK subr_eq0. Qed. Lemma eqr_opp x y : (- x == - y) = (x == y). Proof. exact: can_eq opprK x y. Qed. @@ -1054,7 +1061,7 @@ Lemma commrN x y : comm x y -> comm x (- y). Proof. by move=> com_xy; rewrite /comm mulrN com_xy mulNr. Qed. Lemma commrN1 x : comm x (-1). -Proof. apply: commrN; exact: commr1. Qed. +Proof. exact/commrN/commr1. Qed. Lemma commrD x y z : comm x y -> comm x z -> comm x (y + z). Proof. by rewrite /comm mulrDl mulrDr => -> ->. Qed. @@ -1069,7 +1076,7 @@ Lemma commrM x y z : comm x y -> comm x z -> comm x (y * z). Proof. by move=> com_xy; rewrite /comm mulrA com_xy -!mulrA => ->. Qed. Lemma commr_nat x n : comm x n%:R. -Proof. by apply: commrMn; exact: commr1. Qed. +Proof. exact/commrMn/commr1. Qed. Lemma commrX x y n : comm x y -> comm x (y ^+ n). Proof. @@ -1096,7 +1103,7 @@ Qed. Lemma exprM x m n : x ^+ (m * n) = x ^+ m ^+ n. Proof. elim: m => [|m IHm]; first by rewrite expr1n. -by rewrite mulSn exprD IHm exprS exprMn_comm //; exact: commrX. +by rewrite mulSn exprD IHm exprS exprMn_comm //; apply: commrX. Qed. Lemma exprAC x m n : (x ^+ m) ^+ n = (x ^+ n) ^+ m. @@ -1178,7 +1185,7 @@ Proof. by move=> reg_x reg_y z t; rewrite -!mulrA => /reg_x/reg_y. Qed. Lemma lregX x n : lreg x -> lreg (x ^+ n). Proof. -by move=> reg_x; elim: n => [|n]; [exact: lreg1 | rewrite exprS; exact: lregM]. +by move=> reg_x; elim: n => [|n]; [apply: lreg1 | rewrite exprS; apply: lregM]. Qed. Lemma lreg_sign n : lreg ((-1) ^+ n : R). @@ -1320,14 +1327,14 @@ Qed. Lemma Frobenius_autMn x n : (x *+ n)^f = x^f *+ n. Proof. elim: n => [|n IHn]; first exact: Frobenius_aut0. -rewrite !mulrS Frobenius_autD_comm ?IHn //; exact: commrMn. +by rewrite !mulrS Frobenius_autD_comm ?IHn //; apply: commrMn. Qed. Lemma Frobenius_aut_nat n : (n%:R)^f = n%:R. Proof. by rewrite Frobenius_autMn Frobenius_aut1. Qed. Lemma Frobenius_autM_comm x y : comm x y -> (x * y)^f = x^f * y^f. -Proof. by exact: exprMn_comm. Qed. +Proof. exact: exprMn_comm. Qed. Lemma Frobenius_autX x n : (x ^+ n)^f = x^f ^+ n. Proof. by rewrite !fE -!exprM mulnC. Qed. @@ -1445,7 +1452,7 @@ Lemma rreg1 : rreg (1 : R). Proof. exact: (@lreg1 Rc). Qed. Lemma rregM x y : rreg x -> rreg y -> rreg (x * y). -Proof. by move=> reg_x reg_y; exact: (@lregM Rc). Qed. +Proof. by move=> reg_x reg_y; apply: (@lregM Rc). Qed. Lemma revrX x n : (x : Rc) ^+ n = (x : R) ^+ n. Proof. by elim: n => // n IHn; rewrite exprS exprSr IHn. Qed. @@ -2415,7 +2422,7 @@ Lemma lrmorphismP : lrmorphism f. Proof. exact: LRMorphism.class. Qed. Lemma can2_lrmorphism f' : cancel f f' -> cancel f' f -> lrmorphism f'. Proof. -move=> fK f'K; split; [exact: (can2_rmorphism fK) | exact: (can2_linear fK)]. +by move=> fK f'K; split; [apply: (can2_rmorphism fK) | apply: (can2_linear fK)]. Qed. Lemma bij_lrmorphism : @@ -2496,7 +2503,7 @@ Lemma mulrAC : @right_commutative R R *%R. Proof. exact: mulmAC. Qed. Lemma mulrACA : @interchange R *%R *%R. Proof. exact: mulmACA. Qed. Lemma exprMn n : {morph (fun x => x ^+ n) : x y / x * y}. -Proof. move=> x y; apply: exprMn_comm; exact: mulrC. Qed. +Proof. by move=> x y; apply: exprMn_comm; apply: mulrC. Qed. Lemma prodrXl n I r (P : pred I) (F : I -> R) : \prod_(i <- r | P i) F i ^+ n = (\prod_(i <- r | P i) F i) ^+ n. @@ -2508,16 +2515,16 @@ Proof. exact: big_undup_iterop_count. Qed. Lemma exprDn x y n : (x + y) ^+ n = \sum_(i < n.+1) (x ^+ (n - i) * y ^+ i) *+ 'C(n, i). -Proof. by rewrite exprDn_comm //; exact: mulrC. Qed. +Proof. by rewrite exprDn_comm //; apply: mulrC. Qed. Lemma exprBn x y n : (x - y) ^+ n = \sum_(i < n.+1) ((-1) ^+ i * x ^+ (n - i) * y ^+ i) *+ 'C(n, i). -Proof. by rewrite exprBn_comm //; exact: mulrC. Qed. +Proof. by rewrite exprBn_comm //; apply: mulrC. Qed. Lemma subrXX x y n : x ^+ n - y ^+ n = (x - y) * (\sum_(i < n) x ^+ (n.-1 - i) * y ^+ i). -Proof. by rewrite -subrXX_comm //; exact: mulrC. Qed. +Proof. by rewrite -subrXX_comm //; apply: mulrC. Qed. Lemma sqrrD x y : (x + y) ^+ 2 = x ^+ 2 + x * y *+ 2 + y ^+ 2. Proof. by rewrite exprDn !big_ord_recr big_ord0 /= add0r mulr1 mul1r. Qed. @@ -2810,10 +2817,10 @@ Proof. by move=> x Ux y; rewrite -mulrA mulVr ?mulr1. Qed. Definition divrK := mulrVK. Lemma mulrI : {in @unit R, right_injective *%R}. -Proof. by move=> x Ux; exact: can_inj (mulKr Ux). Qed. +Proof. by move=> x Ux; apply: can_inj (mulKr Ux). Qed. Lemma mulIr : {in @unit R, left_injective *%R}. -Proof. by move=> x Ux; exact: can_inj (mulrK Ux). Qed. +Proof. by move=> x Ux; apply: can_inj (mulrK Ux). Qed. (* Due to noncommutativity, fractions are inverted. *) Lemma telescope_prodr n m (f : nat -> R) : @@ -2868,10 +2875,14 @@ Proof. by rewrite dvdn_eq => /eqP def_m unit_d; rewrite -{2}def_m natrM mulrK. Qed. +Lemma divrI : {in unit, right_injective (fun x y => x / y)}. +Proof. by move=> x /mulrI/inj_comp; apply; apply: invr_inj. Qed. + +Lemma divIr : {in unit, left_injective (fun x y => x / y)}. +Proof. by move=> x; rewrite -unitrV => /mulIr. Qed. + Lemma unitr0 : (0 \is a @unit R) = false. -Proof. -by apply/unitrP=> [[x [_]]]; apply/eqP; rewrite mul0r eq_sym oner_neq0. -Qed. +Proof. by apply/unitrP=> [[x [_ /esym/eqP]]]; rewrite mul0r oner_eq0. Qed. Lemma invr0 : 0^-1 = 0 :> R. Proof. by rewrite invr_out ?unitr0. Qed. @@ -3031,7 +3042,7 @@ Fact mulC_mulrV : {in unit, right_inverse 1 inv *%R}. Proof. by move=> x Ux /=; rewrite mulrC mulVx. Qed. Fact mulC_unitP x y : y * x = 1 /\ x * y = 1 -> unit x. -Proof. case=> yx _; exact: unitPl yx. Qed. +Proof. by case=> yx _; apply: unitPl yx. Qed. Definition Mixin := UnitRingMixin mulVx mulC_mulrV mulC_unitP. @@ -3175,13 +3186,16 @@ Variable R : comUnitRingType. Implicit Types x y : R. Lemma unitrM x y : (x * y \in unit) = (x \in unit) && (y \in unit). -Proof. by apply: unitrM_comm; exact: mulrC. Qed. +Proof. by apply: unitrM_comm; apply: mulrC. Qed. Lemma unitrPr x : reflect (exists y, x * y = 1) (x \in unit). Proof. by apply: (iffP (unitrP x)) => [[y []] | [y]]; exists y; rewrite // mulrC. Qed. +Lemma divKr x : x \is a unit -> {in unit, involutive (fun y => x / y)}. +Proof. by move=> Ux y Uy; rewrite /= invrM ?unitrV // invrK mulrC divrK. Qed. + Lemma expr_div_n x y n : (x / y) ^+ n = x ^+ n / y ^+ n. Proof. by rewrite exprMn exprVn. Qed. @@ -3485,7 +3499,7 @@ Proof. by case: rpred0D. Qed. Lemma rpred_sum I r (P : pred I) F : (forall i, P i -> F i \in kS) -> \sum_(i <- r | P i) F i \in kS. -Proof. by move=> IH; elim/big_ind: _; [exact: rpred0 | exact: rpredD |]. Qed. +Proof. by move=> IH; elim/big_ind: _; [apply: rpred0 | apply: rpredD |]. Qed. Lemma rpredMn n : {in kS, forall u, u *+ n \in kS}. Proof. by move=> u Su; rewrite -(card_ord n) -sumr_const rpred_sum. Qed. @@ -3558,7 +3572,7 @@ Proof. by case: rpred1M. Qed. Lemma rpred_prod I r (P : pred I) F : (forall i, P i -> F i \in kS) -> \prod_(i <- r | P i) F i \in kS. -Proof. by move=> IH; elim/big_ind: _; [exact: rpred1 | exact: rpredM |]. Qed. +Proof. by move=> IH; elim/big_ind: _; [apply: rpred1 | apply: rpredM |]. Qed. Lemma rpredX n : {in kS, forall u, u ^+ n \in kS}. Proof. by move=> u Su; rewrite -(card_ord n) -prodr_const rpred_prod. Qed. @@ -3952,7 +3966,7 @@ suffices eq0_ring t1: rformula (eq0_rform t1) by elim: f => //= => f1 ->. rewrite /eq0_rform; move: (ub_var t1) => m; set tr := _ m. suffices: all rterm (tr.1 :: tr.2). case: tr => {t1} t1 r /= /andP[t1_r]. - by elim: r m => [|t r IHr] m; rewrite /= ?andbT // => /andP[->]; exact: IHr. + by elim: r m => [|t r IHr] m; rewrite /= ?andbT // => /andP[->]; apply: IHr. have: all rterm [::] by []. rewrite {}/tr; elim: t1 [::] => //=. - move=> t1 IHt1 t2 IHt2 r. @@ -3978,11 +3992,11 @@ suffices{e f} equal0_equiv e t1 t2: - elim: f e => /=; try tauto. + move=> t1 t2 e. by split; [move/equal0_equiv/eqP | move/eqP/equal0_equiv]. - + move=> t1 e; rewrite unitrE; exact: equal0_equiv. - + move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. - + move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. - + move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. - + move=> f1 IHf1 e; move: (IHf1 e); tauto. + + by move=> t1 e; rewrite unitrE; apply: equal0_equiv. + + by move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. + + by move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. + + by move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. + + by move=> f1 IHf1 e; move: (IHf1 e); tauto. + by move=> n f1 IHf1 e; split=> [] [x] /IHf1; exists x. + by move=> n f1 IHf1 e; split=> Hx x; apply/IHf1. rewrite -(add0r (eval e t2)) -(can2_eq (subrK _) (addrK _)). @@ -4086,8 +4100,8 @@ Definition qf_eval e := fix loop (f : formula R) : bool := (* qf_eval is equivalent to holds *) Lemma qf_evalP e f : qf_form f -> reflect (holds e f) (qf_eval e f). Proof. -elim: f => //=; try by move=> *; exact: idP. -- move=> t1 t2 _; exact: eqP. +elim: f => //=; try by move=> *; apply: idP. +- by move=> t1 t2 _; apply: eqP. - move=> f1 IHf1 f2 IHf2 /= /andP[/IHf1[] f1T]; last by right; case. by case/IHf2; [left | right; case]. - move=> f1 IHf1 f2 IHf2 /= /andP[/IHf1[] f1F]; first by do 2 left. @@ -4183,7 +4197,7 @@ Lemma qf_to_dnf_rterm f b : rformula f -> all dnf_rterm (qf_to_dnf f b). Proof. set ok := all dnf_rterm. have cat_ok bcs1 bcs2: ok bcs1 -> ok bcs2 -> ok (bcs1 ++ bcs2). - by move=> ok1 ok2; rewrite [ok _]all_cat; exact/andP. + by move=> ok1 ok2; rewrite [ok _]all_cat; apply/andP. have and_ok bcs1 bcs2: ok bcs1 -> ok bcs2 -> ok (and_dnf bcs1 bcs2). rewrite /and_dnf unlock; elim: bcs1 => //= cl1 bcs1 IH1; rewrite -andbA. case/and3P=> ok11 ok12 ok1 ok2; rewrite cat_ok ?{}IH1 {bcs1 ok1}//. @@ -4283,7 +4297,7 @@ Lemma foldForallP I e : <-> holds e (foldr Forall f I). Proof. elim: I e => /= [|i I IHi] e. - by split=> [|f_e e' eq_e]; [exact | apply: eq_holds f_e => i; rewrite eq_e]. + by split=> [|f_e e' eq_e]; [apply | apply: eq_holds f_e => i; rewrite eq_e]. split=> [f_e' x | f_e e' eq_e]; first set e_x := set_nth 0 e i x. apply/IHi=> e' eq_e; apply: f_e' => j. by have:= eq_e j; rewrite nth_set_nth /= !inE; case: eqP. @@ -4387,12 +4401,12 @@ Lemma prodf_seq_eq0 I r (P : pred I) (F : I -> R) : Proof. by rewrite (big_morph _ mulf_eq0 (oner_eq0 _)) big_has_cond. Qed. Lemma mulf_neq0 x y : x != 0 -> y != 0 -> x * y != 0. -Proof. move=> x0 y0; rewrite mulf_eq0; exact/norP. Qed. +Proof. by move=> x0 y0; rewrite mulf_eq0; apply/norP. Qed. Lemma prodf_neq0 (I : finType) (P : pred I) (F : I -> R) : reflect (forall i, P i -> (F i != 0)) (\prod_(i | P i) F i != 0). Proof. -by rewrite (sameP (prodf_eq0 _ _) exists_inP) negb_exists_in; exact: forall_inP. +by rewrite (sameP (prodf_eq0 _ _) exists_inP) negb_exists_in; apply: forall_inP. Qed. Lemma prodf_seq_neq0 I r (P : pred I) (F : I -> R) : @@ -4450,13 +4464,18 @@ Proof. by rewrite -subr_eq0 subr_sqr mulf_eq0 subr_eq0 addr_eq0. Qed. Lemma mulfI x : x != 0 -> injective ( *%R x). Proof. -move=> nz_x y z; rewrite -[x * z]add0r; move/(canLR (addrK _))/eqP. -rewrite -mulrN -mulrDr mulf_eq0 (negbTE nz_x) /=. -by move/eqP/(canRL (subrK _)); rewrite add0r. +move=> nz_x y z; apply: contra_eq => neq_yz. +by rewrite -subr_eq0 -mulrBr mulf_neq0 ?subr_eq0. Qed. Lemma mulIf x : x != 0 -> injective ( *%R^~ x). -Proof. by move=> nz_x y z; rewrite -!(mulrC x); exact: mulfI. Qed. +Proof. by move=> nz_x y z; rewrite -!(mulrC x); apply: mulfI. Qed. + +Lemma divfI x : x != 0 -> injective (fun y => x / y). +Proof. by move/mulfI/inj_comp; apply; apply: invr_inj. Qed. + +Lemma divIf y : y != 0 -> injective (fun x => x / y). +Proof. by rewrite -invr_eq0; apply: mulIf. Qed. Lemma sqrf_eq1 x : (x ^+ 2 == 1) = (x == 1) || (x == -1). Proof. by rewrite -subr_eq0 subr_sqr_1 mulf_eq0 subr_eq0 addr_eq0. Qed. @@ -4589,18 +4608,18 @@ Lemma unitfE x : (x \in unit) = (x != 0). Proof. by apply/idP/idP=> [/(memPn _)-> | /fieldP]; rewrite ?unitr0. Qed. Lemma mulVf x : x != 0 -> x^-1 * x = 1. -Proof. by rewrite -unitfE; exact: mulVr. Qed. +Proof. by rewrite -unitfE; apply: mulVr. Qed. Lemma divff x : x != 0 -> x / x = 1. -Proof. by rewrite -unitfE; exact: divrr. Qed. +Proof. by rewrite -unitfE; apply: divrr. Qed. Definition mulfV := divff. Lemma mulKf x : x != 0 -> cancel ( *%R x) ( *%R x^-1). -Proof. by rewrite -unitfE; exact: mulKr. Qed. +Proof. by rewrite -unitfE; apply: mulKr. Qed. Lemma mulVKf x : x != 0 -> cancel ( *%R x^-1) ( *%R x). -Proof. by rewrite -unitfE; exact: mulVKr. Qed. +Proof. by rewrite -unitfE; apply: mulVKr. Qed. Lemma mulfK x : x != 0 -> cancel ( *%R^~ x) ( *%R^~ x^-1). -Proof. by rewrite -unitfE; exact: mulrK. Qed. +Proof. by rewrite -unitfE; apply: mulrK. Qed. Lemma mulfVK x : x != 0 -> cancel ( *%R^~ x^-1) ( *%R^~ x). -Proof. by rewrite -unitfE; exact: divrK. Qed. +Proof. by rewrite -unitfE; apply: divrK. Qed. Definition divfK := mulfVK. Lemma invfM : {morph @inv F : x y / x * y}. @@ -4613,6 +4632,9 @@ Qed. Lemma invf_div x y : (x / y)^-1 = y / x. Proof. by rewrite invfM invrK mulrC. Qed. +Lemma divKf x : x != 0 -> involutive (fun y => x / y). +Proof. by move=> nz_x y; rewrite invf_div mulrC divfK. Qed. + Lemma expfB_cond m n x : (x == 0) + n <= m -> x ^+ (m - n) = x ^+ m / x ^+ n. Proof. move/subnK=> <-; rewrite addnA addnK !exprD. @@ -4687,7 +4709,7 @@ Variables (R : unitRingType) (f : {rmorphism F -> R}). Lemma fmorph_unit x : (f x \in unit) = (x != 0). Proof. have [-> |] := altP (x =P _); first by rewrite rmorph0 unitr0. -by rewrite -unitfE; exact: rmorph_unit. +by rewrite -unitfE; apply: rmorph_unit. Qed. Lemma fmorphV : {morph f: x / x^-1}. @@ -4712,10 +4734,10 @@ Lemma scalerK a : a != 0 -> cancel ( *:%R a : V -> V) ( *:%R a^-1). Proof. by move=> nz_a v; rewrite scalerA mulVf // scale1r. Qed. Lemma scalerKV a : a != 0 -> cancel ( *:%R a^-1 : V -> V) ( *:%R a). -Proof. by rewrite -invr_eq0 -{3}[a]invrK; exact: scalerK. Qed. +Proof. by rewrite -invr_eq0 -{3}[a]invrK; apply: scalerK. Qed. Lemma scalerI a : a != 0 -> injective ( *:%R a : V -> V). -Proof. move=> nz_a; exact: can_inj (scalerK nz_a). Qed. +Proof. by move=> nz_a; apply: can_inj (scalerK nz_a). Qed. Lemma scaler_eq0 a v : (a *: v == 0) = (a == 0) || (v == 0). Proof. @@ -4740,16 +4762,16 @@ Section Predicates. Context (S : pred_class) (divS : @divrPred F S) (kS : keyed_pred divS). Lemma fpredMl x y : x \in kS -> x != 0 -> (x * y \in kS) = (y \in kS). -Proof. by rewrite -!unitfE; exact: rpredMl. Qed. +Proof. by rewrite -!unitfE; apply: rpredMl. Qed. Lemma fpredMr x y : x \in kS -> x != 0 -> (y * x \in kS) = (y \in kS). -Proof. by rewrite -!unitfE; exact: rpredMr. Qed. +Proof. by rewrite -!unitfE; apply: rpredMr. Qed. Lemma fpred_divl x y : x \in kS -> x != 0 -> (x / y \in kS) = (y \in kS). -Proof. by rewrite -!unitfE; exact: rpred_divl. Qed. +Proof. by rewrite -!unitfE; apply: rpred_divl. Qed. Lemma fpred_divr x y : x \in kS -> x != 0 -> (y / x \in kS) = (y \in kS). -Proof. by rewrite -!unitfE; exact: rpred_divr. Qed. +Proof. by rewrite -!unitfE; apply: rpred_divr. Qed. End Predicates. @@ -4943,7 +4965,7 @@ suffices or_wf fs : let ofs := foldr Or False fs in - apply: or_wf. suffices map_proj_wf bcs: let mbcs := map (proj n) bcs in all dnf_rterm bcs -> all (@qf_form _) mbcs && all (@rformula _) mbcs. - by apply: map_proj_wf; exact: qf_to_dnf_rterm. + by apply: map_proj_wf; apply: qf_to_dnf_rterm. elim: bcs => [|bc bcs ihb] bcsr //= /andP[rbc rbcs]. by rewrite andbAC andbA wf_proj //= andbC ihb. elim: fs => //= g gs ihg; rewrite -andbA => /and4P[-> qgs -> rgs] /=. @@ -4960,7 +4982,7 @@ have auxP f0 e0 n0: qf_form f0 && rformula f0 -> apply: (@iffP (rc e0 n0 (dnf_to_form bcs))); last first. - by case=> x; rewrite -qf_to_dnfP //; exists x. - by case=> x; rewrite qf_to_dnfP //; exists x. - have: all dnf_rterm bcs by case/andP: cf => _; exact: qf_to_dnf_rterm. + have: all dnf_rterm bcs by case/andP: cf => _; apply: qf_to_dnf_rterm. elim: {f0 cf}bcs => [|bc bcs IHbcs] /=; first by right; case. case/andP=> r_bc /IHbcs {IHbcs}bcsP. have f_qf := dnf_to_form_qf [:: bc]. @@ -4972,8 +4994,8 @@ have auxP f0 e0 n0: qf_form f0 && rformula f0 -> case/orP => [bc_x|]; last by exists x. by case: no_x; exists x; apply/(qf_evalP _ f_qf); rewrite /= bc_x. elim: f e => //. -- move=> b e _; exact: idP. -- move=> t1 t2 e _; exact: eqP. +- by move=> b e _; apply: idP. +- by move=> t1 t2 e _; apply: eqP. - move=> f1 IH1 f2 IH2 e /= /andP[/IH1[] f1e]; last by right; case. by case/IH2; [left | right; case]. - move=> f1 IH1 f2 IH2 e /= /andP[/IH1[] f1e]; first by do 2!left. @@ -4982,7 +5004,7 @@ elim: f e => //. by case/IH2; [left | right; move/(_ f1e)]. - by move=> f IHf e /= /IHf[]; [right | left]. - move=> n f IHf e /= rf; have rqf := quantifier_elim_wf rf. - by apply: (iffP (auxP _ _ _ rqf)) => [] [x]; exists x; exact/IHf. + by apply: (iffP (auxP _ _ _ rqf)) => [] [x]; exists x; apply/IHf. move=> n f IHf e /= rf; have rqf := quantifier_elim_wf rf. case: auxP => // [f_x|no_x]; first by right=> no_x; case: f_x => x /IHf[]. by left=> x; apply/IHf=> //; apply/idPn=> f_x; case: no_x; exists x. @@ -5299,10 +5321,15 @@ Definition addNKr := addNKr. Definition addrK := addrK. Definition addrNK := addrNK. Definition subrK := subrK. +Definition subKr := subKr. Definition addrI := @addrI. Definition addIr := @addIr. +Definition subrI := @subrI. +Definition subIr := @subIr. Implicit Arguments addrI [[V] x1 x2]. Implicit Arguments addIr [[V] x1 x2]. +Implicit Arguments subrI [[V] x1 x2]. +Implicit Arguments subIr [[V] x1 x2]. Definition opprK := opprK. Definition oppr_inj := @oppr_inj. Implicit Arguments oppr_inj [[V] x1 x2]. @@ -5313,6 +5340,7 @@ Definition opprB := opprB. Definition subr0 := subr0. Definition sub0r := sub0r. Definition subr_eq := subr_eq. +Definition subr0_eq := subr0_eq. Definition subr_eq0 := subr_eq0. Definition addr_eq0 := addr_eq0. Definition eqr_opp := eqr_opp. @@ -5480,6 +5508,8 @@ Definition mulrVK := mulrVK. Definition divrK := divrK. Definition mulrI := mulrI. Definition mulIr := mulIr. +Definition divrI := divrI. +Definition divIr := divIr. Definition telescope_prodr := telescope_prodr. Definition commrV := commrV. Definition unitrE := unitrE. @@ -5554,6 +5584,7 @@ Definition holds_fsubst := holds_fsubst. Definition unitrM := unitrM. Definition unitrPr {R x} := @unitrPr R x. Definition expr_div_n := expr_div_n. +Definition divKr := divKr. Definition mulf_eq0 := mulf_eq0. Definition prodf_eq0 := prodf_eq0. Definition prodf_seq_eq0 := prodf_seq_eq0. @@ -5570,6 +5601,8 @@ Definition charf0P := charf0P. Definition eqf_sqr := eqf_sqr. Definition mulfI := mulfI. Definition mulIf := mulIf. +Definition divfI := divfI. +Definition divIf := divIf. Definition sqrf_eq1 := sqrf_eq1. Definition expfS_eq1 := expfS_eq1. Definition fieldP := fieldP. @@ -5582,6 +5615,7 @@ Definition mulVKf := mulVKf. Definition mulfK := mulfK. Definition mulfVK := mulfVK. Definition divfK := divfK. +Definition divKf := divKf. Definition invfM := invfM. Definition invf_div := invf_div. Definition expfB_cond := expfB_cond. diff --git a/mathcomp/algebra/ssrint.v b/mathcomp/algebra/ssrint.v index 48c78dd..5a0bafa 100644 --- a/mathcomp/algebra/ssrint.v +++ b/mathcomp/algebra/ssrint.v @@ -1,11 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import choice fintype finfun bigop. -Require Import ssralg ssrnum poly. - +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat choice seq. +From mathcomp +Require Import fintype finfun bigop ssralg ssrnum poly. Import GRing.Theory Num.Theory. (******************************************************************************) @@ -594,11 +592,11 @@ Proof. by rewrite -scalezrE scaler_nat. Qed. Lemma mulrz_sumr : forall x I r (P : pred I) F, x *~ (\sum_(i <- r | P i) F i) = \sum_(i <- r | P i) x *~ F i. -Proof. by rewrite -/M^z; exact: scaler_suml. Qed. +Proof. by rewrite -/M^z; apply: scaler_suml. Qed. Lemma mulrz_suml : forall n I r (P : pred I) (F : I -> M), (\sum_(i <- r | P i) F i) *~ n= \sum_(i <- r | P i) F i *~ n. -Proof. by rewrite -/M^z; exact: scaler_sumr. Qed. +Proof. by rewrite -/M^z; apply: scaler_sumr. Qed. Canonical intmul_additive x := Additive (@mulrzBr x). @@ -1751,7 +1749,7 @@ Canonical Znat_keyd := KeyedQualifier Znat_key. Lemma Znat_def n : (n \is a Znat) = (0 <= n). Proof. by []. Qed. Lemma Znat_semiring_closed : semiring_closed Znat. -Proof. by do 2?split => //; [exact: addr_ge0 | exact: mulr_ge0]. Qed. +Proof. by do 2?split => //; [apply: addr_ge0 | apply: mulr_ge0]. Qed. Canonical Znat_addrPred := AddrPred Znat_semiring_closed. Canonical Znat_mulrPred := MulrPred Znat_semiring_closed. Canonical Znat_semiringPred := SemiringPred Znat_semiring_closed. diff --git a/mathcomp/algebra/ssrnum.v b/mathcomp/algebra/ssrnum.v index aecb910..997602c 100644 --- a/mathcomp/algebra/ssrnum.v +++ b/mathcomp/algebra/ssrnum.v @@ -1,12 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype bigop finset. -From mathcomp.fingroup -Require Import fingroup. -Require Import ssralg zmodp poly. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq div choice fintype. +From mathcomp +Require Import bigop ssralg finset fingroup zmodp poly. (******************************************************************************) (* *) @@ -873,7 +870,7 @@ Canonical nneg_mulrPred := MulrPred nneg_divr_closed. Canonical nneg_divrPred := DivrPred nneg_divr_closed. Fact nneg_addr_closed : addr_closed (@nneg R). -Proof. by split; [exact: lerr | exact: addr_ge0]. Qed. +Proof. by split; [apply: lerr | apply: addr_ge0]. Qed. Canonical nneg_addrPred := AddrPred nneg_addr_closed. Canonical nneg_semiringPred := SemiringPred nneg_divr_closed. @@ -1019,8 +1016,8 @@ Lemma le0r x : (0 <= x) = (x == 0) || (0 < x). Proof. exact: le0r. Qed. Lemma lt0r_neq0 (x : R) : 0 < x -> x != 0. Proof. by rewrite lt0r; case/andP. Qed. -Lemma ltr0_neq0 (x : R) : 0 < x -> x != 0. -Proof. by rewrite lt0r; case/andP. Qed. +Lemma ltr0_neq0 (x : R) : x < 0 -> x != 0. +Proof. by rewrite ltr_neqAle; case/andP. Qed. Lemma gtr_eqF x y : y < x -> x == y = false. Proof. by rewrite ltr_def; case/andP; move/negPf=> ->. Qed. @@ -1037,7 +1034,7 @@ by rewrite !le0r mulf_eq0; case: eqP => // [-> /negPf[] | _ /pmulr_rgt0->]. Qed. (* Integer comparisons and characteristic 0. *) -Lemma ler01 : 0 <= 1 :> R. Proof. exact: ler01. Qed. +Lemma ler01 : 0 <= 1 :> R. Proof. exact: ler01. Qed. Lemma ltr01 : 0 < 1 :> R. Proof. exact: ltr01. Qed. Lemma ler0n n : 0 <= n%:R :> R. Proof. by rewrite -nnegrE rpred_nat. Qed. Hint Resolve ler01 ltr01 ler0n. @@ -2163,19 +2160,23 @@ elim/big_rec2: _ => // i x2 x1 /leE12/andP[le0Ei leEi12] [x1ge0 le_x12]. by rewrite mulr_ge0 // ler_pmul. Qed. -Lemma ltr_prod (E1 E2 : nat -> R) (n m : nat) : +Lemma ltr_prod I r (P : pred I) (E1 E2 : I -> R) : + has P r -> (forall i, P i -> 0 <= E1 i < E2 i) -> + \prod_(i <- r | P i) E1 i < \prod_(i <- r | P i) E2 i. +Proof. +elim: r => //= i r IHr; rewrite !big_cons; case: ifP => {IHr}// Pi _ ltE12. +have /andP[le0E1i ltE12i] := ltE12 i Pi; set E2r := \prod_(j <- r | P j) E2 j. +apply: ler_lt_trans (_ : E1 i * E2r < E2 i * E2r). + by rewrite ler_wpmul2l ?ler_prod // => j /ltE12/andP[-> /ltrW]. +by rewrite ltr_pmul2r ?prodr_gt0 // => j /ltE12/andP[le0E1j /ler_lt_trans->]. +Qed. + +Lemma ltr_prod_nat (E1 E2 : nat -> R) (n m : nat) : (m < n)%N -> (forall i, (m <= i < n)%N -> 0 <= E1 i < E2 i) -> \prod_(m <= i < n) E1 i < \prod_(m <= i < n) E2 i. Proof. -elim: n m => // n ihn m; rewrite ltnS leq_eqVlt; case/orP => [/eqP -> | ltnm hE]. - by move/(_ n) => /andb_idr; rewrite !big_nat1 leqnn ltnSn /=; case/andP. -rewrite big_nat_recr ?[X in _ < X]big_nat_recr ?(ltnW ltnm) //=. -move/andb_idr: (hE n); rewrite leqnn ltnW //=; case/andP => h1n h12n. -rewrite big_nat_cond [X in _ < X * _]big_nat_cond; apply: ltr_pmul => //=. -- apply: prodr_ge0 => i; rewrite andbT; case/andP=> hm hn. - by move/andb_idr: (hE i); rewrite hm /= ltnS ltnW //=; case/andP. -rewrite -!big_nat_cond; apply: ihn => // i /andP [hm hn]; apply: hE. -by rewrite hm ltnW. +move=> lt_mn ltE12; rewrite !big_nat ltr_prod {ltE12}//. +by apply/hasP; exists m; rewrite ?mem_index_iota leqnn. Qed. (* real of mul *) @@ -2981,7 +2982,7 @@ by move=> xR; rewrite ger0_def eq_sym; apply: lerif_eq; rewrite real_ler_norm. Qed. Lemma lerif_pmul x1 x2 y1 y2 C1 C2 : - 0 <= x1 -> 0 <= x2 -> x1 <= y1 ?= iff C1 -> x2 <= y2 ?= iff C2 -> + 0 <= x1 -> 0 <= x2 -> x1 <= y1 ?= iff C1 -> x2 <= y2 ?= iff C2 -> x1 * x2 <= y1 * y2 ?= iff (y1 * y2 == 0) || C1 && C2. Proof. move=> x1_ge0 x2_ge0 le_xy1 le_xy2; have [y_0 | ] := altP (_ =P 0). @@ -2998,7 +2999,7 @@ by apply: lerif_trans; rewrite (mono_lerif _ (ler_pmul2r _)) // ltr_def x2nz. Qed. Lemma lerif_nmul x1 x2 y1 y2 C1 C2 : - y1 <= 0 -> y2 <= 0 -> x1 <= y1 ?= iff C1 -> x2 <= y2 ?= iff C2 -> + y1 <= 0 -> y2 <= 0 -> x1 <= y1 ?= iff C1 -> x2 <= y2 ?= iff C2 -> y1 * y2 <= x1 * x2 ?= iff (x1 * x2 == 0) || C1 && C2. Proof. rewrite -!oppr_ge0 -mulrNN -[x1 * x2]mulrNN => y1le0 y2le0 le_xy1 le_xy2. @@ -3082,7 +3083,7 @@ rewrite prodrMn exprMn_n -/n' ler_pmuln2r ?expn_gt0; last by case: (n'). have ->: \prod_(k in A') E' k = E' j * pi. by rewrite (bigD1 j) //=; congr *%R; apply: eq_bigr => k /andP[_ /negPf->]. rewrite -(ler_pmul2l mu_gt0) -exprS -Dn mulrA; apply: ltr_le_trans. -rewrite ltr_pmul2r //= eqxx -addrA mulrDr mulrC -ltr_subl_addl -mulrBl. +rewrite ltr_pmul2r //= eqxx -addrA mulrDr mulrC -ltr_subl_addl -mulrBl. by rewrite mulrC ltr_pmul2r ?subr_gt0. Qed. diff --git a/mathcomp/algebra/vector.v b/mathcomp/algebra/vector.v index 01f2f53..aca5f86 100644 --- a/mathcomp/algebra/vector.v +++ b/mathcomp/algebra/vector.v @@ -1,10 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import choice fintype bigop finfun tuple. -Require Import ssralg matrix mxalgebra zmodp. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype bigop. +From mathcomp +Require Import finfun tuple ssralg matrix mxalgebra zmodp. (******************************************************************************) (* * Finite dimensional vector spaces *) @@ -648,7 +647,7 @@ Proof. by rewrite memv_cap; apply: andP. Qed. Lemma vspace_modl U V W : (U <= W -> U + (V :&: W) = (U + V) :&: W)%VS. Proof. -by move=> sUV; apply/vs2mxP; rewrite !(vs2mxD, vs2mxI); exact/eqmxP/matrix_modl. +by move=> sUV; apply/vs2mxP; rewrite !(vs2mxD, vs2mxI); apply/eqmxP/matrix_modl. Qed. Lemma vspace_modr U V W : (W <= U -> (U :&: V) + W = U :&: (V + W))%VS. @@ -1919,7 +1918,7 @@ Canonical vsproj_unlockable := [unlockable fun vsproj]. Lemma vsprojK : {in U, cancel vsproj vsval}. Proof. by rewrite unlock; apply: projv_id. Qed. Lemma vsvalK : cancel vsval vsproj. -Proof. by move=> w; exact/val_inj/vsprojK/subvsP. Qed. +Proof. by move=> w; apply/val_inj/vsprojK/subvsP. Qed. Lemma vsproj_is_linear : linear vsproj. Proof. by move=> k w1 w2; apply: val_inj; rewrite unlock /= linearP. Qed. diff --git a/mathcomp/algebra/zmodp.v b/mathcomp/algebra/zmodp.v index 9dd3ff5..0a9bf5e 100644 --- a/mathcomp/algebra/zmodp.v +++ b/mathcomp/algebra/zmodp.v @@ -1,12 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import div fintype bigop finset prime. -From mathcomp.fingroup -Require Import fingroup. -Require Import ssralg finalg. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq div. +From mathcomp +Require Import fintype bigop finset prime fingroup ssralg finalg. (******************************************************************************) (* Definition of the additive group and ring Zp, represented as 'I_p *) @@ -170,7 +167,7 @@ Lemma Zp_expg x n : x ^+ n = inZp (x * n). Proof. exact: Zp_mulrn. Qed. Lemma Zp1_expgz x : Zp1 ^+ x = x. -Proof. by rewrite Zp_expg; exact: Zp_mul1z. Qed. +Proof. by rewrite Zp_expg; apply: Zp_mul1z. Qed. Lemma Zp_cycle : setT = <[Zp1]>. Proof. by apply/setP=> x; rewrite -[x]Zp1_expgz inE groupX ?mem_gen ?set11. Qed. @@ -185,13 +182,13 @@ Implicit Arguments Zp1 [[p']]. Implicit Arguments inZp [[p']]. Lemma ord1 : all_equal_to (0 : 'I_1). -Proof. by case=> [[] // ?]; exact: val_inj. Qed. +Proof. by case=> [[] // ?]; apply: val_inj. Qed. Lemma lshift0 m n : lshift m (0 : 'I_n.+1) = (0 : 'I_(n + m).+1). Proof. exact: val_inj. Qed. Lemma rshift1 n : @rshift 1 n =1 lift (0 : 'I_n.+1). -Proof. by move=> i; exact: val_inj. Qed. +Proof. by move=> i; apply: val_inj. Qed. Lemma split1 n i : split (i : 'I_(1 + n)) = oapp (@inr _ _) (inl _ 0) (unlift 0 i). @@ -286,7 +283,7 @@ by move=> p_gt1; rewrite qualifE /= val_Zp_nat ?Zp_cast ?coprime_modr. Qed. Lemma Zp_group_set : group_set Zp. -Proof. rewrite /Zp; case: (p > 1); exact: groupP. Qed. +Proof. by rewrite /Zp; case: (p > 1); apply: groupP. Qed. Canonical Zp_group := Group Zp_group_set. Lemma card_Zp : p > 0 -> #|Zp| = p. @@ -307,7 +304,7 @@ by rewrite totient_count_coprime big_mkord. Qed. Lemma units_Zp_abelian : abelian units_Zp. -Proof. apply/centsP=> u _ v _; exact: unit_Zp_mulgC. Qed. +Proof. by apply/centsP=> u _ v _; apply: unit_Zp_mulgC. Qed. End Groups. diff --git a/mathcomp/all/all.v b/mathcomp/all/all.v index 20c76bf..b860e17 100644 --- a/mathcomp/all/all.v +++ b/mathcomp/all/all.v @@ -1,7 +1,7 @@ -Require Export mathcomp.ssreflect.all. -Require Export mathcomp.discrete.all. -Require Export mathcomp.algebra.all. -Require Export mathcomp.field.all. -Require Export mathcomp.character.all. -Require Export mathcomp.fingroup.all. -Require Export mathcomp.solvable.all. +Require Export mathcomp.ssreflect.all_ssreflect. +Require Export mathcomp.basic.all_basic. +Require Export mathcomp.algebra.all_algebra. +Require Export mathcomp.field.all_field. +Require Export mathcomp.character.all_character. +Require Export mathcomp.fingroup.all_fingroup. +Require Export mathcomp.solvable.all_solvable. diff --git a/mathcomp/attic/algnum_basic.v b/mathcomp/attic/algnum_basic.v index a302e7a..53f9016 100644 --- a/mathcomp/attic/algnum_basic.v +++ b/mathcomp/attic/algnum_basic.v @@ -1,7 +1,12 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype tuple div. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype tuple div. +From mathcomp Require Import bigop prime finset fingroup ssralg finalg zmodp abelian. +From mathcomp Require Import matrix vector falgebra finfield action poly ssrint cyclotomic. +From mathcomp Require Import fieldext mxalgebra mxpoly. (************************************************************************************************) @@ -120,7 +125,7 @@ Qed. Lemma int_clos_incl a : a \in A -> integral a. Proof. move=> ainA; exists ('X - a%:P); rewrite monicXsubC root_XsubC. -rewrite polyOverXsubC => //; by exact Asubr. +rewrite polyOverXsubC => //; by apply Asubr. Qed. Lemma intPl (I : eqType) G (r : seq I) l : has (fun x => G x != 0) r -> @@ -156,7 +161,7 @@ have rs : size r = n.-1 by rewrite /r size_takel // size_opp leq_pred. exists r; split. apply/eqP => /nilP; rewrite /nilp /r size_takel; last by rewrite size_opp leq_pred. by rewrite -subn1 subn_eq0 leqNgt ps. - have : - p \is a polyOver A by rewrite rpredN //; exact Asubr. + have : - p \is a polyOver A by rewrite rpredN //; apply Asubr. by move => /allP-popA; apply/allP => x /mem_take /popA. move: pr => /rootP; rewrite horner_coef -(prednK (n := size p)); last by rewrite ltnW. rewrite big_ord_recr /= rs; have := monicP pm; rewrite /lead_coef => ->; rewrite mul1r => /eqP. @@ -283,9 +288,9 @@ pose Y := map_tuple finv Xdual; exists Y => Und Xb. have Ydef (i : 'I_m) : Y`_i = finv Xdual`_i by rewrite -!tnth_nth tnth_map. have XiU (i : 'I_m) : X`_i \in U by apply/(basis_mem Xb)/mem_nth; rewrite size_tuple. have Xii (i : 'I_m) : coord X i X`_i = 1%:R. - by rewrite coord_free ?eqxx //; exact (basis_free Xb). + by rewrite coord_free ?eqxx //; apply (basis_free Xb). have Xij (i j : 'I_m) : j != i -> coord X i X`_j = 0%:R. - by rewrite coord_free; [move => /negbTE -> | exact (basis_free Xb)]. + by rewrite coord_free; [move => /negbTE -> | apply (basis_free Xb)]. have Xdualb : basis_of fullv Xdual. suffices Xdualf : free Xdual. rewrite /basis_of Xdualf andbC /= -dimv_leqif_eq ?subvf // eq_sym HomVdim. @@ -326,7 +331,7 @@ Hypothesis Und : ndeg tr U. Hypothesis Xb : basis_of U X. Lemma dualb_basis : basis_of U dual_basis. -Proof. have [Yb _] := svalP dual_basis_def Und Xb; exact Yb. Qed. +Proof. by have [Yb _] := svalP dual_basis_def Und Xb; apply Yb. Qed. Lemma dualb_orth : forall (i : 'I_m), tr X`_i dual_basis`_i = 1 /\ @@ -433,7 +438,7 @@ suffices FisK (F : fieldType) (L0 : fieldExtType F) (A : pred L0) (L : {subfield by apply/eqP/XfL0; rewrite -{3}kS => {i}; apply/eq_bigr => i _; rewrite -[RHS]mulr_algl kk'. move => Asubr Aint Apid Afrac1 Lnd; pose n := \dim L; have Amulr : mulr_closed A := Asubr. have [A0 _] : zmod_closed A := Asubr; have [Asub1 _] := Afrac1. -have AsubL : {subset A <= L} by move => a /Asub1; exact (subvP (sub1v L) a). +have AsubL : {subset A <= L} by move => a /Asub1; apply (subvP (sub1v L) a). have [b1 [b1B b1H]] : exists (b1 : n.-tuple L0), [/\ basis_of L b1 & forall i : 'I_n, integral A b1`_i]. pose b0 := vbasis L; have [f /all_and3-[fH0 fHa fHi]] := frac_field_alg_int Asubr Afrac1. diff --git a/mathcomp/attic/all.v b/mathcomp/attic/all.v deleted file mode 100644 index 41badbb..0000000 --- a/mathcomp/attic/all.v +++ /dev/null @@ -1,9 +0,0 @@ -Require Export algnum_basic. -Require Export amodule. -Require Export fib. -Require Export forms. -Require Export galgebra. -Require Export multinom. -Require Export mxtens. -Require Export quote. -Require Export tutorial. diff --git a/mathcomp/attic/amodule.v b/mathcomp/attic/amodule.v index f23ed60..e9d799e 100644 --- a/mathcomp/attic/amodule.v +++ b/mathcomp/attic/amodule.v @@ -1,6 +1,10 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype fintype finfun finset ssralg. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype fintype finfun finset ssralg. +From mathcomp Require Import bigop seq tuple choice ssrnat prime ssralg fingroup pgroup. +From mathcomp Require Import zmodp matrix vector falgebra galgebra. (*****************************************************************************) @@ -103,7 +107,7 @@ Implicit Types v w: M. Implicit Types c: F. Lemma rmulD x: {morph (rmul^~ x): v1 v2 / v1 + v2}. -Proof. move=> *; exact: linearD. Qed. +Proof. by move=> *; apply: linearD. Qed. Lemma rmul_linear_proof : forall v, linear (rmul v). Proof. by rewrite /rmul /rmorph; case: M => s [] b []. Qed. @@ -113,10 +117,10 @@ Lemma rmulA v x y: v :* (x * y) = (v :* x) :* y. Proof. exact: AModuleType.action_morph. Qed. Lemma rmulZ : forall c v x, (c *: v) :* x = c *: (v :* x). -Proof. move=> c v x; exact: linearZZ. Qed. +Proof. by move=> c v x; apply: linearZZ. Qed. Lemma rmul0 : left_zero 0 rmul. -Proof. move=> x; exact: linear0. Qed. +Proof. by move=> x; apply: linear0. Qed. Lemma rmul1 : forall v , v :* 1 = v. Proof. by rewrite /rmul /rmorph; case: M => s [] b []. Qed. @@ -154,7 +158,7 @@ Lemma eprodvP : forall vs1 ws vs2, Proof. move=> vs1 ws vs2; apply: (iffP idP). move=> Hs a b Ha Hb. - by apply: subv_trans Hs; exact: memv_eprod. + by apply: subv_trans Hs; apply: memv_eprod. move=> Ha; apply/subvP=> v. move/coord_span->; apply: memv_suml=> i _ /=. apply: memvZ. @@ -210,8 +214,8 @@ move=> vs1 vs2 ws; apply subv_anti; apply/andP; split. by rewrite rmulD; apply: memv_add; apply: memv_eprod. apply/subvP=> v; case/memv_addP=> v1 Hv1 [v2 Hv2 ->]. apply: memvD. - move: v1 Hv1; apply/subvP; apply: eprodvSl; exact: addvSl. -move: v2 Hv2; apply/subvP; apply: eprodvSl; exact: addvSr. + by move: v1 Hv1; apply/subvP; apply: eprodvSl; apply: addvSl. +by move: v2 Hv2; apply/subvP; apply: eprodvSl; apply: addvSr. Qed. Lemma eprodv_sumr vs ws1 ws2 : (vs :* (ws1 + ws2) = vs :* ws1 + vs :* ws2)%VS. @@ -221,8 +225,8 @@ apply subv_anti; apply/andP; split. by rewrite linearD; apply: memv_add; apply: memv_eprod. apply/subvP=> v; case/memv_addP=> v1 Hv1 [v2 Hv2 ->]. apply: memvD. - move: v1 Hv1; apply/subvP; apply: eprodvSr; exact: addvSl. -move: v2 Hv2; apply/subvP; apply: eprodvSr; exact: addvSr. + by move: v1 Hv1; apply/subvP; apply: eprodvSr; apply: addvSl. +by move: v2 Hv2; apply/subvP; apply: eprodvSr; apply: addvSr. Qed. Definition modv (vs: {vspace M}) (al: {aspace A}) := @@ -235,7 +239,7 @@ Lemma modv1 : forall vs, modv vs (aspace1 A). Proof. by move=> vs; rewrite /modv eprodv1 subvv. Qed. Lemma modfv : forall al, modv fullv al. -Proof. by move=> al; exact: subvf. Qed. +Proof. by move=> al; apply: subvf. Qed. Lemma memv_mod_mul : forall ms al m a, modv ms al -> m \in ms -> a \in al -> m :* a \in ms. @@ -273,7 +277,7 @@ Definition completely_reducible ms al := Lemma completely_reducible0 : forall al, completely_reducible 0 al. Proof. move=> al ms1 Hms1; rewrite subv0; move/eqP->. -by exists 0%VS; split; [exact: mod0v | exact: cap0v | exact: add0v]. +by exists 0%VS; split; [apply: mod0v | apply: cap0v | apply: add0v]. Qed. End AModuleDef. @@ -394,22 +398,22 @@ have If: limg f = ms1. by apply/subvP=> v Hv; rewrite -(Himf _ Hv) memv_img // memvf. apply/subvP=> i; case/memv_imgP=> x _ ->. rewrite !lfunE memvZ //= sum_lfunE memv_suml=> // j Hj. - rewrite lfunE /= lfunE (memv_mod_mul Hms1) //; first by exact: memv_proj. + rewrite lfunE /= lfunE (memv_mod_mul Hms1) //; first by apply: memv_proj. by rewrite memvE /= /gvspace (bigD1 (j^-1)%g) ?addvSl // groupV. exists (ms :&: lker f)%VS; split. - - apply: modv_ker=> //; apply/modfP=> *; exact: Cf. + - by apply: modv_ker=> //; apply/modfP=> *; apply: Cf. apply/eqP; rewrite -subv0; apply/subvP=> v; rewrite memv0. rewrite !memv_cap; case/andP=> Hv1; case/andP=> Hv2 Hv3. by move: Hv3; rewrite memv_ker Himf. -apply: subv_anti; rewrite subv_add Hsub capvSl. +apply: subv_anti; rewrite subv_add Hsub capvSl. apply/subvP=> v Hv. have->: v = f v + (v - f v) by rewrite addrC -addrA addNr addr0. apply: memv_add; first by rewrite -If memv_img // memvf. rewrite memv_cap; apply/andP; split. apply: memvB=> //; apply: subv_trans Hsub. - by rewrite -If; apply: memv_img; exact: memvf. + by rewrite -If; apply: memv_img; apply: memvf. rewrite memv_ker linearB /= (Himf (f v)) ?subrr // /in_mem /= -If. -by apply: memv_img; exact: memvf. +by apply: memv_img; apply: memvf. Qed. End ModuleRepresentation. diff --git a/mathcomp/attic/fib.v b/mathcomp/attic/fib.v index e002a72..b94a06e 100644 --- a/mathcomp/attic/fib.v +++ b/mathcomp/attic/fib.v @@ -1,5 +1,8 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp Require Import bigop div prime finfun tuple ssralg zmodp matrix binomial. (*****************************************************************************) @@ -106,7 +109,7 @@ Proof. move=> m n; elim: n=> [|[|n] IH]; first by case: m. by case: m{IH}=> [|[]]. rewrite fibSS leq_eqVlt; case/orP=>[|Hm]; first by move/eqP->. -by apply: (leq_trans (IH _)) => //; exact: leq_addr. +by apply: (leq_trans (IH _)) => //; apply: leq_addr. Qed. Lemma fib_eq1 : forall n, (fib n == 1) = ((n == 1) || (n == 2)). @@ -119,7 +122,7 @@ Lemma fib_eq : forall m n, (fib m == fib n) = [|| m == n, (m == 1) && (n == 2) | (m == 2) && (n == 1)]. Proof. move=> m n; wlog: m n/ m <= n=> [HH|]. - case/orP: (leq_total m n)=> Hm; first by exact: HH. + case/orP: (leq_total m n)=> Hm; first by apply: HH. by rewrite eq_sym HH // eq_sym ![(_ == 1) && _]andbC [(_ && _) || _] orbC. rewrite leq_eqVlt; case/orP=>[|]; first by move/eqP->; rewrite !eqxx. case: m=> [|[|m]] Hm. @@ -170,8 +173,8 @@ Qed. Lemma gcdn_fib: forall m n, gcdn (fib m) (fib n) = fib (gcdn m n). Proof. move=> m n; apply: gcdn_def. -- by apply: dvdn_fib; exact: dvdn_gcdl. -- by apply: dvdn_fib; exact: dvdn_gcdr. +- by apply: dvdn_fib; apply: dvdn_gcdl. +- by apply: dvdn_fib; apply: dvdn_gcdr. move=> d' Hdm Hdn. case: m Hdm=> [|m Hdm]; first by rewrite gcdnE eqxx. have F: 0 < m.+1 by []. diff --git a/mathcomp/attic/forms.v b/mathcomp/attic/forms.v index 1c88af5..5cb7662 100644 --- a/mathcomp/attic/forms.v +++ b/mathcomp/attic/forms.v @@ -1,5 +1,8 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype bigop. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype bigop. +From mathcomp Require Import finfun tuple ssralg matrix zmodp vector. Set Implicit Arguments. diff --git a/mathcomp/attic/galgebra.v b/mathcomp/attic/galgebra.v index 411fb6a..3b4004b 100644 --- a/mathcomp/attic/galgebra.v +++ b/mathcomp/attic/galgebra.v @@ -1,5 +1,8 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq choice fintype finfun. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq choice fintype finfun. +From mathcomp Require Import bigop finset ssralg fingroup zmodp matrix vector falgebra. (*****************************************************************************) @@ -213,7 +216,7 @@ case/memv_sumP: Hu => u_ Hu ->; rewrite big_distrl /=. apply: memv_suml=> i Hi. case/memv_sumP: Hv => v_ Hv ->; rewrite big_distrr /=. apply: memv_suml=> j Hj. -rewrite /gvspace (bigD1 (i*j)%g) /=; last by exact: groupM. +rewrite /gvspace (bigD1 (i*j)%g) /=; last by apply: groupM. apply: subv_trans (addvSl _ _). case/vlineP: (Hu _ Hi)=> k ->; case/vlineP: (Hv _ Hj)=> l ->. apply/vlineP; exists (k * l). diff --git a/mathcomp/attic/multinom.v b/mathcomp/attic/multinom.v index b203381..e17d959 100644 --- a/mathcomp/attic/multinom.v +++ b/mathcomp/attic/multinom.v @@ -1,5 +1,8 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp Require Import tuple finfun bigop ssralg poly generic_quotient bigenough. (* We build the ring of multinomials with an arbitrary (countable) *) diff --git a/mathcomp/attic/quote.v b/mathcomp/attic/quote.v index bde7fac..85c211a 100644 --- a/mathcomp/attic/quote.v +++ b/mathcomp/attic/quote.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq. Set Implicit Arguments. Unset Strict Implicit. @@ -99,10 +101,10 @@ Lemma interp_wf_eval : forall y0 e t y, interp e t = Some y -> wf t /\ eval y0 e t = y. Proof. move=> y0 e t; elim/term_ind': t => [i|s a IHa] y /=. - by elim: i e => [|i IHi] [|z e] //=; [case | elim: i {IHi} | exact: IHi]. + by elim: i e => [|i IHi] [|z e] //=; [case | elim: i {IHi} | apply: IHi]. case: symop => /= n x1. elim: n x1 a IHa => [|n IHn] x1 [|f a] //=; first by move=> _ []. -case: (interp e f) => //= x []; case/(_ x)=> // -> ->; exact: IHn. +by case: (interp e f) => //= x []; case/(_ x)=> // -> ->; apply: IHn. Qed. Definition var_val := @id T. @@ -128,7 +130,7 @@ Canonical Structure var_form i x P s := Form (@var_form_subproof i x P s). Lemma op_form_subproof : forall e t tP (f : form e t tP) x, x = @fval e t tP f -> interp e t = Some (op_val x). -Proof. by move=> e t tP f _ ->; exact: formP. Qed. +Proof. by move=> e t tP f _ ->; apply: formP. Qed. Canonical Structure op_form e t tP f x := Form (@op_form_subproof e t tP f x). @@ -199,7 +201,7 @@ Lemma simp_form : forall e t y ptP, (@postProp_statement (close (Env e) /\ simp e t = Some y) ptP), fval f = y. Proof. -move=> e t y [tP [_ def_y]] [x /= def_x]; apply: simpP def_y; exact: def_x. +by move=> e t y [tP [_ def_y]] [x /= def_x]; apply: simpP def_y; apply: def_x. Qed. End GenSimp. @@ -360,6 +362,6 @@ Time rewrite bsimp. Time rewrite !bsimp. by []. Qed. -Print try_bsimp. + diff --git a/mathcomp/attic/tutorial.v b/mathcomp/attic/tutorial.v index 3326c90..52d31cd 100644 --- a/mathcomp/attic/tutorial.v +++ b/mathcomp/attic/tutorial.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq. Section HilbertSaxiom. @@ -26,7 +28,7 @@ Qed. Check (hAiB hA). Lemma HilbertS3 : C. -Proof. by apply: hAiBiC; last exact: hAiB. Qed. +Proof. by apply: hAiBiC; last apply: hAiB. Qed. Lemma HilbertS4 : C. Proof. exact: (hAiBiC _ (hAiB _)). Qed. @@ -226,12 +228,12 @@ move=> m [|d] //=; rewrite -{1}[m]/(0 * d.+1 + m). elim: m {-2}m 0 (leqnn m) => [|n IHn] [|m] q //=; rewrite ltnS => le_mn. rewrite subn_if_gt; case: ltnP => // le_dm. rewrite -{1}(subnKC le_dm) -addSn addnA -mulSnr; apply: IHn. -apply: leq_trans le_mn; exact: leq_subr. +by apply: leq_trans le_mn; apply: leq_subr. Qed. Lemma edivn_eq : forall d q r, r < d -> edivn (q * d + r) d = (q, r). Proof. -move=> d q r lt_rd; have d_gt0: 0 < d by exact: leq_trans lt_rd. +move=> d q r lt_rd; have d_gt0: 0 < d by apply: leq_trans lt_rd. case: edivnP lt_rd => q' r'; rewrite d_gt0 /=. wlog: q q' r r' / q <= q' by case (ltnP q q'); last symmetry; eauto. rewrite leq_eqVlt; case: eqP => [-> _|_] /=; first by move/addnI->. @@ -254,7 +256,7 @@ move=> m [|d] //=; rewrite -{1}[m]/(0 * d.+1 + m). elim: m {-2}m 0 (leqnn m) => [|n IHn] [|m] q //=; rewrite ltnS => le_mn. rewrite subn_if_gt; case: ltnP => // le_dm. rewrite -{1}(subnKC le_dm) -addSn addnA -mulSnr; apply: IHn. -apply: leq_trans le_mn; exact: leq_subr. +by apply: leq_trans le_mn; apply: leq_subr. Qed. Lemma edivnP_right : forall m d, edivn_spec_right m d (edivn m d). @@ -263,12 +265,12 @@ move=> m [|d] //=; rewrite -{1}[m]/(0 * d.+1 + m). elim: m {-2}m 0 (leqnn m) => [|n IHn] [|m] q //=; rewrite ltnS => le_mn. rewrite subn_if_gt; case: ltnP => // le_dm. rewrite -{1}(subnKC le_dm) -addSn addnA -mulSnr; apply: IHn. -apply: leq_trans le_mn; exact: leq_subr. +by apply: leq_trans le_mn; apply: leq_subr. Qed. Lemma edivn_eq_right : forall d q r, r < d -> edivn (q * d + r) d = (q, r). Proof. -move=> d q r lt_rd; have d_gt0: 0 < d by exact: leq_trans lt_rd. +move=> d q r lt_rd; have d_gt0: 0 < d by apply: leq_trans lt_rd. set m := q * d + r; have: m = q * d + r by []. set d' := d; have: d' = d by []. case: (edivnP_right m d') => {m d'} m d' q' r' -> lt_r'd' d'd q'd'r'. @@ -282,7 +284,7 @@ Qed. Lemma edivn_eq_left : forall d q r, r < d -> edivn (q * d + r) d = (q, r). Proof. -move=> d q r lt_rd; have d_gt0: 0 < d by exact: leq_trans lt_rd. +move=> d q r lt_rd; have d_gt0: 0 < d by apply: leq_trans lt_rd. case: (edivnP_left (q * d + r) d) lt_rd; rewrite d_gt0 /=. set q':= (edivn (q * d + r) d).1; set r':= (edivn (q * d + r) d).2. rewrite (surjective_pairing (edivn (q * d + r) d)) -/q' -/r'. diff --git a/mathcomp/discrete/AUTHORS b/mathcomp/basic/AUTHORS index b55a98d..b55a98d 120000 --- a/mathcomp/discrete/AUTHORS +++ b/mathcomp/basic/AUTHORS diff --git a/mathcomp/discrete/CeCILL-B b/mathcomp/basic/CeCILL-B index 83e22fd..83e22fd 120000 --- a/mathcomp/discrete/CeCILL-B +++ b/mathcomp/basic/CeCILL-B diff --git a/mathcomp/discrete/INSTALL b/mathcomp/basic/INSTALL index 6aa7ec5..6aa7ec5 120000 --- a/mathcomp/discrete/INSTALL +++ b/mathcomp/basic/INSTALL diff --git a/mathcomp/discrete/Make b/mathcomp/basic/Make index bfbbfe2..cd88088 100644 --- a/mathcomp/discrete/Make +++ b/mathcomp/basic/Make @@ -1,4 +1,4 @@ -all.v +all_basic.v bigop.v binomial.v choice.v diff --git a/mathcomp/discrete/Makefile b/mathcomp/basic/Makefile index e872352..e872352 100644 --- a/mathcomp/discrete/Makefile +++ b/mathcomp/basic/Makefile diff --git a/mathcomp/discrete/README b/mathcomp/basic/README index e4e30e8..e4e30e8 120000 --- a/mathcomp/discrete/README +++ b/mathcomp/basic/README diff --git a/mathcomp/discrete/all.v b/mathcomp/basic/all_basic.v index dfa8536..dfa8536 100644 --- a/mathcomp/discrete/all.v +++ b/mathcomp/basic/all_basic.v diff --git a/mathcomp/discrete/bigop.v b/mathcomp/basic/bigop.v index 8a25624..08c7e0a 100644 --- a/mathcomp/discrete/bigop.v +++ b/mathcomp/basic/bigop.v @@ -1,19 +1,24 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -Require Import path div fintype tuple finfun. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div fintype. +From mathcomp +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. *) +(* set of indices (bigop); 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 *) +(* and which support a natural use of big operators. *) +(* To improve performance of the Coq typechecker on large expressions, the *) +(* bigop constant is OPAQUE. It can however be unlocked to reveal the *) +(* transparent constant reducebig, to let Coq expand summation on an explicit *) +(* sequence with an explicit test. *) (* The lemmas can be classified according to the operator being iterated: *) -(* 1. results independent of the operator: extensionality with respect to *) +(* 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 *) @@ -21,7 +26,7 @@ Require Import path div fintype tuple finfun. (* 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: *) +(* 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 *) @@ -34,13 +39,13 @@ Require Import path div fintype tuple finfun. (* - <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 <- 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) *) +(* 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 *) @@ -54,13 +59,13 @@ Require Import path div fintype tuple finfun. (* 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. *) +(* the %R scope; in mxalgebra, vector & falgebra in the %MS and %VS scopes; *) +(* "\prod" is also overloaded in fingroup, in 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 *) +(* 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. *) @@ -689,7 +694,7 @@ Lemma big_rec3 I r (P : pred I) F1 F2 F3 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. +Proof. by rewrite unlock; elim: r => //= i r; case: ifP => //; apply: 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). @@ -698,7 +703,7 @@ Lemma big_ind3 I r (P : pred I) F1 F2 F3 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. +Proof. by apply: big_rec3 => i x1 x2 x3 /K_F; apply: Kop. Qed. End Elim3. @@ -717,13 +722,13 @@ 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. +Proof. by rewrite unlock; elim: r => //= i r; case: ifP => //; apply: 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. +Proof. by apply: big_rec2 => i x1 x2 /K_F; apply: 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 : @@ -746,12 +751,12 @@ 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. +Proof. by rewrite unlock; elim: r => //= i r; case: ifP => //; apply: 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. +Proof. by apply: big_rec => // i x /K_F /Kop; apply. 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)) : @@ -789,7 +794,7 @@ Lemma big_filter_cond r (P1 P2 : pred I) F : = \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. +by rewrite -filter_predI; apply: eq_filter => i; apply: andbC. Qed. Lemma eq_bigl r (P1 P2 : pred I) F : @@ -801,7 +806,7 @@ Proof. by move=> eqP12; rewrite -!(big_filter r) (eq_filter eqP12). Qed. 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. +Proof. by apply: eq_bigl => i; apply: 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. @@ -815,7 +820,7 @@ 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. +Proof. by move=> <-{r2}; apply: eq_big. Qed. Lemma big_nil (P : pred I) F : \big[op/idx]_(i <- [::] | P i) F i = idx. Proof. by rewrite unlock. Qed. @@ -846,7 +851,7 @@ 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. +Proof. by move/eq_bigl->; apply: 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 @@ -911,8 +916,8 @@ Lemma congr_big_nat m1 n1 m2 n2 P1 P2 F1 F2 : 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. + by apply: andb_id2l; apply: eqP12. +by rewrite andbC; apply: eqF12. Qed. Lemma eq_big_nat m n F1 F2 : @@ -934,7 +939,7 @@ 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. +Proof. by move=> lt_mn; apply: 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 = @@ -984,7 +989,7 @@ 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. +Proof. by move=> le_n12; apply: (big_ord_widen_cond (predT)). Qed. Lemma big_ord_widen_leq n1 n2 (P : pred 'I_(n1.+1)) F : n1 < n2 -> @@ -1047,7 +1052,7 @@ 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. +Proof. exact: (big_ord_narrow_cond_leq (predT)). Qed. Lemma big_ord_recl n F : \big[op/idx]_(i < n.+1) F i = @@ -1115,12 +1120,12 @@ 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. +by rewrite big_const_seq; elim: (count _ _) => //= n ->; apply: 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. +Proof. by move/(eq_bigr _)->; apply: big1_eq. Qed. Lemma big1_seq (I : eqType) r (P : pred I) F : (forall i, P i && (i \in r) -> F i = 1) -> @@ -1159,7 +1164,7 @@ 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. +Proof. by move/(eq_bigl _ _)->; apply: 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 = @@ -1233,7 +1238,7 @@ elim: r1 r2 => [|i r1 IHr1] r2 eq_r12. 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. +by move/(_ a): eq_r12; rewrite !count_cat /= addnCA; apply: addnI. Qed. Lemma big_uniq (I : finType) (r : seq I) F : @@ -1371,7 +1376,7 @@ 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. +Proof. by move=> injh; apply: reindex (onW_bij _ (injF_bij injh)). Qed. Implicit Arguments reindex_inj [I h P F]. Lemma big_nat_rev m n P F : @@ -1414,7 +1419,7 @@ 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. +by case/andP; apply: PQxQ. Qed. Implicit Arguments exchange_big_dep [I J rI rJ P Q F]. @@ -1522,11 +1527,11 @@ 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. +Proof. by rewrite (big_endo ( *%M^~ a)) ?mul0m // => x y; apply: 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. +Proof. by rewrite big_endo ?mulm0 // => x y; apply: 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) @@ -1542,11 +1547,11 @@ 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. +have: uniq r by apply: 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. + by apply/ffunP=> i; rewrite ffunE; apply/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. @@ -1595,7 +1600,7 @@ 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. +Proof. by rewrite bigA_distr_big; apply: eq_bigl => ?; apply/familyP. Qed. End Distributivity. @@ -1677,13 +1682,13 @@ Lemma leqif_sum (I : finType) (P C : pred I) (E1 E2 : I -> nat) : \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. +by elim/big_rec3: _ => // i Ci m1 m2 /leE12; apply: 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. +Proof. by move=> leE12; elim/big_ind2: _ => // m1 m2 n1 n2; apply: 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]. @@ -1695,7 +1700,7 @@ 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. +Proof. by move=> Fpos; apply: 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. @@ -1710,20 +1715,20 @@ 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 apply: leq_trans leFm; apply: 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. +Proof. by move=> Pi0 le_m_Fi0; apply: 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. +by apply/eqP; rewrite eqn_leq leq_bigmax_cond // andbT; apply/bigmax_leqP. Qed. Implicit Arguments bigmax_eq_arg [I P F]. @@ -1731,7 +1736,7 @@ 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]. +by exists [arg max_(i > i0 in A) F i]; [case: arg_maxP | apply: bigmax_eq_arg]. Qed. Lemma eq_bigmax (I : finType) F : #|I| > 0 -> {i0 : I | \max_i F i = F i0}. diff --git a/mathcomp/discrete/binomial.v b/mathcomp/basic/binomial.v index 51b9635..a959bc7 100644 --- a/mathcomp/discrete/binomial.v +++ b/mathcomp/basic/binomial.v @@ -1,19 +1,20 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -Require Import path div fintype tuple finfun bigop prime finset. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq path div. +From mathcomp +Require Import fintype tuple finfun bigop prime finset. (******************************************************************************) (* This files contains the definition of: *) +(* 'C(n, m) == the binomial coeficient n choose m. *) (* 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 *) +(* the product n * (n - 1) * ... * (n - m + 1). *) +(* Note that n ^_ m = 0 if m > n, and 'C(n, m) = n ^_ m %/ m/!. *) (* *) -(* In additions to the properties of these functions, triangular_sum, Wilson *) -(* and Pascal are examples of how to manipulate expressions with bigops. *) +(* In additions to the properties of these functions, we prove a few seminal *) +(* results such as triangular_sum, Wilson and Pascal; their proofs are good *) +(* examples of how to manipulate expressions with bigops. *) (******************************************************************************) Set Implicit Arguments. @@ -55,7 +56,7 @@ have dFact n: 0 < n -> (n.-1)`! = \prod_(0 <= i < n | i != 0) i. 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. + apply/primeP; split=> // d dv_dp; have: d <= p by apply: 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). @@ -65,7 +66,7 @@ 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 Fp_mod (i : 'I_p) : i %% p = i by apply: 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. @@ -358,7 +359,7 @@ Lemma card_inj_ffuns_on D T (R : pred T) : Proof. rewrite -card_uniq_tuples. have bijFF: {on (_ : pred _), bijective (@Finfun D T)}. - by exists val => // x _; exact: val_inj. + by exists val => // x _; apply: 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. @@ -371,17 +372,24 @@ 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). +Lemma cards_draws T (B : {set T}) k : + #|[set A : {set T} | A \subset B & #|A| == k]| = 'C(#|B|, k). Proof. -have [ltTk | lekT] := ltnP #|T| k. +have [ltTk | lekT] := ltnP #|B| k. rewrite bin_small // eq_card0 // => A. - by rewrite inE eqn_leq andbC leqNgt (leq_ltn_trans (max_card _)). + rewrite inE eqn_leq [k <= _]leqNgt. + have [AsubB /=|//] := boolP (A \subset B). + by rewrite (leq_ltn_trans (subset_leq_card AsubB)) ?andbF. 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. +rewrite -sum_nat_dep_const -{1 3}(card_ord k). +rewrite -card_inj_ffuns_on -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. +rewrite (partition_big imIk (fun A => (A \subset B) && (#|A| == k))) /= + => [|f]; last first. + move=> /andP [/ffun_onP f_ffun /injectiveP inj_f]. + rewrite card_imset ?card_ord // eqxx andbT. + by apply/subsetP => x /imsetP [i _ ->]; rewrite f_ffun. +apply/eqP; apply: eq_bigr => A /andP [AsubB /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. @@ -392,18 +400,27 @@ rewrite (reindex (fun p : {ffun _} => [ffun i => f0 (p i)])) /=; last first. 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]. + rewrite -im_f0 => /andP[/andP[/ffun_onP f_ffun /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]. +rewrite -andbA. +apply/and3P/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. +have injf: injective f by move=> i j; rewrite !ffunE => /inj_f0; apply: inj_p. +have imIkf : imIk f == A. + rewrite eqEcard card_imset // cardAk card_ord leqnn andbT -im_f0. + by apply/subsetP=> x /imsetP[i _ ->]; rewrite ffunE mem_imset. +split; [|exact/injectiveP|exact: imIkf]. +apply/ffun_onP => x; apply: (subsetP AsubB). +by rewrite -(eqP imIkf) mem_imset. +Qed. + +Lemma card_draws T k : #|[set A : {set T} | #|A| == k]| = 'C(#|T|, k). +Proof. +by rewrite -cardsT -cards_draws; apply: eq_card => A; rewrite !inE subsetT. Qed. Lemma card_ltn_sorted_tuples m n : diff --git a/mathcomp/discrete/choice.v b/mathcomp/basic/choice.v index 5909752..378387e 100644 --- a/mathcomp/discrete/choice.v +++ b/mathcomp/basic/choice.v @@ -1,12 +1,12 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq. (******************************************************************************) (* This file contains the definitions of: *) -(* choiceType == interface for types with a choice operator *) -(* countType == interface for countable types *) +(* choiceType == interface for types with a choice operator. *) +(* countType == interface for countable types (implies choiceType). *) (* 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 *) @@ -35,10 +35,14 @@ Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. (* 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. *) +(* T-labeled leaves, for example GenTree.Leaf (x : T), *) +(* GenTree.Node 5 [:: t; t']. GenTree.tree is equipped *) +(* with canonical eqType, choiceType, and countType *) +(* instances, and so simple datatypes can be similarly *) +(* equipped by encoding into GenTree.tree and using *) +(* the mixins above. *) +(* CodeSeq.code == bijection from seq nat to nat. *) +(* CodeSeq.decode == bijection inverse to CodeSeq.code. *) (* 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.). *) @@ -48,7 +52,7 @@ Set Implicit Arguments. Unset Strict Implicit. Unset Printing Implicit Defensive. -(* Technical definitions about coding and decoding of nat sequances, which *) +(* Technical definitions about coding and decoding of nat sequences, which *) (* are used below to define various Canonical instances of the choice and *) (* countable interfaces. *) @@ -212,18 +216,18 @@ Canonical tree_eqType (T : eqType) := EqType (GenTree.tree T) (tree_eqMixin T). (* 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 *) +(* [x]choose does. This is not a restriction in a constructive theory, where *) +(* all types are concrete and hence countable. In the case of an axiomatic *) +(* theory, such as that of 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. *) +(* get 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}. *) +(* be 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 *) @@ -332,12 +336,12 @@ 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. +Proof. by move=> exP; exists (xchoose exP); apply: 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. +by have [x Px Qx] := exPQ; exists x; apply/andP. Qed. Lemma sig_eqW (vT : eqType) (lhs rhs : T -> vT) : @@ -363,13 +367,13 @@ 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. +Proof. by move=> Px0 Py0; rewrite /choose !insubT /=; apply: 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. +by rewrite insubT; apply: eq_xchoose. Qed. Section CanChoice. diff --git a/mathcomp/discrete/div.v b/mathcomp/basic/div.v index 8ffa036..8e9a3cf 100644 --- a/mathcomp/discrete/div.v +++ b/mathcomp/basic/div.v @@ -1,22 +1,23 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. +From mathcomp +Require Import 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 %/ d == quotient of the Euclidean division of m by d. *) +(* m %% d == remainder of the Euclidean division 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. *) +(* egcdn m n == the extended GCD (Bezout coefficient pair) of m and n. *) +(* If egcdn m n = (u, v), then gcdn m n = m * u - n * v. *) (* 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. *) @@ -44,12 +45,12 @@ 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. +by rewrite -{1}(subnKC le_dm) -addSn addnA -mulSnr; apply: 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. +move=> lt_rd; have d_gt0: 0 < d by apply: 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-> |] //=. @@ -78,7 +79,7 @@ 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. +by apply: IHn; apply: leq_trans le_mn; apply: leq_subr. Qed. Lemma edivn_def m d : edivn m d = (m %/ d, m %% d). @@ -334,7 +335,7 @@ 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. +Proof. by move=> d_dv_n /dvdnP[n1 ->]; apply: dvdn_mull. Qed. Lemma dvdn_eq d m : (d %| m) = (m %/ d * d == m). Proof. @@ -433,7 +434,7 @@ 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. +Proof. by rewrite addnC; apply: dvdn_addr. Qed. Lemma dvdn_add d m n : d %| m -> d %| n -> d %| m + n. Proof. by move/dvdn_addr->. Qed. @@ -520,7 +521,7 @@ 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. +by case: m n => [|m] [|n] //; apply: (@dvdn_gt0 _ m.+1) => //; apply: dvdn_gcdl. Qed. Lemma gcdnMDl k m n : gcdn m (k * m + n) = gcdn m n. @@ -628,7 +629,7 @@ 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. +Proof. by rewrite gcdnC; apply: Bezoutl. Qed. (* Back to the gcd. *) @@ -687,7 +688,7 @@ 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 lcm0n : left_zero 0 lcmn. Proof. by move=> n; apply: div0n. Qed. Lemma lcmn0 : right_zero 0 lcmn. Proof. by move=> n; rewrite lcmnC lcm0n. Qed. Lemma lcm1n : left_id 1 lcmn. @@ -866,13 +867,13 @@ 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. +Proof. by rewrite !(coprime_sym m); apply: 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. +Proof. by rewrite !(coprime_sym p); apply: coprime_dvdl. Qed. Lemma coprime_egcdn n m : n > 0 -> coprime (egcdn n m).1 (egcdn n m).2. Proof. diff --git a/mathcomp/discrete/finfun.v b/mathcomp/basic/finfun.v index 849bbe2..44c1ceb 100644 --- a/mathcomp/discrete/finfun.v +++ b/mathcomp/basic/finfun.v @@ -1,8 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -Require Import choice fintype tuple. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype tuple. (******************************************************************************) (* This file implements a type for functions with a finite domain: *) @@ -207,7 +206,7 @@ Lemma pffun_onP y D R f : (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=> _ /imageP[x Ax ->]; apply: f_fam. by move=> x Ax; apply: f_fam; apply/imageP; exists x. Qed. diff --git a/mathcomp/discrete/fingraph.v b/mathcomp/basic/fingraph.v index 34167d0..617188c 100644 --- a/mathcomp/discrete/fingraph.v +++ b/mathcomp/basic/fingraph.v @@ -1,8 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -Require Import path fintype. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path fintype. (******************************************************************************) (* This file develops the theory of finite graphs represented by an "edge" *) @@ -66,7 +65,7 @@ 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. +by apply: subset_trans (IHn _ _); apply/subsetP=> y; apply: predU1r. Qed. Inductive dfs_path v x y : Prop := @@ -84,7 +83,7 @@ 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. + by rewrite sub_dfs ?mem_head //; left; apply: 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]. @@ -94,13 +93,13 @@ apply: (@equivP (exists2 x1, x1 \in a & dfs_path v1 x1 y)); last first. 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{neq_yx not_vy}: y \notin v1 by apply/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]. + by rewrite sub_dfs //; left; exists x; [apply: mem_head | apply: 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]]. @@ -109,7 +108,7 @@ split=> [] [x1 a_x1 [p g_p p_y not_pv]]. 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. + suffices not_vx: x \notin v by apply/IHn; last apply: 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. @@ -156,7 +155,7 @@ 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. +Proof. by move->; apply: 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. @@ -188,10 +187,10 @@ 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. +Proof. by move/connect1; apply: 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. +Proof. by move/connect1; apply: same_connect_r. Qed. Lemma rootP x y : reflect (root x = root y) (connect x y). Proof. @@ -245,7 +244,7 @@ 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. +by move/IHp; apply: connect_trans; apply: e'e. Qed. Lemma relU_sym e e' : @@ -295,7 +294,7 @@ 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. + by move=> x y; rewrite sym_e; apply/idP/idP; apply: 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. @@ -305,18 +304,18 @@ 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. +by elim: p y a_y e_p => //= y p IHp x a_x /andP[/cl_a/(_ a_x)]; apply: 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. +by elim: p x e_p => //= y p IHp x /andP[/cl_a->]; apply: IHp. Qed. Lemma connect_closed x : closed e (connect e x). -Proof. by move=> y z /connect1/same_connect_r; exact. Qed. +Proof. by move=> y z /connect1/same_connect_r; apply. Qed. Lemma predC_closed a : closed e a -> closed e [predC a]. Proof. by move=> cl_a x y /cl_a; rewrite !inE => ->. Qed. @@ -324,14 +323,14 @@ 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. +by apply: eq_disjoint; apply: 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. +Proof. by apply/subsetP; apply: mem_closure. Qed. Lemma n_comp_closure2 x y : n_comp e (closure e (pred2 x y)) = (~~ connect e x y).+1. @@ -367,7 +366,7 @@ 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 ]. +by exists (traject f (f x) n); [apply: fpath_traject | rewrite last_traject]. Qed. Lemma fconnect1 x : fconnect f x (f x). @@ -386,13 +385,13 @@ 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. +by apply/subsetP=> _ /trajectP[i _ ->]; apply: 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. + by rewrite last_traject; apply/loopingP/looping_order. exact: fconnect_iter. Qed. @@ -452,7 +451,7 @@ 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. +Proof. by rewrite -(card_uniqP Up); apply (eq_card fconnect_cycle). Qed. Lemma orbit_rot_cycle : {i : nat | orbit x = rot i p}. Proof. @@ -478,10 +477,10 @@ 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. +Proof. by exists finv; [apply finv_f | apply f_finv]. Qed. Lemma finv_bij : bijective finv. -Proof. exists f; [ exact f_finv | exact finv_f ]. Qed. +Proof. by exists f; [apply f_finv | apply finv_f]. Qed. Lemma finv_inj : injective finv. Proof. exact (can_inj f_finv). Qed. @@ -496,7 +495,7 @@ 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. +Proof. by rewrite -orderSpred iterS; apply (f_finv x). Qed. Lemma iter_finv n x : n <= order x -> iter n finv x = iter (order x - n) f x. Proof. @@ -538,7 +537,7 @@ symmetry; transitivity #|preim (froot f) b|. 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). + by rewrite eq_card0 // => x; apply: (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 //]]. @@ -548,7 +547,7 @@ 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. +Proof. by move=> cl_a x; apply: cl_a (eqxx _). Qed. Lemma same_fconnect1 x : fconnect f x =1 fconnect f (f x). Proof. by apply: same_connect1 => /=. Qed. @@ -577,7 +576,7 @@ 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. +Proof. by apply/eqP; apply: froots_id. Qed. Lemma fcard_id (a : pred T) : fcard id a = #|a|. Proof. by apply: eq_card => x; rewrite inE froots_id. Qed. @@ -590,7 +589,8 @@ 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))). +move=> fK; have inj_f := can_inj fK. +by apply: bij_can_eq fK; [apply: fin_inj_bij | apply: finv_f]. Qed. Hypothesis eq_f : f =1 f'. @@ -624,7 +624,7 @@ 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. +Proof. by move=> x; apply: 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. diff --git a/mathcomp/discrete/finset.v b/mathcomp/basic/finset.v index 78c1753..ef63a24 100644 --- a/mathcomp/discrete/finset.v +++ b/mathcomp/basic/finset.v @@ -1,8 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -Require Import bigop choice fintype div finfun. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat div seq choice fintype. +From mathcomp +Require Import finfun bigop. (******************************************************************************) (* This file defines a type for sets over a finite Type, similar to the type *) @@ -316,7 +317,7 @@ 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. +Proof. by move=> sAB; rewrite eqEsubset sAB; apply: subset_leqif_card. Qed. Lemma in_set0 x : x \in set0 = false. Proof. by rewrite inE. Qed. @@ -336,7 +337,7 @@ 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. +by apply/setP=> x; rewrite inE; apply: A0. Qed. Lemma enum_set0 : enum set0 = [::] :> seq T. @@ -356,7 +357,7 @@ 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. +Proof. by rewrite inE; apply: eqP. Qed. Lemma enum_setT : enum [set: T] = Finite.enum T. Proof. by rewrite (eq_enum (in_set _)) enumT. Qed. @@ -374,7 +375,7 @@ 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. +Proof. by rewrite !inE; apply: predU1P. Qed. Lemma in_setU1 x a B : (x \in a |: B) = (x == a) || (x \in B). Proof. by rewrite !inE. Qed. @@ -403,7 +404,7 @@ 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. +Proof. by rewrite !inE; apply: andP. Qed. Lemma in_setD1 x A b : (x \in A :\ b) = (x != b) && (x \in A) . Proof. by rewrite !inE. Qed. @@ -420,7 +421,7 @@ 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. +Proof. by rewrite !inE; apply: pred2P. Qed. Lemma in_set2 x a b : (x \in [set a; b]) = (x == a) || (x == b). Proof. by rewrite !inE. Qed. @@ -432,7 +433,7 @@ 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. +Proof. by rewrite !inE; apply: orP. Qed. Lemma in_setU x A B : (x \in A :|: B) = (x \in A) || (x \in B). Proof. exact: in_set. Qed. @@ -443,7 +444,7 @@ 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). +by case: (x \in C) => //; apply: (subsetP sAB). Qed. Lemma setSU A B C : A \subset B -> A :|: C \subset B :|: C. @@ -489,11 +490,11 @@ Proof. by rewrite !(setUC A) setUUl. Qed. (* 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. +Proof. by rewrite !inE; apply: 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. +Proof. by rewrite !inE; apply: 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. @@ -510,7 +511,7 @@ 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). +by case: (x \in C) => //; apply: (subsetP sAB). Qed. Lemma setSI A B C : A \subset B -> A :&: C \subset B :&: C. @@ -582,7 +583,7 @@ 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. +Proof. by rewrite !inE; apply: negP. Qed. Lemma in_setC x A : (x \in ~: A) = (x \notin A). Proof. exact: in_set. Qed. @@ -626,7 +627,7 @@ 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. +Proof. by rewrite inE andbC; apply: andP. Qed. Lemma in_setD A B x : (x \in A :\: B) = (x \notin B) && (x \in A). Proof. exact: in_set. Qed. @@ -635,10 +636,10 @@ 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. +Proof. by rewrite !setDE; apply: setSI. Qed. Lemma setDS A B C : A \subset B -> C :\: B \subset C :\: A. -Proof. by rewrite !setDE -setCS; exact: setIS. Qed. +Proof. by rewrite !setDE -setCS; apply: 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. @@ -711,7 +712,7 @@ 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. +Proof. exact/eq_card/in_set. Qed. Lemma sum1dep_card pA : \sum_(x | pA x) 1 = #|[set x | pA x]|. Proof. by rewrite sum1_card cardsE. Qed. @@ -726,7 +727,7 @@ 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. +Proof. by rewrite -cards_eq0; apply: existsP. Qed. Lemma card_gt0 A : (0 < #|A|) = (A != set0). Proof. by rewrite lt0n cards_eq0. Qed. @@ -823,26 +824,26 @@ 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. +by apply/setP=> x; rewrite inE; apply/andb_idr/sAB. Qed. Implicit Arguments setIidPl [A B]. Lemma setIidPr A B : reflect (A :&: B = B) (B \subset A). -Proof. rewrite setIC; exact: setIidPl. Qed. +Proof. by rewrite setIC; apply: 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. +by rewrite -setCS (sameP setIidPl eqP) -setCU (inj_eq setC_inj); apply: eqP. Qed. Lemma setUidPr A B : reflect (A :|: B = B) (A \subset B). -Proof. rewrite setUC; exact: setUidPl. Qed. +Proof. by rewrite setUC; apply: setUidPl. Qed. Lemma setDidPl A B : reflect (A :\: B = A) [disjoint A & B]. -Proof. rewrite setDE disjoints_subset; exact: setIidPl. Qed. +Proof. by rewrite setDE disjoints_subset; apply: 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. @@ -854,7 +855,7 @@ 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. +Proof. by rewrite subsetI; apply: andP. Qed. Lemma subsetIidl A B : (A \subset A :&: B) = (A \subset B). Proof. by rewrite subsetI subxx. Qed. @@ -869,10 +870,10 @@ 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. +Proof. by rewrite -!(setCS _ A) setCU; apply: 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. +Proof. by rewrite subUset; apply: andP. Qed. Lemma subsetC A B : (A \subset ~: B) = (B \subset ~: A). Proof. by rewrite -setCS setCK. Qed. @@ -892,7 +893,7 @@ Qed. Lemma subsetDP A B C : reflect (A \subset B /\ [disjoint A & C]) (A \subset B :\: C). -Proof. by rewrite subsetD; exact: andP. Qed. +Proof. by rewrite subsetD; apply: andP. Qed. Lemma setU_eq0 A B : (A :|: B == set0) = (A == set0) && (B == set0). Proof. by rewrite -!subset0 subUset. Qed. @@ -910,7 +911,7 @@ 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. +Proof. by rewrite subsetD1; apply: andP. Qed. Lemma properD1 A x : x \in A -> A :\ x \proper A. Proof. @@ -1009,7 +1010,7 @@ 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. +Proof. by rewrite inE; apply: andP. Qed. Lemma cardsX : #|setX| = #|A1| * #|A2|. Proof. by rewrite cardsE cardX. Qed. @@ -1193,7 +1194,7 @@ 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. +Proof. by rewrite [@imset]unlock inE; apply: 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. @@ -1239,7 +1240,7 @@ 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. +Proof. by move=> sAB; apply/subsetP=> y; rewrite !inE; apply: subsetP. Qed. Lemma preimset0 : f @^-1: set0 = set0. Proof. by apply/setP=> x; rewrite !inE. Qed. @@ -1274,7 +1275,7 @@ 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. +by case/imsetP=> y Ay /injf-> //; apply: subsetP sAB y Ay. Qed. Lemma preimset_proper (A B : {set rT}) : @@ -1419,7 +1420,7 @@ 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. +rewrite (eq_bigl hA) => [|j]; last by apply/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. @@ -1483,7 +1484,7 @@ 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. +Proof. by rewrite [@imset]unlock cardsE; apply: image_injP. Qed. Lemma can2_in_imset_pre : {in D, cancel f g} -> {on D, cancel g & f} -> f @: D = g @^-1: D. @@ -1493,7 +1494,7 @@ 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. +Proof. by move=> fK gK; apply: can2_in_imset_pre; apply: in1W. Qed. End CardFunImage. @@ -1511,7 +1512,7 @@ Lemma can_imset_pre (T : finType) f g (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. +exact/(subset_cardP (card_codom (can_inj fK)))/subsetP. Qed. Lemma imset_id (T : finType) (A : {set T}) : [set x | x in A] = A. @@ -1521,7 +1522,7 @@ 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. +have ontof: _ \in codom f by apply/(subset_cardP (card_codom injf))/subsetP. by exists (fun x => iinv (ontof x)) => x; rewrite (f_iinv, iinv_f). Qed. @@ -1529,7 +1530,7 @@ 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 sDA: (D \subset A) = (sf \subset A) by apply: eq_subset; apply: 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. @@ -1613,13 +1614,13 @@ 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. +Proof. by move=> Pj sUF; apply: 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 apply: subsetP x; apply: bigcup_sup. by elim/big_rec: _ => [|i _ Pi _ /setUP[|//]]; [rewrite inE | exists i]. Qed. @@ -1627,8 +1628,8 @@ 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)). + by apply: subset_trans sFU; apply: bigcup_sup. +by apply/subsetP=> x /bigcupP[i Pi]; apply: (subsetP (sFU i Pi)). Qed. Lemma bigcup_disjoint U P F : @@ -1663,13 +1664,13 @@ 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. +Proof. by move=> Pj; apply: 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. + by apply: subset_trans sUF _; apply: bigcap_inf. elim/big_rec: _ => [|i V Pi sUV]; apply/subsetP=> x Ux; rewrite inE //. by rewrite !(subsetP _ x Ux) ?sUF. Qed. @@ -1793,8 +1794,8 @@ 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]; +apply: (iffP IHP) => [] tI B C PB PC; last by apply: tI; apply: setU1r. +by case/setU1P: PC PB => [->|PC] /setU1P[->|PB]; try by [apply: tI | case/eqP]; first rewrite disjoint_sym; rewrite disjoints_subset disjA. Qed. @@ -1832,7 +1833,7 @@ 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. +by case/andP=> PA _{y} /= Ax; apply: def_pblock. Qed. Lemma eq_pblock P x y : @@ -1852,14 +1853,14 @@ 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)]. + by [apply: (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. +by exists (F i); first apply: mem_imset. Qed. Lemma trivIimset J F (P := F @: J) : @@ -1882,7 +1883,7 @@ 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. +by move=> uniP /card_partition->; rewrite -sum_nat_const; apply: eq_bigr. Qed. Section BigOps. @@ -1898,23 +1899,23 @@ 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. +by case/andP=> Px _; apply: 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. +by rewrite -biginT big_trivIset_cond //; apply: eq_bigr => A _; apply: 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. +Proof. by case/and3P=> /eqP <- tI_P _; apply: 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. +Proof. by case/and3P=> /eqP <- tI_P _; apply: big_trivIset. Qed. Lemma partition_disjoint_bigcup (F : I -> {set T}) E : (forall i j, i != j -> [disjoint F i & F j]) -> @@ -2063,7 +2064,7 @@ 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. +Proof. by move=> x0; apply: (can_in_inj (pblockK x0)). Qed. Lemma pblock_transversal : pblock P @: X = P. Proof. @@ -2074,12 +2075,12 @@ 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. +Proof. by rewrite -pblock_transversal card_in_imset //; apply: 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. +by apply: eq_in_imset; apply: pblockK. Qed. End Transversals. @@ -2137,7 +2138,7 @@ 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. +by apply/implyP=> sBA; apply/eqP; apply/eqP/idP=> [-> // | /minA]; apply. Qed. Implicit Arguments minsetP [P A]. @@ -2145,7 +2146,7 @@ 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. +Proof. by case/minsetP=> _; apply. Qed. Lemma ex_minset P : (exists A, P A) -> {A | minset P A}. Proof. @@ -2192,7 +2193,7 @@ 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. +Proof. by case/maxsetP=> _; apply. Qed. Lemma ex_maxset P : (exists A, P A) -> {A | maxset P A}. Proof. diff --git a/mathcomp/discrete/fintype.v b/mathcomp/basic/fintype.v index 347c823..66ac880 100644 --- a/mathcomp/discrete/fintype.v +++ b/mathcomp/basic/fintype.v @@ -1,8 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -Require Import choice. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice. (******************************************************************************) (* The Finite interface describes Types with finitely many elements, *) @@ -20,7 +19,7 @@ Require Import choice. (* 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. *) +(* 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 *) @@ -37,12 +36,15 @@ Require Import choice. (* [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 *) +(* 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. *) +(* seq_sub s == the subType of all x \in s, where s : seq T for some *) +(* eqType T; seq_sub s has a canonical finType instance *) +(* when T is a choiceType. *) +(* adhoc_seq_sub_choiceType s, adhoc_seq_sub_finType s == *) +(* non-canonical instances for seq_sub s, s : seq T, *) +(* which can be used when T is not a choiceType. *) (* 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, *) @@ -487,14 +489,14 @@ CoInductive pick_spec : option T -> Type := Lemma pickP : pick_spec (pick P). Proof. rewrite /pick; case: (enum _) (mem_enum P) => [|x s] Pxs /=. - by right; exact: fsym. + by right; apply: 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. +Proof. by move=> eqPQ; apply: 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. @@ -506,7 +508,7 @@ 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. +Proof. by move <-; apply: eq_card. Qed. Lemma card0 : #|@pred0 T| = 0. Proof. by rewrite cardE enum0. Qed. @@ -567,12 +569,12 @@ 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. +by rewrite cardU1 /=; case: (~~ _) => //; apply: leqW. Qed. Lemma card_uniqP s : reflect (#|s| = size s) (uniq s). Proof. -elim: s => [|x s IHs]; first by left; exact card0. +elim: s => [|x s IHs]; first by left; apply: card0. rewrite cardU1 /= /addn; case: {+}(x \in s) => /=. by right=> card_Ssz; have:= card_size s; rewrite card_Ssz ltnn. by apply: (iffP IHs) => [<-| [<-]]. @@ -582,7 +584,7 @@ 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. +Proof. by apply: (iffP eqP); [apply: card0_eq | apply: eq_card0]. Qed. Lemma pred0Pn P : reflect (exists x, P x) (~~ pred0b P). Proof. @@ -592,7 +594,7 @@ 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. +Proof. by rewrite lt0n; apply: pred0Pn. Qed. Lemma subsetE A B : (A \subset B) = pred0b [predD A & B]. Proof. by rewrite unlock. Qed. @@ -601,7 +603,7 @@ 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. +by rewrite andbC -negb_imply; apply/negbF/implyP; apply: sAB. Qed. Lemma subsetPn A B : @@ -616,7 +618,7 @@ 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. +by rewrite !inE andbC; case Ax: (x \in A) => //; apply: subsetP Ax. Qed. Lemma subxx_hint (mA : mem_pred T) : subset mA mA. @@ -649,10 +651,10 @@ 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. +Proof. by move/subsetP=> allA x; apply: 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. +Proof. by apply/subsetP/idP=> [-> // | Ax y /eqP-> //]; apply: eqxx. Qed. Lemma subset_eqP A B : reflect (A =i B) ((A \subset B) && (B \subset A)). Proof. @@ -678,11 +680,11 @@ 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. +by move/subsetP=> sAB /subsetP=> sBC; apply/subsetP=> x /sAB; apply: sBC. Qed. Lemma subset_all s A : (s \subset A) = all (mem A) s. -Proof. by exact (sameP (subsetP _ _) allP). Qed. +Proof. exact: (sameP (subsetP _ _) allP). Qed. Lemma properE A B : A \proper B = (A \subset B) && ~~(B \subset A). Proof. by []. Qed. @@ -703,7 +705,7 @@ 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. +by apply: contra nBy; apply: subsetP. Qed. Lemma proper_sub_trans A B C : A \proper B -> B \subset C -> A \proper C. @@ -715,7 +717,7 @@ 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. +by apply/subsetPn; exists x => //; apply: contra nBx; apply: subsetP. Qed. Lemma proper_card A B : A \proper B -> #|A| < #|B|. @@ -743,7 +745,7 @@ 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. +Proof. by congr (_ == 0); apply: eq_card => x; apply: andbC. Qed. Lemma eq_disjoint A1 A2 : A1 =i A2 -> disjoint (mem A1) =1 disjoint (mem A2). Proof. @@ -766,13 +768,13 @@ 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. +Proof. by rewrite 2!disjoint_subset; apply: 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. +Proof. by move/eq_disjoint->; apply: disjoint0. Qed. Lemma disjoint1 x A : [disjoint pred1 x & A] = (x \notin A). Proof. @@ -783,7 +785,7 @@ 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. +Proof. by move/eq_disjoint->; apply: disjoint1. Qed. Lemma disjointU A B C : [disjoint predU A B & C] = [disjoint A & C] && [disjoint B & C]. @@ -944,7 +946,7 @@ 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. +Let exFP : exists n, FP n. Proof. by exists (F i0); apply: FP_F. Qed. Lemma arg_minP : extremum_spec leq arg_min. Proof. @@ -1036,7 +1038,7 @@ 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. + by case/eqP: neqxy; apply: 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. @@ -1049,7 +1051,7 @@ apply: (iffP (dinjectivePn _)) => [[x _ [y nxy eqfxy]] | [x [y nxy eqfxy]]]; Qed. Lemma injectiveP : reflect (injective f) injectiveb. -Proof. apply: (iffP (dinjectiveP _)) => injf x y => [|_ _]; exact: injf. Qed. +Proof. by apply: (iffP (dinjectiveP _)) => injf x y => [|_ _]; apply: injf. Qed. End Injectiveb. @@ -1117,7 +1119,7 @@ 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]. +by move=> injf x fAfx Ax; apply: injf => //; [apply: mem_iinv | apply: f_iinv]. Qed. Lemma preim_iinv A B y fAy : preim f B (@iinv A y fAy) = B y. @@ -1127,10 +1129,10 @@ 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. +Proof. by apply: image_f. Qed. Lemma image_codom A : {subset image f A <= codom f}. -Proof. by move=> _ /imageP[x _ ->]; exact: codom_f. Qed. +Proof. by move=> _ /imageP[x _ ->]; apply: codom_f. Qed. Lemma image_pred0 : image f pred0 =i pred0. Proof. by move=> x; rewrite /image_mem /= enum0. Qed. @@ -1150,7 +1152,7 @@ Lemma image_iinv A y (fTy : y \in codom f) : 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. +Proof. by apply: in_iinv_f; first apply: 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. @@ -1174,7 +1176,7 @@ Fixpoint preim_seq s := 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 rewrite /= fx_y IHs // => z s_z; apply: fTs; apply: predU1r. by case/imageP: (fTs y (mem_head y s)) => x _ fx_y; case/eqP: (nfTy x). Qed. @@ -1202,7 +1204,7 @@ 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. +by rewrite map_inj_in_uniq ?enum_uniq // => x y; rewrite !mem_enum; apply: injf. Qed. Lemma image_injP A : reflect {in A &, injective f} (#|image f A| == #|A|). @@ -1214,7 +1216,7 @@ Qed. Hypothesis injf : injective f. Lemma card_image A : #|image f A| = #|A|. -Proof. apply: card_in_image; exact: in2W. Qed. +Proof. by apply: card_in_image; apply: in2W. Qed. Lemma card_codom : #|codom f| = #|T|. Proof. exact: card_image. Qed. @@ -1249,8 +1251,8 @@ 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 invF_f : cancel f invF. Proof. by move=> x; apply: iinv_f. Qed. +Lemma f_invF : cancel invF f. Proof. by move=> y; apply: f_iinv. Qed. Lemma injF_bij : bijective f. Proof. exact: inj_card_bij. Qed. End Inv. @@ -1296,56 +1298,6 @@ 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. @@ -1428,9 +1380,6 @@ End SubFinType. 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). @@ -1446,7 +1395,7 @@ 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. +by apply: eq_filter => x; apply: isSome_insub. Qed. (* We can't declare a canonical structure here because we've already *) @@ -1504,6 +1453,73 @@ Proof. exact: card_sub. Qed. End CardSig. +(* Subtype for an explicit enumeration. *) +Section SeqSubType. + +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_countMixin := CountMixin seq_sub_pickleK. +Fact seq_sub_axiom : Finite.axiom seq_sub_enum. +Proof. exact: Finite.uniq_enumP (undup_uniq _) mem_seq_sub_enum. Qed. +Definition seq_sub_finMixin := Finite.Mixin seq_sub_countMixin seq_sub_axiom. + +(* Beware: these are not the canonical instances, as they are not consistent *) +(* the generic sub_choiceType canonical instance. *) +Definition adhoc_seq_sub_choiceMixin := PcanChoiceMixin seq_sub_pickleK. +Definition adhoc_seq_sub_choiceType := + Eval hnf in ChoiceType seq_sub adhoc_seq_sub_choiceMixin. +Definition adhoc_seq_sub_finType := + [finType of seq_sub for FinType adhoc_seq_sub_choiceType seq_sub_finMixin]. + +End SeqSubType. + +Section SeqFinType. + +Variables (T : choiceType) (s : seq T). +Local Notation sT := (seq_sub s). + +Definition seq_sub_choiceMixin := [choiceMixin of sT by <:]. +Canonical seq_sub_choiceType := Eval hnf in ChoiceType sT seq_sub_choiceMixin. + +Canonical seq_sub_countType := Eval hnf in CountType sT (seq_sub_countMixin s). +Canonical seq_sub_subCountType := Eval hnf in [subCountType of sT]. +Canonical seq_sub_finType := Eval hnf in FinType sT (seq_sub_finMixin s). +Canonical seq_sub_subFinType := Eval hnf in [subFinType of sT]. + +Lemma card_seq_sub : uniq s -> #|{:sT}| = size s. +Proof. +by move=> Us; rewrite cardE enumT -(size_map val) unlock val_seq_sub_enum. +Qed. + +End SeqFinType. + + (**********************************************************************) (* *) (* Ordinal finType : {0, ... , n-1} *) @@ -1577,7 +1593,7 @@ 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. +Proof. by apply: val_inj; apply: nth_enum_ord. Qed. Lemma index_enum_ord (i : 'I_n) : index i (enum 'I_n) = i. Proof. @@ -1604,11 +1620,11 @@ 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. +Proof. by move=> i; apply: 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. +Proof. by move=> i; apply: 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. @@ -1836,7 +1852,7 @@ 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. +by rewrite negbK in def_i; apply/eqP. Qed. Lemma neq_lift n (h : 'I_n) i : h != lift h i. @@ -1866,7 +1882,7 @@ 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. +Proof. by apply: leq_trans (valP i) _; apply: leq_addr. Qed. Lemma rshift_subproof m n (i : 'I_n) : m + i < m + n. Proof. by rewrite ltn_add2l. Qed. diff --git a/mathcomp/discrete/generic_quotient.v b/mathcomp/basic/generic_quotient.v index 38696d8..d835b32 100644 --- a/mathcomp/discrete/generic_quotient.v +++ b/mathcomp/basic/generic_quotient.v @@ -1,9 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) (* -*- coding : utf-8 -*- *) + Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -Require Import choice fintype. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat choice seq fintype. (*****************************************************************************) (* Provided a base type T, this files defines an interface for quotients Q *) diff --git a/mathcomp/discrete/opam b/mathcomp/basic/opam index 3646845..c72ed39 100644 --- a/mathcomp/discrete/opam +++ b/mathcomp/basic/opam @@ -1,5 +1,5 @@ opam-version: "1.2" -name: "coq:mathcomp:discrete" +name: "coq:mathcomp:basic" version: "1.5" maintainer: "Ssreflect <ssreflect@msr-inria.inria.fr>" authors: "Ssreflect <ssreflect@msr-inria.inria.fr>" @@ -8,5 +8,5 @@ bug-reports: "ssreflect@msr-inria.inria.fr" license: "CeCILL-B" build: [ make "-j" "%{jobs}%" ] install: [ make "install" ] -remove: [ "sh" "-c" "rm -rf '%{lib}%/coq/user-contrib/mathcomp/discrete'" ] +remove: [ "sh" "-c" "rm -rf '%{lib}%/coq/user-contrib/mathcomp/basic'" ] depends: [ "coq:mathcomp:ssreflect" { = "1.5" } ] diff --git a/mathcomp/discrete/path.v b/mathcomp/basic/path.v index cea608f..1ef7724 100644 --- a/mathcomp/discrete/path.v +++ b/mathcomp/basic/path.v @@ -1,7 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq. (******************************************************************************) (* The basic theory of paths over an eqType; this file is essentially a *) @@ -119,7 +119,7 @@ Lemma pathP x p x0 : 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]. +by split; [apply: (e_p 0) | apply/(IHp y) => i; apply: e_p i.+1]. Qed. Definition cycle p := if p is x :: p' then path x (rcons p' x) else true. @@ -144,7 +144,7 @@ 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. +Proof. by move=> ee' [|x p] //=; apply: 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. @@ -182,7 +182,7 @@ CoInductive splitl x1 x : seq T -> Type := 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. +by split; apply: last_rcons. Qed. CoInductive splitr x : seq T -> Type := @@ -214,7 +214,7 @@ Lemma next_nth p 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. + by case: ifP => // _; apply: IHp. Qed. Lemma prev_nth p x : @@ -224,7 +224,7 @@ Lemma prev_nth p 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. +by case: ifP => // _; apply: IHp. Qed. Lemma mem_next p x : (next p x \in p) = (x \in p). @@ -361,13 +361,13 @@ CoInductive shorten_spec x p : T -> seq T -> Type := 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. +move=> e_p; have: x \in x :: p by apply: 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. + by split=> // y /p'p; apply: 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. @@ -392,7 +392,7 @@ 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. +Proof. by case: s => //= y s -> //; apply: mem_head. Qed. Section Transitive. @@ -448,8 +448,8 @@ Lemma eq_sorted_irr : irreflexive leT -> 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. + by move=> m n /andP[? ltnm]; case/idP: (leT_irr m); apply: leT_tr ltnm. +by move/eq_sorted; apply=> //; apply: uniq_perm_eq => //; apply: sorted_uniq. Qed. End Transitive. @@ -525,10 +525,10 @@ have: sorted s -> sorted (merge_sort_pop s ss). 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]). +have: sorted s1 by apply: (@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 _ _)]. +by case: {1}s2=> /= [|_ _]; [rewrite ord_s1 | apply: IHss (merge_sorted _ _)]. Qed. Lemma perm_sort s : perm_eql (sort s) s. @@ -579,7 +579,7 @@ 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)). +by rewrite eqn_leq lenm; apply: (allP (order_path_min leq_trans le_ms)). Qed. Lemma iota_sorted i n : sorted leq (iota i n). @@ -622,7 +622,7 @@ 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. +by case: ltngtP le_i_n => [? _||->] //; apply: IHn. Qed. End Trajectory. @@ -729,7 +729,7 @@ 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). +by apply: eq_path; rewrite -def_p; apply: (can2_eq prev_next next_prev). Qed. Lemma cycle_from_next : (forall x, x \in p -> e x (next p x)) -> cycle e p. diff --git a/mathcomp/discrete/prime.v b/mathcomp/basic/prime.v index 4c3bdf8..88b9229 100644 --- a/mathcomp/discrete/prime.v +++ b/mathcomp/basic/prime.v @@ -1,8 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -Require Import path fintype div bigop. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq path fintype. +From mathcomp +Require Import div bigop. (******************************************************************************) (* This file contains the definitions of: *) @@ -65,7 +66,7 @@ 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. +by rewrite addSnnS; apply: IHr. Qed. Fixpoint elogn2 e q r {struct q} := @@ -82,9 +83,9 @@ 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. + by move/ltnW; apply: IHq. clear 1; rewrite subn1 -[_.-1.+1]doubleS -mul2n mulnA -expnSr. -rewrite -{1}(addKn q q) addnn; exact: IHq. +by rewrite -{1}(addKn q q) addnn; apply: IHq. Qed. Definition ifnz T n (x y : T) := if n is 0 then y else x. @@ -186,7 +187,7 @@ have leq_pd_ok m p q pd: q <= p -> pd_ok p m pd -> pd_ok q m pd. 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 mul1n; apply: leq_pd_ok; apply: 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. @@ -237,7 +238,7 @@ 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]. + by case: ifnzP; [move/ltn_predK->; apply: 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). @@ -355,7 +356,7 @@ 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. +Proof. by move/prime_gt1; apply: ltnW. Qed. Hint Resolve prime_gt1 prime_gt0. @@ -439,7 +440,7 @@ Qed. Lemma sorted_primes n : sorted ltn (primes n). Proof. -by case: (posnP n) => [-> // | /prime_decomp_correct[_ _]]; exact: path_sorted. +by case: (posnP n) => [-> // | /prime_decomp_correct[_ _]]; apply: path_sorted. Qed. Lemma eq_primes m n : (primes m =i primes n) <-> (primes m = primes n). @@ -503,7 +504,7 @@ 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. +by move/(allP (order_path_min ltn_trans ord_pm)); apply: ltnW. Qed. Lemma max_pdiv_max n p : p \in \pi(n) -> p <= max_pdiv n. @@ -742,7 +743,7 @@ 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. +have d_gt0: d > 0 by apply: 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. @@ -851,7 +852,7 @@ Section PnatTheory. Implicit Types (n p : nat) (pi rho : nat_pred). Lemma negnK pi : pi^'^' =i pi. -Proof. move=> p; exact: negbK. Qed. +Proof. by move=> p; apply: 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. @@ -884,7 +885,7 @@ 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. +Proof. by apply: eq_partn; apply: negnK. Qed. Lemma widen_partn m pi n : n <= m -> n`_pi = \prod_(0 <= p < m.+1 | p \in pi) p ^ logn p n. @@ -903,7 +904,7 @@ 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. +have le_pmul m' n': m' > 0 -> n' <= m' * n' by move/prednK <-; apply: 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. @@ -964,7 +965,7 @@ 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. +by case=> _ n_gt0 dv_p_n; apply: leq_ltn_trans lt_n_m; apply: dvdn_leq. Qed. Lemma partn_pi n : n > 0 -> n`_\pi(n) = n. @@ -1048,7 +1049,7 @@ 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). +by apply/allP=> p pr_p; apply: subpi => //; apply: (allP pi_n). Qed. Lemma eq_in_pnat pi rho n : {in \pi(n), pi =i rho} -> pi.-nat n = rho.-nat n. @@ -1069,7 +1070,7 @@ 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)]. +by rewrite primes_mul // => /orP[]; [apply: (allP pi_m) | apply: (allP pi_n)]. Qed. Lemma pnat_exp pi m n : pi.-nat (m ^ n) = pi.-nat m || (n == 0). @@ -1093,7 +1094,7 @@ 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. +Proof. by rewrite /pnat => ->; apply/allP. Qed. Lemma pi_of_dvd m n : m %| n -> n > 0 -> {subset \pi(m) <= \pi(n)}. Proof. @@ -1102,7 +1103,7 @@ 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. +Proof. by move=> m_gt0 n_gt0 p; apply: 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. @@ -1123,7 +1124,7 @@ 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. +Proof. by case/andP=> _ /allP; apply. 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. @@ -1136,9 +1137,8 @@ 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. +case/andP=> m_gt0 pi_m /andP[n_gt0 pi'_n]; rewrite coprime_has_primes //. +by apply/hasPn=> p /(allP pi'_n); apply/contra/allP. Qed. Lemma p'nat_coprime pi m n : pi^'.-nat m -> pi.-nat n -> coprime m n. @@ -1147,11 +1147,11 @@ 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. +by move=> pi'rho pi_m; move/(sub_in_pnat (in1W pi'rho)); apply: pnat_coprime. Qed. Lemma coprime_partC pi m n : coprime m`_pi n`_pi^'. -Proof. by apply: (@pnat_coprime pi); exact: part_pnat. Qed. +Proof. by apply: (@pnat_coprime pi); apply: part_pnat. Qed. Lemma pnat_1 pi n : pi.-nat n -> pi^'.-nat n -> n = 1. Proof. @@ -1211,7 +1211,7 @@ 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. +by apply: sub_in_pnat (part_pnat _ _) => q _; apply/contra/pi_sub_rho. Qed. Lemma partnI pi rho n : n`_[predI pi & rho] = n`_pi`_rho. @@ -1242,7 +1242,7 @@ 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 lt0p: 0 < p by apply: 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. @@ -1295,13 +1295,13 @@ Proof. by move/dvdn_divisors <-. Qed. 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. +Proof. by move=> dF; elim/big_ind: _ => //; apply: 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. + by apply: dvdn_trans n_dvd_m; apply: 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 /=. @@ -1365,7 +1365,7 @@ 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. +pose p := pdiv n; have p_pr: prime p by apply: 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. @@ -1384,7 +1384,7 @@ have ->: totient np = #|[pred d : 'I_np | coprime np 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. + by rewrite eqn_pmul2l // => eq_op12; apply/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]]. diff --git a/mathcomp/discrete/tuple.v b/mathcomp/basic/tuple.v index 450106a..a3adfe7 100644 --- a/mathcomp/discrete/tuple.v +++ b/mathcomp/basic/tuple.v @@ -1,8 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -Require Import choice fintype. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. Set Implicit Arguments. Unset Strict Implicit. @@ -15,7 +14,7 @@ Unset Printing Implicit Defensive. (* 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] == the empty tuple. *) (* [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). *) @@ -226,7 +225,7 @@ CoInductive tuple1_spec : n.+1.-tuple T -> Type := 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. +by rewrite (_ : Tuple _ = [tuple of x :: t]) //; apply: val_inj. Qed. Lemma tnth_map f t i : tnth [tuple of map f t] i = f (tnth t i) :> rT. @@ -239,7 +238,7 @@ Lemma tnth_behead n T (t : n.+1.-tuple T) i : 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. +Proof. by case/tupleP: t => x t; apply: val_inj. Qed. Section TupleQuantifiers. diff --git a/mathcomp/character/Make b/mathcomp/character/Make index 646641d..6d1f18c 100644 --- a/mathcomp/character/Make +++ b/mathcomp/character/Make @@ -1,4 +1,4 @@ -all.v +all_character.v character.v classfun.v finfield.v diff --git a/mathcomp/character/all.v b/mathcomp/character/all_character.v index 927d9a0..927d9a0 100644 --- a/mathcomp/character/all.v +++ b/mathcomp/character/all_character.v diff --git a/mathcomp/character/character.v b/mathcomp/character/character.v index 166fcc5..146d965 100644 --- a/mathcomp/character/character.v +++ b/mathcomp/character/character.v @@ -1,27 +1,17 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype tuple finfun bigop prime finset. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism quotient action gproduct. -From mathcomp.algebra -Require Import ssralg poly finalg zmodp cyclic. -From mathcomp.algebra -Require Import matrix mxalgebra mxpoly vector ssrnum. -From mathcomp.solvable -Require Import commutator center pgroup nilpotent sylow abelian. -From mathcomp.field -Require Import algC. -Require Import mxrepresentation classfun. - -Set Implicit Arguments. -Unset Strict Implicit. -Unset Printing Implicit Defensive. - -Import GroupScope GRing.Theory Num.Theory. -Local Open Scope ring_scope. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp +Require Import fintype tuple finfun bigop prime ssralg poly finset gproduct. +From mathcomp +Require Import fingroup morphism perm automorphism quotient finalg action. +From mathcomp +Require Import zmodp commutator cyclic center pgroup nilpotent sylow abelian. +From mathcomp +Require Import matrix mxalgebra mxpoly mxrepresentation vector ssrnum algC. +From mathcomp +Require Import classfun. (******************************************************************************) (* This file contains the basic notions of character theory, based on Isaacs. *) @@ -78,6 +68,13 @@ Local Open Scope ring_scope. (* class_Iirr xG == the index of xG \in classes G, in Iirr G. *) (******************************************************************************) +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GroupScope GRing.Theory Num.Theory. +Local Open Scope ring_scope. + Local Notation algCF := [fieldType of algC]. Section AlgC. @@ -333,9 +330,9 @@ rewrite !big_cons /= in dxW defW *. rewrite 2!(big_nth i) !big_mkord /= in IHe dxW defW. set Wi := (\sum_i _)%MS in defW dxW IHe. rewrite -mxdirectE mxdirect_addsE !mxdirectE eqxx /= -/Wi in dxW. -have modWi: mxmodule rG Wi by exact: sumsmx_module. +have modWi: mxmodule rG Wi by apply: sumsmx_module. case/andP: dxW; move/(IHe Wi modWi) {IHe}; move/(_ (eqmx_refl _))=> rsimWi. -by move/eqP; move/mxdirect_addsP=> dxUiWi; exact: mx_rsim_dadd (rsimU i) rsimWi. +by move/eqP; move/mxdirect_addsP=> dxUiWi; apply: mx_rsim_dadd (rsimU i) rsimWi. Qed. Definition muln_grepr rW k := \big[dadd_grepr/grepr0]_(i < k) rW. @@ -348,7 +345,7 @@ Proof. set M := socle_base W => modW rsimM. have simM: mxsimple rG M := socle_simple W. have rankM_gt0: (\rank M > 0)%N by rewrite lt0n mxrank_eq0; case: simM. -have [I /= U_I simU]: mxsemisimple rG W by exact: component_mx_semisimple. +have [I /= U_I simU]: mxsemisimple rG W by apply: component_mx_semisimple. pose U (i : 'I_#|I|) := U_I (enum_val i). have reindexI := reindex _ (onW_bij I (enum_val_bij I)). rewrite mxdirectE /= !reindexI -mxdirectE /= => defW dxW. @@ -360,8 +357,8 @@ have ->: socle_mult W = #|I|. rewrite -defW (mxdirectP dxW) /= -sum_nat_const reindexI /=. by apply: eq_bigr => i _; rewrite -(mxrank_iso (isoU i)). have modU: mxmodule rG (U _) := mxsimple_module (simU _). -suff: mx_rsim (submod_repr (modU _)) rW by exact: mx_rsim_dsum defW dxW. -by move=> i; apply: mx_rsim_trans (mx_rsim_sym _) rsimM; exact/mx_rsim_iso. +suff: mx_rsim (submod_repr (modU _)) rW by apply: mx_rsim_dsum defW dxW. +by move=> i; apply: mx_rsim_trans (mx_rsim_sym _) rsimM; apply/mx_rsim_iso. Qed. End DsumRepr. @@ -466,7 +463,7 @@ have [defS dxS]: (S :=: 1%:M)%MS /\ mxdirect S. rewrite /soc; case: pickP => [W' | /(_ Wi)] /= /eqP // eqWi. apply/eqP/socle_rsimP. apply: mx_rsim_trans (rsim_irr_comp iG C'G (socle_irr _)) (mx_rsim_sym _). - by rewrite [irr_comp _ _]eqWi; exact: rsim_irr_comp (socle_irr _). + by rewrite [irr_comp _ _]eqWi; apply: rsim_irr_comp (socle_irr _). have bij_irr: {on [pred i | soc i], bijective standard_irr}. exists (odflt W0 \o soc) => [Wi _ | i]; first by rewrite /= irrK. by rewrite inE /soc /=; case: pickP => //= Wi; move/eqP. @@ -485,7 +482,7 @@ apply: mx_rsim_dsum (modW) _ defS dxS _ => i. rewrite /W /standard_irr_coef /modW /soc; case: pickP => [Wi|_] /=; last first. rewrite /muln_grepr big_ord0. by exists 0 => [||x _]; rewrite ?mxrank0 ?mulmx0 ?mul0mx. -by move/eqP=> <-; apply: mx_rsim_socle; exact: rsim_irr_comp (socle_irr Wi). +by move/eqP=> <-; apply: mx_rsim_socle; apply: rsim_irr_comp (socle_irr Wi). Qed. End StandardRepr. @@ -560,7 +557,7 @@ Lemma NirrE : Nirr G = #|classes G|. Proof. by rewrite /pred_Nirr (cardD1 [1]) classes1. Qed. Fact Iirr_cast : Nirr G = #|sG|. -Proof. by rewrite NirrE ?card_irr ?algC'G //; exact: groupC. Qed. +Proof. by rewrite NirrE ?card_irr ?algC'G //; apply: groupC. Qed. Let offset := cast_ord (esym Iirr_cast) (enum_rank [1 sG]%irr). @@ -790,7 +787,7 @@ Lemma irr_reprP xi : (xi \in irr G). Proof. apply: (iffP (irrP xi)) => [[i ->] | [[n rG] irr_rG ->]]. - by exists (Representation 'Chi_i); [exact: socle_irr | rewrite irrRepr]. + by exists (Representation 'Chi_i); [apply: socle_irr | rewrite irrRepr]. exists (irr_of_socle (irr_comp sG rG)); rewrite -irrRepr irr_of_socleK /=. exact/cfRepr_sim/rsim_irr_comp. Qed. @@ -799,7 +796,7 @@ Qed. Lemma Wedderburn_id_expansion i : 'e_i = #|G|%:R^-1 *: \sum_(x in G) 'chi_i 1%g * 'chi_i x^-1%g *: aG x. Proof. -have Rei: ('e_i \in 'R_i)%MS by exact: Wedderburn_id_mem. +have Rei: ('e_i \in 'R_i)%MS by apply: Wedderburn_id_mem. have /envelop_mxP[a def_e]: ('e_i \in R_G)%MS; last rewrite -/aG in def_e. by move: Rei; rewrite genmxE mem_sub_gring => /andP[]. apply: canRL (scalerK (neq0CG _)) _; rewrite def_e linear_sum /=. @@ -932,20 +929,21 @@ Section AutChar. Variables (gT : finGroupType) (G : {group gT}). Implicit Type u : {rmorphism algC -> algC}. +Implicit Type chi : 'CF(G). Lemma cfRepr_map u n (rG : mx_representation algCF G n) : cfRepr (map_repr u rG) = cfAut u (cfRepr rG). Proof. by apply/cfun_inP=> x Gx; rewrite !cfunE Gx map_reprE trace_map_mx. Qed. -Lemma cfAut_char u (chi : 'CF(G)) : - chi \is a character -> cfAut u chi \is a character. +Lemma cfAut_char u chi : (cfAut u chi \is a character) = (chi \is a character). Proof. -case/char_reprP=> rG ->; apply/char_reprP. +without loss /char_reprP[rG ->]: u chi / chi \is a character. + by move=> IHu; apply/idP/idP=> ?; first rewrite -(cfAutK u chi); rewrite IHu. +rewrite cfRepr_char; apply/char_reprP. by exists (Representation (map_repr u rG)); rewrite cfRepr_map. Qed. -Lemma cfConjC_char (chi : 'CF(G)) : - chi \is a character -> chi^*%CF \is a character. +Lemma cfConjC_char chi : (chi^*%CF \is a character) = (chi \is a character). Proof. exact: cfAut_char. Qed. Lemma cfAut_char1 u (chi : 'CF(G)) : @@ -1037,7 +1035,7 @@ Qed. Lemma lin_char_irr : xi \in irr G. Proof. case/andP: CFxi => /char_reprP[rG ->]; rewrite cfRepr1 pnatr_eq1 => /eqP n1. -by apply/irr_reprP; exists rG => //; exact/mx_abs_irrW/linear_mx_abs_irr. +by apply/irr_reprP; exists rG => //; apply/mx_abs_irrW/linear_mx_abs_irr. Qed. Lemma mul_conjC_lin_char : xi * xi^*%CF = 1. @@ -1052,12 +1050,6 @@ Proof. by apply/unitrPr; exists xi^*%CF; apply: mul_conjC_lin_char. Qed. Lemma invr_lin_char : xi^-1 = xi^*%CF. Proof. by rewrite -[_^-1]mulr1 -mul_conjC_lin_char mulKr ?lin_char_unitr. Qed. -Lemma cfAut_lin_char u : cfAut u xi \is a linear_char. -Proof. by rewrite qualifE cfunE lin_char1 rmorph1 cfAut_char ?lin_charW /=. Qed. - -Lemma cfConjC_lin_char : xi^*%CF \is a linear_char. -Proof. exact: cfAut_lin_char. Qed. - Lemma fful_lin_char_inj : cfaithful xi -> {in G &, injective xi}. Proof. move=> fful_phi x y Gx Gy xi_xy; apply/eqP; rewrite eq_mulgV1 -in_set1. @@ -1068,6 +1060,14 @@ Qed. End OneChar. +Lemma cfAut_lin_char u (xi : 'CF(G)) : + (cfAut u xi \is a linear_char) = (xi \is a linear_char). +Proof. by rewrite qualifE cfAut_char; apply/andb_id2l=> /cfAut_char1->. Qed. + +Lemma cfConjC_lin_char (xi : 'CF(G)) : + (xi^*%CF \is a linear_char) = (xi \is a linear_char). +Proof. exact: cfAut_lin_char. Qed. + Lemma card_Iirr_abelian : abelian G -> #|Iirr G| = #|G|. Proof. by rewrite card_ord NirrE card_classes_abelian => /eqP. Qed. @@ -1079,7 +1079,7 @@ Lemma char_abelianP : Proof. apply: (iffP idP) => [cGG i | CF_G]. rewrite qualifE irr_char /= irr1_degree. - by rewrite irr_degree_abelian //; last exact: groupC. + by rewrite irr_degree_abelian //; last apply: groupC. rewrite card_classes_abelian -NirrE -eqC_nat -irr_sum_square //. rewrite -{1}[Nirr G]card_ord -sumr_const; apply/eqP/eq_bigr=> i _. by rewrite lin_char1 ?expr1n ?CF_G. @@ -1115,173 +1115,6 @@ End Linear. Prenex Implicits linear_char. -Section Restrict. - -Variable (gT : finGroupType) (G H : {group gT}). - -Lemma cfRepr_sub n (rG : mx_representation algCF G n) (sHG : H \subset G) : - cfRepr (subg_repr rG sHG) = 'Res[H] (cfRepr rG). -Proof. -by apply/cfun_inP => x Hx; rewrite cfResE // !cfunE Hx (subsetP sHG). -Qed. - -Lemma cfRes_char chi : chi \is a character -> 'Res[H, G] chi \is a character. -Proof. -have [sHG | not_sHG] := boolP (H \subset G). - by case/char_reprP=> rG ->; rewrite -(cfRepr_sub rG sHG) cfRepr_char. -by move/Cnat_char1=> Nchi1; rewrite cfResEout // rpredZ_Cnat ?rpred1. -Qed. - -Lemma cfRes_eq0 phi : phi \is a character -> ('Res[H, G] phi == 0) = (phi == 0). -Proof. by move=> Nchi; rewrite -!char1_eq0 ?cfRes_char // cfRes1. Qed. - -Lemma cfRes_lin_char chi : - chi \is a linear_char -> 'Res[H, G] chi \is a linear_char. -Proof. by case/andP=> Nchi; rewrite qualifE cfRes_char ?cfRes1. Qed. - -Lemma Res_irr_neq0 i : 'Res[H, G] 'chi_i != 0. -Proof. by rewrite cfRes_eq0 ?irr_neq0 ?irr_char. Qed. - -Lemma cfRes_lin_lin (chi : 'CF(G)) : - chi \is a character -> 'Res[H] chi \is a linear_char -> chi \is a linear_char. -Proof. by rewrite !qualifE cfRes1 => -> /andP[]. Qed. - -Lemma cfRes_irr_irr chi : - chi \is a character -> 'Res[H] chi \in irr H -> chi \in irr G. -Proof. -have [sHG /char_reprP[rG ->] | not_sHG Nchi] := boolP (H \subset G). - rewrite -(cfRepr_sub _ sHG) => /irr_reprP[rH irrH def_rH]; apply/irr_reprP. - suffices /subg_mx_irr: mx_irreducible (subg_repr rG sHG) by exists rG. - by apply: mx_rsim_irr irrH; exact/cfRepr_rsimP/eqP. -rewrite cfResEout // => /irrP[j Dchi_j]; apply/lin_char_irr/cfRes_lin_lin=> //. -suffices j0: j = 0 by rewrite cfResEout // Dchi_j j0 irr0 rpred1. -apply: contraNeq (irr1_neq0 j) => nz_j. -have:= xcfun_id j 0; rewrite -Dchi_j cfunE xcfunZl -irr0 xcfun_id eqxx => ->. -by rewrite (negPf nz_j). -Qed. - -Definition Res_Iirr (A B : {set gT}) i := cfIirr ('Res[B, A] 'chi_i). - -Lemma Res_Iirr0 : Res_Iirr H (0 : Iirr G) = 0. -Proof. by rewrite /Res_Iirr irr0 rmorph1 -irr0 irrK. Qed. - -Lemma lin_Res_IirrE i : 'chi[G]_i 1%g = 1 -> 'chi_(Res_Iirr H i) = 'Res 'chi_i. -Proof. -move=> chi1; rewrite cfIirrE ?lin_char_irr ?cfRes_lin_char //. -by rewrite qualifE irr_char /= chi1. -Qed. - -End Restrict. - -Arguments Scope Res_Iirr [_ group_scope group_scope ring_scope]. - -Section Morphim. - -Variables (aT rT : finGroupType) (G D : {group aT}) (f : {morphism D >-> rT}). -Implicit Type chi : 'CF(f @* G). - -Lemma cfRepr_morphim n (rfG : mx_representation algCF (f @* G) n) sGD : - cfRepr (morphim_repr rfG sGD) = cfMorph (cfRepr rfG). -Proof. -apply/cfun_inP=> x Gx; have Dx: x \in D := subsetP sGD x Gx. -by rewrite cfMorphE // !cfunE ?mem_morphim ?Gx. -Qed. - -Lemma cfMorph_char chi : chi \is a character -> cfMorph chi \is a character. -Proof. -have [sGD /char_reprP[rG ->] | outGD Nchi] := boolP (G \subset D); last first. - by rewrite cfMorphEout // rpredZ_Cnat ?rpred1 ?Cnat_char1. -apply/char_reprP; exists (Representation (morphim_repr rG sGD)). -by rewrite cfRepr_morphim. -Qed. - -Lemma cfMorph_lin_char chi : - chi \is a linear_char -> cfMorph chi \is a linear_char. -Proof. by case/andP=> Nchi; rewrite qualifE cfMorph_char ?cfMorph1. Qed. - -Lemma cfMorph_irr chi : - G \subset D -> chi \in irr (f @* G) -> cfMorph chi \in irr G. -Proof. -move=> sGD /irr_reprP[rG irrG ->]; apply/irr_reprP. -exists (Representation (morphim_repr rG sGD)); first exact/morphim_mx_irr. -apply/cfun_inP=> x Gx; rewrite !cfunElock /= sGD Gx. -by rewrite mem_morphim ?(subsetP sGD). -Qed. - -Definition morph_Iirr i := cfIirr (cfMorph 'chi[f @* G]_i). - -Lemma morph_Iirr0 : morph_Iirr 0 = 0. -Proof. by rewrite /morph_Iirr irr0 rmorph1 -irr0 irrK. Qed. - -Hypothesis sGD : G \subset D. - -Lemma morph_IirrE i : 'chi_(morph_Iirr i) = cfMorph 'chi_i. -Proof. by rewrite cfIirrE ?cfMorph_irr ?mem_irr. Qed. - -Lemma morph_Iirr_inj : injective morph_Iirr. -Proof. -by move=> i j eq_ij; apply/irr_inj/cfMorph_inj; rewrite // -!morph_IirrE eq_ij. -Qed. - -Lemma morph_Iirr_eq0 i : (morph_Iirr i == 0) = (i == 0). -Proof. by rewrite -!irr_eq1 morph_IirrE cfMorph_eq1. Qed. - -End Morphim. - -Section Isom. - -Variables (aT rT : finGroupType) (G : {group aT}) (f : {morphism G >-> rT}). -Variables (R : {group rT}) (isoGR : isom G R f). -Implicit Type chi : 'CF(G). - -Lemma cfIsom_char chi : chi \is a character -> cfIsom isoGR chi \is a character. -Proof. -by move=> Nchi; rewrite [cfIsom _]locked_withE cfMorph_char ?cfRes_char. -Qed. - -Lemma cfIsom_lin_char chi : - chi \is a linear_char -> cfIsom isoGR chi \is a linear_char. -Proof. by case/andP=> Nchi; rewrite qualifE cfIsom_char ?cfIsom1. Qed. - -Lemma cfIsom_irr chi : chi \in irr G -> cfIsom isoGR chi \in irr R. -Proof. -move=> irr_chi; rewrite [cfIsom _]locked_withE cfMorph_irr //. -by rewrite (isom_im (isom_sym isoGR)) cfRes_id. -Qed. - -Definition isom_Iirr i := cfIirr (cfIsom isoGR 'chi_i). - -Lemma isom_IirrE i : 'chi_(isom_Iirr i) = cfIsom isoGR 'chi_i. -Proof. by rewrite cfIirrE ?cfIsom_irr ?mem_irr. Qed. - -Lemma isom_Iirr_inj : injective isom_Iirr. -Proof. -by move=> i j eqij; apply/irr_inj/(cfIsom_inj isoGR); rewrite -!isom_IirrE eqij. -Qed. - -Lemma isom_Iirr_eq0 i : (isom_Iirr i == 0) = (i == 0). -Proof. by rewrite -!irr_eq1 isom_IirrE cfIsom_eq1. Qed. - -Lemma isom_Iirr0 : isom_Iirr 0 = 0. -Proof. by apply/eqP; rewrite isom_Iirr_eq0. Qed. - -End Isom. - -Implicit Arguments isom_Iirr_inj [aT rT G f R x1 x2]. - -Section IsomInv. - -Variables (aT rT : finGroupType) (G : {group aT}) (f : {morphism G >-> rT}). -Variables (R : {group rT}) (isoGR : isom G R f). - -Lemma isom_IirrK : cancel (isom_Iirr isoGR) (isom_Iirr (isom_sym isoGR)). -Proof. by move=> i; apply: irr_inj; rewrite !isom_IirrE cfIsomK. Qed. - -Lemma isom_IirrKV : cancel (isom_Iirr (isom_sym isoGR)) (isom_Iirr isoGR). -Proof. by move=> i; apply: irr_inj; rewrite !isom_IirrE cfIsomKV. Qed. - -End IsomInv. - Section OrthogonalityRelations. Variables aT gT : finGroupType. @@ -1303,9 +1136,9 @@ have [I U W simU W1 dxW]: mxsemisimple rG 1%:M. rewrite -(reducible_Socle1 (DecSocleType rG) (mx_Maschke _ (algC'G G))). exact: Socle_semisimple. have linU i: \rank (U i) = 1%N. - by apply: mxsimple_abelian_linear cGG (simU i); exact: groupC. + by apply: mxsimple_abelian_linear cGG (simU i); apply: groupC. have castI: f = #|I|. - by rewrite -(mxrank1 algCF f) -W1 (eqnP dxW) /= -sum1_card; exact/eq_bigr. + by rewrite -(mxrank1 algCF f) -W1 (eqnP dxW) /= -sum1_card; apply/eq_bigr. pose B := \matrix_j nz_row (U (enum_val (cast_ord castI j))). have rowU i: (nz_row (U i) :=: U i)%MS. apply/eqmxP; rewrite -(geq_leqif (mxrank_leqif_eq (nz_row_sub _))) linU. @@ -1417,8 +1250,8 @@ Proof. by move=> xG GxG; rewrite /c cast_ordKV enum_rankK_in. Qed. Lemma reindex_irr_class R idx (op : @Monoid.com_law R idx) F : \big[op/idx]_(xG in classes G) F xG = \big[op/idx]_i F (c i). Proof. -rewrite (reindex c); first by apply: eq_bigl => i; exact: enum_valP. -by exists iC; [apply: in1W; exact: irr_classK | exact: class_IirrK]. +rewrite (reindex c); first by apply: eq_bigl => i; apply: enum_valP. +by exists iC; [apply: in1W; apply: irr_classK | apply: class_IirrK]. Qed. (* The explicit value of the inverse is needed for the proof of the second *) @@ -1451,10 +1284,10 @@ transitivity ((#|'C_G[repr (y ^: G)]|%:R *: (X' *m X)) i_y i_x). rewrite scalemxAl !mxE; apply: eq_bigr => k _; rewrite !mxE mulrC -!mulrA. by rewrite !class_IirrK ?mem_classes // !cfun_repr mulVKf ?neq0CG. rewrite mulmx1C // !mxE -!divg_index !(index_cent1, =^~ indexgI). -rewrite (class_transr (mem_repr y _)) ?class_refl // mulr_natr. +rewrite (class_eqP (mem_repr y _)) ?class_refl // mulr_natr. rewrite (can_in_eq class_IirrK) ?mem_classes //. have [-> | not_yGx] := altP eqP; first by rewrite class_refl. -by rewrite [x \in _](contraNF _ not_yGx) // => /class_transr->. +by rewrite [x \in _](contraNF _ not_yGx) // => /class_eqP->. Qed. Lemma eq_irr_mem_classP x y : @@ -1638,24 +1471,435 @@ Qed. End InnerProduct. +Section IrrConstt. + +Variable (gT : finGroupType) (G H : {group gT}). + +Lemma char1_ge_norm (chi : 'CF(G)) x : + chi \is a character -> `|chi x| <= chi 1%g. +Proof. +case/char_reprP=> rG ->; case Gx: (x \in G); last first. + by rewrite cfunE cfRepr1 Gx normr0 ler0n. +by have [e [_ _ []]] := repr_rsim_diag rG Gx. +Qed. + +Lemma max_cfRepr_norm_scalar n (rG : mx_representation algCF G n) x : + x \in G -> `|cfRepr rG x| = cfRepr rG 1%g -> + exists2 c, `|c| = 1 & rG x = c%:M. +Proof. +move=> Gx; have [e [[B uB def_x] [_ e1] [-> _] _]] := repr_rsim_diag rG Gx. +rewrite cfRepr1 -[n in n%:R]card_ord -sumr_const -(eq_bigr _ (in1W e1)). +case/normC_sum_eq1=> [i _ | c /eqP norm_c_1 def_e]; first by rewrite e1. +have{def_e} def_e: e = const_mx c by apply/rowP=> i; rewrite mxE def_e ?andbT. +by exists c => //; rewrite def_x def_e diag_const_mx scalar_mxC mulmxKV. +Qed. + +Lemma max_cfRepr_mx1 n (rG : mx_representation algCF G n) x : + x \in G -> cfRepr rG x = cfRepr rG 1%g -> rG x = 1%:M. +Proof. +move=> Gx kerGx; have [|c _ def_x] := @max_cfRepr_norm_scalar n rG x Gx. + by rewrite kerGx cfRepr1 normr_nat. +move/eqP: kerGx; rewrite cfRepr1 cfunE Gx {rG}def_x mxtrace_scalar. +case: n => [_|n]; first by rewrite ![_%:M]flatmx0. +rewrite mulrb -subr_eq0 -mulrnBl -mulr_natl mulf_eq0 pnatr_eq0 /=. +by rewrite subr_eq0 => /eqP->. +Qed. + +Definition irr_constt (B : {set gT}) phi := [pred i | '[phi, 'chi_i]_B != 0]. + +Lemma irr_consttE i phi : (i \in irr_constt phi) = ('[phi, 'chi_i]_G != 0). +Proof. by []. Qed. + +Lemma constt_charP (i : Iirr G) chi : + chi \is a character -> + reflect (exists2 chi', chi' \is a character & chi = 'chi_i + chi') + (i \in irr_constt chi). +Proof. +move=> Nchi; apply: (iffP idP) => [i_in_chi| [chi' Nchi' ->]]; last first. + rewrite inE /= cfdotDl cfdot_irr eqxx -(eqP (Cnat_cfdot_char_irr i Nchi')). + by rewrite -natrD pnatr_eq0. +exists (chi - 'chi_i); last by rewrite addrC subrK. +apply/forallP=> j; rewrite coord_cfdot cfdotBl cfdot_irr. +have [<- | _] := eqP; last by rewrite subr0 Cnat_cfdot_char_irr. +have := i_in_chi; rewrite inE /= -(eqP (Cnat_cfdot_char_irr i Nchi)) pnatr_eq0. +by case: (truncC _) => // n _; rewrite mulrSr addrK ?isNatC_nat. +Qed. + +Lemma cfun_sum_constt (phi : 'CF(G)) : + phi = \sum_(i in irr_constt phi) '[phi, 'chi_i] *: 'chi_i. +Proof. +rewrite {1}[phi]cfun_sum_cfdot (bigID [pred i | '[phi, 'chi_i] == 0]) /=. +by rewrite big1 ?add0r // => i /eqP->; rewrite scale0r. +Qed. + +Lemma neq0_has_constt (phi : 'CF(G)) : + phi != 0 -> exists i, i \in irr_constt phi. +Proof. +move=> nz_phi; apply/existsP; apply: contra nz_phi => /pred0P phi0. +by rewrite [phi]cfun_sum_constt big_pred0. +Qed. + +Lemma constt_irr i : irr_constt 'chi[G]_i =i pred1 i. +Proof. +by move=> j; rewrite !inE cfdot_irr pnatr_eq0 (eq_sym j); case: (i == j). +Qed. + +Lemma char1_ge_constt (i : Iirr G) chi : + chi \is a character -> i \in irr_constt chi -> 'chi_i 1%g <= chi 1%g. +Proof. +move=> {chi} _ /constt_charP[// | chi Nchi ->]. +by rewrite cfunE addrC -subr_ge0 addrK char1_ge0. +Qed. + +Lemma constt_ortho_char (phi psi : 'CF(G)) i j : + phi \is a character -> psi \is a character -> + i \in irr_constt phi -> j \in irr_constt psi -> + '[phi, psi] = 0 -> '['chi_i, 'chi_j] = 0. +Proof. +move=> _ _ /constt_charP[//|phi1 Nphi1 ->] /constt_charP[//|psi1 Npsi1 ->]. +rewrite cfdot_irr; case: eqP => // -> /eqP/idPn[]. +rewrite cfdotDl !cfdotDr cfnorm_irr -addrA gtr_eqF ?ltr_paddr ?ltr01 //. +by rewrite Cnat_ge0 ?rpredD ?Cnat_cfdot_char ?irr_char. +Qed. + +End IrrConstt. + +Arguments Scope irr_constt [_ group_scope cfun_scope]. + +Section Kernel. + +Variable (gT : finGroupType) (G : {group gT}). +Implicit Types (phi chi xi : 'CF(G)) (H : {group gT}). + +Lemma cfker_repr n (rG : mx_representation algCF G n) : + cfker (cfRepr rG) = rker rG. +Proof. +apply/esym/setP=> x; rewrite inE mul1mx /=. +case Gx: (x \in G); last by rewrite inE Gx. +apply/eqP/idP=> Kx; last by rewrite max_cfRepr_mx1 // cfker1. +rewrite inE Gx; apply/forallP=> y; rewrite !cfunE !mulrb groupMl //. +by case: ifP => // Gy; rewrite repr_mxM // Kx mul1mx. +Qed. + +Lemma cfkerEchar chi : + chi \is a character -> cfker chi = [set x in G | chi x == chi 1%g]. +Proof. +move=> Nchi; apply/setP=> x; apply/idP/setIdP=> [Kx | [Gx /eqP chi_x]]. + by rewrite (subsetP (cfker_sub chi)) // cfker1. +case/char_reprP: Nchi => rG -> in chi_x *; rewrite inE Gx; apply/forallP=> y. +rewrite !cfunE groupMl // !mulrb; case: ifP => // Gy. +by rewrite repr_mxM // max_cfRepr_mx1 ?mul1mx. +Qed. + +Lemma cfker_nzcharE chi : + chi \is a character -> chi != 0 -> cfker chi = [set x | chi x == chi 1%g]. +Proof. +move=> Nchi nzchi; apply/setP=> x; rewrite cfkerEchar // !inE andb_idl //. +by apply: contraLR => /cfun0-> //; rewrite eq_sym char1_eq0. +Qed. + +Lemma cfkerEirr i : cfker 'chi[G]_i = [set x | 'chi_i x == 'chi_i 1%g]. +Proof. by rewrite cfker_nzcharE ?irr_char ?irr_neq0. Qed. + +Lemma cfker_irr0 : cfker 'chi[G]_0 = G. +Proof. by rewrite irr0 cfker_cfun1. Qed. + +Lemma cfaithful_reg : cfaithful (cfReg G). +Proof. +apply/subsetP=> x; rewrite cfkerEchar ?cfReg_char // !inE !cfRegE eqxx. +by case/andP=> _; apply: contraLR => /negbTE->; rewrite eq_sym neq0CG. +Qed. + +Lemma cfkerE chi : + chi \is a character -> + cfker chi = G :&: \bigcap_(i in irr_constt chi) cfker 'chi_i. +Proof. +move=> Nchi; rewrite cfkerEchar //; apply/setP=> x; rewrite !inE. +apply: andb_id2l => Gx; rewrite {1 2}[chi]cfun_sum_constt !sum_cfunE. +apply/eqP/bigcapP=> [Kx i Ci | Kx]; last first. + by apply: eq_bigr => i /Kx Kx_i; rewrite !cfunE cfker1. +rewrite cfkerEirr inE /= -(inj_eq (mulfI Ci)). +have:= (normC_sum_upper _ Kx) i; rewrite !cfunE => -> // {i Ci} i _. +have chi_i_ge0: 0 <= '[chi, 'chi_i]. + by rewrite Cnat_ge0 ?Cnat_cfdot_char_irr. +by rewrite !cfunE normrM (normr_idP _) ?ler_wpmul2l ?char1_ge_norm ?irr_char. +Qed. + +Lemma TI_cfker_irr : \bigcap_i cfker 'chi[G]_i = [1]. +Proof. +apply/trivgP; apply: subset_trans cfaithful_reg; rewrite cfkerE ?cfReg_char //. +rewrite subsetI (bigcap_min 0) //=; last by rewrite cfker_irr0. +by apply/bigcapsP=> i _; rewrite bigcap_inf. +Qed. + +Lemma cfker_constt i chi : + chi \is a character -> i \in irr_constt chi -> + cfker chi \subset cfker 'chi[G]_i. +Proof. by move=> Nchi Ci; rewrite cfkerE ?subIset ?(bigcap_min i) ?orbT. Qed. + +Section KerLin. + +Variable xi : 'CF(G). +Hypothesis lin_xi : xi \is a linear_char. +Let Nxi: xi \is a character. Proof. by have [] := andP lin_xi. Qed. + +Lemma lin_char_der1 : G^`(1)%g \subset cfker xi. +Proof. +rewrite gen_subG /=; apply/subsetP=> _ /imset2P[x y Gx Gy ->]. +rewrite cfkerEchar // inE groupR //= !lin_charM ?lin_charV ?in_group //. +by rewrite mulrCA mulKf ?mulVf ?lin_char_neq0 // lin_char1. +Qed. + +Lemma cforder_lin_char : #[xi]%CF = exponent (G / cfker xi)%g. +Proof. +apply/eqP; rewrite eqn_dvd; apply/andP; split. + apply/dvdn_cforderP=> x Gx; rewrite -lin_charX // -cfQuoEker ?groupX //. + rewrite morphX ?(subsetP (cfker_norm xi)) //= expg_exponent ?mem_quotient //. + by rewrite cfQuo1 ?cfker_normal ?lin_char1. +have abGbar: abelian (G / cfker xi) := sub_der1_abelian lin_char_der1. +have [_ /morphimP[x Nx Gx ->] ->] := exponent_witness (abelian_nil abGbar). +rewrite order_dvdn -morphX //= coset_id cfkerEchar // !inE groupX //=. +by rewrite lin_charX ?lin_char1 // (dvdn_cforderP _ _ _). +Qed. + +Lemma cforder_lin_char_dvdG : #[xi]%CF %| #|G|. +Proof. +by rewrite cforder_lin_char (dvdn_trans (exponent_dvdn _)) ?dvdn_morphim. +Qed. + +Lemma cforder_lin_char_gt0 : (0 < #[xi]%CF)%N. +Proof. by rewrite cforder_lin_char exponent_gt0. Qed. + +End KerLin. + +End Kernel. + +Section Restrict. + +Variable (gT : finGroupType) (G H : {group gT}). + +Lemma cfRepr_sub n (rG : mx_representation algCF G n) (sHG : H \subset G) : + cfRepr (subg_repr rG sHG) = 'Res[H] (cfRepr rG). +Proof. +by apply/cfun_inP => x Hx; rewrite cfResE // !cfunE Hx (subsetP sHG). +Qed. + +Lemma cfRes_char chi : chi \is a character -> 'Res[H, G] chi \is a character. +Proof. +have [sHG | not_sHG] := boolP (H \subset G). + by case/char_reprP=> rG ->; rewrite -(cfRepr_sub rG sHG) cfRepr_char. +by move/Cnat_char1=> Nchi1; rewrite cfResEout // rpredZ_Cnat ?rpred1. +Qed. + +Lemma cfRes_eq0 phi : phi \is a character -> ('Res[H, G] phi == 0) = (phi == 0). +Proof. by move=> Nchi; rewrite -!char1_eq0 ?cfRes_char // cfRes1. Qed. + +Lemma cfRes_lin_char chi : + chi \is a linear_char -> 'Res[H, G] chi \is a linear_char. +Proof. by case/andP=> Nchi; rewrite qualifE cfRes_char ?cfRes1. Qed. + +Lemma Res_irr_neq0 i : 'Res[H, G] 'chi_i != 0. +Proof. by rewrite cfRes_eq0 ?irr_neq0 ?irr_char. Qed. + +Lemma cfRes_lin_lin (chi : 'CF(G)) : + chi \is a character -> 'Res[H] chi \is a linear_char -> chi \is a linear_char. +Proof. by rewrite !qualifE cfRes1 => -> /andP[]. Qed. + +Lemma cfRes_irr_irr chi : + chi \is a character -> 'Res[H] chi \in irr H -> chi \in irr G. +Proof. +have [sHG /char_reprP[rG ->] | not_sHG Nchi] := boolP (H \subset G). + rewrite -(cfRepr_sub _ sHG) => /irr_reprP[rH irrH def_rH]; apply/irr_reprP. + suffices /subg_mx_irr: mx_irreducible (subg_repr rG sHG) by exists rG. + by apply: mx_rsim_irr irrH; apply/cfRepr_rsimP/eqP. +rewrite cfResEout // => /irrP[j Dchi_j]; apply/lin_char_irr/cfRes_lin_lin=> //. +suffices j0: j = 0 by rewrite cfResEout // Dchi_j j0 irr0 rpred1. +apply: contraNeq (irr1_neq0 j) => nz_j. +have:= xcfun_id j 0; rewrite -Dchi_j cfunE xcfunZl -irr0 xcfun_id eqxx => ->. +by rewrite (negPf nz_j). +Qed. + +Definition Res_Iirr (A B : {set gT}) i := cfIirr ('Res[B, A] 'chi_i). + +Lemma Res_Iirr0 : Res_Iirr H (0 : Iirr G) = 0. +Proof. by rewrite /Res_Iirr irr0 rmorph1 -irr0 irrK. Qed. + +Lemma lin_Res_IirrE i : 'chi[G]_i 1%g = 1 -> 'chi_(Res_Iirr H i) = 'Res 'chi_i. +Proof. +move=> chi1; rewrite cfIirrE ?lin_char_irr ?cfRes_lin_char //. +by rewrite qualifE irr_char /= chi1. +Qed. + +End Restrict. + +Arguments Scope Res_Iirr [_ group_scope group_scope ring_scope]. + +Section MoreConstt. + +Variables (gT : finGroupType) (G H : {group gT}). + +Lemma constt_Ind_Res i j : + i \in irr_constt ('Ind[G] 'chi_j) = (j \in irr_constt ('Res[H] 'chi_i)). +Proof. by rewrite !irr_consttE cfdotC conjC_eq0 -cfdot_Res_l. Qed. + +Lemma cfdot_Res_ge_constt i j psi : + psi \is a character -> j \in irr_constt psi -> + '['Res[H, G] 'chi_j, 'chi_i] <= '['Res[H] psi, 'chi_i]. +Proof. +move=> {psi} _ /constt_charP[// | psi Npsi ->]. +rewrite linearD cfdotDl addrC -subr_ge0 addrK Cnat_ge0 //=. +by rewrite Cnat_cfdot_char_irr // cfRes_char. +Qed. + +Lemma constt_Res_trans j psi : + psi \is a character -> j \in irr_constt psi -> + {subset irr_constt ('Res[H, G] 'chi_j) <= irr_constt ('Res[H] psi)}. +Proof. +move=> Npsi Cj i; apply: contraNneq; rewrite eqr_le => {1}<-. +rewrite cfdot_Res_ge_constt ?Cnat_ge0 ?Cnat_cfdot_char_irr //. +by rewrite cfRes_char ?irr_char. +Qed. + +End MoreConstt. + +Section Morphim. + +Variables (aT rT : finGroupType) (G D : {group aT}) (f : {morphism D >-> rT}). +Implicit Type chi : 'CF(f @* G). + +Lemma cfRepr_morphim n (rfG : mx_representation algCF (f @* G) n) sGD : + cfRepr (morphim_repr rfG sGD) = cfMorph (cfRepr rfG). +Proof. +apply/cfun_inP=> x Gx; have Dx: x \in D := subsetP sGD x Gx. +by rewrite cfMorphE // !cfunE ?mem_morphim ?Gx. +Qed. + +Lemma cfMorph_char chi : chi \is a character -> cfMorph chi \is a character. +Proof. +have [sGD /char_reprP[rfG ->] | outGD Nchi] := boolP (G \subset D); last first. + by rewrite cfMorphEout // rpredZ_Cnat ?rpred1 ?Cnat_char1. +apply/char_reprP; exists (Representation (morphim_repr rfG sGD)). +by rewrite cfRepr_morphim. +Qed. + +Lemma cfMorph_lin_char chi : + chi \is a linear_char -> cfMorph chi \is a linear_char. +Proof. by case/andP=> Nchi; rewrite qualifE cfMorph1 cfMorph_char. Qed. + +Lemma cfMorph_charE chi : + G \subset D -> (cfMorph chi \is a character) = (chi \is a character). +Proof. +move=> sGD; apply/idP/idP=> [/char_reprP[[n rG] /=Dfchi] | /cfMorph_char//]. +pose H := 'ker_G f; have kerH: H \subset rker rG. + by rewrite -cfker_repr -Dfchi cfker_morph // setIS // ker_sub_pre. +have nHG: G \subset 'N(H) by rewrite normsI // (subset_trans sGD) ?ker_norm. +have [h injh im_h] := first_isom_loc f sGD; rewrite -/H in h injh im_h. +have DfG: invm injh @*^-1 (G / H) == (f @* G)%g by rewrite morphpre_invm im_h. +pose rfG := eqg_repr (morphpre_repr _ (quo_repr kerH nHG)) DfG. +apply/char_reprP; exists (Representation rfG). +apply/cfun_inP=> _ /morphimP[x Dx Gx ->]; rewrite -cfMorphE // Dfchi !cfunE Gx. +pose xH := coset H x; have GxH: xH \in (G / H)%g by apply: mem_quotient. +suffices Dfx: f x = h xH by rewrite mem_morphim //= Dfx invmE ?quo_repr_coset. +by apply/set1_inj; rewrite -?morphim_set1 ?im_h ?(subsetP nHG) ?sub1set. +Qed. + +Lemma cfMorph_lin_charE chi : + G \subset D -> (cfMorph chi \is a linear_char) = (chi \is a linear_char). +Proof. by rewrite qualifE cfMorph1 => /cfMorph_charE->. Qed. + +Lemma cfMorph_irr chi : + G \subset D -> (cfMorph chi \in irr G) = (chi \in irr (f @* G)). +Proof. by move=> sGD; rewrite !irrEchar cfMorph_charE // cfMorph_iso. Qed. + +Definition morph_Iirr i := cfIirr (cfMorph 'chi[f @* G]_i). + +Lemma morph_Iirr0 : morph_Iirr 0 = 0. +Proof. by rewrite /morph_Iirr irr0 rmorph1 -irr0 irrK. Qed. + +Hypothesis sGD : G \subset D. + +Lemma morph_IirrE i : 'chi_(morph_Iirr i) = cfMorph 'chi_i. +Proof. by rewrite cfIirrE ?cfMorph_irr ?mem_irr. Qed. + +Lemma morph_Iirr_inj : injective morph_Iirr. +Proof. +by move=> i j eq_ij; apply/irr_inj/cfMorph_inj; rewrite // -!morph_IirrE eq_ij. +Qed. + +Lemma morph_Iirr_eq0 i : (morph_Iirr i == 0) = (i == 0). +Proof. by rewrite -!irr_eq1 morph_IirrE cfMorph_eq1. Qed. + +End Morphim. + +Section Isom. + +Variables (aT rT : finGroupType) (G : {group aT}) (f : {morphism G >-> rT}). +Variables (R : {group rT}) (isoGR : isom G R f). +Implicit Type chi : 'CF(G). + +Lemma cfIsom_char chi : + (cfIsom isoGR chi \is a character) = (chi \is a character). +Proof. +rewrite [cfIsom _]locked_withE cfMorph_charE //. +by rewrite (isom_im (isom_sym _)) cfRes_id. +Qed. + +Lemma cfIsom_lin_char chi : + (cfIsom isoGR chi \is a linear_char) = (chi \is a linear_char). +Proof. by rewrite qualifE cfIsom_char cfIsom1. Qed. + +Lemma cfIsom_irr chi : (cfIsom isoGR chi \in irr R) = (chi \in irr G). +Proof. by rewrite !irrEchar cfIsom_char cfIsom_iso. Qed. + +Definition isom_Iirr i := cfIirr (cfIsom isoGR 'chi_i). + +Lemma isom_IirrE i : 'chi_(isom_Iirr i) = cfIsom isoGR 'chi_i. +Proof. by rewrite cfIirrE ?cfIsom_irr ?mem_irr. Qed. + +Lemma isom_Iirr_inj : injective isom_Iirr. +Proof. +by move=> i j eqij; apply/irr_inj/(cfIsom_inj isoGR); rewrite -!isom_IirrE eqij. +Qed. + +Lemma isom_Iirr_eq0 i : (isom_Iirr i == 0) = (i == 0). +Proof. by rewrite -!irr_eq1 isom_IirrE cfIsom_eq1. Qed. + +Lemma isom_Iirr0 : isom_Iirr 0 = 0. +Proof. by apply/eqP; rewrite isom_Iirr_eq0. Qed. + +End Isom. + +Implicit Arguments isom_Iirr_inj [aT rT G f R x1 x2]. + +Section IsomInv. + +Variables (aT rT : finGroupType) (G : {group aT}) (f : {morphism G >-> rT}). +Variables (R : {group rT}) (isoGR : isom G R f). + +Lemma isom_IirrK : cancel (isom_Iirr isoGR) (isom_Iirr (isom_sym isoGR)). +Proof. by move=> i; apply: irr_inj; rewrite !isom_IirrE cfIsomK. Qed. + +Lemma isom_IirrKV : cancel (isom_Iirr (isom_sym isoGR)) (isom_Iirr isoGR). +Proof. by move=> i; apply: irr_inj; rewrite !isom_IirrE cfIsomKV. Qed. + +End IsomInv. + Section Sdprod. Variables (gT : finGroupType) (K H G : {group gT}). Hypothesis defG : K ><| H = G. +Let nKG: G \subset 'N(K). Proof. by have [/andP[]] := sdprod_context defG. Qed. Lemma cfSdprod_char chi : - chi \is a character -> cfSdprod defG chi \is a character. -Proof. by move=> Nchi; rewrite unlock cfMorph_char ?cfIsom_char. Qed. + (cfSdprod defG chi \is a character) = (chi \is a character). +Proof. by rewrite unlock cfMorph_charE // cfIsom_char. Qed. Lemma cfSdprod_lin_char chi : - chi \is a linear_char -> cfSdprod defG chi \is a linear_char. -Proof. by move=> Nphi; rewrite unlock cfMorph_lin_char ?cfIsom_lin_char. Qed. + (cfSdprod defG chi \is a linear_char) = (chi \is a linear_char). +Proof. by rewrite qualifE cfSdprod_char cfSdprod1. Qed. -Lemma cfSdprod_irr chi : chi \in irr H -> cfSdprod defG chi \in irr G. -Proof. -have [/andP[_ nKG] _ _ _ _] := sdprod_context defG. -by move=> Nphi; rewrite unlock cfMorph_irr ?cfIsom_irr. -Qed. +Lemma cfSdprod_irr chi : (cfSdprod defG chi \in irr G) = (chi \in irr H). +Proof. by rewrite !irrEchar cfSdprod_char cfSdprod_iso. Qed. Definition sdprod_Iirr j := cfIirr (cfSdprod defG 'chi_j). @@ -1707,17 +1951,17 @@ Lemma cfDprodKr_abelian i : abelian K -> cancel (cfDprod KxH 'chi_i) 'Res. Proof. by move=> cKK; apply: cfDprodKr; apply/lin_char1/char_abelianP. Qed. Lemma cfDprodl_char phi : - phi \is a character -> cfDprodl KxH phi \is a character. + (cfDprodl KxH phi \is a character) = (phi \is a character). Proof. exact: cfSdprod_char. Qed. Lemma cfDprodr_char psi : - psi \is a character -> cfDprodr KxH psi \is a character. + (cfDprodr KxH psi \is a character) = (psi \is a character). Proof. exact: cfSdprod_char. Qed. Lemma cfDprod_char phi psi : phi \is a character -> psi \is a character -> cfDprod KxH phi psi \is a character. -Proof. by move=> /cfDprodl_char Nphi /cfDprodr_char; apply: rpredM. Qed. +Proof. by move=> Nphi Npsi; rewrite rpredM ?cfDprodl_char ?cfDprodr_char. Qed. Lemma cfDprod_eq1 phi psi : phi \is a character -> psi \is a character -> @@ -1732,22 +1976,22 @@ by rewrite !rmorph1. Qed. Lemma cfDprodl_lin_char phi : - phi \is a linear_char -> cfDprodl KxH phi \is a linear_char. + (cfDprodl KxH phi \is a linear_char) = (phi \is a linear_char). Proof. exact: cfSdprod_lin_char. Qed. Lemma cfDprodr_lin_char psi : - psi \is a linear_char -> cfDprodr KxH psi \is a linear_char. + (cfDprodr KxH psi \is a linear_char) = (psi \is a linear_char). Proof. exact: cfSdprod_lin_char. Qed. Lemma cfDprod_lin_char phi psi : phi \is a linear_char -> psi \is a linear_char -> cfDprod KxH phi psi \is a linear_char. -Proof. by move=> /cfDprodl_lin_char Lphi /cfDprodr_lin_char; apply: rpredM. Qed. +Proof. by move=> Nphi Npsi; rewrite rpredM ?cfSdprod_lin_char. Qed. -Lemma cfDprodl_irr chi : chi \in irr K -> cfDprodl KxH chi \in irr G. +Lemma cfDprodl_irr chi : (cfDprodl KxH chi \in irr G) = (chi \in irr K). Proof. exact: cfSdprod_irr. Qed. -Lemma cfDprodr_irr chi : chi \in irr H -> cfDprodr KxH chi \in irr G. +Lemma cfDprodr_irr chi : (cfDprodr KxH chi \in irr G) = (chi \in irr H). Proof. exact: cfSdprod_irr. Qed. Definition dprodl_Iirr i := cfIirr (cfDprodl KxH 'chi_i). @@ -1834,10 +2078,10 @@ Qed. Definition inv_dprod_Iirr i := iinv (dprod_Iirr_onto i). Lemma dprod_IirrK : cancel dprod_Iirr inv_dprod_Iirr. -Proof. by move=> p; exact: (iinv_f dprod_Iirr_inj). Qed. +Proof. by move=> p; apply: (iinv_f dprod_Iirr_inj). Qed. Lemma inv_dprod_IirrK : cancel inv_dprod_Iirr dprod_Iirr. -Proof. by move=> i; exact: f_iinv. Qed. +Proof. by move=> i; apply: f_iinv. Qed. Lemma inv_dprod_Iirr0 : inv_dprod_Iirr 0 = (0, 0). Proof. by apply/(canLR dprod_IirrK); rewrite dprod_Iirr0. Qed. @@ -1862,7 +2106,11 @@ Proof. by move=> Pi; rewrite -(bigdprodWY defG) (bigD1 i) ?joing_subl. Qed. Lemma cfBigdprodi_char i (phi : 'CF(A i)) : phi \is a character -> cfBigdprodi defG phi \is a character. -Proof. by move=> Nphi; rewrite cfDprodl_char ?cfRes_char. Qed. +Proof. by move=> Nphi; rewrite cfDprodl_char cfRes_char. Qed. + +Lemma cfBigdprodi_charE i (phi : 'CF(A i)) : + P i -> (cfBigdprodi defG phi \is a character) = (phi \is a character). +Proof. by move=> Pi; rewrite cfDprodl_char Pi cfRes_id. Qed. Lemma cfBigdprod_char phi : (forall i, P i -> phi i \is a character) -> @@ -1875,6 +2123,10 @@ Lemma cfBigdprodi_lin_char i (phi : 'CF(A i)) : phi \is a linear_char -> cfBigdprodi defG phi \is a linear_char. Proof. by move=> Lphi; rewrite cfDprodl_lin_char ?cfRes_lin_char. Qed. +Lemma cfBigdprodi_lin_charE i (phi : 'CF(A i)) : + P i -> (cfBigdprodi defG phi \is a linear_char) = (phi \is a linear_char). +Proof. by move=> Pi; rewrite qualifE cfBigdprodi_charE // cfBigdprodi1. Qed. + Lemma cfBigdprod_lin_char phi : (forall i, P i -> phi i \is a linear_char) -> cfBigdprod defG phi \is a linear_char. @@ -1883,8 +2135,8 @@ by move=> Lphi; apply/rpred_prod=> i /Lphi; apply: cfBigdprodi_lin_char. Qed. Lemma cfBigdprodi_irr i chi : - P i -> chi \in irr (A i) -> cfBigdprodi defG chi \in irr G. -Proof. by move=> Pi Nchi; rewrite cfDprodl_irr // Pi cfRes_id. Qed. + P i -> (cfBigdprodi defG chi \in irr G) = (chi \in irr (A i)). +Proof. by move=> Pi; rewrite !irrEchar cfBigdprodi_charE ?cfBigdprodi_iso. Qed. Lemma cfBigdprod_irr chi : (forall i, P i -> chi i \in irr (A i)) -> cfBigdprod defG chi \in irr G. @@ -1946,7 +2198,7 @@ Lemma conjC_charAut u (chi : 'CF(G)) x : Proof. have [Gx | /cfun0->] := boolP (x \in G); last by rewrite !rmorph0. case/char_reprP=> rG ->; have [e [_ [en1 _] [-> _] _]] := repr_rsim_diag rG Gx. -by rewrite !rmorph_sum; apply: eq_bigr => i _; exact: aut_unity_rootC (en1 i). +by rewrite !rmorph_sum; apply: eq_bigr => i _; apply: aut_unity_rootC (en1 i). Qed. Lemma conjC_irrAut u i x : (u ('chi[G]_i x))^* = u ('chi_i x)^*. @@ -1960,17 +2212,17 @@ Lemma cfdot_aut_irr u phi i : '[cfAut u phi, cfAut u 'chi[G]_i] = u '[phi, 'chi_i]. Proof. exact: cfdot_aut_char (irr_char i). Qed. -Lemma cfAut_irr u chi : chi \in irr G -> cfAut u chi \in irr G. +Lemma cfAut_irr u chi : (cfAut u chi \in irr G) = (chi \in irr G). Proof. -case/irrP=> i ->; rewrite irrEchar cfAut_char ?irr_char //=. -by rewrite cfdot_aut_irr // cfdot_irr eqxx rmorph1. +rewrite !irrEchar cfAut_char; apply/andb_id2l=> /cfdot_aut_char->. +exact: fmorph_eq1. Qed. Lemma cfConjC_irr i : (('chi_i)^*)%CF \in irr G. -Proof. by rewrite cfAut_irr ?mem_irr. Qed. +Proof. by rewrite cfAut_irr mem_irr. Qed. Lemma irr_aut_closed u : cfAut_closed u (irr G). -Proof. exact: cfAut_irr. Qed. +Proof. by move=> chi; rewrite /= cfAut_irr. Qed. Definition aut_Iirr u i := cfIirr (cfAut u 'chi[G]_i). @@ -2002,255 +2254,10 @@ Proof. by move=> i j eq_ij; apply/irr_inj/(cfAut_inj u); rewrite -!aut_IirrE eq_ij. Qed. -Lemma char_aut u (chi : 'CF(G)) : - (cfAut u chi \is a character) = (chi \is a character). -Proof. -apply/idP/idP=> [Nuchi|]; last exact: cfAut_char. -rewrite [chi]cfun_sum_cfdot rpred_sum // => i _. -rewrite rpredZ_Cnat ?irr_char // -(Cnat_aut u) -cfdot_aut_irr. -by rewrite -aut_IirrE Cnat_cfdot_char_irr. -Qed. - -Lemma irr_aut u chi : (cfAut u chi \in irr G) = (chi \in irr G). -Proof. -rewrite !irrEchar char_aut; apply/andb_id2l=> /cfdot_aut_char->. -by rewrite fmorph_eq1. -Qed. - End Aut. -Section IrrConstt. - -Variable (gT : finGroupType) (G H : {group gT}). - -Lemma char1_ge_norm (chi : 'CF(G)) x : - chi \is a character -> `|chi x| <= chi 1%g. -Proof. -case/char_reprP=> rG ->; case Gx: (x \in G); last first. - by rewrite cfunE cfRepr1 Gx normr0 ler0n. -by have [e [_ _ []]] := repr_rsim_diag rG Gx. -Qed. - -Lemma max_cfRepr_norm_scalar n (rG : mx_representation algCF G n) x : - x \in G -> `|cfRepr rG x| = cfRepr rG 1%g -> - exists2 c, `|c| = 1 & rG x = c%:M. -Proof. -move=> Gx; have [e [[B uB def_x] [_ e1] [-> _] _]] := repr_rsim_diag rG Gx. -rewrite cfRepr1 -[n in n%:R]card_ord -sumr_const -(eq_bigr _ (in1W e1)). -case/normC_sum_eq1=> [i _ | c /eqP norm_c_1 def_e]; first by rewrite e1. -have{def_e} def_e: e = const_mx c by apply/rowP=> i; rewrite mxE def_e ?andbT. -by exists c => //; rewrite def_x def_e diag_const_mx scalar_mxC mulmxKV. -Qed. - -Lemma max_cfRepr_mx1 n (rG : mx_representation algCF G n) x : - x \in G -> cfRepr rG x = cfRepr rG 1%g -> rG x = 1%:M. -Proof. -move=> Gx kerGx; have [|c _ def_x] := @max_cfRepr_norm_scalar n rG x Gx. - by rewrite kerGx cfRepr1 normr_nat. -move/eqP: kerGx; rewrite cfRepr1 cfunE Gx {rG}def_x mxtrace_scalar. -case: n => [_|n]; first by rewrite ![_%:M]flatmx0. -rewrite mulrb -subr_eq0 -mulrnBl -mulr_natl mulf_eq0 pnatr_eq0 /=. -by rewrite subr_eq0 => /eqP->. -Qed. - -Definition irr_constt (B : {set gT}) phi := [pred i | '[phi, 'chi_i]_B != 0]. - -Lemma irr_consttE i phi : (i \in irr_constt phi) = ('[phi, 'chi_i]_G != 0). -Proof. by []. Qed. - -Lemma constt_charP (i : Iirr G) chi : - chi \is a character -> - reflect (exists2 chi', chi' \is a character & chi = 'chi_i + chi') - (i \in irr_constt chi). -Proof. -move=> Nchi; apply: (iffP idP) => [i_in_chi| [chi' Nchi' ->]]; last first. - rewrite inE /= cfdotDl cfdot_irr eqxx -(eqP (Cnat_cfdot_char_irr i Nchi')). - by rewrite -natrD pnatr_eq0. -exists (chi - 'chi_i); last by rewrite addrC subrK. -apply/forallP=> j; rewrite coord_cfdot cfdotBl cfdot_irr. -have [<- | _] := eqP; last by rewrite subr0 Cnat_cfdot_char_irr. -have := i_in_chi; rewrite inE /= -(eqP (Cnat_cfdot_char_irr i Nchi)) pnatr_eq0. -by case: (truncC _) => // n _; rewrite mulrSr addrK ?isNatC_nat. -Qed. - -Lemma cfun_sum_constt (phi : 'CF(G)) : - phi = \sum_(i in irr_constt phi) '[phi, 'chi_i] *: 'chi_i. -Proof. -rewrite {1}[phi]cfun_sum_cfdot (bigID [pred i | '[phi, 'chi_i] == 0]) /=. -by rewrite big1 ?add0r // => i /eqP->; rewrite scale0r. -Qed. - -Lemma neq0_has_constt (phi : 'CF(G)) : - phi != 0 -> exists i, i \in irr_constt phi. -Proof. -move=> nz_phi; apply/existsP; apply: contra nz_phi => /pred0P phi0. -by rewrite [phi]cfun_sum_constt big_pred0. -Qed. - -Lemma constt_irr i : irr_constt 'chi[G]_i =i pred1 i. -Proof. -by move=> j; rewrite !inE cfdot_irr pnatr_eq0 (eq_sym j); case: (i == j). -Qed. - -Lemma char1_ge_constt (i : Iirr G) chi : - chi \is a character -> i \in irr_constt chi -> 'chi_i 1%g <= chi 1%g. -Proof. -move=> {chi} _ /constt_charP[// | chi Nchi ->]. -by rewrite cfunE addrC -subr_ge0 addrK char1_ge0. -Qed. - -Lemma constt_ortho_char (phi psi : 'CF(G)) i j : - phi \is a character -> psi \is a character -> - i \in irr_constt phi -> j \in irr_constt psi -> - '[phi, psi] = 0 -> '['chi_i, 'chi_j] = 0. -Proof. -move=> _ _ /constt_charP[//|phi1 Nphi1 ->] /constt_charP[//|psi1 Npsi1 ->]. -rewrite cfdot_irr; case: eqP => // -> /eqP/idPn[]. -rewrite cfdotDl !cfdotDr cfnorm_irr -addrA gtr_eqF ?ltr_paddr ?ltr01 //. -by rewrite Cnat_ge0 ?rpredD ?Cnat_cfdot_char ?irr_char. -Qed. - -End IrrConstt. - -Arguments Scope irr_constt [_ group_scope cfun_scope]. Implicit Arguments aut_Iirr_inj [gT G x1 x2]. -Section MoreConstt. - -Variables (gT : finGroupType) (G H : {group gT}). - -Lemma constt_Ind_Res i j : - i \in irr_constt ('Ind[G] 'chi_j) = (j \in irr_constt ('Res[H] 'chi_i)). -Proof. by rewrite !irr_consttE cfdotC conjC_eq0 -cfdot_Res_l. Qed. - -Lemma cfdot_Res_ge_constt i j psi : - psi \is a character -> j \in irr_constt psi -> - '['Res[H, G] 'chi_j, 'chi_i] <= '['Res[H] psi, 'chi_i]. -Proof. -move=> {psi} _ /constt_charP[// | psi Npsi ->]. -rewrite linearD cfdotDl addrC -subr_ge0 addrK Cnat_ge0 //=. -by rewrite Cnat_cfdot_char_irr // cfRes_char. -Qed. - -Lemma constt_Res_trans j psi : - psi \is a character -> j \in irr_constt psi -> - {subset irr_constt ('Res[H, G] 'chi_j) <= irr_constt ('Res[H] psi)}. -Proof. -move=> Npsi Cj i; apply: contraNneq; rewrite eqr_le => {1}<-. -rewrite cfdot_Res_ge_constt ?Cnat_ge0 ?Cnat_cfdot_char_irr //. -by rewrite cfRes_char ?irr_char. -Qed. - -End MoreConstt. - -Section Kernel. - -Variable (gT : finGroupType) (G : {group gT}). -Implicit Types (phi chi xi : 'CF(G)) (H : {group gT}). - -Lemma cfker_repr n (rG : mx_representation algCF G n) : - cfker (cfRepr rG) = rker rG. -Proof. -apply/esym/setP=> x; rewrite inE mul1mx /=. -case Gx: (x \in G); last by rewrite inE Gx. -apply/eqP/idP=> Kx; last by rewrite max_cfRepr_mx1 // cfker1. -rewrite inE Gx; apply/forallP=> y; rewrite !cfunE !mulrb groupMl //. -by case: ifP => // Gy; rewrite repr_mxM // Kx mul1mx. -Qed. - -Lemma cfkerEchar chi : - chi \is a character -> cfker chi = [set x in G | chi x == chi 1%g]. -Proof. -move=> Nchi; apply/setP=> x; apply/idP/setIdP=> [Kx | [Gx /eqP chi_x]]. - by rewrite (subsetP (cfker_sub chi)) // cfker1. -case/char_reprP: Nchi => rG -> in chi_x *; rewrite inE Gx; apply/forallP=> y. -rewrite !cfunE groupMl // !mulrb; case: ifP => // Gy. -by rewrite repr_mxM // max_cfRepr_mx1 ?mul1mx. -Qed. - -Lemma cfker_nzcharE chi : - chi \is a character -> chi != 0 -> cfker chi = [set x | chi x == chi 1%g]. -Proof. -move=> Nchi nzchi; apply/setP=> x; rewrite cfkerEchar // !inE andb_idl //. -by apply: contraLR => /cfun0-> //; rewrite eq_sym char1_eq0. -Qed. - -Lemma cfkerEirr i : cfker 'chi[G]_i = [set x | 'chi_i x == 'chi_i 1%g]. -Proof. by rewrite cfker_nzcharE ?irr_char ?irr_neq0. Qed. - -Lemma cfker_irr0 : cfker 'chi[G]_0 = G. -Proof. by rewrite irr0 cfker_cfun1. Qed. - -Lemma cfaithful_reg : cfaithful (cfReg G). -Proof. -apply/subsetP=> x; rewrite cfkerEchar ?cfReg_char // !inE !cfRegE eqxx. -by case/andP=> _; apply: contraLR => /negbTE->; rewrite eq_sym neq0CG. -Qed. - -Lemma cfkerE chi : - chi \is a character -> - cfker chi = G :&: \bigcap_(i in irr_constt chi) cfker 'chi_i. -Proof. -move=> Nchi; rewrite cfkerEchar //; apply/setP=> x; rewrite !inE. -apply: andb_id2l => Gx; rewrite {1 2}[chi]cfun_sum_constt !sum_cfunE. -apply/eqP/bigcapP=> [Kx i Ci | Kx]; last first. - by apply: eq_bigr => i /Kx Kx_i; rewrite !cfunE cfker1. -rewrite cfkerEirr inE /= -(inj_eq (mulfI Ci)). -have:= (normC_sum_upper _ Kx) i; rewrite !cfunE => -> // {i Ci} i _. -have chi_i_ge0: 0 <= '[chi, 'chi_i]. - by rewrite Cnat_ge0 ?Cnat_cfdot_char_irr. -by rewrite !cfunE normrM (normr_idP _) ?ler_wpmul2l ?char1_ge_norm ?irr_char. -Qed. - -Lemma TI_cfker_irr : \bigcap_i cfker 'chi[G]_i = [1]. -Proof. -apply/trivgP; apply: subset_trans cfaithful_reg; rewrite cfkerE ?cfReg_char //. -rewrite subsetI (bigcap_min 0) //=; last by rewrite cfker_irr0. -by apply/bigcapsP=> i _; rewrite bigcap_inf. -Qed. - -Lemma cfker_constt i chi : - chi \is a character -> i \in irr_constt chi -> - cfker chi \subset cfker 'chi[G]_i. -Proof. by move=> Nchi Ci; rewrite cfkerE ?subIset ?(bigcap_min i) ?orbT. Qed. - -Section KerLin. - -Variable xi : 'CF(G). -Hypothesis lin_xi : xi \is a linear_char. -Let Nxi: xi \is a character. Proof. by have [] := andP lin_xi. Qed. - -Lemma lin_char_der1 : G^`(1)%g \subset cfker xi. -Proof. -rewrite gen_subG /=; apply/subsetP=> _ /imset2P[x y Gx Gy ->]. -rewrite cfkerEchar // inE groupR //= !lin_charM ?lin_charV ?in_group //. -by rewrite mulrCA mulKf ?mulVf ?lin_char_neq0 // lin_char1. -Qed. - -Lemma cforder_lin_char : #[xi]%CF = exponent (G / cfker xi)%g. -Proof. -apply/eqP; rewrite eqn_dvd; apply/andP; split. - apply/dvdn_cforderP=> x Gx; rewrite -lin_charX // -cfQuoEker ?groupX //. - rewrite morphX ?(subsetP (cfker_norm xi)) //= expg_exponent ?mem_quotient //. - by rewrite cfQuo1 ?cfker_normal ?lin_char1. -have abGbar: abelian (G / cfker xi) := sub_der1_abelian lin_char_der1. -have [_ /morphimP[x Nx Gx ->] ->] := exponent_witness (abelian_nil abGbar). -rewrite order_dvdn -morphX //= coset_id cfkerEchar // !inE groupX //=. -by rewrite lin_charX ?lin_char1 // (dvdn_cforderP _ _ _). -Qed. - -Lemma cforder_lin_char_dvdG : #[xi]%CF %| #|G|. -Proof. -by rewrite cforder_lin_char (dvdn_trans (exponent_dvdn _)) ?dvdn_morphim. -Qed. - -Lemma cforder_lin_char_gt0 : (0 < #[xi]%CF)%N. -Proof. by rewrite cforder_lin_char exponent_gt0. Qed. - -End KerLin. - -End Kernel. - Section Coset. Variable (gT : finGroupType). @@ -2260,27 +2267,18 @@ Implicit Types G H : {group gT}. Lemma cfQuo_char G H (chi : 'CF(G)) : chi \is a character -> (chi / H)%CF \is a character. Proof. -move=> Nchi; case KchiH: (H \subset cfker chi); last first. +move=> Nchi; without loss kerH: / H \subset cfker chi. + move/contraNF=> IHchi; apply/wlog_neg=> N'chiH. suffices ->: (chi / H)%CF = (chi 1%g)%:A. by rewrite rpredZ_Cnat ?Cnat_char1 ?rpred1. - by apply/cfunP=> x; rewrite cfunE cfun1E mulr_natr cfunElock KchiH. -have sHG := subset_trans KchiH (cfker_sub _). -pose N := 'N_G(H); pose phi := 'Res[N] chi. -have nsHN: H <| N by [rewrite normal_subnorm]; have [sHN nHN] := andP nsHN. -have{Nchi} Nphi: phi \is a character by apply: cfRes_char. -have KphiH: H \subset cfker phi. - apply/subsetP=> x Hx; have [Kx Nx] := (subsetP KchiH x Hx, subsetP sHN x Hx). - by rewrite cfkerEchar // inE Nx cfRes1 cfResE ?subsetIl //= cfker1. -pose psi := 'Res[(N / H)%g] (chi / H)%CF. -have ->: (chi / H)%CF = 'Res psi by rewrite /psi quotientInorm !cfRes_id. -have{KchiH} ->: psi = (phi / H)%CF. - apply/cfun_inP => _ /morphimP[x nHx Nx ->]; have [Gx _] := setIP Nx. - rewrite cfResE ?mem_quotient ?quotientS ?subsetIl // cfQuoEnorm //. - by rewrite cfQuoE ?cfResE ?subsetIl. -have [rG Dphi] := char_reprP Nphi; rewrite {phi Nphi}Dphi cfker_repr in KphiH *. -apply/cfRes_char/char_reprP; exists (Representation (quo_repr KphiH nHN)). -apply/cfun_inP=> _ /morphimP[x nHx Nx ->]; rewrite cfQuoE ?cfker_repr //=. -by rewrite !cfunE Nx quo_repr_coset ?mem_quotient. + by apply/cfunP=> x; rewrite cfunE cfun1E mulr_natr cfunElock IHchi. +without loss nsHG: G chi Nchi kerH / H <| G. + move=> IHchi; have nsHN := normalSG (subset_trans kerH (cfker_sub chi)). + by rewrite cfQuoInorm ?(cfRes_char, IHchi) ?sub_cfker_Res // ?normal_sub. +have [rG Dchi] := char_reprP Nchi; rewrite Dchi cfker_repr in kerH. +apply/char_reprP; exists (Representation (quo_repr kerH (normal_norm nsHG))). +apply/cfun_inP=> _ /morphimP[x nHx Gx ->]; rewrite Dchi cfQuoE ?cfker_repr //=. +by rewrite !cfunE Gx quo_repr_coset ?mem_quotient. Qed. Lemma cfQuo_lin_char G H (chi : 'CF(G)) : @@ -2295,8 +2293,26 @@ Lemma cfMod_lin_char G H (chi : 'CF(G / H)) : chi \is a linear_char -> (chi %% H)%CF \is a linear_char. Proof. exact: cfMorph_lin_char. Qed. +Lemma cfMod_charE G H (chi : 'CF(G / H)) : + H <| G -> (chi %% H \is a character)%CF = (chi \is a character). +Proof. by case/andP=> _; apply: cfMorph_charE. Qed. + +Lemma cfMod_lin_charE G H (chi : 'CF(G / H)) : + H <| G -> (chi %% H \is a linear_char)%CF = (chi \is a linear_char). +Proof. by case/andP=> _; apply: cfMorph_lin_charE. Qed. + +Lemma cfQuo_charE G H (chi : 'CF(G)) : + H <| G -> H \subset cfker chi -> + (chi / H \is a character)%CF = (chi \is a character). +Proof. by move=> nsHG kerH; rewrite -cfMod_charE ?cfQuoK. Qed. + +Lemma cfQuo_lin_charE G H (chi : 'CF(G)) : + H <| G -> H \subset cfker chi -> + (chi / H \is a linear_char)%CF = (chi \is a linear_char). +Proof. by move=> nsHG kerH; rewrite -cfMod_lin_charE ?cfQuoK. Qed. + Lemma cfMod_irr G H chi : - H <| G -> chi \in irr (G / H) -> (chi %% H)%CF \in irr G. + H <| G -> (chi %% H \in irr G)%CF = (chi \in irr (G / H)). Proof. by case/andP=> _; apply: cfMorph_irr. Qed. Definition mod_Iirr G H i := cfIirr ('chi[G / H]_i %% H)%CF. @@ -2312,16 +2328,9 @@ Lemma mod_Iirr_eq0 G H i : Proof. by case/andP=> _ /morph_Iirr_eq0->. Qed. Lemma cfQuo_irr G H chi : - H <| G -> H \subset cfker chi -> chi \in irr G -> - (chi / H)%CF \in irr (G / H). -Proof. -move=> nsHG sHK /irr_reprP[rG irrG Dchi]; have [sHG nHG] := andP nsHG. -have sHKr: H \subset rker rG by rewrite -cfker_repr -Dchi. -apply/irr_reprP; exists (Representation (quo_repr sHKr nHG)). - exact/quo_mx_irr. -apply/cfun_inP=> _ /morphimP[x Nx Gx ->]. -by rewrite cfQuoE //= Dchi !cfunE Gx quo_repr_coset ?mem_quotient. -Qed. + H <| G -> H \subset cfker chi -> + ((chi / H)%CF \in irr (G / H)) = (chi \in irr G). +Proof. by move=> nsHG kerH; rewrite -cfMod_irr ?cfQuoK. Qed. Definition quo_Iirr G H i := cfIirr ('chi[G]_i / H)%CF. @@ -2359,7 +2368,7 @@ Qed. Lemma mod_Iirr_bij H G : H <| G -> {on [pred i | H \subset cfker 'chi_i], bijective (@mod_Iirr G H)}. Proof. -by exists (quo_Iirr H) => [i _ | i]; [exact: mod_IirrK | exact: quo_IirrK]. +by exists (quo_Iirr H) => [i _ | i]; [apply: mod_IirrK | apply: quo_IirrK]. Qed. Lemma sum_norm_irr_quo H G x : @@ -2394,7 +2403,7 @@ Qed. End Coset. -Section Derive. +Section DerivedGroup. Variable gT : finGroupType. Implicit Types G H : {group gT}. @@ -2405,7 +2414,7 @@ Proof. apply/idP/idP=> [|sG'K]; first by apply: lin_char_der1. have nsG'G: G^`(1) <| G := der_normal 1 G. rewrite qualifE irr_char -[i](quo_IirrK nsG'G) // mod_IirrE //=. -by rewrite cfModE // morph1 lin_char1 //; exact/char_abelianP/der_abelian. +by rewrite cfModE // morph1 lin_char1 //; apply/char_abelianP/der_abelian. Qed. Lemma subGcfker G i : (G \subset cfker 'chi[G]_i) = (i == 0). @@ -2447,7 +2456,7 @@ by apply: eq_card => i; rewrite !inE mod_IirrE ?cfker_mod. (* Alternative: use the equivalent result in modular representation theory transitivity #|@socle_of_Iirr _ G @^-1: linear_irr _|; last first. rewrite (on_card_preimset (socle_of_Iirr_bij _)). - by rewrite card_linear_irr ?algC'G; last exact: groupC. + by rewrite card_linear_irr ?algC'G; last apply: groupC. by apply: eq_card => i; rewrite !inE /lin_char irr_char irr1_degree -eqC_nat. *) Qed. @@ -2466,7 +2475,7 @@ Qed. (* A combinatorial group isommorphic to the linear characters. *) Lemma lin_char_group G : - {linG : finGroupType & {cF : linG -> 'CF(G) | + {linG : finGroupType & {cF : linG -> 'CF(G) | [/\ injective cF, #|linG| = #|G : G^`(1)|, forall u, cF u \is a linear_char & forall phi, phi \is a linear_char -> exists u, phi = cF u] @@ -2542,7 +2551,7 @@ rewrite (bigID (fun i => H \subset cfker 'chi[G]_i)) //= addrC addKr. by apply: sumr_ge0 => i _; rewrite normCK mul_conjC_ge0. Qed. -End Derive. +End DerivedGroup. Implicit Arguments irr_prime_injP [gT G i]. @@ -2789,7 +2798,7 @@ have [-> | nz_chi] := eqVneq chi 0. by apply/subsetP=> x /setIdP[Gx _]; rewrite inE Gx /= !cfunE. have [xi Lxi def_chi] := cfcenter_Res chi. set Z := ('Z(_))%CF in xi Lxi def_chi *. -have sZG: Z \subset G by exact: cfcenter_sub. +have sZG: Z \subset G by apply: cfcenter_sub. have ->: cfker chi = cfker xi. rewrite -(setIidPr (normal_sub (cfker_center_normal _))) -/Z. rewrite !cfkerEchar // ?lin_charW //= -/Z. @@ -2797,7 +2806,7 @@ have ->: cfker chi = cfker xi. rewrite (subsetP sZG) //= -!(cfResE chi sZG) ?group1 // def_chi !cfunE. by rewrite (inj_eq (mulfI _)) ?char1_eq0. have: abelian (Z / cfker xi) by rewrite sub_der1_abelian ?lin_char_der1. -have [rG irrG ->] := irr_reprP _ (lin_char_irr Lxi); rewrite cfker_repr. +have /irr_reprP[rG irrG ->] := lin_char_irr Lxi; rewrite cfker_repr. apply: mx_faithful_irr_abelian_cyclic (kquo_mx_faithful rG) _. exact/quo_mx_irr. Qed. @@ -2824,7 +2833,7 @@ apply/subsetP=> _ /setIP[/morphimP[x /= _ Gx ->] cGx]; rewrite mem_quotient //=. rewrite -irrRepr cfker_repr cfcenter_repr inE Gx in cGx *. apply: mx_abs_irr_cent_scalar 'Chi_i _ _ _. by apply: groupC; apply: socle_irr. -have nKG: G \subset 'N(rker 'Chi_i) by exact: rker_norm. +have nKG: G \subset 'N(rker 'Chi_i) by apply: rker_norm. (* GG -- locking here is critical to prevent Coq kernel divergence. *) apply/centgmxP=> y Gy; rewrite [eq]lock -2?(quo_repr_coset (subxx _) nKG) //. move: (quo_repr _ _) => rG; rewrite -2?repr_mxM ?mem_quotient // -lock. @@ -2915,10 +2924,10 @@ Lemma pgroup_cyclic_faithful (p : nat) : p.-group G -> cyclic 'Z(G) -> exists i, cfaithful 'chi[G]_i. Proof. pose Z := 'Ohm_1('Z(G)) => pG cycZG; have nilG := pgroup_nil pG. -have [-> | ntG] := eqsVneq G [1]; first by exists 0; exact: cfker_sub. +have [-> | ntG] := eqsVneq G [1]; first by exists 0; apply: cfker_sub. have{pG} [[p_pr _ _] pZ] := (pgroup_pdiv pG ntG, pgroupS (center_sub G) pG). have ntZ: 'Z(G) != [1] by rewrite center_nil_eq1. -have{pZ} oZ: #|Z| = p by exact: Ohm1_cyclic_pgroup_prime. +have{pZ} oZ: #|Z| = p by apply: Ohm1_cyclic_pgroup_prime. apply/existsP; apply: contraR ntZ; rewrite negb_exists => /forallP-not_ffulG. rewrite -Ohm1_eq1 -subG1 /= -/Z -(TI_cfker_irr G); apply/bigcapsP=> i _. rewrite prime_meetG ?oZ // setIC meet_Ohm1 // meet_center_nil ?cfker_normal //. @@ -2973,7 +2982,7 @@ move=> sHG Nchi nzchi; rewrite !cfker_nzcharE ?cfInd_char ?cfInd_eq0 //. apply/setP=> x; rewrite inE cfIndE // (can2_eq (mulVKf _) (mulKf _)) ?neq0CG //. rewrite cfInd1 // mulrA -natrM Lagrange // mulr_natl -sumr_const. apply/eqP/bigcapP=> [/normC_sum_upper ker_chiG_x y Gy | ker_chiG_x]. - by rewrite mem_conjg inE ker_chiG_x ?groupV // => z _; exact: char1_ge_norm. + by rewrite mem_conjg inE ker_chiG_x ?groupV // => z _; apply: char1_ge_norm. by apply: eq_bigr => y /groupVr/ker_chiG_x; rewrite mem_conjgV inE => /eqP. Qed. diff --git a/mathcomp/character/classfun.v b/mathcomp/character/classfun.v index 32eeabb..3b3d894 100644 --- a/mathcomp/character/classfun.v +++ b/mathcomp/character/classfun.v @@ -1,25 +1,23 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype tuple finfun bigop prime finset. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism quotient action gproduct. -From mathcomp.algebra -Require Import ssralg poly finalg zmodp cyclic vector ssrnum matrix vector. -From mathcomp.solvable -Require Import commutator center pgroup sylow. -From mathcomp.field -Require Import falgebra algC algnum. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp +Require Import fintype tuple finfun bigop prime ssralg poly finset. +From mathcomp +Require Import fingroup morphism perm automorphism quotient finalg action. +From mathcomp +Require Import gproduct zmodp commutator cyclic center pgroup sylow. +From mathcomp +Require Import matrix vector falgebra ssrnum algC algnum. (******************************************************************************) (* This file contains the basic theory of class functions: *) (* 'CF(G) == the type of class functions on G : {group gT}, i.e., *) (* which map gT to the type algC of complex algebraics, *) (* have support in G, and are constant on each conjugacy *) -(* class of G. 'CF(G) implements the algebraType interface *) -(* of finite-dimensional F-algebras. *) +(* class of G. 'CF(G) implements the FalgType interface of *) +(* finite-dimensional F-algebras. *) (* The identity 1 : 'CF(G) is the indicator function of G, *) (* and (later) the principal character. *) (* --> The %CF scope (cfun_scope) is bound to the 'CF(_) types. *) @@ -28,7 +26,7 @@ Require Import falgebra algC algnum. (* phi x == the image of x : gT under phi : 'CF(G). *) (* #[phi]%CF == the multiplicative order of phi : 'CF(G). *) (* cfker phi == the kernel of phi : 'CF(G); note that cfker phi <| G. *) -(* cfaithful phi <=> phi : 'CF(G) is faithful (has a trivial kernel). *) +(* cfaithful phi <=> phi : 'CF(G) is faithful (has a trivial kernel). *) (* '1_A == the indicator function of A as a function of 'CF(G). *) (* (Provided A <| G; G is determined by the context.) *) (* phi^*%CF == the function conjugate to phi : 'CF(G). *) @@ -52,10 +50,10 @@ Require Import falgebra algC algnum. (* whose range is contained in CR. *) (* cfReal phi <=> phi is real, i.e., phi^* == phi. *) (* cfAut_closed u S <-> S : seq 'CF(G) is closed under conjugation by u. *) -(* conjC_closed S <-> S : seq 'CF(G) is closed under complex conjugation. *) +(* cfConjC_closed S <-> S : seq 'CF(G) is closed under complex conjugation. *) (* conjC_subset S1 S2 <-> S1 : seq 'CF(G) represents a subset of S2 closed *) (* under complex conjugation. *) -(* := [/\ uniq S1, {subset S1 <= S2} & conjC_closed S1]. *) +(* := [/\ uniq S1, {subset S1 <= S2} & cfConjC_closed S1]. *) (* 'Res[H] phi == the restriction of phi : 'CF(G) to a function of 'CF(H) *) (* 'Res[H, G] phi 'Res[H] phi x = phi x if x \in H (when H \subset G), *) (* 'Res phi 'Res[H] phi x = 0 if x \notin H. The syntax variants *) @@ -88,7 +86,7 @@ Require Import falgebra algC algnum. (* cfBigdprodi defG phi == for phi : 'CF(A i) s.t. P i, the class function *) (* of 'CF(G) that maps x to phi x_i, where x_i is the *) (* (A i)-component of x : G. *) -(* cfBigdprodi defG phi == for phi : forall i, 'CF(A i), the class function *) +(* cfBigdprod defG phi == for phi : forall i, 'CF(A i), the class function *) (* of 'CF(G) that maps x to \prod_(i | P i) phi i x_i, *) (* where x_i is the (A i)-component of x : G. *) (******************************************************************************) @@ -181,7 +179,7 @@ Lemma cfunP phi psi : phi =1 psi <-> phi = psi. Proof. by split=> [/ffunP/val_inj | ->]. Qed. Lemma cfun0gen phi x : x \notin G -> phi x = 0. -Proof. by case: phi => f fP; case: (andP fP) => _ /supportP; exact. Qed. +Proof. by case: phi => f fP; case: (andP fP) => _ /supportP; apply. Qed. Lemma cfun_in_genP phi psi : {in G, phi =1 psi} -> phi = psi. Proof. @@ -335,8 +333,8 @@ Proof. by apply/cfunP=> x; rewrite !cfunE rmorphM. Qed. Lemma cfAut_is_rmorphism : rmorphism cfAut. Proof. -by do 2?split=> [phi psi|]; last exact: cfAut_cfun1i; - apply/cfunP=> x; rewrite !cfunE (rmorphB, rmorphM). +by do 2?split=> [phi psi|]; apply/cfunP=> x; + rewrite ?cfAut_cfun1i // !cfunE (rmorphB, rmorphM). Qed. Canonical cfAut_additive := Additive cfAut_is_rmorphism. Canonical cfAut_rmorphism := RMorphism cfAut_is_rmorphism. @@ -409,7 +407,7 @@ Notation "''CF' ( G , A )" := (classfun_on G A) : ring_scope. Notation "1" := (@GRing.one (cfun_ringType _)) (only parsing) : cfun_scope. Notation "phi ^*" := (cfAut conjC phi) : cfun_scope. -Notation conjC_closed := (cfAut_closed conjC). +Notation cfConjC_closed := (cfAut_closed conjC). Prenex Implicits cfReal. (* Workaround for overeager projection reduction. *) Notation eqcfP := (@eqP (cfun_eqType _) _ _) (only parsing). @@ -482,7 +480,7 @@ Lemma cfunJ phi x y : y \in G -> phi (x ^ y) = phi x. Proof. by rewrite -{1}(genGid G) => /(cfunJgen phi)->. Qed. Lemma cfun_repr phi x : phi (repr (x ^: G)) = phi x. -Proof. by have [y Gy ->] := repr_class G x; exact: cfunJ. Qed. +Proof. by have [y Gy ->] := repr_class G x; apply: cfunJ. Qed. Lemma cfun_inP phi psi : {in G, phi =1 psi} -> phi = psi. Proof. by rewrite -{1}genGid => /cfun_in_genP. Qed. @@ -500,7 +498,7 @@ Lemma eq_mul_cfuni A phi : A <| G -> {in A, phi * '1_A =1 phi}. Proof. by move=> nsAG x Ax; rewrite cfunE cfuniE // Ax mulr1. Qed. Lemma eq_cfuni A : A <| G -> {in A, '1_A =1 (1 : 'CF(G))}. -Proof. by rewrite -['1_A]mul1r; exact: eq_mul_cfuni. Qed. +Proof. by rewrite -['1_A]mul1r; apply: eq_mul_cfuni. Qed. Lemma cfuniG : '1_G = 1. Proof. by rewrite -[G in '1_G]genGid. Qed. @@ -556,12 +554,12 @@ apply/cfun_inP=> x Gx; rewrite sum_cfunE (bigD1 (x ^: G)) ?mem_classes //=. rewrite cfunE cfun_repr cfun_classE Gx class_refl mulr1. rewrite big1 ?addr0 // => _ /andP[/imsetP[y Gy ->]]; apply: contraNeq. rewrite cfunE cfun_repr cfun_classE Gy mulf_eq0 => /norP[_]. -by rewrite pnatr_eq0 -lt0n lt0b => /class_transr->. +by rewrite pnatr_eq0 -lt0n lt0b => /class_eqP->. Qed. Implicit Arguments cfun_onP [A phi]. Lemma cfun_on0 A phi x : phi \in 'CF(G, A) -> x \notin A -> phi x = 0. -Proof. by move/cfun_onP; exact. Qed. +Proof. by move/cfun_onP; apply. Qed. Lemma sum_by_classes (R : ringType) (F : gT -> R) : {in G &, forall g h, F (g ^ h) = F g} -> @@ -571,7 +569,7 @@ move=> FJ; rewrite {1}(partition_big _ _ ((@mem_classes gT)^~ G)) /=. apply: eq_bigr => _ /imsetP[x Gx ->]; have [y Gy ->] := repr_class G x. rewrite mulr_natl -sumr_const FJ {y Gy}//; apply/esym/eq_big=> y /=. apply/idP/andP=> [xGy | [Gy /eqP<-]]; last exact: class_refl. - by rewrite (class_transr xGy) (subsetP (class_subG Gx (subxx _))). + by rewrite (class_eqP xGy) (subsetP (class_subG Gx (subxx _))). by case/imsetP=> z Gz ->; rewrite FJ. Qed. @@ -587,7 +585,7 @@ apply: contraNeq; rewrite b_i !cfunE mulf_eq0 => /norP[_]. rewrite -(inj_eq enum_val_inj). have /setIdP[/imsetP[x _ ->] _] := enum_valP i; rewrite cfun_repr. have /setIdP[/imsetP[y Gy ->] _] := enum_valP j; rewrite cfun_classE Gy. -by rewrite pnatr_eq0 -lt0n lt0b => /class_transr->. +by rewrite pnatr_eq0 -lt0n lt0b => /class_eqP->. Qed. Lemma dim_cfun : \dim 'CF(G) = #|classes G|. @@ -895,7 +893,7 @@ Proof. by rewrite -cfdot_conjC cfConjCK. Qed. Lemma cfnorm_ge0 phi : 0 <= '[phi]. Proof. -by rewrite mulr_ge0 ?invr_ge0 ?ler0n ?sumr_ge0 // => x _; exact: mul_conjC_ge0. +by rewrite mulr_ge0 ?invr_ge0 ?ler0n ?sumr_ge0 // => x _; apply: mul_conjC_ge0. Qed. Lemma cfnorm_eq0 phi : ('[phi] == 0) = (phi == 0). @@ -903,7 +901,7 @@ Proof. apply/idP/eqP=> [|->]; last by rewrite cfdot0r. rewrite mulf_eq0 invr_eq0 (negbTE (neq0CG G)) /= => /eqP/psumr_eq0P phi0. apply/cfun_inP=> x Gx; apply/eqP; rewrite cfunE -mul_conjC_eq0. -by rewrite phi0 // => y _; exact: mul_conjC_ge0. +by rewrite phi0 // => y _; apply: mul_conjC_ge0. Qed. Lemma cfnorm_gt0 phi : ('[phi] > 0) = (phi != 0). @@ -995,7 +993,7 @@ Lemma orthogonal_cons phi R S : Proof. by rewrite /orthogonal /= andbT. Qed. Lemma orthoP phi psi : reflect ('[phi, psi] = 0) (orthogonal phi psi). -Proof. by rewrite /orthogonal /= !andbT; exact: eqP. Qed. +Proof. by rewrite /orthogonal /= !andbT; apply: eqP. Qed. Lemma orthogonalP S R : reflect {in S & R, forall phi psi, '[phi, psi] = 0} (orthogonal S R). @@ -1028,7 +1026,7 @@ Lemma eq_orthogonal R1 R2 S1 S2 : R1 =i R2 -> S1 =i S2 -> orthogonal R1 S1 = orthogonal R2 S2. Proof. move=> eqR eqS; rewrite [orthogonal _ _](eq_all_r eqR). -by apply: eq_all => psi /=; exact: eq_all_r. +by apply: eq_all => psi /=; apply: eq_all_r. Qed. Lemma orthogonal_catl R1 R2 S : @@ -1083,7 +1081,7 @@ Qed. Lemma orthogonal_oppr S R : orthogonal S (map -%R R) = orthogonal S R. Proof. wlog suffices IH: S R / orthogonal S R -> orthogonal S (map -%R R). - apply/idP/idP=> /IH; rewrite ?mapK //; exact: opprK. + by apply/idP/idP=> /IH; rewrite ?mapK //; apply: opprK. move/orthogonalP=> oSR; apply/orthogonalP=> xi1 _ Sxi1 /mapP[xi2 Rxi2 ->]. by rewrite cfdotNr oSR ?oppr0. Qed. @@ -1104,7 +1102,7 @@ have [opS | not_opS] := allP; last first. by rewrite opS ?mem_head 1?mem_behead // (memPnC notSp). rewrite (contra (opS _)) /= ?cfnorm_eq0 //. apply: (iffP IH) => [] [uniqS oSS]; last first. - by split=> //; apply: sub_in2 oSS => psi Spsi; exact: mem_behead. + by split=> //; apply: sub_in2 oSS => psi Spsi; apply: mem_behead. split=> // psi xi; rewrite !inE => /predU1P[-> // | Spsi]. by case/predU1P=> [-> | /opS] /eqP. case/predU1P=> [-> _ | Sxi /oSS-> //]. @@ -1134,14 +1132,14 @@ Lemma sub_pairwise_orthogonal S1 S2 : Proof. move=> sS12 uniqS1 /pairwise_orthogonalP[/andP[notS2_0 _] oS2]. apply/pairwise_orthogonalP; rewrite /= (contra (sS12 0)) //. -by split=> //; exact: sub_in2 oS2. +by split=> //; apply: sub_in2 oS2. Qed. Lemma orthogonal_free S : pairwise_orthogonal S -> free S. Proof. case/pairwise_orthogonalP=> [/=/andP[notS0 uniqS] oSS]. rewrite -(in_tupleE S); apply/freeP => a aS0 i. -have S_i: S`_i \in S by exact: mem_nth. +have S_i: S`_i \in S by apply: mem_nth. have /eqP: '[S`_i, 0]_G = 0 := cfdot0r _. rewrite -{2}aS0 raddf_sum /= (bigD1 i) //= big1 => [|j neq_ji]; last 1 first. by rewrite cfdotZr oSS ?mulr0 ?mem_nth // eq_sym nth_uniq. @@ -1202,7 +1200,7 @@ Lemma sub_orthonormal S1 S2 : {subset S1 <= S2} -> uniq S1 -> orthonormal S2 -> orthonormal S1. Proof. move=> sS12 uniqS1 /orthonormalP[_ oS1]. -by apply/orthonormalP; split; last exact: sub_in2 sS12 _ _. +by apply/orthonormalP; split; last apply: sub_in2 sS12 _ _. Qed. Lemma orthonormal2P phi psi : @@ -1214,7 +1212,7 @@ by apply: (iffP and3P) => [] []; do 3!move/eqP->. Qed. Lemma conjC_pair_orthogonal S chi : - conjC_closed S -> ~~ has cfReal S -> pairwise_orthogonal S -> chi \in S -> + cfConjC_closed S -> ~~ has cfReal S -> pairwise_orthogonal S -> chi \in S -> pairwise_orthogonal (chi :: chi^*%CF). Proof. move=> ccS /hasPn nrS oSS Schi; apply: sub_pairwise_orthogonal oSS. @@ -1225,6 +1223,16 @@ Qed. Lemma cfdot_real_conjC phi psi : cfReal phi -> '[phi, psi^*]_G = '[phi, psi]^*. Proof. by rewrite -cfdot_conjC => /eqcfP->. Qed. +Lemma extend_cfConjC_subset S X phi : + cfConjC_closed S -> ~~ has cfReal S -> phi \in S -> phi \notin X -> + cfConjC_subset X S -> cfConjC_subset [:: phi, phi^* & X]%CF S. +Proof. +move=> ccS nrS Sphi X'phi [uniqX /allP-sXS ccX]. +split; last 1 [by apply/allP; rewrite /= Sphi ccS | apply/allP]; rewrite /= inE. + by rewrite negb_or X'phi eq_sym (hasPn nrS) // (contra (ccX _)) ?cfConjCK. +by rewrite cfConjCK !mem_head orbT; apply/allP=> xi Xxi; rewrite !inE ccX ?orbT. +Qed. + (* Note: other isometry lemmas, and the dot product lemmas for orthogonal *) (* and orthonormal sequences are in vcharacter, because we need the 'Z[S] *) (* notation for the isometry domains. Alternatively, this could be moved to *) @@ -1291,7 +1299,22 @@ Lemma sub_iso_to U1 U2 W1 W2 tau : {subset U2 <= U1} -> {subset W1 <= W2} -> {in U1, isometry tau, to W1} -> {in U2, isometry tau, to W2}. Proof. -by move=> sU sW [Itau Wtau]; split=> [|u /sU/Wtau/sW //]; exact: sub_in2 Itau. +by move=> sU sW [Itau Wtau]; split=> [|u /sU/Wtau/sW //]; apply: sub_in2 Itau. +Qed. + +Lemma isometry_of_free S f : + free S -> {in S &, isometry f} -> + {tau : {linear 'CF(L) -> 'CF(G)} | + {in S, tau =1 f} & {in <<S>>%VS &, isometry tau}}. +Proof. +move=> freeS If; have defS := free_span freeS. +have [tau /(_ freeS (size_map f S))Dtau] := linear_of_free S (map f S). +have{Dtau} Dtau: {in S, tau =1 f}. + by move=> _ /(nthP 0)[i ltiS <-]; rewrite -!(nth_map 0 0) ?Dtau. +exists tau => // _ _ /defS[a -> _] /defS[b -> _]. +rewrite !{1}linear_sum !{1}cfdot_suml; apply/eq_big_seq=> xi1 Sxi1. +rewrite !{1}cfdot_sumr; apply/eq_big_seq=> xi2 Sxi2. +by rewrite !linearZ /= !Dtau // !cfdotZl !cfdotZr If. Qed. Lemma isometry_of_cfnorm S tauS : @@ -1302,16 +1325,15 @@ Lemma isometry_of_cfnorm S tauS : Proof. move=> oS oT eq_nST; have freeS := orthogonal_free oS. have eq_sz: size tauS = size S by have:= congr1 size eq_nST; rewrite !size_map. -have [tau /(_ freeS eq_sz) defT] := linear_of_free S tauS. -rewrite -[S]/(tval (in_tuple S)). -exists tau => // u v /coord_span-> /coord_span->; rewrite !raddf_sum /=. +have [tau defT] := linear_of_free S tauS; rewrite -[S]/(tval (in_tuple S)). +exists tau => [|u v /coord_span-> /coord_span->]; rewrite ?raddf_sum ?defT //=. apply: eq_bigr => i _ /=; rewrite linearZ !cfdotZr !cfdot_suml; congr (_ * _). apply: eq_bigr => j _ /=; rewrite linearZ !cfdotZl; congr (_ * _). -rewrite -!((nth_map _ 0) tau) // defT; have [-> | neq_ji] := eqVneq j i. - by rewrite -!['[_]]((nth_map _ 0) cfnorm) ?eq_sz // eq_nST. +rewrite -!(nth_map 0 0 tau) ?{}defT //; have [-> | neq_ji] := eqVneq j i. + by rewrite -!['[_]](nth_map 0 0 cfnorm) ?eq_sz // eq_nST. have{oS} [/=/andP[_ uS] oS] := pairwise_orthogonalP oS. have{oT} [/=/andP[_ uT] oT] := pairwise_orthogonalP oT. -by rewrite oS ?oT ?mem_nth ? nth_uniq ?eq_sz. +by rewrite oS ?oT ?mem_nth ?nth_uniq ?eq_sz. Qed. Lemma isometry_raddf_inj U (tau : {additive 'CF(L) -> 'CF(G)}) : @@ -1628,7 +1650,7 @@ Local Notation "phi %% 'B'" := (cfMod phi) (at level 40) : cfun_scope. (* stronger results are needed. *) Lemma cfModE phi x : B <| G -> x \in G -> (phi %% B)%CF x = phi (coset B x). -Proof. by move/normal_norm=> nBG; exact: cfMorphE. Qed. +Proof. by move/normal_norm=> nBG; apply: cfMorphE. Qed. Lemma cfMod1 phi : (phi %% B)%CF 1%g = phi 1%g. Proof. exact: cfMorph1. Qed. @@ -2115,7 +2137,7 @@ apply: eq_bigr => _ /morphimP[x Dx Gx ->]. rewrite -(card_rcoset _ x) mulr_natl -sumr_const. apply/eq_big => [y | y /andP[Gy /eqP <-]]; last by rewrite !cfMorphE. rewrite mem_rcoset inE groupMr ?groupV // -mem_rcoset. -by apply: andb_id2l => /(subsetP sGD) Dy; exact: sameP eqP (rcoset_kerP f _ _). +by apply: andb_id2l => /(subsetP sGD) Dy; apply: sameP eqP (rcoset_kerP f _ _). Qed. Lemma cfIsom_iso rT G (R : {group rT}) (f : {morphism G >-> rT}) : @@ -2374,11 +2396,14 @@ Proof. exact: raddfZ_Cnat. Qed. Lemma cfAutZ_Cint z phi : z \in Cint -> (z *: phi)^u = z *: phi^u. Proof. exact: raddfZ_Cint. Qed. +Lemma cfAutK : cancel (@cfAut gT G u) (cfAut (algC_invaut_rmorphism u)). +Proof. by move=> phi; apply/cfunP=> x; rewrite !cfunE /= algC_autK. Qed. + +Lemma cfAutVK : cancel (cfAut (algC_invaut_rmorphism u)) (@cfAut gT G u). +Proof. by move=> phi; apply/cfunP=> x; rewrite !cfunE /= algC_invautK. Qed. + Lemma cfAut_inj : injective (@cfAut gT G u). -Proof. -move=> phi psi /cfunP eqfg; apply/cfunP=> x. -by have := eqfg x; rewrite !cfunE => /fmorph_inj. -Qed. +Proof. exact: can_inj cfAutK. Qed. Lemma cfAut_eq1 phi : (cfAut u phi == 1) = (phi == 1). Proof. by rewrite rmorph_eq1 //; apply: cfAut_inj. Qed. @@ -2460,6 +2485,8 @@ Proof. by rewrite rmorphM /= cfAutDprodl cfAutDprodr. Qed. End FieldAutomorphism. +Implicit Arguments cfAutK [[gT] [G]]. +Implicit Arguments cfAutVK [[gT] [G]]. Implicit Arguments cfAut_inj [gT G x1 x2]. Definition conj_cfRes := cfAutRes conjC. diff --git a/mathcomp/character/finfield.v b/mathcomp/character/finfield.v index 05b4a31..467449a 100644 --- a/mathcomp/character/finfield.v +++ b/mathcomp/character/finfield.v @@ -1,23 +1,14 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype tuple bigop prime finset. -From mathcomp.fingroup -Require Import fingroup morphism action. -From mathcomp.algebra -Require Import ssralg poly polydiv finalg zmodp cyclic. -From mathcomp.algebra -Require Import matrix vector. -From mathcomp.solvable -Require Import center pgroup abelian. -From mathcomp.field -Require Import falgebra fieldext separable galois. -Require Import mxabelem. - -From mathcomp.algebra Require ssrnum ssrint. -From mathcomp.field Require algC cyclotomic. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype div. +From mathcomp +Require Import tuple bigop prime finset fingroup ssralg poly polydiv. +From mathcomp +Require Import morphism action finalg zmodp cyclic center pgroup abelian. +From mathcomp +Require Import matrix mxabelem vector falgebra fieldext separable galois. +Require ssrnum ssrint algC cyclotomic. (******************************************************************************) (* Additional constructions and results on finite fields. *) @@ -498,7 +489,7 @@ have [p p_pr charRp]: exists2 p, prime p & p \in [char R]. rewrite big_seq; elim/big_rec: _ => [|[p m] /= n]; first by rewrite oner_eq0. case/mem_prime_decomp=> p_pr _ _ IHn. elim: m => [|m IHm]; rewrite ?mul1n {IHn}// expnS -mulnA natrM. - by case/eqP/domR/orP=> //; exists p; last exact/andP. + by case/eqP/domR/orP=> //; exists p; last apply/andP. pose Rp := PrimeCharType charRp; pose L : {vspace Rp} := fullv. pose G := [set: {unit R}]; pose ofG : {unit R} -> Rp := val. pose projG (E : {vspace Rp}) := [preim ofG of E]. diff --git a/mathcomp/character/inertia.v b/mathcomp/character/inertia.v index 9a795e8..c6ddf44 100644 --- a/mathcomp/character/inertia.v +++ b/mathcomp/character/inertia.v @@ -1,20 +1,17 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype tuple finfun bigop prime finset. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism quotient action gproduct. -From mathcomp.algebra -Require Import ssralg ssrnum zmodp cyclic matrix mxalgebra vector. -From mathcomp.solvable -Require Import center commutator gseries nilpotent pgroup sylow maximal. -From mathcomp.solvable +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq path choice div. +From mathcomp +Require Import fintype tuple finfun bigop prime ssralg ssrnum finset fingroup. +From mathcomp +Require Import morphism perm automorphism quotient action zmodp cyclic center. +From mathcomp +Require Import gproduct commutator gseries nilpotent pgroup sylow maximal. +From mathcomp Require Import frobenius. -From mathcomp.field -Require Import algC. -Require Import mxrepresentation classfun character. +From mathcomp +Require Import matrix mxalgebra mxrepresentation vector algC classfun character. Set Implicit Arguments. Unset Strict Implicit. @@ -105,7 +102,7 @@ Qed. (* Isaacs' 6.1.b *) Lemma cfConjgM L phi : G <| L -> {in L &, forall y z, phi ^ (y * z) = (phi ^ y) ^ z}%CF. -Proof. by case/andP=> _ /subsetP nGL; exact: sub_in2 (cfConjgMnorm phi). Qed. +Proof. by case/andP=> _ /subsetP nGL; apply: sub_in2 (cfConjgMnorm phi). Qed. Lemma cfConjgJ1 phi : (phi ^ 1)%CF = phi. Proof. by apply/cfunP=> x; rewrite cfConjgE ?group1 // invg1 conjg1. Qed. @@ -403,7 +400,7 @@ apply: (iffP imageP) => [[_ /rcosetsP[y Ay ->] ->] | [y Ay ->]]. by case: repr_rcosetP => z /setIdP[Az _]; exists (z * y)%g; rewrite ?groupM. without loss nHy: y Ay / y \in 'N(H). have [nHy | /cfConjgEout->] := boolP (y \in 'N(H)); first exact. - by move/(_ 1%g); rewrite !group1 !cfConjgJ1; exact. + by move/(_ 1%g); rewrite !group1 !cfConjgJ1; apply. exists ('I_A[phi] :* y); first by rewrite -rcosetE mem_imset. case: repr_rcosetP => z /setIP[_ /setIdP[nHz /eqP Tz]]. by rewrite cfConjgMnorm ?Tz. @@ -414,25 +411,25 @@ Proof. move=> xi; apply/cfclassP/cfclassP=> [[x /setIP[Gx _] ->] | [x Gx ->]]. by exists x. have [Nx | /cfConjgEout-> //] := boolP (x \in 'N(H)). - by exists x; first exact/setIP. + by exists x; first apply/setIP. by exists 1%g; rewrite ?group1 ?cfConjgJ1. Qed. Lemma cfclass_refl phi : phi \in (phi ^: G)%CF. Proof. by apply/cfclassP; exists 1%g => //; rewrite cfConjgJ1. Qed. -Lemma cfclass_transl phi psi : +Lemma cfclass_transr phi psi : (psi \in phi ^: G)%CF -> (phi ^: G =i psi ^: G)%CF. Proof. rewrite -cfclassInorm; case/cfclassP=> x Gx -> xi; rewrite -!cfclassInorm. -have nHN: {subset 'N_G(H) <= 'N(H)} by apply/subsetP; exact: subsetIr. +have nHN: {subset 'N_G(H) <= 'N(H)} by apply/subsetP; apply: subsetIr. apply/cfclassP/cfclassP=> [[y Gy ->] | [y Gy ->]]. by exists (x^-1 * y)%g; rewrite -?cfConjgMnorm ?groupM ?groupV ?nHN // mulKVg. by exists (x * y)%g; rewrite -?cfConjgMnorm ?groupM ?nHN. Qed. Lemma cfclass_sym phi psi : (psi \in phi ^: G)%CF = (phi \in psi ^: G)%CF. -Proof. by apply/idP/idP=> /cfclass_transl <-; exact: cfclass_refl. Qed. +Proof. by apply/idP/idP=> /cfclass_transr <-; apply: cfclass_refl. Qed. Lemma cfclass_uniq phi : H <| G -> uniq (phi ^: G)%CF. Proof. @@ -440,7 +437,7 @@ move=> nsHG; rewrite map_inj_in_uniq ?enum_uniq // => Ty Tz; rewrite !mem_enum. move=> {Ty}/rcosetsP[y Gy ->] {Tz}/rcosetsP[z Gz ->] /eqP. case: repr_rcosetP => u Iphi_u; case: repr_rcosetP => v Iphi_v. have [[Gu _] [Gv _]] := (setIdP Iphi_u, setIdP Iphi_v). -rewrite cfConjg_eqE ?groupM // => /rcoset_transl. +rewrite cfConjg_eqE ?groupM // => /rcoset_eqP. by rewrite !rcosetM (rcoset_id Iphi_v) (rcoset_id Iphi_u). Qed. @@ -466,7 +463,7 @@ Lemma eq_cfclass_IirrE i j : (cfclass_Iirr G j == cfclass_Iirr G i) = (j \in cfclass_Iirr G i). Proof. apply/eqP/idP=> [<- | iGj]; first by rewrite cfclass_IirrE cfclass_refl. -by apply/setP=> k; rewrite !cfclass_IirrE in iGj *; apply/esym/cfclass_transl. +by apply/setP=> k; rewrite !cfclass_IirrE in iGj *; apply/esym/cfclass_transr. Qed. Lemma im_cfclass_Iirr i : @@ -1581,7 +1578,7 @@ apply: contraR => notKx; apply/cards1P; exists 1%g; apply/esym/eqP. rewrite eqEsubset !(sub1set, inE) classes1 /= conjs1g eqxx /=. apply/subsetP=> _ /setIP[/imsetP[y Ky ->] /afix1P /= cyKx]. have /imsetP[z Kz def_yx]: y ^ x \in y ^: K. - by rewrite -cyKx; apply: mem_imset; exact: class_refl. + by rewrite -cyKx; apply: mem_imset; apply: class_refl. rewrite inE classG_eq1; apply: contraR notKx => nty. rewrite -(groupMr x (groupVr Kz)). apply: (subsetP (regK y _)); first exact/setD1P. @@ -1606,7 +1603,7 @@ Proof. have [_ _ nsKG _] := Frobenius_kerP frobGK; have [sKG nKG] := andP nsKG. apply: (iffP idP) => [not_chijK1 | [i nzi ->]]; last first. by rewrite cfker_Ind_irr ?sub_gcore // subGcfker. -have /neq0_has_constt[i chijKi]: 'Res[K] 'chi_j != 0 by exact: Res_irr_neq0. +have /neq0_has_constt[i chijKi]: 'Res[K] 'chi_j != 0 by apply: Res_irr_neq0. have nz_i: i != 0. by apply: contraNneq not_chijK1 => i0; rewrite constt0_Res_cfker // -i0. have /irrP[k def_chik] := irr_induced_Frobenius_ker nz_i. diff --git a/mathcomp/character/integral_char.v b/mathcomp/character/integral_char.v index 14ae4fc..1770335 100644 --- a/mathcomp/character/integral_char.v +++ b/mathcomp/character/integral_char.v @@ -1,19 +1,18 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype tuple finfun bigop prime finset. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism quotient action gproduct. -From mathcomp.algebra -Require Import ssralg poly polydiv finalg zmodp cyclic matrix mxalgebra mxpoly. -From mathcomp.solvable -Require Import commutator center pgroup sylow gseries nilpotent abelian. -From mathcomp.algebra -Require Import ssrnum ssrint polydiv rat matrix mxalgebra intdiv mxpoly vector. -From mathcomp.field -Require Import fieldext separable galois algC cyclotomic algnum falgebra. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp +Require Import fintype tuple finfun bigop prime ssralg poly finset. +From mathcomp +Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +From mathcomp +Require Import commutator cyclic center pgroup sylow gseries nilpotent abelian. +From mathcomp +Require Import ssrnum ssrint polydiv rat matrix mxalgebra intdiv mxpoly. +From mathcomp +Require Import vector falgebra fieldext separable galois algC cyclotomic algnum. +From mathcomp Require Import mxrepresentation classfun character. (******************************************************************************) @@ -31,7 +30,7 @@ Require Import mxrepresentation classfun character. (* gring_classM_coef_set A B z == the set of all (x, y) in setX A B such *) (* that x * y = z; if A and B are respectively the ith and jth *) (* conjugacy class of G, and z is in the kth conjugacy class, then *) -(* gring_classM_coef i j k is exactly the cadinal of this set. *) +(* gring_classM_coef i j k is exactly the cardinal of this set. *) (* 'omega_i[A] == the mode of 'chi[G]_i on (A \in 'Z(group_ring algC G))%MS, *) (* i.e., the z such that gring_op 'Chi_i A = z%:M. *) (******************************************************************************) @@ -73,7 +72,7 @@ exists (SplittingFieldType _ _ Qn_ax). apply: separable_Xn_sub_1; rewrite -(fmorph_eq0 QnC) rmorph_nat. by rewrite pnatr_eq0 -lt0n cardG_gt0. exists QnC => [// nuQn|]. - by exact: (extend_algC_subfield_aut QnC [rmorphism of nuQn]). + by apply: (extend_algC_subfield_aut QnC [rmorphism of nuQn]). rewrite span_seq1 in genQn. exists w => // hT H phi Nphi x x_dv_n. apply: sig_eqW; have [rH ->] := char_reprP Nphi. @@ -146,7 +145,7 @@ rewrite scaler_sumr; apply: eq_bigr => g Kk_g; rewrite scaler_nat. rewrite (set_gring_classM_coef _ _ Kk_g) -sumr_const; apply: eq_big => [] [x y]. rewrite !inE /= dKi dKj /h1 /h2 /=; apply: andb_id2r => /eqP ->. have /imsetP[zk Gzk dKk] := enum_valP k; rewrite dKk in Kk_g. - by rewrite (class_transr Kk_g) -dKk enum_valK_in eqxx andbT. + by rewrite (class_eqP Kk_g) -dKk enum_valK_in eqxx andbT. by rewrite /h2 /= => /andP[_ /eqP->]. Qed. @@ -179,7 +178,7 @@ exact/unity_rootP. Qed. Lemma Aint_irr i x : 'chi[G]_i x \in Aint. -Proof. by apply: Aint_char; exact: irr_char. Qed. +Proof. exact/Aint_char/irr_char. Qed. Local Notation R_G := (group_ring algCfield G). Local Notation a := gring_classM_coef. @@ -240,9 +239,9 @@ Qed. Lemma Aint_gring_mode_class_sum k : 'omega_i['K_k] \in Aint. Proof. move: k; pose X := [tuple 'omega_i['K_k] | k < #|classes G| ]. -have memX k: 'omega_i['K_k] \in X by apply: map_f; exact: mem_enum. +have memX k: 'omega_i['K_k] \in X by apply: image_f. have S_P := Cint_spanP X; set S := Cint_span X in S_P. -have S_X: {subset X <= S} by exact: mem_Cint_span. +have S_X: {subset X <= S} by apply: mem_Cint_span. have S_1: 1 \in S. apply: S_X; apply/codomP; exists (enum_rank_in (classes1 G) 1%g). rewrite (@gring_mode_class_sum_eq _ 1%g) ?enum_rankK_in ?classes1 //. @@ -383,7 +382,7 @@ rewrite ltnS => leGn piGle2; have [simpleG | ] := boolP (simple G); last first. exact: solvable1. rewrite [N == G]eqEproper sNG eqbF_neg !negbK => ltNG /and3P[grN]. case/isgroupP: grN => {N}N -> in sNG ltNG *; rewrite /= genGid => ntN nNG. - have nsNG: N <| G by exact/andP. + have nsNG: N <| G by apply/andP. have dv_le_pi m: (m %| #|G| -> size (primes m) <= 2)%N. move=> m_dv_G; apply: leq_trans piGle2. by rewrite uniq_leq_size ?primes_uniq //; apply: pi_of_dvd. @@ -498,7 +497,7 @@ move=> a /= Kg1 Kg2 Kg; rewrite mulrAC; apply: canRL (mulfK (neq0CG G)) _. transitivity (\sum_j (#|G| * a j)%:R *+ (j == k) : algC). by rewrite (bigD1 k) //= eqxx -natrM mulnC big1 ?addr0 // => j /negPf->. have defK (j : 'I_#|classes G|) x: x \in enum_val j -> enum_val j = x ^: G. - by have /imsetP[y Gy ->] := enum_valP j => /class_transr. + by have /imsetP[y Gy ->] := enum_valP j => /class_eqP. have Gg: g \in G. by case/imsetP: (enum_valP k) Kg => x Gx -> /imsetP[y Gy ->]; apply: groupJ. transitivity (\sum_j \sum_i 'omega_i['K_j] * 'chi_i 1%g * ('chi_i g)^* *+ a j). @@ -680,7 +679,7 @@ have{Qpi1} Zpi1: pi1 \in Cint. have{pi1 Zpi1} pi2_ge1: 1 <= pi2. have ->: pi2 = `|pi1| ^+ 2. by rewrite (big_morph Num.norm (@normrM _) (@normr1 _)) -prodrXl. - by rewrite Cint_normK // sqr_Cint_ge1 //; exact/prodf_neq0. + by rewrite Cint_normK // sqr_Cint_ge1 //; apply/prodf_neq0. have Sgt0: (#|S| > 0)%N by rewrite (cardD1 g) [g \in S]Sg. rewrite -mulr_natr -ler_pdivl_mulr ?ltr0n //. have n2chi_ge0 s: s \in S -> 0 <= `|chi s| ^+ 2 by rewrite exprn_ge0 ?normr_ge0. @@ -693,7 +692,7 @@ Theorem nonlinear_irr_vanish gT (G : {group gT}) i : 'chi[G]_i 1%g > 1 -> exists2 x, x \in G & 'chi_i x = 0. Proof. move=> chi1gt1; apply/exists_eq_inP; apply: contraFT (ltr_geF chi1gt1). -rewrite negb_exists_in => /forall_inP nz_chi. +rewrite negb_exists_in => /forall_inP-nz_chi. rewrite -(norm_Cnat (Cnat_irr1 i)) -(@expr_le1 _ 2) ?normr_ge0 //. rewrite -(ler_add2r (#|G|%:R * '['chi_i])) {1}cfnorm_irr mulr1. rewrite (cfnormE (cfun_onG _)) mulVKf ?neq0CG // (big_setD1 1%g) //=. diff --git a/mathcomp/character/mxabelem.v b/mathcomp/character/mxabelem.v index 377cb72..8685796 100644 --- a/mathcomp/character/mxabelem.v +++ b/mathcomp/character/mxabelem.v @@ -1,18 +1,15 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype tuple finfun bigop prime finset. -From mathcomp.fingroup +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp +Require Import fintype tuple finfun bigop prime ssralg poly finset. +From mathcomp Require Import fingroup morphism perm automorphism quotient gproduct action. -From mathcomp.algebra -Require Import ssralg poly finalg zmodp cyclic matrix mxalgebra. -From mathcomp.solvable -Require Import commutator center pgroup gseries nilpotent. -From mathcomp.solvable -Require Import sylow maximal abelian. -Require Import mxrepresentation. +From mathcomp +Require Import finalg zmodp commutator cyclic center pgroup gseries nilpotent. +From mathcomp +Require Import sylow maximal abelian matrix mxalgebra mxrepresentation. (******************************************************************************) (* This file completes the theory developed in mxrepresentation.v with the *) @@ -145,7 +142,7 @@ Proof. by rewrite inE. Qed. Fact rowg_group_set m A : group_set (@rowg m A). Proof. -by apply/group_setP; split=> [|u v]; rewrite !inE ?sub0mx //; exact: addmx_sub. +by apply/group_setP; split=> [|u v]; rewrite !inE ?sub0mx //; apply: addmx_sub. Qed. Canonical rowg_group m A := Group (@rowg_group_set m A). @@ -157,7 +154,7 @@ Lemma rowgS m1 m2 (A : 'M_(m1, n)) (B : 'M_(m2, n)) : Proof. apply/subsetP/idP=> sAB => [| u]. by apply/row_subP=> i; have:= sAB (row i A); rewrite !inE row_sub => ->. -by rewrite !inE => suA; exact: submx_trans sAB. +by rewrite !inE => suA; apply: submx_trans sAB. Qed. Lemma eq_rowg m1 m2 (A : 'M_(m1, n)) (B : 'M_(m2, n)) : @@ -214,8 +211,8 @@ Proof. by apply/eqP; rewrite -submx0 -(rowg0 0) rowgK sub0mx. Qed. Lemma rowg_mx_eq0 (L : {group rVn}) : (rowg_mx L == 0) = (L :==: 1%g). Proof. -rewrite -trivg_rowg; apply/idP/idP=> [|/eqP->]; last by rewrite rowg_mx1 rowg0. -by rewrite !(sameP eqP trivgP); apply: subset_trans; exact: sub_rowg_mx. +rewrite -trivg_rowg; apply/idP/eqP=> [|->]; last by rewrite rowg_mx1 rowg0. +exact/contraTeq/subG1_contra/sub_rowg_mx. Qed. Lemma rowgI m1 m2 (A : 'M_(m1, n)) (B : 'M_(m2, n)) : @@ -312,7 +309,7 @@ Lemma astabs_rowg_repr m (A : 'M_(m, n)) : 'N(rowg A | 'MR rG) = rstabs rG A. Proof. apply/setP=> x; rewrite !inE /=; apply: andb_id2l => Gx. apply/subsetP/idP=> nAx => [|u]; last first. - rewrite !inE mx_repr_actE // => Au; exact: (submx_trans (submxMr _ Au)). + by rewrite !inE mx_repr_actE // => Au; apply: (submx_trans (submxMr _ Au)). apply/row_subP=> i; move/implyP: (nAx (row i A)). by rewrite !inE row_sub mx_repr_actE //= row_mul. Qed. @@ -449,13 +446,13 @@ Variables p n : nat. Local Notation rVn := 'rV['F_p]_n. Lemma rowg_mxK (L : {group rVn}) : rowg (rowg_mx L) = L. -Proof. by apply: stable_rowg_mxK; exact: mx_Fp_stable. Qed. +Proof. by apply: stable_rowg_mxK; apply: mx_Fp_stable. Qed. Lemma rowg_mxSK (L : {set rVn}) (M : {group rVn}) : (rowg_mx L <= rowg_mx M)%MS = (L \subset M). Proof. apply/idP/idP; last exact: rowg_mxS. -by rewrite -rowgS rowg_mxK; apply: subset_trans; exact: sub_rowg_mx. +by rewrite -rowgS rowg_mxK; apply/subset_trans/sub_rowg_mx. Qed. Lemma mxrank_rowg (L : {group rVn}) : @@ -508,7 +505,7 @@ Proof. by case/misomP: (xchooseP ab_rV_P). Qed. Lemma abelem_rV_injm : 'injm ErV. Proof. by case/isomP: abelem_rV_isom. Qed. Lemma abelem_rV_inj : {in E &, injective ErV}. -Proof. by apply/injmP; exact: abelem_rV_injm. Qed. +Proof. by apply/injmP; apply: abelem_rV_injm. Qed. Lemma im_abelem_rV : ErV @* E = setT. Proof. by case/isomP: abelem_rV_isom. Qed. @@ -701,7 +698,7 @@ Qed. Lemma abelem_mx_irrP : reflect (mx_irreducible rG) (minnormal E G). Proof. -by rewrite -[E in minnormal E G]im_rVabelem -rowg1; exact: mxsimple_abelemP. +by rewrite -[E in minnormal E G]im_rVabelem -rowg1; apply: mxsimple_abelemP. Qed. Lemma rfix_abelem (H : {set gT}) : @@ -850,7 +847,7 @@ Let p_gt1 := prime_gt1 p_pr. Let oZp := card_center_extraspecial pS esS. Let modIp' (i : 'I_p.-1) : (i.+1 %% p = i.+1)%N. -Proof. by case: i => i; rewrite /= -ltnS prednK //; exact: modn_small. Qed. +Proof. by case: i => i; rewrite /= -ltnS prednK //; apply: modn_small. Qed. (* This is Aschbacher (34.9), parts (1)-(4). *) Theorem extraspecial_repr_structure (sS : irrType F S) : @@ -873,7 +870,7 @@ have nb_irr: #|sS| = (p ^ n.*2 + p.-1)%N. pose Zcl := classes S ::&: 'Z(S). have cardZcl: #|Zcl| = p. transitivity #|[set [set z] | z in 'Z(S)]|; last first. - by rewrite card_imset //; exact: set1_inj. + by rewrite card_imset //; apply: set1_inj. apply: eq_card => zS; apply/setIdP/imsetP=> [[] | [z]]. case/imsetP=> z Sz ->{zS} szSZ. have Zz: z \in 'Z(S) by rewrite (subsetP szSZ) ?class_refl. @@ -900,7 +897,7 @@ have nb_irr: #|sS| = (p ^ n.*2 + p.-1)%N. apply: contra notZxS => Zxy; rewrite -{1}(lcoset_id Sy) class_lcoset. rewrite ((_ ^: _ =P [set x ^ y])%g _) ?sub1set // eq_sym eqEcard. rewrite sub1set class_refl cards1 -index_cent1 (setIidPl _) ?indexgg //. - by rewrite sub_cent1; apply: subsetP Zxy; exact: subsetIr. + by rewrite sub_cent1; apply: subsetP Zxy; apply: subsetIr. rewrite sum_nat_dep_const mulnC eqn_pmul2l //; move/eqP <-. rewrite addSnnS prednK // -cardZcl -[card _](cardsID Zcl) /= addnC. by congr (_ + _)%N; apply: eq_card => t; rewrite !inE andbC // andbAC andbb. @@ -954,11 +951,10 @@ have alpha_i_z i: ((alpha ^+ ephi i) z = z ^+ i.+1)%g. by rewrite autE ?cycle_id //= val_Zp_nat ozp ?modIp'. have rphiP i: S :==: autm (groupX (ephi i) Aut_alpha) @* S by rewrite im_autm. pose rphi i := morphim_repr (eqg_repr phi (rphiP i)) (subxx S). -have rphi_irr i: mx_irreducible (rphi i). - by apply/morphim_mx_irr; exact/eqg_mx_irr. +have rphi_irr i: mx_irreducible (rphi i) by apply/morphim_mx_irr/eqg_mx_irr. have rphi_fful i: mx_faithful (rphi i). rewrite /mx_faithful rker_morphim rker_eqg. - by rewrite (trivgP (fful_nlin _ nlin_i0)) morphpreIdom; exact: injm_autm. + by rewrite (trivgP (fful_nlin _ nlin_i0)) morphpreIdom; apply: injm_autm. have rphi_z i: rphi i z = (w ^+ i.+1)%:M. by rewrite /rphi [phi]lock /= /morphim_mx autmE alpha_i_z -lock phi_ze. pose iphi i := irr_comp sS (rphi i); pose phi_ i := irr_repr (iphi i). @@ -1021,7 +1017,7 @@ suffices IH V: mxsimple rS V -> mx_iso rZ U V -> - split=> [|/= V simV isoUV]. by case/andP: (IH U simU (mx_iso_refl _ _)) => /eqP. by case/andP: (IH V simV isoUV) => _ /(mxsimple_isoP simU). -move=> simV isoUV; wlog sS: / irrType F S by exact: socle_exists. +move=> simV isoUV; wlog sS: / irrType F S by apply: socle_exists. have [[_ defS'] prZ] := esS. have{prZ} ntZ: 'Z(S) :!=: 1%g by case: eqP prZ => // ->; rewrite cards1. have [_ [iphi]] := extraspecial_repr_structure sS. @@ -1030,13 +1026,13 @@ have [modU nzU _]:= simU; pose rU := submod_repr modU. have nlinU: \rank U != 1%N. apply/eqP=> /(rker_linear rU); apply/negP; rewrite /rker rstab_submod. by rewrite (eqmx_rstab _ (val_submod1 _)) (eqP ffulU) defS' subG1. -have irrU: mx_irreducible rU by exact/submod_mx_irr. +have irrU: mx_irreducible rU by apply/submod_mx_irr. have rsimU := rsim_irr_comp sS F'S irrU. set iU := irr_comp sS rU in rsimU; have [_ degU _ _]:= rsimU. have phiUP: iU \in codom iphi by rewrite im_phi !inE -degU. rewrite degU -(f_iinv phiUP) dim_phi eqxx /=; apply/(mxsimple_isoP simU). have [modV _ _]:= simV; pose rV := submod_repr modV. -have irrV: mx_irreducible rV by exact/submod_mx_irr. +have irrV: mx_irreducible rV by apply/submod_mx_irr. have rsimV := rsim_irr_comp sS F'S irrV. set iV := irr_comp sS rV in rsimV; have [_ degV _ _]:= rsimV. have phiVP: iV \in codom iphi by rewrite im_phi !inE -degV -(mxrank_iso isoUV). @@ -1045,7 +1041,7 @@ have [z Zz ntz]:= trivgPn _ ntZ. have [|w prim_w phi_z] := phiZ z; first by rewrite 2!inE ntz. suffices eqjUV: jU == jV. apply/(mx_rsim_iso modU modV); apply: mx_rsim_trans rsimU _. - by rewrite -(f_iinv phiUP) -/jU (eqP eqjUV) f_iinv; exact: mx_rsim_sym. + by rewrite -(f_iinv phiUP) -/jU (eqP eqjUV) f_iinv; apply: mx_rsim_sym. have rsimUV: mx_rsim (subg_repr (phi jU) sZS) (subg_repr (phi jV) sZS). have [bU _ bUfree bUhom] := mx_rsim_sym rsimU. have [bV _ bVfree bVhom] := rsimV. diff --git a/mathcomp/character/mxrepresentation.v b/mathcomp/character/mxrepresentation.v index e947fe0..7ceae6e 100644 --- a/mathcomp/character/mxrepresentation.v +++ b/mathcomp/character/mxrepresentation.v @@ -1,15 +1,13 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype tuple finfun bigop prime finset. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism quotient action gproduct. -From mathcomp.algebra -Require Import ssralg poly polydiv finalg zmodp cyclic matrix mxalgebra mxpoly. -From mathcomp.solvable -Require Import commutator center pgroup. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp +Require Import fintype tuple finfun bigop prime ssralg poly polydiv finset. +From mathcomp +Require Import fingroup morphism perm automorphism quotient finalg action zmodp. +From mathcomp +Require Import commutator cyclic center pgroup matrix mxalgebra mxpoly. (******************************************************************************) (* This file provides linkage between classic Group Theory and commutative *) @@ -308,7 +306,7 @@ Qed. Lemma repr_mxKV m x : x \in G -> cancel ((@mulmx _ m n n)^~ (rG x^-1)) (mulmx^~ (rG x)). -Proof. by rewrite -groupV -{3}[x]invgK; exact: repr_mxK. Qed. +Proof. by rewrite -groupV -{3}[x]invgK; apply: repr_mxK. Qed. Lemma repr_mx_unit x : x \in G -> rG x \in unitmx. Proof. by move=> Gx; case/mulmx1_unit: (repr_mxKV Gx 1%:M). Qed. @@ -923,7 +921,7 @@ Proof. by case/setIdP=> _ /eqP cUx /submxP[w ->]; rewrite -mulmxA cUx. Qed. Lemma rstabs_act x m1 (W : 'M_(m1, n)) : x \in rstabs -> (W <= U)%MS -> (W *m rG x <= U)%MS. Proof. -by case/setIdP=> [_ nUx] sWU; apply: submx_trans nUx; exact: submxMr. +by case/setIdP=> [_ nUx] sWU; apply: submx_trans nUx; apply: submxMr. Qed. Definition mxmodule := G \subset rstabs. @@ -964,7 +962,7 @@ Proof. by apply/mxmoduleP=> x _; rewrite submx1. Qed. Lemma mxmodule_trans m1 m2 (U : 'M_(m1, n)) (W : 'M_(m2, n)) x : mxmodule U -> x \in G -> (W <= U -> W *m rG x <= U)%MS. Proof. -by move=> modU Gx sWU; apply: submx_trans (mxmoduleP modU x Gx); exact: submxMr. +by move=> modU Gx sWU; apply: submx_trans (mxmoduleP modU x Gx); apply: submxMr. Qed. Lemma mxmodule_eigenvector m (U : 'M_(m, n)) : @@ -989,7 +987,7 @@ Qed. Lemma sumsmx_module I r (P : pred I) U : (forall i, P i -> mxmodule (U i)) -> mxmodule (\sum_(i <- r | P i) U i)%MS. Proof. -by move=> modU; elim/big_ind: _; [exact: mxmodule0 | exact: addsmx_module | ]. +by move=> modU; elim/big_ind: _; [apply: mxmodule0 | apply: addsmx_module | ]. Qed. Lemma capmx_module m1 m2 U V : @@ -1002,7 +1000,7 @@ Qed. Lemma bigcapmx_module I r (P : pred I) U : (forall i, P i -> mxmodule (U i)) -> mxmodule (\bigcap_(i <- r | P i) U i)%MS. Proof. -by move=> modU; elim/big_ind: _; [exact: mxmodule1 | exact: capmx_module | ]. +by move=> modU; elim/big_ind: _; [apply: mxmodule1 | apply: capmx_module | ]. Qed. (* Sub- and factor representations induced by a (sub)module. *) @@ -1023,7 +1021,7 @@ Lemma in_submodE m W : @in_submod m W = W *m in_submod 1%:M. Proof. by rewrite mulmxA mulmx1. Qed. Lemma val_submod1 : (val_submod 1%:M :=: U)%MS. -Proof. by rewrite /val_submod /= mul1mx; exact: eq_row_base. Qed. +Proof. by rewrite /val_submod /= mul1mx; apply: eq_row_base. Qed. Lemma val_submodP m W : (@val_submod m W <= U)%MS. Proof. by rewrite mulmx_sub ?eq_row_base. Qed. @@ -1128,7 +1126,7 @@ Lemma in_factmod_addsK m (W : 'M_(m, n)) : (in_factmod (U + W)%MS :=: in_factmod W)%MS. Proof. apply: eqmx_trans (addsmxMr _ _ _) _. -by rewrite ((_ *m _ =P 0) _) ?in_factmod_eq0 //; exact: adds0mx. +by rewrite ((_ *m _ =P 0) _) ?in_factmod_eq0 //; apply: adds0mx. Qed. Lemma add_sub_fact_mod m (W : 'M_(m, n)) : @@ -1151,7 +1149,7 @@ Lemma in_factmodsK m (W : 'M_(m, n)) : (U <= W)%MS -> (U + val_factmod (in_factmod W) :=: W)%MS. Proof. move/addsmx_idPr; apply: eqmx_trans (eqmx_sym _). -by rewrite -{1}[W]add_sub_fact_mod; apply: addsmx_addKl; exact: val_submodP. +by rewrite -{1}[W]add_sub_fact_mod; apply: addsmx_addKl; apply: val_submodP. Qed. Lemma mxrank_in_factmod m (W : 'M_(m, n)) : @@ -1245,7 +1243,7 @@ Lemma envelop_mxP A : reflect (exists a, A = \sum_(x in G) a x *: rG x) (A \in E_G)%MS. Proof. have G_1 := group1 G; have bijG := enum_val_bij_in G_1. -set h := enum_val in bijG; have Gh: h _ \in G by exact: enum_valP. +set h := enum_val in bijG; have Gh: h _ \in G by apply: enum_valP. apply: (iffP submxP) => [[u defA] | [a ->]]. exists (fun x => u 0 (enum_rank_in G_1 x)); apply: (can_inj mxvecK). rewrite defA mulmx_sum_row linear_sum (reindex h) //=. @@ -1374,7 +1372,7 @@ Proof. exact/rfix_mxP. Qed. Lemma rfix_mxS (H K : {set gT}) : H \subset K -> (rfix_mx K <= rfix_mx H)%MS. Proof. -by move=> sHK; apply/rfix_mxP=> x Hx; exact: rfix_mxP (subsetP sHK x Hx). +by move=> sHK; apply/rfix_mxP=> x Hx; apply: rfix_mxP (subsetP sHK x Hx). Qed. Lemma rfix_mx_conjsg (H : {set gT}) x : @@ -1437,7 +1435,7 @@ Proof. by apply/cyclic_mxP; exists 1%:M; rewrite ?mulmx1 ?envelop_mx1. Qed. Lemma cyclic_mx_eq0 u : (cyclic_mx u == 0) = (u == 0). Proof. rewrite -!submx0; apply/idP/idP. - by apply: submx_trans; exact: cyclic_mx_id. + by apply: submx_trans; apply: cyclic_mx_id. move/submx0null->; rewrite genmxE; apply/row_subP=> i. by rewrite row_mul mul_rV_lin1 /= mul0mx ?sub0mx. Qed. @@ -1529,7 +1527,7 @@ Lemma mx_iso_sym U V : mx_iso U V -> mx_iso V U. Proof. case=> f injf homUf defV; exists (invmx f); first by rewrite unitmx_inv. by rewrite dom_hom_invmx // -defV submxMr. -by rewrite -[U](mulmxK injf); exact: eqmxMr (eqmx_sym _). +by rewrite -[U](mulmxK injf); apply: eqmxMr (eqmx_sym _). Qed. Lemma mx_iso_trans U V W : mx_iso U V -> mx_iso V W -> mx_iso U W. @@ -1537,7 +1535,7 @@ Proof. case=> f injf homUf defV [g injg homVg defW]. exists (f *m g); first by rewrite unitmx_mul injf. by apply/hom_mxP=> x Gx; rewrite !mulmxA 2?(hom_mxP _) ?defV. -by rewrite mulmxA; exact: eqmx_trans (eqmxMr g defV) defW. +by rewrite mulmxA; apply: eqmx_trans (eqmxMr g defV) defW. Qed. Lemma mxrank_iso U V : mx_iso U V -> \rank U = \rank V. @@ -1545,7 +1543,7 @@ Proof. by case=> f injf _ <-; rewrite mxrankMfree ?row_free_unit. Qed. Lemma mx_iso_module U V : mx_iso U V -> mxmodule U -> mxmodule V. Proof. -by case=> f _ homUf defV; rewrite -(eqmx_module defV); exact: hom_mxmodule. +by case=> f _ homUf defV; rewrite -(eqmx_module defV); apply: hom_mxmodule. Qed. (* Simple modules (we reserve the term "irreducible" for representations). *) @@ -1562,7 +1560,7 @@ Lemma mxsimpleP U : Proof. do [split => [] [modU nzU simU]; split] => // [V modV sVU nzV | [V]]. apply/idPn; rewrite -(ltn_leqif (mxrank_leqif_sup sVU)) => ltVU. - by case: simU; exists V; exact/and4P. + by case: simU; exists V; apply/and4P. by case/and4P=> modV sVU nzV; apply/negP; rewrite -leqNgt mxrankS ?simU. Qed. @@ -1578,7 +1576,7 @@ have genU := genmxE U; apply simU; exists <<U>>%MS; last by rewrite genU. apply/mxsimpleP; split; rewrite ?(eqmx_eq0 genU) ?(eqmx_module genU) //. case=> V; rewrite !genU=> /and4P[modV sVU nzV ltVU]; case: notF. apply: IHr nzV _ => // [|[W simW sWV]]; first exact: leq_trans ltVU _. -by apply: simU; exists W => //; exact: submx_trans sWV sVU. +by apply: simU; exists W => //; apply: submx_trans sWV sVU. Qed. Lemma mx_iso_simple U V : mx_iso U V -> mxsimple U -> mxsimple V. @@ -1663,19 +1661,19 @@ apply: (iffP and3P) => [[modV] | isoUV]; last first. by case: (mx_iso_simple isoUV simU). have [f injf homUf defV] := isoUV; apply/rowV0Pn; exists (u *m f). rewrite sub_capmx -defV submxMr //. - by apply/row_hom_mxP; exists f; first exact: (submx_trans Uu). + by apply/row_hom_mxP; exists f; first apply: (submx_trans Uu). by rewrite -(mul0mx _ f) (can_eq (mulmxK injf)) nz_u. case/rowV0Pn=> v; rewrite sub_capmx => /andP[Vv]. case/row_hom_mxP => f homMf def_v nz_v eqrUV. pose uG := cyclic_mx u; pose vG := cyclic_mx v. -have def_vG: (uG *m f :=: vG)%MS by rewrite /vG -def_v; exact: hom_cyclic_mx. -have defU: (U :=: uG)%MS by exact: mxsimple_cyclic. +have def_vG: (uG *m f :=: vG)%MS by rewrite /vG -def_v; apply: hom_cyclic_mx. +have defU: (U :=: uG)%MS by apply: mxsimple_cyclic. have mod_uG: mxmodule uG by rewrite cyclic_mx_module. have homUf: (U <= dom_hom_mx f)%MS. by rewrite defU cyclic_mx_sub ?dom_hom_mx_module. have isoUf: mx_iso U (U *m f). apply: mx_Schur_inj_iso => //; apply: contra nz_v; rewrite -!submx0. - by rewrite (eqmxMr f defU) def_vG; exact: submx_trans (cyclic_mx_id v). + by rewrite (eqmxMr f defU) def_vG; apply: submx_trans (cyclic_mx_id v). apply: mx_iso_trans (isoUf) (eqmx_iso _); apply/eqmxP. have sUfV: (U *m f <= V)%MS by rewrite (eqmxMr f defU) def_vG cyclic_mx_sub. by rewrite -mxrank_leqif_eq ?eqn_leq 1?mxrankS // -(mxrank_iso isoUf). @@ -1684,7 +1682,7 @@ Qed. Lemma mxsimple_iso_simple U V : mxsimple_iso U V -> mxsimple U -> mxsimple V. Proof. -by move=> isoUV simU; apply: mx_iso_simple (simU); exact/mxsimple_isoP. +by move=> isoUV simU; apply: mx_iso_simple (simU); apply/mxsimple_isoP. Qed. (* For us, "semisimple" means "sum of simple modules"; this is classically, *) @@ -1713,7 +1711,7 @@ exists J; split=> //; apply/eqmxP; rewrite addsmx_sub sUV; apply/andP; split. by apply/sumsmx_subP=> i Ji; rewrite (sumsmx_sup i). rewrite -/(V_ J); apply/sumsmx_subP=> i _. case Ji: (i \in J). - by apply: submx_trans (addsmxSr _ _); exact: (sumsmx_sup i). + by apply: submx_trans (addsmxSr _ _); apply: (sumsmx_sup i). have [modWi nzWi simWi] := simW i. rewrite (sameP capmx_idPl eqmxP) simWi ?capmxSl ?capmx_module ?addsmx_module //. apply: contraFT (Ji); rewrite negbK => dxWiUVJ. @@ -1754,7 +1752,7 @@ rewrite addsmxC big1 ?adds0mx_id => [|i /andP[_ /eqP] //]. set tI := 'I_(_); set r_ := nth _ _ => defV. have{simW} simWr (i : tI) : mxsimple (W (r_ i)). case: i => m /=; set Pr := fun i => _ => lt_m_r /=. - suffices: (Pr (r_ m)) by case/andP; exact: simW. + suffices: (Pr (r_ m)) by case/andP; apply: simW. apply: all_nthP m lt_m_r; apply/all_filterP. by rewrite -filter_predI; apply: eq_filter => i; rewrite /= andbb. have [J []] := sum_mxsimple_direct_sub simWr defV. @@ -1763,7 +1761,7 @@ case: (set_0Vmem J) => [-> V0 | [j0 Jj0]]. pose K := {j | j \in J}; pose k0 : K := Sub j0 Jj0. have bij_KJ: {on J, bijective (sval : K -> _)}. by exists (insubd k0) => [k _ | j Jj]; rewrite ?valKd ?insubdK. -have J_K (k : K) : sval k \in J by exact: valP k. +have J_K (k : K) : sval k \in J by apply: valP k. rewrite mxdirectE /= !(reindex _ bij_KJ) !(eq_bigl _ _ J_K) -mxdirectE /= -/tI. exact: MxSemisimple. Qed. @@ -1779,7 +1777,7 @@ Lemma addsmx_semisimple U V : Proof. case=> [I W /= simW defU _] [J T /= simT defV _]. have defUV: (\sum_ij sum_rect (fun _ => 'M_n) W T ij :=: U + V)%MS. - by rewrite big_sumType /=; exact: adds_eqmx. + by rewrite big_sumType /=; apply: adds_eqmx. by apply: intro_mxsemisimple defUV _; case=> /=. Qed. @@ -1792,7 +1790,7 @@ Qed. Lemma eqmx_semisimple U V : (U :=: V)%MS -> mxsemisimple U -> mxsemisimple V. Proof. -by move=> eqUV [I W S simW defU dxS]; exists I W => //; exact: eqmx_trans eqUV. +by move=> eqUV [I W S simW defU dxS]; exists I W => //; apply: eqmx_trans eqUV. Qed. Lemma hom_mxsemisimple (V f : 'M_n) : @@ -1800,7 +1798,7 @@ Lemma hom_mxsemisimple (V f : 'M_n) : Proof. case=> I W /= simW defV _; rewrite -defV => /sumsmx_subP homWf. have{defV} defVf: (\sum_i W i *m f :=: V *m f)%MS. - by apply: eqmx_trans (eqmx_sym _) (eqmxMr f defV); exact: sumsmxMr. + by apply: eqmx_trans (eqmx_sym _) (eqmxMr f defV); apply: sumsmxMr. apply: (intro_mxsemisimple defVf) => i _ nzWf. by apply: mx_iso_simple (simW i); apply: mx_Schur_inj_iso; rewrite ?homWf. Qed. @@ -1872,7 +1870,7 @@ Proof. move=> modV redV [] // nssimV; move: {-1}_.+1 (ltnSn (\rank V)) => r leVr. elim: r => // r IHr in V leVr modV redV nssimV. have [V0 | nzV] := eqVneq V 0. - by rewrite nssimV ?V0 //; exact: mxsemisimple0. + by rewrite nssimV ?V0 //; apply: mxsemisimple0. apply (mxsimple_exists modV nzV) => [[U simU sUV]]; have [modU nzU _] := simU. have [W modW defUW dxUW] := redV U modU sUV. have sWV: (W <= V)%MS by rewrite -defUW addsmxSr. @@ -1964,7 +1962,7 @@ rewrite [in compU]unlock; apply/genmxP/andP; split; last first. by apply/sumsmx_subP => i _; rewrite (sumsmx_sup (sval i)). apply/sumsmx_subP => i _. case i0: (r_nz i); first by rewrite (sumsmx_sup (Sub i i0)). -by move/negbFE: i0; rewrite -cyclic_mx_eq0 => /eqP->; exact: sub0mx. +by move/negbFE: i0; rewrite -cyclic_mx_eq0 => /eqP->; apply: sub0mx. Qed. Lemma component_mx_semisimple : mxsemisimple compU. @@ -2027,7 +2025,7 @@ move=> simU simV; apply: (iffP eqP) => isoUV. by apply: component_mx_iso; rewrite ?isoUV ?component_mx_id. rewrite -(genmx_component U) -(genmx_component V); apply/genmxP. wlog suffices: U V simU simV isoUV / (component_mx U <= component_mx V)%MS. - by move=> IH; rewrite !IH //; exact: mx_iso_sym. + by move=> IH; rewrite !IH //; apply: mx_iso_sym. have [I [W isoWU ->]] := component_mx_def simU. apply/sumsmx_subP => i _; apply: mx_iso_component => //. exact: mx_iso_trans (mx_iso_sym isoUV) (isoWU i). @@ -2041,7 +2039,7 @@ move=> simU simV neUV; apply: contraNeq neUV => ntUV. apply: (mxsimple_exists _ ntUV) => [|[W simW]]. by rewrite capmx_module ?component_mx_module. rewrite sub_capmx => /andP[sWU sWV]; apply/component_mx_isoP=> //. -by apply: mx_iso_trans (_ : mx_iso U W) (mx_iso_sym _); exact: component_mx_iso. +by apply: mx_iso_trans (_ : mx_iso U W) (mx_iso_sym _); apply: component_mx_iso. Qed. Section Socle. @@ -2054,7 +2052,7 @@ Record socleType := EnumSocle { Lemma socle_exists : classically socleType. Proof. -pose V : 'M[F]_n := 0; have: mxsemisimple V by exact: mxsemisimple0. +pose V : 'M[F]_n := 0; have: mxsemisimple V by apply: mxsemisimple0. have: n - \rank V < n.+1 by rewrite mxrank0 subn0. elim: _.+1 V => // n' IHn' V; rewrite ltnS => le_nV_n' ssimV. case=> // maxV; apply: (maxV); have [I /= U simU defV _] := ssimV. @@ -2063,12 +2061,12 @@ suffices sMV: (M <= V)%MS. rewrite -defV -(mulmx1 (\sum_i _)%MS) in sMV. have [//| i _] := hom_mxsemisimple_iso simM _ (scalar_mx_hom _ _) sMV. move/mx_iso_sym=> isoM; apply/hasP. - exists (U i); [exact: codom_f | exact/mxsimple_isoP]. + by exists (U i); [apply: codom_f | apply/mxsimple_isoP]. have ssimMV := addsmx_semisimple (mxsimple_semisimple simM) ssimV. apply: contraLR isT => nsMV; apply: IHn' ssimMV _ maxV. apply: leq_trans le_nV_n'; rewrite ltn_sub2l //. rewrite ltn_neqAle rank_leq_row andbT -[_ == _]sub1mx. - apply: contra nsMV; apply: submx_trans; exact: submx1. + by apply: contra nsMV; apply: submx_trans; apply: submx1. rewrite (ltn_leqif (mxrank_leqif_sup _)) ?addsmxSr //. by rewrite addsmx_sub submx_refl andbT. Qed. @@ -2083,7 +2081,7 @@ Lemma component_socle M : mxsimple M -> component_mx M \in socle_enum. Proof. rewrite /socle_enum; case: sG0 => e0 /= sim_e mem_e simM. have /hasP[M' e0M' isoMM'] := mem_e M simM; apply/mapP; exists M' => //. -by apply/eqP/component_mx_isoP; [|exact: sim_e | exact/mxsimple_isoP]. +by apply/eqP/component_mx_isoP; [|apply: sim_e | apply/mxsimple_isoP]. Qed. Inductive socle_sort : predArgType := PackSocle W of W \in socle_enum. @@ -2110,7 +2108,7 @@ Definition socle_repr W := submod_repr (socle_module W). Lemma nz_socle (W : sG) : W != 0 :> 'M_n. Proof. have simW := socle_simple W; have [_ nzW _] := simW; apply: contra nzW. -by rewrite -!submx0; exact: submx_trans (component_mx_id simW). +by rewrite -!submx0; apply: submx_trans (component_mx_id simW). Qed. Lemma socle_mem (W : sG) : (W : 'M_n) \in socle_enum. @@ -2129,7 +2127,7 @@ Definition socle_choiceMixin := Eval hnf in [choiceMixin of sG by <:]. Canonical socle_choiceType := ChoiceType sG socle_choiceMixin. Lemma socleP (W W' : sG) : reflect (W = W') (W == W')%MS. -Proof. by rewrite (sameP genmxP eqP) !{1}genmx_component; exact: (W =P _). Qed. +Proof. by rewrite (sameP genmxP eqP) !{1}genmx_component; apply: (W =P _). Qed. Fact socle_finType_subproof : cancel (fun W => SeqSub (socle_mem W)) (fun s => PackSocle (valP s)). @@ -2154,7 +2152,7 @@ Variable P : pred sG. Notation S := (\sum_(W : sG | P W) socle_val W)%MS. Lemma subSocle_module : mxmodule S. -Proof. by rewrite sumsmx_module // => W _; exact: component_mx_module. Qed. +Proof. by rewrite sumsmx_module // => W _; apply: component_mx_module. Qed. Lemma subSocle_semisimple : mxsemisimple S. Proof. @@ -2203,7 +2201,7 @@ Proof. apply/mxdirect_sumsP=> W _; apply/eqP. rewrite -submx0 capmx_subSocle ?component_mx_module //. apply/sumsmx_subP=> W' /andP[_ neWW']. -by rewrite capmxC component_mx_disjoint //; exact: socle_simple. +by rewrite capmxC component_mx_disjoint //; apply: socle_simple. Qed. Definition Socle := (\sum_(W : sG) W)%MS. @@ -2216,7 +2214,7 @@ Qed. Lemma semisimple_Socle U : mxsemisimple U -> (U <= Socle)%MS. Proof. -by case=> I M /= simM <- _; apply/sumsmx_subP=> i _; exact: simple_Socle. +by case=> I M /= simM <- _; apply/sumsmx_subP=> i _; apply: simple_Socle. Qed. Lemma reducible_Socle U : @@ -2273,7 +2271,7 @@ by rewrite -mulmxA -cGf // mulmxA mulmx_ker mul0mx. Qed. Lemma centgmx_hom m (U : 'M_(m, n)) : centgmx rG f -> (U <= dom_hom_mx f)%MS. -Proof. by rewrite -row_full_dom_hom -sub1mx; exact: submx_trans (submx1 _). Qed. +Proof. by rewrite -row_full_dom_hom -sub1mx; apply: submx_trans (submx1 _). Qed. End CentHom. @@ -2307,7 +2305,7 @@ Lemma mx_abs_irrP : mx_absolutely_irreducible. Proof. have G_1 := group1 G; have bijG := enum_val_bij_in G_1. -set h := enum_val in bijG; have Gh : h _ \in G by exact: enum_valP. +set h := enum_val in bijG; have Gh : h _ \in G by apply: enum_valP. rewrite /mx_absolutely_irreducible; case: (n > 0); last by right; case. apply: (iffP row_fullP) => [[E' E'G] | [_ [a_ a_G]]]. split=> //; exists (fun x B => (mxvec B *m E') 0 (enum_rank_in G_1 x)) => B. @@ -2391,7 +2389,7 @@ Variable rG : mx_representation F G n. Lemma envelop_mx_ring : mxring (enveloping_algebra_mx rG). Proof. -apply/andP; split; first by apply/mulsmx_subP; exact: envelop_mxM. +apply/andP; split; first by apply/mulsmx_subP; apply: envelop_mxM. apply/mxring_idP; exists 1%:M; split=> *; rewrite ?mulmx1 ?mul1mx //. by rewrite -mxrank_eq0 mxrank1. exact: envelop_mx1. @@ -2500,17 +2498,17 @@ Lemma rstabs_subg : rstabs rH U = H :&: rstabs rG U. Proof. by apply/setP=> x; rewrite !inE andbA -in_setI (setIidPl sHG). Qed. Lemma mxmodule_subg : mxmodule rG U -> mxmodule rH U. -Proof. by rewrite /mxmodule rstabs_subg subsetI subxx; exact: subset_trans. Qed. +Proof. by rewrite /mxmodule rstabs_subg subsetI subxx; apply: subset_trans. Qed. End Stabilisers. Lemma mxsimple_subg M : mxmodule rG M -> mxsimple rH M -> mxsimple rG M. Proof. -by move=> modM [_ nzM minM]; split=> // U /mxmodule_subg; exact: minM. +by move=> modM [_ nzM minM]; split=> // U /mxmodule_subg; apply: minM. Qed. Lemma subg_mx_irr : mx_irreducible rH -> mx_irreducible rG. -Proof. by apply: mxsimple_subg; exact: mxmodule1. Qed. +Proof. by apply: mxsimple_subg; apply: mxmodule1. Qed. Lemma subg_mx_abs_irr : mx_absolutely_irreducible rH -> mx_absolutely_irreducible rG. @@ -2734,7 +2732,7 @@ by rewrite val_factmod_module (eqmx_module _ (in_factmod_addsK _ _)). Qed. Lemma rker_submod : rker rU = rstab rG U. -Proof. by rewrite /rker rstab_submod; exact: eqmx_rstab (val_submod1 U). Qed. +Proof. by rewrite /rker rstab_submod; apply: eqmx_rstab (val_submod1 U). Qed. Lemma rstab_norm : G \subset 'N(rstab rG U). Proof. by rewrite -rker_submod rker_norm. Qed. @@ -2807,7 +2805,7 @@ Proof. have Bfree: row_free B by rewrite row_free_unit. split => /mx_irrP[n_gt0 irrG]; apply/mx_irrP; split=> // U. rewrite -[U](mulmxKV uB) -mxmodule_conj -mxrank_eq0 /row_full mxrankMfree //. - by rewrite mxrank_eq0; exact: irrG. + by rewrite mxrank_eq0; apply: irrG. rewrite -mxrank_eq0 /row_full -(mxrankMfree _ Bfree) mxmodule_conj mxrank_eq0. exact: irrG. Qed. @@ -2882,7 +2880,7 @@ Lemma quotient_splitting_field gT (G : {group gT}) (H : {set gT}) : G \subset 'N(H) -> group_splitting_field G -> group_splitting_field (G / H). Proof. move=> nHG splitG n rGH irrGH. -by rewrite -(morphim_mx_abs_irr _ nHG) splitG //; exact/morphim_mx_irr. +by rewrite -(morphim_mx_abs_irr _ nHG) splitG //; apply/morphim_mx_irr. Qed. Lemma coset_splitting_field gT (H : {set gT}) : @@ -2890,7 +2888,7 @@ Lemma coset_splitting_field gT (H : {set gT}) : Proof. move=> split_gT Gbar; have ->: Gbar = (coset H @*^-1 Gbar / H)%G. by apply: val_inj; rewrite /= /quotient morphpreK ?sub_im_coset. -by apply: quotient_splitting_field; [exact: subsetIl | exact: split_gT]. +by apply: quotient_splitting_field; [apply: subsetIl | apply: split_gT]. Qed. End SplittingField. @@ -2958,9 +2956,9 @@ Lemma der1_sub_rker : group_splitting_field G -> mx_irreducible rG -> (G^`(1) \subset rker rG)%g = (n == 1)%N. Proof. -move=> splitG irrG; apply/idP/idP; last by move/eqP; exact: rker_linear. +move=> splitG irrG; apply/idP/idP; last by move/eqP; apply: rker_linear. move/sub_der1_abelian; move/(abelian_abs_irr (kquo_repr rG))=> <-. -by apply: (quotient_splitting_field (rker_norm _) splitG); exact/quo_mx_irr. +by apply: (quotient_splitting_field (rker_norm _) splitG); apply/quo_mx_irr. Qed. End AbelianQuotient. @@ -3086,7 +3084,7 @@ Proof. move=> addUV dxUV. have eqUV: \rank U = \rank (cokermx V). by rewrite mxrank_coker -{3}(mxrank1 F n) -addUV (mxdirectP dxUV) addnK. -have{dxUV} dxUV: (U :&: V = 0)%MS by exact/mxdirect_addsP. +have{dxUV} dxUV: (U :&: V = 0)%MS by apply/mxdirect_addsP. exists (in_submod U (val_factmod 1%:M *m proj_mx U V)) => // [|x Gx]. rewrite /row_free -{6}eqUV -[_ == _]sub1mx -val_submodS val_submod1. rewrite in_submodK ?proj_mx_sub // -{1}[U](proj_mx_id dxUV) //. @@ -3121,13 +3119,13 @@ Variables (gT : finGroupType) (G : {group gT}). Variables (n : nat) (rG : mx_representation F G n) (sG : socleType rG). Lemma socle_irr (W : sG) : mx_irreducible (socle_repr W). -Proof. by apply/submod_mx_irr; exact: socle_simple. Qed. +Proof. by apply/submod_mx_irr; apply: socle_simple. Qed. Lemma socle_rsimP (W1 W2 : sG) : reflect (mx_rsim (socle_repr W1) (socle_repr W2)) (W1 == W2). Proof. have [simW1 simW2] := (socle_simple W1, socle_simple W2). -by apply: (iffP (component_mx_isoP simW1 simW2)); move/mx_rsim_iso; exact. +by apply: (iffP (component_mx_isoP simW1 simW2)); move/mx_rsim_iso; apply. Qed. Local Notation mG U := (mxmodule rG U). @@ -3160,7 +3158,7 @@ Qed. Lemma mxtrace_submod1 U (modU : mG U) : (U :=: 1%:M)%MS -> {in G, forall x, \tr (sr modU x) = \tr (rG x)}. -Proof. by move=> defU; exact: mxtrace_rsim (rsim_submod1 modU defU). Qed. +Proof. by move=> defU; apply: mxtrace_rsim (rsim_submod1 modU defU). Qed. Lemma mxtrace_dadd_mod U V W (modU : mG U) (modV : mG V) (modW : mG W) : (U + V :=: W)%MS -> mxdirect (U + V) -> @@ -3213,9 +3211,9 @@ have isoW i: mx_iso rG U (W i). have ->: (\rank V %/ \rank U)%N = #|I|. symmetry; rewrite -(mulnK #|I| rankU_gt0); congr (_ %/ _)%N. rewrite -defV (mxdirectP dxV) /= -sum_nat_const. - by apply: eq_bigr => i _; exact: mxrank_iso. + by apply: eq_bigr => i _; apply: mxrank_iso. rewrite -sumr_const; apply: eq_bigr => i _; symmetry. -by apply: mxtrace_rsim Gx; apply/mx_rsim_iso; exact: isoW. +by apply: mxtrace_rsim Gx; apply/mx_rsim_iso; apply: isoW. Qed. Lemma mxtrace_Socle : let modS := Socle_module sG in @@ -3271,7 +3269,7 @@ Lemma Clifford_iso2 x U V : Proof. case=> [f injf homUf defV] Gx; have Gx' := groupVr Gx. pose fx := rG (x^-1)%g *m f *m rG x; exists fx; last 1 first. -- by rewrite !mulmxA repr_mxK //; exact: eqmxMr. +- by rewrite !mulmxA repr_mxK //; apply: eqmxMr. - by rewrite !unitmx_mul andbC !repr_mx_unit. apply/hom_mxP=> h Hh; have Gh := subsetP sHG h Hh. rewrite -(mulmxA U) -repr_mxM // conjgCV repr_mxM ?groupJ // !mulmxA. @@ -3287,10 +3285,10 @@ set simH := mxsimple rH; set cH := component_mx rH. have actG: {in G, forall y M, simH M -> cH M *m rG y <= cH (M *m rG y)}%MS. move=> {M} y Gy /= M simM; have [I [U isoU def_cHM]] := component_mx_def simM. rewrite /cH def_cHM sumsmxMr; apply/sumsmx_subP=> i _. - by apply: mx_iso_component; [exact: Clifford_simple | exact: Clifford_iso2]. + by apply: mx_iso_component; [apply: Clifford_simple | apply: Clifford_iso2]. move=> simM Gx; apply/eqmxP; rewrite actG // -/cH. rewrite -{1}[cH _](repr_mxKV rG Gx) submxMr // -{2}[M](repr_mxK rG Gx). -by rewrite actG ?groupV //; exact: Clifford_simple. +by rewrite actG ?groupV //; apply: Clifford_simple. Qed. Hypothesis irrG : mx_irreducible rG. @@ -3300,7 +3298,7 @@ Lemma Clifford_basis M : mxsimple rH M -> let S := \sum_(x in X) M *m rG x in S :=: 1%:M /\ mxdirect S}%MS. Proof. move=> simM. have simMG (g : [subg G]) : mxsimple rH (M *m rG (val g)). - by case: g => x Gx; exact: Clifford_simple. + by case: g => x Gx; apply: Clifford_simple. have [|XG [defX1 dxX1]] := sum_mxsimple_direct_sub simMG (_ : _ :=: 1%:M)%MS. apply/eqmxP; case irrG => _ _ ->; rewrite ?submx1 //; last first. rewrite -submx0; apply/sumsmx_subP; move/(_ 1%g (erefl _)); apply: negP. @@ -3333,7 +3331,8 @@ Proof. split=> [x W W' eqWW' | W x y Gx Gy]. pose Gx := subgP (subg G x); apply/socleP; apply/eqmxP. rewrite -(repr_mxK rG Gx W) -(repr_mxK rG Gx W'); apply: eqmxMr. - apply: eqmx_trans (eqmx_sym _) (valWact _ _); rewrite -eqWW'; exact: valWact. + apply: eqmx_trans (eqmx_sym _) (valWact _ _). + by rewrite -eqWW'; apply: valWact. apply/socleP; rewrite !{1}valWact 2!{1}(eqmxMr _ (valWact _ _)). by rewrite !subgK ?groupM ?repr_mxM ?mulmxA ?andbb. Qed. @@ -3361,7 +3360,7 @@ have Gx := sXG x Xx; apply/imsetP; exists x => //; apply/socleP/eqmxP/eqmx_sym. apply: eqmx_trans (val_Clifford_act _ Gx) _; rewrite PackSocleK. apply: eqmx_trans (eqmx_sym (Clifford_componentJ simM Gx)) _. apply/eqmxP; rewrite (sameP genmxP eqP) !{1}genmx_component. -by apply/component_mx_isoP=> //; exact: Clifford_simple. +by apply/component_mx_isoP=> //; apply: Clifford_simple. Qed. Lemma Clifford_Socle1 : Socle sH = 1%:M. @@ -3431,15 +3430,15 @@ Proof. rewrite join_subG !subsetI sHG subsetIl /=; apply/andP; split. apply/subsetP=> h Hh; have Gh := subsetP sHG h Hh; rewrite inE. apply/subsetP=> W _; have simW := socle_simple W; have [modW _ _] := simW. - have simWh: mxsimple rH (socle_base W *m rG h) by exact: Clifford_simple. + have simWh: mxsimple rH (socle_base W *m rG h) by apply: Clifford_simple. rewrite inE -val_eqE /= PackSocleK eq_sym. apply/component_mx_isoP; rewrite ?subgK //; apply: component_mx_iso => //. by apply: submx_trans (component_mx_id simW); move/mxmoduleP: modW => ->. apply/subsetP=> z cHz; have [Gz _] := setIP cHz; rewrite inE. apply/subsetP=> W _; have simW := socle_simple W; have [modW _ _] := simW. -have simWz: mxsimple rH (socle_base W *m rG z) by exact: Clifford_simple. +have simWz: mxsimple rH (socle_base W *m rG z) by apply: Clifford_simple. rewrite inE -val_eqE /= PackSocleK eq_sym. -by apply/component_mx_isoP; rewrite ?subgK //; exact: Clifford_iso. +by apply/component_mx_isoP; rewrite ?subgK //; apply: Clifford_iso. Qed. Lemma Clifford_astab1 (W : sH) : 'C[W | 'Cl] = rstabs rG W. @@ -3592,7 +3591,7 @@ split=> [ [/andP[modU modV] maxU] | [[modU maxU] modV maxV]]. by split=> // i ltiU; have:= maxU i (ltnW ltiU); rewrite !nth_rcons leqW ltiU. rewrite modV; split=> // i; rewrite !nth_rcons ltnS leq_eqVlt. case: eqP => [-> _ | /= _ ltiU]; first by rewrite ltnn ?eqxx -last_nth. -by rewrite ltiU; exact: maxU. +by rewrite ltiU; apply: maxU. Qed. Theorem mx_Schreier U : @@ -3620,7 +3619,7 @@ set U' := take i U; set U'' := drop _ U; case/and3P=> incU' ltUi incU''. split=> // W [modW ltUW ltWV]; case: notF. apply: {IH_U}(IH_U (U' ++ W :: Ui :: U'')) noV; last 2 first. - by rewrite /mx_subseries -drop_nth // all_cat /= modW -all_cat defU. -- by rewrite cat_path /= -defU'i; exact/and4P. +- by rewrite cat_path /= -defU'i; apply/and4P. - by rewrite -drop_nth // size_cat /= addnS -size_cat defU addSnnS. by rewrite (subseq_trans sU0U) // -defU cat_subseq // -drop_nth ?subseq_cons. Qed. @@ -3686,7 +3685,7 @@ Qed. Lemma section_eqmx U1 U2 V1 V2 modU1 modU2 modV1 modV2 (eqU : (U1 :=: U2)%MS) (eqV : (V1 :=: V2)%MS) : mx_rsim (@section_repr U1 V1 modU1 modV1) (@section_repr U2 V2 modU2 modV2). -Proof. by apply: section_eqmx_add => //; exact: adds_eqmx. Qed. +Proof. by apply: section_eqmx_add => //; apply: adds_eqmx. Qed. Lemma mx_butterfly U V W modU modV modW : ~~ (U == V)%MS -> max_submod U W -> max_submod V W -> @@ -3696,16 +3695,16 @@ Lemma mx_butterfly U V W modU modV modW : Proof. move=> neUV maxU maxV modUV; have{neUV maxU} defW: (U + V :=: W)%MS. wlog{neUV modUV} ltUV: U V modU modV maxU maxV / ~~ (V <= U)%MS. - by case/nandP: neUV => ?; first rewrite addsmxC; exact. + by case/nandP: neUV => ?; first rewrite addsmxC; apply. apply/eqmxP/idPn=> neUVW; case: maxU => ltUW; case/(_ (U + V)%MS). rewrite addsmx_module // ltmxE ltmxEneq neUVW addsmxSl !addsmx_sub. by have [ltVW _] := maxV; rewrite submx_refl andbT ltUV !ltmxW. have sUV_U := capmxSl U V; have sVW: (V <= W)%MS by rewrite -defW addsmxSr. set goal := mx_rsim _ _; suffices{maxV} simUV: goal. split=> //; apply/(max_submodP modUV modU sUV_U). - by apply: mx_rsim_irr simUV _; exact/max_submodP. + by apply: mx_rsim_irr simUV _; apply/max_submodP. apply: {goal}mx_rsim_sym. -by apply: mx_rsim_trans (mx_second_rsim modU modV) _; exact: section_eqmx. +by apply: mx_rsim_trans (mx_second_rsim modU modV) _; apply: section_eqmx. Qed. Lemma mx_JordanHolder_exists U V : @@ -3719,7 +3718,7 @@ case eqUV: (last 0 U == V)%MS. by rewrite andbC; move/eqmx0P->; exists [::]. rewrite last_rcons; move/eqmxP=> eqU'V; case/mx_series_rcons=> compU _ maxUm'. exists (rcons U' V); last by rewrite last_rcons. - apply/mx_series_rcons; split => //; exact: max_submod_eqmx maxUm'. + by apply/mx_series_rcons; split => //; apply: max_submod_eqmx maxUm'. set Um' := last 0 U in maxUm eqUV; have [modU _] := compU. have modUm': modG Um' by rewrite /Um' (last_nth 0) mx_subseries_module'. have [|||W compW lastW] := IHU (V :&: Um')%MS; rewrite ?capmx_module //. @@ -3771,12 +3770,12 @@ have [eqUVm' | neqUVm'] := altP (@eqmxP _ _ _ _ Um' Vm'). split; first by rewrite szV. exists (lift_perm i_m i_m p) => i; case: (unliftP i_m i) => [j|] ->{i}. apply: rsimT (rsimC _) (rsimT (sim_p j) _). - by rewrite lift_max; exact: rsim_rcons. + by rewrite lift_max; apply: rsim_rcons. by rewrite lift_perm_lift lift_max; apply: rsim_rcons; rewrite -szV. have simUVm := section_eqmx modUm' modVm' modUm modVm eqUVm' eqUVm. apply: rsimT (rsimC _) (rsimT simUVm _); first exact: rsim_last. - by rewrite lift_perm_id /= szV; exact: rsim_last. -have maxVUm: max_submod Vm' Um by exact: max_submod_eqmx (eqmx_sym _) maxVm. + by rewrite lift_perm_id /= szV; apply: rsim_last. +have maxVUm: max_submod Vm' Um by apply: max_submod_eqmx (eqmx_sym _) maxVm. have:= mx_butterfly modUm' modVm' modUm neqUVm' maxUm maxVUm. move: (capmx_module _ _); set Wm := (Um' :&: Vm')%MS => modWm [maxWUm simWVm]. have:= mx_butterfly modVm' modUm' modUm _ maxVUm maxUm. @@ -3795,19 +3794,19 @@ rewrite !permM; case: (unliftP i_m i) => [j {simWUm}|] ->{i}; last first. rewrite lift_perm_id tpermL lift_perm_lift lift_max {simWVm}. apply: rsimT (rsimT (pWV j_m) _); last by apply: rsim_rcons; rewrite -szV. apply: rsimT (rsimC _) {simWUm}(rsimT simWUm _); first exact: rsim_last. - by rewrite -lastW in modWm *; exact: rsim_last. + by rewrite -lastW in modWm *; apply: rsim_last. apply: rsimT (rsimC _) {pUW}(rsimT (pUW j) _). - by rewrite lift_max; exact: rsim_rcons. + by rewrite lift_max; apply: rsim_rcons. rewrite lift_perm_lift; case: (unliftP j_m (pU j)) => [k|] ->{j pU}. rewrite tpermD ?(inj_eq (@lift_inj _ _)) ?neq_lift //. rewrite lift_perm_lift !lift_max; set j := lift j_m k. have ltjW: j < size W by have:= ltn_ord k; rewrite -(lift_max k) /= {1 3}szW. apply: rsimT (rsimT (pWV j) _); last by apply: rsim_rcons; rewrite -szV. - by apply: rsimT (rsimC _) (rsim_rcons compW _ _); first exact: rsim_rcons. + by apply: rsimT (rsimC _) (rsim_rcons compW _ _); first apply: rsim_rcons. apply: rsimT {simWVm}(rsimC (rsimT simWVm _)) _. - by rewrite -lastW in modWm *; exact: rsim_last. + by rewrite -lastW in modWm *; apply: rsim_last. rewrite tpermR lift_perm_id /= szV. -by apply: rsimT (rsim_last modVm' modVm _); exact: section_eqmx. +by apply: rsimT (rsim_last modVm' modVm _); apply: section_eqmx. Qed. Lemma mx_JordanHolder_max U (m := size U) V compU modV : @@ -3818,7 +3817,7 @@ rewrite {}/m; set Um := last 0 U => Um1 irrV. have modUm: modG Um := last_mod compU; have simV := rsimC (mx_factmod_sub modV). have maxV: max_submod V Um. move/max_submodP: (mx_rsim_irr simV irrV) => /(_ (submx1 _)). - by apply: max_submod_eqmx; last exact: eqmx_sym. + by apply: max_submod_eqmx; last apply: eqmx_sym. have [W compW lastW] := mx_JordanHolder_exists compU modV maxV. have compWU: mx_series (rcons W Um) by apply/mx_series_rcons; rewrite lastW. have:= mx_JordanHolder compU compWU; rewrite last_rcons size_rcons. @@ -3920,7 +3919,7 @@ Proof. move=> lastU; have [V [modV simGV]] := rsim_regular_factmod. have irrV := mx_rsim_irr simGV irrG. have [i simVU] := mx_JordanHolder_max compU lastU irrV. -exists i; exact: mx_rsim_trans simGV simVU. +by exists i; apply: mx_rsim_trans simGV simVU. Qed. Hypothesis F'G : [char F]^'.-group G. @@ -3943,7 +3942,7 @@ Local Notation tG := #|pred_of_set (classes (gval G))|. Definition classg_base := \matrix_(k < tG) mxvec (gset_mx (enum_val k)). Let groupCl : {in G, forall x, {subset x ^: G <= G}}. -Proof. by move=> x Gx; apply: subsetP; exact: class_subG. Qed. +Proof. by move=> x Gx; apply: subsetP; apply: class_subG. Qed. Lemma classg_base_free : row_free classg_base. Proof. @@ -3960,7 +3959,7 @@ rewrite rowK !linearZ /= mxvecK -(inj_eq enum_val_inj) def_k eq_sym. have [z Gz ->] := imsetP (enum_valP k'). move/eqP=> not_Gxz; rewrite linear_sum big1 ?scaler0 //= => y zGy. rewrite gring_projE ?(groupCl Gz zGy) //. -by case: eqP zGy => // <- /class_transr. +by case: eqP zGy => // <- /class_eqP. Qed. Lemma classg_base_center : (classg_base :=: 'Z(R_G))%MS. @@ -4080,7 +4079,7 @@ Qed. Lemma degree_irr1 : 'n_1 = 1%N. Proof. apply/eqP; rewrite eqn_leq irr_degree_gt0 -rank_irr1. -by rewrite mxrankS ?component_mx_id //; exact: socle_simple. +by rewrite mxrankS ?component_mx_id //; apply: socle_simple. Qed. Definition Wedderburn_subring (i : sG) := <<i *m R_G>>%MS. @@ -4144,7 +4143,7 @@ Qed. Hypothesis F'G : [char F]^'.-group G. Lemma irr_mx_sum : (\sum_(i : sG) i = 1%:M)%MS. -Proof. by apply: reducible_Socle1; exact: mx_Maschke. Qed. +Proof. by apply: reducible_Socle1; apply: mx_Maschke. Qed. Lemma Wedderburn_sum : (\sum_i 'R_i :=: R_G)%MS. Proof. by apply: eqmx_trans sums_R _; rewrite /Socle irr_mx_sum mul1mx. Qed. @@ -4195,7 +4194,7 @@ Qed. Lemma Wedderburn_is_ring i : mxring 'R_i. Proof. rewrite /mxring /left_mx_ideal Wedderburn_closed submx_refl. -by apply/mxring_idP; exists 'e_i; exact: Wedderburn_is_id. +by apply/mxring_idP; exists 'e_i; apply: Wedderburn_is_id. Qed. Lemma Wedderburn_min_ideal m i (E : 'A_(m, nG)) : @@ -4249,7 +4248,7 @@ transitivity (f *m in_submod _ (val_submod 1%:M *m A) *m f'). rewrite (_ : _ *m A = 0) ?(linear0, mul0mx) //. apply/row_matrixP=> i; rewrite row_mul row0 -[row _ _]gring_mxK -gring_row_mul. rewrite (Wedderburn_mulmx0 ne_iG_j) ?linear0 // genmxE mem_gring_mx. -by rewrite (row_subP _) // val_submod1 component_mx_id //; exact: socle_simple. +by rewrite (row_subP _) // val_submod1 component_mx_id //; apply: socle_simple. Qed. Definition irr_comp := odflt 1%irr [pick i | gring_op rG 'e_i != 0]. @@ -4285,7 +4284,7 @@ by move/eqP: (G0 j). Qed. Lemma irr_comp'_op0 j A : j != iG -> (A \in 'R_j)%MS -> gring_op rG A = 0. -Proof. by rewrite eq_sym; exact: not_rsim_op0 rsim_irr_comp. Qed. +Proof. by rewrite eq_sym; apply: not_rsim_op0 rsim_irr_comp. Qed. Lemma irr_comp_envelop : ('R_iG *m lin_mx (gring_op rG) :=: E_G)%MS. Proof. @@ -4311,7 +4310,7 @@ case/sub_sumsmxP=> e ->; rewrite linear_sum mulmx_suml summx_sub // => j _. rewrite -(in_submodK (submxMl _ (M j))); move: (in_submod _ _) => v. have modMj: mxmodule aG (M j) by apply: mx_iso_module (isoM j) _; case: simMi. have rsimMj: mx_rsim rG (submod_repr modMj). - by apply: mx_rsim_trans rsim_irr_comp _; exact/mx_rsim_iso. + by apply: mx_rsim_trans rsim_irr_comp _; apply/mx_rsim_iso. have [f [f' _ hom_f]] := mx_rsim_def (mx_rsim_sym rsimMj); rewrite submx0. have <-: (gring_mx aG (val_submod (v *m (f *m gring_op rG A *m f')))) = 0. by rewrite (eqP opA0) !(mul0mx, linear0). @@ -4352,13 +4351,13 @@ Qed. Lemma irr_reprK i : irr_comp (irr_repr i) = i. Proof. apply/eqP; apply/component_mx_isoP; try exact: socle_simple. -by move/mx_rsim_iso: (rsim_irr_comp (socle_irr i)); exact: mx_iso_sym. +by move/mx_rsim_iso: (rsim_irr_comp (socle_irr i)); apply: mx_iso_sym. Qed. Lemma irr_repr'_op0 i j A : j != i -> (A \in 'R_j)%MS -> gring_op (irr_repr i) A = 0. Proof. -by move=> neq_ij /irr_comp'_op0-> //; [exact: socle_irr | rewrite irr_reprK]. +by move=> neq_ij /irr_comp'_op0-> //; [apply: socle_irr | rewrite irr_reprK]. Qed. Lemma op_Wedderburn_id i : gring_op (irr_repr i) 'e_i = 1%:M. @@ -4380,7 +4379,7 @@ Lemma irr1_repr x : x \in G -> irr_repr 1 x = 1%:M. Proof. move=> Gx; suffices: x \in rker (irr_repr 1) by case/rkerP. apply: subsetP x Gx; rewrite rker_submod rfix_mx_rstabC // -irr1_rfix. -by apply: component_mx_id; exact: socle_simple. +by apply: component_mx_id; apply: socle_simple. Qed. Hypothesis splitG : group_splitting_field G. @@ -4415,7 +4414,7 @@ Qed. Definition linear_irr := [set i | 'n_i == 1%N]. Lemma irr_degree_abelian : abelian G -> forall i, 'n_i = 1%N. -Proof. by move=> cGG i; exact: mxsimple_abelian_linear (socle_simple i). Qed. +Proof. by move=> cGG i; apply: mxsimple_abelian_linear (socle_simple i). Qed. Lemma linear_irr_comp i : 'n_i = 1%N -> (i :=: socle_base i)%MS. Proof. @@ -4539,25 +4538,25 @@ Lemma card_linear_irr (sG : irrType G) : #|linear_irr sG| = #|G : G^`(1)|%g. Proof. move=> F'G splitG; apply/eqP. -wlog sGq: / irrType (G / G^`(1))%G by exact: socle_exists. +wlog sGq: / irrType (G / G^`(1))%G by apply: socle_exists. have [_ nG'G] := andP (der_normal 1 G); apply/eqP; rewrite -card_quotient //. -have cGqGq: abelian (G / G^`(1))%g by exact: sub_der1_abelian. -have F'Gq: [char F]^'.-group (G / G^`(1))%g by exact: morphim_pgroup. +have cGqGq: abelian (G / G^`(1))%g by apply: sub_der1_abelian. +have F'Gq: [char F]^'.-group (G / G^`(1))%g by apply: morphim_pgroup. have splitGq: group_splitting_field (G / G^`(1))%G. exact: quotient_splitting_field. rewrite -(sum_irr_degree sGq) // -sum1_card. pose rG (j : sGq) := morphim_repr (socle_repr j) nG'G. -have irrG j: mx_irreducible (rG j) by apply/morphim_mx_irr; exact: socle_irr. +have irrG j: mx_irreducible (rG j) by apply/morphim_mx_irr; apply: socle_irr. rewrite (reindex (fun j => irr_comp sG (rG j))) /=. apply: eq_big => [j | j _]; last by rewrite irr_degree_abelian. have [_ lin_j _ _] := rsim_irr_comp sG F'G (irrG j). by rewrite inE -lin_j -irr_degreeE irr_degree_abelian. pose sGlin := {i | i \in linear_irr sG}. have sG'k (i : sGlin) : G^`(1)%g \subset rker (irr_repr (val i)). - by case: i => i /=; rewrite !inE => lin; rewrite rker_linear //=; exact/eqP. + by case: i => i /=; rewrite !inE => lin; rewrite rker_linear //=; apply/eqP. pose h' u := irr_comp sGq (quo_repr (sG'k u) nG'G). have irrGq u: mx_irreducible (quo_repr (sG'k u) nG'G). - by apply/quo_mx_irr; exact: socle_irr. + by apply/quo_mx_irr; apply: socle_irr. exists (fun i => oapp h' [1 sGq]%irr (insub i)) => [j | i] lin_i. rewrite (insubT (mem _) lin_i) /=; apply/esym/eqP/socle_rsimP. apply: mx_rsim_trans (rsim_irr_comp sGq F'Gq (irrGq _)). @@ -4585,7 +4584,7 @@ case: (pickP [pred x in G | ~~ is_scalar_mx (rG x)]) => [x | scalG]. rewrite big_nat_recr //= rmorphM mxrankMfree {IHi}//. rewrite row_free_unit rmorphB /= horner_mx_X horner_mx_C. rewrite (mx_Schur irrG) ?subr_eq0 //; last first. - by apply: contraNneq nscal_rGx => ->; exact: scalar_mx_is_scalar. + by apply: contraNneq nscal_rGx => ->; apply: scalar_mx_is_scalar. rewrite -memmx_cent_envelop linearB. rewrite addmx_sub ?eqmx_opp ?scalar_mx_cent //= memmx_cent_envelop. by apply/centgmxP=> j Zh_j; rewrite -!repr_mxM // (centsP cGG). @@ -4628,13 +4627,13 @@ have /hasP[w _ prim_w]: has #[x].-primitive_root (map r (enum sG)). by rewrite size_map -cardE card_sG. by apply/allP=> _ /mapP[W _ ->]; rewrite unity_rootE rx1. have iphi'P := prim_rootP prim_w (rx1 _); pose iphi' := sval (iphi'P _). -have def_r W: r W = w ^+ iphi' W by exact: svalP (iphi'P W). +have def_r W: r W = w ^+ iphi' W by apply: svalP (iphi'P W). have inj_iphi': injective iphi'. by move=> i j eq_ij; apply: inj_r; rewrite !def_r eq_ij. have iphiP: codom iphi' =i 'I_#[x]. by apply/subset_cardP; rewrite ?subset_predT // card_ord card_image. pose iphi i := iinv (iphiP i); exists w => //; exists iphi. -have iphiK: cancel iphi iphi' by move=> i; exact: f_iinv. +have iphiK: cancel iphi iphi' by move=> i; apply: f_iinv. have r_iphi i: r (iphi i) = w ^+ i by rewrite def_r iphiK. split=> // [|i]; last by rewrite scalX r_iphi. by exists iphi' => // W; rewrite /iphi iinv_f. @@ -4645,7 +4644,7 @@ Lemma splitting_cyclic_primitive_root : classically {z : F | #|G|.-primitive_root z}. Proof. case/cyclicP=> x defG F'G splitF; case=> // IH. -wlog sG: / irrType G by exact: socle_exists. +wlog sG: / irrType G by apply: socle_exists. have [w prim_w _] := cycle_repr_structure sG defG F'G splitF. by apply: IH; exists w. Qed. @@ -4761,10 +4760,10 @@ Lemma dec_mxsimple_exists (U : 'M_n) : Proof. elim: {U}_.+1 {-2}U (ltnSn (\rank U)) => // m IHm U leUm modU nzU. have [nsimU | simU] := mxnonsimpleP nzU; last first. - by exists U; first exact/mxsimpleP. + by exists U; first apply/mxsimpleP. move: (xchooseP nsimU); move: (xchoose _) => W /and4P[modW sWU nzW ltWU]. case: (IHm W) => // [|V simV sVW]; first exact: leq_trans ltWU _. -by exists V; last exact: submx_trans sVW sWU. +by exists V; last apply: submx_trans sVW sWU. Qed. Lemma dec_mx_reducible_semisimple U : @@ -4815,7 +4814,7 @@ case haveU: (mx_sat (exU 0%N (fun U => mxmodule_form rG U /\ ~ meetUVf _ U)%T)). by rewrite andbC -sub_capmx. case/andP=> nzU tiUV; have [M simM sMU] := dec_mxsimple_exists modU nzU. apply: (IHm (M :: Ms)) => [|M']; last first. - by case/predU1P=> [-> //|]; exact: simMs. + by case/predU1P=> [-> //|]; apply: simMs. have [_ nzM _] := simM. suffices ltVMV: \rank V < \rank (span (M :: Ms)). rewrite (leq_trans _ Ms_ge_n) // ltn_sub2l ?(leq_trans ltVMV) //. @@ -4834,12 +4833,12 @@ have sMV: (M <= V)%MS. apply: contra not_sMV => /and3P[nzW Vw Mw]. have{Vw Mw} [sWV sWM]: (W <= V /\ W <= M)%MS. rewrite !cyclic_mx_sub ?(submx_trans (nz_row_sub _)) //. - by rewrite sumsmx_module // => M' _; exact: component_mx_module. + by rewrite sumsmx_module // => M' _; apply: component_mx_module. by rewrite (submx_trans _ sWV) // minM ?cyclic_mx_module. -wlog sG: / socleType rG by exact: socle_exists. +wlog sG: / socleType rG by apply: socle_exists. have sVS: (V <= \sum_(W : sG | has (fun Mi => Mi <= W) Ms) W)%MS. rewrite [V](big_nth 0) big_mkord; apply/sumsmx_subP=> i _. - set Mi := Ms`_i; have MsMi: Mi \in Ms by exact: mem_nth. + set Mi := Ms`_i; have MsMi: Mi \in Ms by apply: mem_nth. have simMi := simMs _ MsMi; have S_Mi := component_socle sG simMi. rewrite (sumsmx_sup (PackSocle S_Mi)) ?PackSocleK //. by apply/hasP; exists Mi; rewrite ?component_mx_id. @@ -4874,12 +4873,12 @@ Qed. Lemma rcent_map A : rcent rGf A^f = rcent rG A. Proof. -by apply/setP=> x; rewrite !inE -!map_mxM inj_eq //; exact: map_mx_inj. +by apply/setP=> x; rewrite !inE -!map_mxM inj_eq //; apply: map_mx_inj. Qed. Lemma rstab_map m (U : 'M_(m, n)) : rstab rGf U^f = rstab rG U. Proof. -by apply/setP=> x; rewrite !inE -!map_mxM inj_eq //; exact: map_mx_inj. +by apply/setP=> x; rewrite !inE -!map_mxM inj_eq //; apply: map_mx_inj. Qed. Lemma rstabs_map m (U : 'M_(m, n)) : rstabs rGf U^f = rstabs rG U. @@ -4967,7 +4966,7 @@ have mf_i V: nth 0^f (mf V) i = (V`_i)^f. case: (ltnP i (size V)) => [ltiV | leVi]; first exact: nth_map. by rewrite !nth_default ?size_map. rewrite -(map_mx0 f) mf_i (mf_i (0 :: U)) => modUi'f modUif modUi' modUi. -by apply: map_section_repr; exact: map_regular_repr. +by apply: map_section_repr; apply: map_regular_repr. Qed. Lemma extend_group_splitting_field : @@ -4978,7 +4977,7 @@ have modU0: all ((mxmodule (regular_repr aF G)) #|G|) [::] by []. apply: (mx_Schreier modU0 _) => // [[U [compU lastU _]]]; have [modU _]:= compU. pose Uf := map ((map_mx f) _ _) U. have{lastU} lastUf: (last 0 Uf :=: 1%:M)%MS. - by rewrite -(map_mx0 f) -(map_mx1 f) last_map; exact/map_eqmx. + by rewrite -(map_mx0 f) -(map_mx1 f) last_map; apply/map_eqmx. have modUf: mx_subseries (regular_repr rF G) Uf. rewrite /mx_subseries all_map; apply: etrans modU; apply: eq_all => Ui /=. rewrite -mxmodule_map; apply: eq_subset_r => x. @@ -4986,7 +4985,7 @@ have modUf: mx_subseries (regular_repr rF G) Uf. have absUf i: i < size U -> mx_absolutely_irreducible (subseries_repr i modUf). move=> lt_i_U; rewrite -(mx_rsim_abs_irr (map_regular_subseries i modU _)). rewrite map_mx_abs_irr; apply: splitG. - by apply: mx_rsim_irr (mx_series_repr_irr compU lt_i_U); exact: section_eqmx. + by apply: mx_rsim_irr (mx_series_repr_irr compU lt_i_U); apply: section_eqmx. have compUf: mx_composition_series (regular_repr rF G) Uf. split=> // i; rewrite size_map => ltiU. move/max_submodP: (mx_abs_irrW (absUf i ltiU)); apply. @@ -4994,7 +4993,7 @@ have compUf: mx_composition_series (regular_repr rF G) Uf. by rewrite map_submx // ltmxW // (pathP _ (mx_series_lt compU)). have [[i ltiU] simUi] := rsim_regular_series irrG compUf lastUf. have{simUi} simUi: mx_rsim rG (subseries_repr i modUf). - by apply: mx_rsim_trans simUi _; exact: section_eqmx. + by apply: mx_rsim_trans simUi _; apply: section_eqmx. by rewrite (mx_rsim_abs_irr simUi) absUf; rewrite size_map in ltiU. Qed. @@ -5138,7 +5137,7 @@ Proof. by move=> x y; rewrite mxvalD mxvalN. Qed. Canonical mxval_additive := Additive mxval_sub. Lemma mxval_is_multiplicative : multiplicative mxval. -Proof. by split; [exact: mxvalM | exact: mxval1]. Qed. +Proof. by split; [apply: mxvalM | apply: mxval1]. Qed. Canonical mxval_rmorphism := AddRMorphism mxval_is_multiplicative. Lemma mxval_centg x : centgmx rG (mxval x). @@ -5196,7 +5195,7 @@ Lemma map_mxminpoly_groot : (map_poly gen pA).[groot] = 0. Proof. (* The [_ groot] prevents divergence of simpl. *) apply: mxval_inj; rewrite -horner_map [_ groot]/= mxval_groot mxval0. rewrite -(mx_root_minpoly A); congr ((_ : {poly _}).[A]). -by apply/polyP=> i; rewrite 3!coef_map; exact: genK. +by apply/polyP=> i; rewrite 3!coef_map; apply: genK. Qed. (* Plugging the extension morphism gen into the ext_repr construction *) @@ -5666,11 +5665,11 @@ Lemma sat_gen_form e f : GRing.rformula f -> Proof. have ExP := Exists_rowP; have set_val := set_nth_map_rVval. elim: f e => //. -- by move=> b e _; exact: (iffP satP). +- by move=> b e _; apply: (iffP satP). - rewrite /gen_form => t1 t2 e rt_t; set t := (_ - _)%T. have:= GRing.qf_evalP (gen_env e) (mxrank_form_qf 0 (gen_term t)). rewrite eval_mxrank mxrank_eq0 eval_gen_term // => tP. - by rewrite (sameP satP tP) /= subr_eq0 val_eqE; exact: eqP. + by rewrite (sameP satP tP) /= subr_eq0 val_eqE; apply: eqP. - move=> f1 IH1 f2 IH2 s /= /andP[/(IH1 s)f1P /(IH2 s)f2P]. by apply: (iffP satP) => [[/satP/f1P ? /satP/f2P] | [/f1P/satP ? /f2P/satP]]. - move=> f1 IH1 f2 IH2 s /= /andP[/(IH1 s)f1P /(IH2 s)f2P]. @@ -5785,7 +5784,7 @@ exists F => //; case=> [|n] rG irrG; first by case/mx_irrP: irrG. apply/idPn=> nabsG; pose cG := ('C(enveloping_algebra_mx rG))%MS. have{nabsG} [A]: exists2 A, (A \in cG)%MS & ~~ is_scalar_mx A. apply/has_non_scalar_mxP; rewrite ?scalar_mx_cent // ltnNge. - by apply: contra nabsG; exact: cent_mx_scalar_abs_irr. + by apply: contra nabsG; apply: cent_mx_scalar_abs_irr. rewrite {cG}memmx_cent_envelop -mxminpoly_linear_is_scalar -ltnNge => cGA. move/(non_linear_gen_reducible irrG cGA). set F' := gen_fieldType _ _; set rG' := @map_repr _ F' _ _ _ _ rG. @@ -5813,7 +5812,7 @@ have{sUV'} defV': V' = U'; last rewrite {V'}defV' in compV'. suffices{irrG'}: mx_irreducible rG' by case/mxsimpleP=> _ _ []. have ltiU': i < size U' by rewrite size_map. apply: mx_rsim_irr (mx_rsim_sym _ ) (mx_series_repr_irr compV' ltiU'). -apply: mx_rsim_trans (mx_rsim_map f' rsimVi) _; exact: map_regular_subseries. +by apply: mx_rsim_trans (mx_rsim_map f' rsimVi) _; apply: map_regular_subseries. Qed. Lemma group_closure_field_exists gT F : @@ -5824,17 +5823,17 @@ set n := #|{group gT}|. suffices: classically {Fs : fieldType & {rmorphism F -> Fs} & forall G : {group gT}, enum_rank G < n -> group_splitting_field Fs G}. - apply: classic_bind => [[Fs f splitFs]] _ -> //. - by exists Fs => // G; exact: splitFs. + by exists Fs => // G; apply: splitFs. elim: (n) => [|i IHi]; first by move=> _ -> //; exists F => //; exists id. apply: classic_bind IHi => [[F' f splitF']]. have [le_n_i _ -> // | lt_i_n] := leqP n i. - by exists F' => // G _; apply: splitF'; exact: leq_trans le_n_i. + by exists F' => // G _; apply: splitF'; apply: leq_trans le_n_i. have:= @group_splitting_field_exists _ (enum_val (Ordinal lt_i_n)) F'. apply: classic_bind => [[Fs f' splitFs]] _ -> //. exists Fs => [|G]; first exact: [rmorphism of (f' \o f)]. rewrite ltnS leq_eqVlt -{1}[i]/(val (Ordinal lt_i_n)) val_eqE. case/predU1P=> [defG | ltGi]; first by rewrite -[G]enum_rankK defG. -by apply: (extend_group_splitting_field f'); exact: splitF'. +by apply: (extend_group_splitting_field f'); apply: splitF'. Qed. Lemma group_closure_closed_field (F : closedFieldType) gT : @@ -5854,7 +5853,7 @@ have [a]: exists a, eigenvalue A a. by rewrite coefN Pd1 mulN1r /= subrr. case/negP; rewrite kermx_eq0 row_free_unit (mx_Schur irrG) ?subr_eq0 //. by rewrite -memmx_cent_envelop -raddfN linearD addmx_sub ?scalar_mx_cent. -by apply: contraNneq nscalA => ->; exact: scalar_mx_is_scalar. +by apply: contraNneq nscalA => ->; apply: scalar_mx_is_scalar. Qed. End BuildSplittingField. diff --git a/mathcomp/character/vcharacter.v b/mathcomp/character/vcharacter.v index 9071740..922a73c 100644 --- a/mathcomp/character/vcharacter.v +++ b/mathcomp/character/vcharacter.v @@ -1,17 +1,16 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype tuple finfun bigop prime finset. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism quotient action gproduct. -From mathcomp.algebra -Require Import ssralg poly finalg zmodp cyclic vector ssrnum ssrint intdiv. -From mathcomp.solvable -Require Import sylow pgroup center frobenius. -From mathcomp.field -Require Import algnum algC. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp +Require Import fintype tuple finfun bigop prime ssralg poly finset. +From mathcomp +Require Import fingroup morphism perm automorphism quotient finalg action. +From mathcomp +Require Import gproduct zmodp commutator cyclic center pgroup sylow frobenius. +From mathcomp +Require Import vector ssrnum ssrint intdiv algC algnum. +From mathcomp Require Import classfun character integral_char. Set Implicit Arguments. @@ -195,11 +194,11 @@ Qed. Lemma zchar_subseq S1 S2 A : subseq S1 S2 -> {subset 'Z[S1, A] <= 'Z[S2, A]}. -Proof. move=> sS12; exact: zchar_subset (mem_subseq sS12). Qed. +Proof. by move/mem_subseq; apply: zchar_subset. Qed. Lemma zchar_filter S A (p : pred 'CF(G)) : {subset 'Z[filter p S, A] <= 'Z[S, A]}. -Proof. by apply: zchar_subset=> f; rewrite mem_filter => /andP[]. Qed. +Proof. by apply: zchar_subset=> f; apply/mem_subseq/filter_subseq. Qed. End Zchar. @@ -297,7 +296,7 @@ Hypotheses (Inu : {in 'Z[S] &, isometry nu}) (oSS : pairwise_orthogonal S). Let freeS := orthogonal_free oSS. Let uniqS : uniq S := free_uniq freeS. -Let Z_S : {subset S <= 'Z[S]}. Proof. by move=> phi; exact: mem_zchar. Qed. +Let Z_S : {subset S <= 'Z[S]}. Proof. by move=> phi; apply: mem_zchar. Qed. Let notS0 : 0 \notin S. Proof. by case/andP: oSS. Qed. Let dotSS := proj2 (pairwise_orthogonalP oSS). @@ -311,7 +310,7 @@ have notSnu0: 0 \notin map nu S. by rewrite -cfnorm_eq0 Inu ?Z_S // cfnorm_eq0 => /eqP <-. apply/pairwise_orthogonalP; split; first by rewrite /= notSnu0 map_inj_in_uniq. move=>_ _ /mapP[phi Sphi ->] /mapP[psi Spsi ->]. -by rewrite (inj_in_eq inj_nu) // Inu ?Z_S //; exact: dotSS. +by rewrite (inj_in_eq inj_nu) // Inu ?Z_S //; apply: dotSS. Qed. Lemma cfproj_sum_orthogonal P z phi : @@ -410,7 +409,7 @@ Lemma cfnorm_orthonormal S : orthonormal S -> '[\sum_(xi <- S) xi] = (size S)%:R. Proof. exact: cfnorm_map_orthonormal. Qed. -Lemma zchar_orthonormalP S : +Lemma vchar_orthonormalP S : {subset S <= 'Z[irr G]} -> reflect (exists I : {set Iirr G}, exists b : Iirr G -> bool, perm_eq S [seq (-1) ^+ b i *: 'chi_i | i in I]) @@ -457,7 +456,7 @@ Lemma vchar_norm1P phi : Proof. move=> Zphi phiN1. have: orthonormal phi by rewrite /orthonormal/= phiN1 eqxx. -case/zchar_orthonormalP=> [xi /predU1P[->|] // | I [b def_phi]]. +case/vchar_orthonormalP=> [xi /predU1P[->|] // | I [b def_phi]]. have: phi \in (phi : seq _) := mem_head _ _. by rewrite (perm_eq_mem def_phi) => /mapP[i _ ->]; exists (b i), i. Qed. @@ -484,7 +483,7 @@ have orthS: orthonormal S. rewrite (leq_exp2r _ 1) // -ltnS -(@ltn_exp2r _ _ 2) //. apply: leq_ltn_trans lt_n_4; rewrite -leC_nat -def_n natrX. rewrite cfdot_sum_irr (bigD1 i) //= -normCK def_m addrC -subr_ge0 addrK. - by rewrite sumr_ge0 // => ? _; exact: mul_conjC_ge0. + by rewrite sumr_ge0 // => ? _; apply: mul_conjC_ge0. have <-: size S = n. by apply/eqP; rewrite -eqC_nat -def_n def_phi cfnorm_orthonormal. exists (in_tuple S); split=> // _ /mapP[i _ ->]. @@ -535,24 +534,14 @@ by rewrite linearZ scale_zchar ?Z_T // -defT map_f ?mem_nth. Qed. Lemma Zisometry_of_iso f : - pairwise_orthogonal S -> {in S, isometry f, to 'Z[irr G]} -> + free S -> {in S, isometry f, to 'Z[irr G]} -> {tau : {linear 'CF(L) -> 'CF(G)} | {in S, tau =1 f} & {in 'Z[S], isometry tau, to 'Z[irr G]}}. Proof. -move=> oS [If Zf]; have [/=/andP[S'0 uS] oSS] := pairwise_orthogonalP oS. -have injf: {in S &, injective f}. - move=> xi1 xi2 Sxi1 Sxi2 /=/(congr1 (cfdot (f xi1)))/eqP; rewrite !If //. - by apply: contraTeq => /oSS-> //; rewrite cfnorm_eq0 (memPn S'0). -have{injf} oSf: pairwise_orthogonal (map f S). - apply/pairwise_orthogonalP; split=> /=. - rewrite map_inj_in_uniq // uS (contra _ S'0) // => /mapP[chi Schi /eqP]. - by rewrite eq_sym -cfnorm_eq0 If // cfnorm_eq0 => /eqP <-. - move=> _ _ /mapP[xi1 Xxi1 ->] /mapP[xi2 Xxi2 ->]. - by rewrite If ?(inj_in_eq injf) // => /oSS->. -have{If} nSf: map cfnorm (map f S) = map cfnorm S. - by rewrite -map_comp; apply/eq_in_map=> xi Sxi; rewrite /= If. -have{Zf} ZSf: {subset map f S <= 'Z[irr G]} by move=> _ /mapP[xi /Zf Zfxi ->]. -by have [tau /eq_in_map] := Zisometry_of_cfnorm oS oSf nSf ZSf; exists tau. +move=> freeS [If Zf]; have [tau Dtau Itau] := isometry_of_free freeS If. +exists tau => //; split; first by apply: sub_in2 Itau; apply: zchar_span. +move=> _ /zchar_nth_expansion[a Za ->]; rewrite linear_sum rpred_sum // => i _. +by rewrite linearZ rpredZ_Cint ?Dtau ?Zf ?mem_nth. Qed. Lemma Zisometry_inj A nu : @@ -584,7 +573,7 @@ by rewrite cfAutZ_Cint // scale_zchar // mem_zchar ?SuS ?mem_nth. Qed. Lemma cfAut_vchar A psi : psi \in 'Z[irr G, A] -> psi^u \in 'Z[irr G, A]. -Proof. by apply: cfAut_zchar; exact: irr_aut_closed. Qed. +Proof. by apply: cfAut_zchar; apply: irr_aut_closed. Qed. Lemma sub_aut_zchar S A psi : {subset S <= 'Z[irr G]} -> psi \in 'Z[S, A] -> psi^u \in 'Z[S, A] -> @@ -797,11 +786,9 @@ Proof. by rewrite /dchi scale1r irr0. Qed. Lemma dirr_dchi i : dchi i \in dirr G. Proof. by apply/dirrP; exists i.1; exists i.2. Qed. -Lemma dIrrP (phi : 'CF(G)) : - reflect (exists i , phi = dchi i) (phi \in dirr G). +Lemma dIrrP phi : reflect (exists i, phi = dchi i) (phi \in dirr G). Proof. -by apply: (iffP idP)=> [/dirrP [b [i ->]]| [i ->]]; - [exists (b, i) | exact: dirr_dchi]. +by apply: (iffP idP)=> [/dirrP[b]|] [i ->]; [exists (b, i) | apply: dirr_dchi]. Qed. Lemma dchi_ndirrE (i : dIirr G) : dchi (ndirr i) = - dchi i. @@ -840,7 +827,7 @@ Qed. Lemma dirr_dIirrE J (f : J -> 'CF(G)) : (forall j, f j \in dirr G) -> forall j, dchi (dirr_dIirr f j) = f j. -Proof. by move=> dirrGf j; exact: (@dirr_dIirrPE _ _ xpredT). Qed. +Proof. by move=> dirrGf j; apply: (@dirr_dIirrPE _ _ xpredT). Qed. Definition dirr_constt (B : {set gT}) (phi: 'CF(B)) : {set (dIirr B)} := [set i | 0 < '[phi, dchi i]]. @@ -909,7 +896,7 @@ Proof. move=> PiZ; rewrite [X in X = _]cfun_sum_constt. *) move=> PiZ; rewrite {1}[phi]cfun_sum_constt. rewrite (reindex (to_dirr phi))=> [/= |]; last first. - by exists (@of_irr _)=> //; exact: of_irrK . + by exists (@of_irr _)=> //; apply: of_irrK . by apply: eq_big=> i; rewrite ?irr_constt_to_dirr // cfdot_todirrE. Qed. diff --git a/mathcomp/field/Make b/mathcomp/field/Make index 6ff83ec..66b6d6a 100644 --- a/mathcomp/field/Make +++ b/mathcomp/field/Make @@ -1,4 +1,4 @@ -all.v +all_field.v algC.v algebraics_fundamentals.v algnum.v diff --git a/mathcomp/field/algC.v b/mathcomp/field/algC.v index 511d69e..a9d1258 100644 --- a/mathcomp/field/algC.v +++ b/mathcomp/field/algC.v @@ -1,12 +1,13 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype bigop finset prime generic_quotient. -From mathcomp.algebra -Require Import ssralg poly polydiv mxpoly ssrnum ssrint rat intdiv. -Require Import countalg algebraics_fundamentals. +From mathcomp +Require Import ssrbool ssrfun ssrnat eqtype seq choice div fintype. +From mathcomp +Require Import path bigop finset prime ssralg poly polydiv mxpoly. +From mathcomp +Require Import generic_quotient countalg ssrnum ssrint rat intdiv. +From mathcomp +Require Import algebraics_fundamentals. (******************************************************************************) (* This file provides an axiomatic construction of the algebraic numbers. *) @@ -69,7 +70,7 @@ have nz2: 2%:R != 0 :> L. apply/eqP=> char2; apply: conj_nt => e; apply/eqP/idPn=> eJ. have opp_id x: - x = x :> L. by apply/esym/eqP; rewrite -addr_eq0 -mulr2n -mulr_natl char2 mul0r. - have{char2} char2: 2 \in [char L] by exact/eqP. + have{char2} char2: 2 \in [char L] by apply/eqP. without loss{eJ} eJ: e / conj e = e + 1. move/(_ (e / (e + conj e))); apply. rewrite fmorph_div rmorphD conjK -{1}[conj e](addNKr e) mulrDl. @@ -1082,7 +1083,7 @@ Proof. by move/(rootCX 0)/(_ ler01). Qed. Lemma rootCpX n x k : (k > 0)%N -> 0 <= x -> n.-root (x ^+ k) = n.-root x ^+ k. Proof. -by case: n => [|n] k_gt0; [rewrite !root0C expr0n gtn_eqF | exact: rootCX]. +by case: n => [|n] k_gt0; [rewrite !root0C expr0n gtn_eqF | apply: rootCX]. Qed. Lemma rootCV n x : (n > 0)%N -> 0 <= x -> n.-root x^-1 = (n.-root x)^-1. @@ -1253,7 +1254,7 @@ Lemma normC_sum_upper (I : finType) (P : pred I) (F G : I -> algC) : forall i, P i -> F i = G i. Proof. set sumF := \sum_(i | _) _; set sumG := \sum_(i | _) _ => leFG eq_sumFG. -have posG i: P i -> 0 <= G i by move/leFG; apply: ler_trans; exact: normr_ge0. +have posG i: P i -> 0 <= G i by move/leFG; apply: ler_trans; apply: normr_ge0. have norm_sumG: `|sumG| = sumG by rewrite ger0_norm ?sumr_ge0. have norm_sumF: `|sumF| = \sum_(i | P i) `|F i|. apply/eqP; rewrite eqr_le ler_norm_sum eq_sumFG norm_sumG -subr_ge0 -sumrB. @@ -1655,7 +1656,9 @@ Proof. by rewrite !(addrC x) eqCmodDr. Qed. Lemma eqCmodD e x1 x2 y1 y2 : (x1 == x2 %[mod e] -> y1 == y2 %[mod e] -> x1 + y1 == x2 + y2 %[mod e])%C. -Proof. rewrite -(eqCmodDl e x2 y1) -(eqCmodDr e y1); exact: eqCmod_trans. Qed. +Proof. +by rewrite -(eqCmodDl e x2 y1) -(eqCmodDr e y1); apply: eqCmod_trans. +Qed. Lemma eqCmod_nat (e m n : nat) : (m == n %[mod e])%C = (m == n %[mod e]). Proof. @@ -1832,7 +1835,7 @@ Section AutLmodC. Variables (U V : lmodType algC) (f : {additive U -> V}). Lemma raddfZ_Cnat a u : a \in Cnat -> f (a *: u) = a *: f u. -Proof. by case/CnatP=> n ->; exact: raddfZnat. Qed. +Proof. by case/CnatP=> n ->; apply: raddfZnat. Qed. Lemma raddfZ_Cint a u : a \in Cint -> f (a *: u) = a *: f u. Proof. by case/CintP=> m ->; rewrite !scaler_int raddfMz. Qed. diff --git a/mathcomp/field/algebraics_fundamentals.v b/mathcomp/field/algebraics_fundamentals.v index 7381b82..f72f67e 100644 --- a/mathcomp/field/algebraics_fundamentals.v +++ b/mathcomp/field/algebraics_fundamentals.v @@ -1,18 +1,15 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div path choice fintype tuple bigop finset prime. -From mathcomp.fingroup -Require Import fingroup. -From mathcomp.algebra -Require Import ssralg poly polydiv mxpoly finalg zmodp cyclic. -From mathcomp.algebra -Require Import ssrnum ssrint rat intdiv vector. -From mathcomp.solvable -Require Import pgroup sylow. -Require Import countalg falgebra fieldext separable galois. +From mathcomp +Require Import ssrbool ssrfun ssrnat eqtype seq choice div fintype. +From mathcomp +Require Import path tuple bigop finset prime ssralg poly polydiv mxpoly. +From mathcomp +Require Import countalg ssrnum ssrint rat intdiv. +From mathcomp +Require Import fingroup finalg zmodp cyclic pgroup sylow. +From mathcomp +Require Import vector falgebra fieldext separable galois. (******************************************************************************) (* The main result in this file is the existence theorem that underpins the *) diff --git a/mathcomp/field/algnum.v b/mathcomp/field/algnum.v index cddc383..1dd7fb9 100644 --- a/mathcomp/field/algnum.v +++ b/mathcomp/field/algnum.v @@ -1,14 +1,13 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype tuple finfun bigop prime. -From mathcomp.algebra -Require Import ssralg finalg zmodp poly ssrnum ssrint rat polydiv intdiv. -From mathcomp.algebra -Require Import matrix mxalgebra mxpoly vector. -Require Import algC falgebra fieldext separable galois cyclotomic. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp +Require Import fintype tuple finfun bigop prime ssralg finalg zmodp poly. +From mathcomp +Require Import ssrnum ssrint rat polydiv intdiv algC matrix mxalgebra mxpoly. +From mathcomp +Require Import vector falgebra fieldext separable galois cyclotomic. (******************************************************************************) (* This file provides a few basic results and constructions in algebraic *) @@ -81,7 +80,7 @@ suffices /sig_eqW[[n [|px [|pz []]]]// [Dpx Dpz]]: by rewrite map_comp_poly horner_comp -Dpz. have [rx nz_rx rx0] := r_exists x. have [rz nz_rz rz0] := r_exists (- z). -have char0_Q: [char rat] =i pred0 by exact: char_num. +have char0_Q: [char rat] =i pred0 by apply: char_num. have [n [[pz Dpz] [px Dpx]]] := char0_PET nz_rz rz0 nz_rx rx0 char0_Q. by exists (n, [:: px; - pz]); rewrite /= !raddfN hornerN -[z]opprK Dpz Dpx. Qed. @@ -197,7 +196,7 @@ by rewrite mulr_sumr; apply: eq_bigr => i _; rewrite ffunE mulrA -rmorphM. Qed. Lemma Crat_spanM b : {in Crat & Crat_span b, forall a x, a * x \in Crat_span b}. -Proof. by move=> _ x /CratP[a ->]; exact: Crat_spanZ. Qed. +Proof. by move=> _ x /CratP[a ->]; apply: Crat_spanZ. Qed. (* In principle CtoQn could be taken to be additive and Q-linear, but this *) (* would require a limit construction. *) @@ -234,7 +233,7 @@ Lemma map_Qnum_poly (nu : {rmorphism algC -> algC}) p : p \in polyOver 1%VS -> map_poly (nu \o QnC) p = (map_poly QnC p). Proof. move=> Qp; apply/polyP=> i; rewrite /= !coef_map /=. -have /vlineP[a ->]: p`_i \in 1%VS by exact: polyOverP. +have /vlineP[a ->]: p`_i \in 1%VS by apply: polyOverP. by rewrite alg_num_field !fmorph_rat. Qed. @@ -483,7 +482,7 @@ have pzn_zk0: root (map_poly \1%VF (minPoly 1 zn)) (zn ^+ k). set p1 := map_poly _ _. have [q1 Dp1]: exists q1, p1 = pQtoC q1. have aP i: (minPoly 1 zn)`_i \in 1%VS. - by apply/polyOverP; exact: minPolyOver. + by apply/polyOverP; apply: minPolyOver. have{aP} a_ i := sig_eqW (vlineP _ _ (aP i)). exists (\poly_(i < size (minPoly 1 zn)) sval (a_ i)). apply/polyP=> i; rewrite coef_poly coef_map coef_poly /=. @@ -584,7 +583,7 @@ have IY_0: 0 \in IY by apply/familyP=> // i; rewrite ffunE. pose inIY := enum_rank_in IY_0. pose Y := [seq \prod_(i < m) X`_i ^+ (f : 'I_N ^ m) i | f in IY]. have S_P := Cint_spanP [tuple of Y]; set S := Cint_span _ in S_P. -have sYS: {subset Y <= S} by exact: mem_Cint_span. +have sYS: {subset Y <= S} by apply: mem_Cint_span. have S_1: 1 \in S. by apply/sYS/imageP; exists 0 => //; rewrite big1 // => i; rewrite ffunE. have SmulX (i : 'I_m): {in S, forall x, x * X`_i \in S}. @@ -596,7 +595,7 @@ have SmulX (i : 'I_m): {in S, forall x, x * X`_i \in S}. by rewrite inordK // ltnS (bigmax_sup i). exists (finfun [eta f with i |-> inord (f i).+1]). apply/familyP=> i1; rewrite inE ffunE /= fun_if fiK. - by case: eqP => [-> // | _]; exact: fP. + by case: eqP => [-> // | _]; apply: fP. rewrite (bigD1 i isT) ffunE /= eqxx fiK; congr (_ * _). by apply: eq_bigr => i1; rewrite ffunE /= => /negPf->. have [/monicP ] := (minCpoly_monic X`_i, root_minCpoly X`_i). @@ -609,7 +608,7 @@ have SmulX (i : 'I_m): {in S, forall x, x * X`_i \in S}. have eK: (inord e : 'I_N) = e :> nat by rewrite inordK // ltnS (bigmax_sup i). exists (finfun [eta f with i |-> inord e]). apply/familyP=> i1; rewrite inE ffunE /= fun_if eK. - by case: eqP => [-> // | _]; exact: fP. + by case: eqP => [-> // | _]; apply: fP. rewrite (bigD1 i isT) ffunE /= eqxx eK; congr (_ * _). by apply: eq_bigr => i1; rewrite ffunE /= => /negPf->. exists S; last by exists (Tagged (fun n => n.-tuple _) [tuple of Y]). @@ -733,7 +732,9 @@ Proof. by rewrite !(addrC x) eqAmodDr. Qed. Lemma eqAmodD e x1 x2 y1 y2 : (x1 == x2 %[mod e] -> y1 == y2 %[mod e] -> x1 + y1 == x2 + y2 %[mod e])%A. -Proof. rewrite -(eqAmodDl e x2 y1) -(eqAmodDr e y1); exact: eqAmod_trans. Qed. +Proof. +by rewrite -(eqAmodDl e x2 y1) -(eqAmodDr e y1); apply: eqAmod_trans. +Qed. Lemma eqAmodm0 e : (e == 0 %[mod e])%A. Proof. by rewrite /eqAmod subr0 unfold_in; case: ifPn => // /divff->. Qed. diff --git a/mathcomp/field/all.v b/mathcomp/field/all_field.v index c26dbba..c26dbba 100644 --- a/mathcomp/field/all.v +++ b/mathcomp/field/all_field.v diff --git a/mathcomp/field/closed_field.v b/mathcomp/field/closed_field.v index b96fc38..94df28b 100644 --- a/mathcomp/field/closed_field.v +++ b/mathcomp/field/closed_field.v @@ -1,11 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import bigop. -From mathcomp.algebra -Require Import ssralg poly polydiv. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq. +From mathcomp +Require Import bigop ssralg poly polydiv. (******************************************************************************) (* A proof that algebraically closed field enjoy quantifier elimination, *) @@ -313,7 +311,7 @@ Lemma redivpT_qf (p : polyF) (k : nat * polyF * polyF -> formula F) (q : polyF) Proof. move=> kP rp rq; rewrite /redivpT; apply: isnull_qf=> // [[]]; first exact: kP. apply: sizeT_qf => // sq; apply: sizeT_qf=> // sp. -apply: lead_coefT_qf=> // lq rlq; exact: redivp_rec_loopT_qf. +by apply: lead_coefT_qf=> // lq rlq; apply: redivp_rec_loopT_qf. Qed. Definition rmodpT (p : polyF) (k : polyF -> fF) (q : polyF) : fF := @@ -603,7 +601,7 @@ rewrite -!map_comp. \big[(@gcdp F)/0]_(j <- l) P j %= \big[(@rgcdp F)/0]_(j <- l) P j. elim: l => [| u l ihl] /=; first by rewrite !big_nil eqpxx. rewrite !big_cons; move: ihl; move/(eqp_gcdr (P u)) => h. - apply: eqp_trans h _; rewrite eqp_sym; exact: eqp_rgcd_gcd. + by apply: eqp_trans h _; rewrite eqp_sym; apply: eqp_rgcd_gcd. case g0: (\big[(@rgcdp F)/0%:P]_(j <- map (eval_poly e \o abstrX i) ps) j == 0). rewrite (eqP g0) rgdcop0. case m0 : (_ == 0)=> //=; rewrite ?(size_poly1,size_poly0) //=. diff --git a/mathcomp/field/countalg.v b/mathcomp/field/countalg.v index de6f01e..dfac24b 100644 --- a/mathcomp/field/countalg.v +++ b/mathcomp/field/countalg.v @@ -1,14 +1,13 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import choice fintype bigop generic_quotient. -From mathcomp.algebra -Require Import ssralg finalg zmodp matrix mxalgebra poly polydiv mxpoly. -From mathcomp.algebra -Require Import ring_quotient ssrint rat mxpoly polyXY. -Require Import closed_field. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp +Require Import bigop ssralg finalg zmodp matrix mxalgebra. +From mathcomp +Require Import poly polydiv mxpoly generic_quotient ring_quotient closed_field. +From mathcomp +Require Import ssrint rat. (*****************************************************************************) (* This file clones part of ssralg hierachy for countable types; it does not *) @@ -800,7 +799,13 @@ Export Zmodule.Exports Ring.Exports ComRing.Exports UnitRing.Exports. Export ComUnitRing.Exports IntegralDomain.Exports. Export Field.Exports DecidableField.Exports ClosedField.Exports. +From mathcomp +Require Import poly polydiv generic_quotient ring_quotient. +From mathcomp +Require Import mxpoly polyXY. Import GRing.Theory. +From mathcomp +Require Import closed_field. Canonical Zp_countZmodType m := [countZmodType of 'I_m.+1]. Canonical Zp_countRingType m := [countRingType of 'I_m.+2]. @@ -915,7 +920,7 @@ have EmulV: GRing.Field.axiom Einv. by rewrite piE -{1}[z]reprK -Quotient.idealrBE subr0 in nz_z. apply/eqP; case: sig_eqW => {ex_uv} [uv uv1]; set i := _.+1 in uv1 *. rewrite piE /= -[z]reprK -(rmorphM PtoE) -Quotient.idealrBE. - by rewrite -uv1 opprD addNKr -mulNr; apply/memI; exists i; exact: dvdp_mull. + by rewrite -uv1 opprD addNKr -mulNr; apply/memI; exists i; apply: dvdp_mull. pose EringU := [comUnitRingType of UnitRingType _ (FieldUnitMixin EmulV Einv0)]. have Eunitf := @FieldMixin _ _ EmulV Einv0. pose Efield := FieldType (IdomainType EringU (FieldIdomainMixin Eunitf)) Eunitf. diff --git a/mathcomp/field/cyclotomic.v b/mathcomp/field/cyclotomic.v index 39faf38..8080dd3 100644 --- a/mathcomp/field/cyclotomic.v +++ b/mathcomp/field/cyclotomic.v @@ -1,16 +1,15 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div path choice fintype tuple finfun bigop finset prime. -From mathcomp.fingroup -Require Import fingroup. -From mathcomp.algebra -Require Import ssralg poly finalg zmodp cyclic. -From mathcomp.algebra -Require Import ssrnum ssrint polydiv rat intdiv mxpoly vector. -Require Import falgebra fieldext separable galois algC. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp +Require Import fintype tuple finfun bigop prime ssralg poly finset. +From mathcomp +Require Import fingroup finalg zmodp cyclic. +From mathcomp +Require Import ssrnum ssrint polydiv rat intdiv. +From mathcomp +Require Import mxpoly vector falgebra fieldext separable galois algC. (******************************************************************************) (* This file provides few basic properties of cyclotomic polynomials. *) @@ -160,7 +159,7 @@ rewrite /'Phi_n; case: (C_prim_root_exists _) => z /=. have n_gt0 := prim_order_gt0 prim_z0; rewrite prednK // => prim_z. have [uDn _ inDn] := divisors_correct n_gt0. pose q := \prod_(d <- rem n (divisors n)) 'Phi_d. -have mon_q: q \is monic by apply: monic_prod => d _; exact: Cyclotomic_monic. +have mon_q: q \is monic by apply: monic_prod => d _; apply: Cyclotomic_monic. have defXn1: cyclotomic z n * pZtoC q = 'X^n - 1. rewrite (prod_cyclotomic prim_z) (big_rem n) ?inDn //=. rewrite divnn n_gt0 rmorph_prod /=; congr (_ * _). @@ -229,7 +228,7 @@ Lemma minCpoly_cyclotomic n z : Proof. move=> prim_z; have n_gt0 := prim_order_gt0 prim_z. have Dpz := Cintr_Cyclotomic prim_z; set pz := cyclotomic z n in Dpz *. -have mon_pz: pz \is monic by exact: cyclotomic_monic. +have mon_pz: pz \is monic by apply: cyclotomic_monic. have pz0: root pz z by rewrite root_cyclotomic. have [pf [Dpf mon_pf] dv_pf] := minCpolyP z. have /dvdpP_rat_int[f [af nz_af Df] [g /esym Dfg]]: pf %| pZtoQ 'Phi_n. diff --git a/mathcomp/field/falgebra.v b/mathcomp/field/falgebra.v index ba96824..32acb2c 100644 --- a/mathcomp/field/falgebra.v +++ b/mathcomp/field/falgebra.v @@ -1,11 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype tuple finfun bigop. -From mathcomp.algebra -Require Import ssralg finalg zmodp matrix vector poly. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq path choice fintype. +From mathcomp +Require Import div tuple finfun bigop ssralg finalg zmodp matrix vector poly. (******************************************************************************) (* Finite dimensional free algebras, usually known as F-algebras. *) diff --git a/mathcomp/field/fieldext.v b/mathcomp/field/fieldext.v index 8ec0943..6d63965 100644 --- a/mathcomp/field/fieldext.v +++ b/mathcomp/field/fieldext.v @@ -1,14 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype tuple finfun bigop generic_quotient. -From mathcomp.algebra -Require Import ssralg finalg zmodp matrix vector poly polydiv mxpoly. -Require Import falgebra. - - +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq div choice fintype. +From mathcomp +Require Import tuple finfun bigop ssralg finalg zmodp matrix vector falgebra. +From mathcomp +Require Import poly polydiv mxpoly generic_quotient. (******************************************************************************) (* * Finite dimensional field extentions *) @@ -847,9 +844,9 @@ have v2rP x: {r : 'rV[K_F]_n | x = r2v r}. pose v2r x := sval (v2rP x). have v2rK: cancel v2r (Linear r2v_lin) by rewrite /v2r => x; case: (v2rP x). suffices r2vK: cancel r2v v2r. - by exists n, v2r; [exact: can2_linear v2rK | exists r2v]. + by exists n, v2r; [apply: can2_linear v2rK | exists r2v]. move=> r; apply/rowP=> i; apply/val_inj/(mulIf (nz_bLi i))/eqP; move: i isT. -by apply/forall_inP; move/directv_sum_unique: dxSbL => <- //; exact/eqP/v2rK. +by apply/forall_inP; move/directv_sum_unique: dxSbL => <- //; apply/eqP/v2rK. Qed. Canonical fieldOver_vectType := VectType K_F L_F fieldOver_vectMixin. @@ -885,7 +882,7 @@ Qed. Fact aspaceOver_suproof E : is_aspace (vspaceOver E). Proof. rewrite /is_aspace has_algid1; last by rewrite mem_vspaceOver (@mem1v _ L). -by apply/prodvP=> u v; rewrite !mem_vspaceOver; exact: memvM. +by apply/prodvP=> u v; rewrite !mem_vspaceOver; apply: memvM. Qed. Canonical aspaceOver E := ASpace (aspaceOver_suproof E). @@ -907,7 +904,7 @@ by rewrite scaler_eq0=> /predU1P[] // /idPn[]; rewrite (memPn nz_b) ?memt_nth. Qed. Lemma dim_aspaceOver E : (F <= E)%VS -> \dim (vspaceOver E) = \dim_F E. -Proof. by rewrite -sup_field_module; exact: dim_vspaceOver. Qed. +Proof. by rewrite -sup_field_module; apply: dim_vspaceOver. Qed. Lemma vspaceOverP V_F : {V | [/\ V_F = vspaceOver V, (F * V <= V)%VS & V_F =i V]}. @@ -932,8 +929,8 @@ Proof. have [V [defEF modV memV]] := vspaceOverP E_F. have algE: has_algid V && (V * V <= V)%VS. rewrite has_algid1; last by rewrite -memV mem1v. - by apply/prodvP=> u v; rewrite -!memV; exact: memvM. -by exists (ASpace algE); rewrite -sup_field_module; split; first exact: val_inj. + by apply/prodvP=> u v; rewrite -!memV; apply: memvM. +by exists (ASpace algE); rewrite -sup_field_module; split; first apply: val_inj. Qed. End FieldOver. @@ -1057,7 +1054,7 @@ Qed. Fact baseAspace_suproof (E : {subfield L}) : is_aspace (baseVspace E). Proof. rewrite /is_aspace has_algid1; last by rewrite mem_baseVspace (mem1v E). -by apply/prodvP=> u v; rewrite !mem_baseVspace; exact: memvM. +by apply/prodvP=> u v; rewrite !mem_baseVspace; apply: memvM. Qed. Canonical baseAspace E := ASpace (baseAspace_suproof E). @@ -1072,7 +1069,7 @@ Proof. by rewrite [F1]unlock dim_baseVspace dimv1 mul1n. Qed. Lemma baseVspace_module V (V0 := baseVspace V) : (F1 * V0 <= V0)%VS. Proof. apply/prodvP=> u v; rewrite [F1]unlock !mem_baseVspace => /vlineP[x ->] Vv. -by rewrite -(@scalerAl F L) mul1r; exact: memvZ. +by rewrite -(@scalerAl F L) mul1r; apply: memvZ. Qed. Lemma sub_baseField (E : {subfield L}) : (F1 <= baseVspace E)%VS. @@ -1106,7 +1103,7 @@ Lemma module_baseAspace (E0 : {subfield L0}) : (F1 <= E0)%VS -> {E | E0 = baseAspace E & E0 =i E}. Proof. rewrite -sup_field_module => /module_baseVspace[E defE0 memE0]. -suffices algE: is_aspace E by exists (ASpace algE); first exact: val_inj. +suffices algE: is_aspace E by exists (ASpace algE); first apply: val_inj. rewrite /is_aspace has_algid1 -?memE0 ?mem1v //. by apply/prodvP=> u v; rewrite -!memE0; apply: memvM. Qed. @@ -1387,7 +1384,7 @@ Proof. by rewrite /subfx_scale rmorphM mulrA. Qed. Fact subfx_scaler1r : left_id 1 subfx_scale. Proof. by move=> x; rewrite /subfx_scale rmorph1 mul1r. Qed. Fact subfx_scalerDr : right_distributive subfx_scale +%R. -Proof. by move=> a; exact: mulrDr. Qed. +Proof. by move=> a; apply: mulrDr. Qed. Fact subfx_scalerDl x : {morph subfx_scale^~ x : a b / a + b}. Proof. by move=> a b; rewrite /subfx_scale rmorphD mulrDl. Qed. Definition subfx_lmodMixin := @@ -1606,7 +1603,7 @@ have unitE: GRing.Field.mixin_of urL. have nz_q: q != 0 by rewrite -(can_eq (@rVpolyK _ _)) raddf0 in nz_x. have /Bezout_eq1_coprimepP[u upq1]: coprimep p q. have /contraR := irr_p _ _ (dvdp_gcdl p q); apply. - have: size (gcdp p q) <= size q by exact: leq_gcdpr. + have: size (gcdp p q) <= size q by apply: leq_gcdpr. rewrite leqNgt;apply:contra;move/eqp_size ->. by rewrite (polySpred nz_p) ltnS size_poly. suffices: x * toL u.2 = 1 by exists (toL u.2); rewrite mulrC. diff --git a/mathcomp/field/galois.v b/mathcomp/field/galois.v index f0bcb0b..8ad9c8c 100644 --- a/mathcomp/field/galois.v +++ b/mathcomp/field/galois.v @@ -1,14 +1,13 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype tuple finfun bigop finset. -From mathcomp.fingroup -Require Import fingroup morphism quotient perm action. -From mathcomp.algebra -Require Import ssralg poly polydiv zmodp cyclic matrix mxalgebra vector. -Require Import falgebra fieldext separable. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div choice fintype. +From mathcomp +Require Import tuple finfun bigop ssralg poly polydiv. +From mathcomp +Require Import finset fingroup morphism quotient perm action zmodp cyclic. +From mathcomp +Require Import matrix mxalgebra vector falgebra fieldext separable. (******************************************************************************) (* This file develops some basic Galois field theory, defining: *) diff --git a/mathcomp/field/separable.v b/mathcomp/field/separable.v index 40e3f93..ba6e3ce 100644 --- a/mathcomp/field/separable.v +++ b/mathcomp/field/separable.v @@ -1,16 +1,13 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype tuple finfun bigop prime finset binomial. -From mathcomp.fingroup -Require Import fingroup morphism perm quotient action gproduct. -From mathcomp.algebra -Require Import ssralg poly polydiv finalg zmodp cyclic matrix mxalgebra mxpoly. -From mathcomp.algebra -Require Import matrix mxalgebra mxpoly polyXY vector. -Require Import falgebra fieldext. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div choice fintype. +From mathcomp +Require Import tuple finfun bigop finset prime binomial ssralg poly polydiv. +From mathcomp +Require Import fingroup perm morphism quotient gproduct finalg zmodp cyclic. +From mathcomp +Require Import matrix mxalgebra mxpoly polyXY vector falgebra fieldext. (******************************************************************************) (* This file provides a theory of separable and inseparable field extensions. *) diff --git a/mathcomp/fingroup/Make b/mathcomp/fingroup/Make index ece5714..90d7618 100644 --- a/mathcomp/fingroup/Make +++ b/mathcomp/fingroup/Make @@ -1,5 +1,5 @@ action.v -all.v +all_fingroup.v automorphism.v fingroup.v gproduct.v diff --git a/mathcomp/fingroup/action.v b/mathcomp/fingroup/action.v index 3e92be4..685bc18 100644 --- a/mathcomp/fingroup/action.v +++ b/mathcomp/fingroup/action.v @@ -1,10 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div fintype bigop finset. -Require Import fingroup morphism perm automorphism quotient. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat div seq fintype. +From mathcomp +Require Import bigop finset fingroup morphism perm automorphism quotient. (******************************************************************************) (* Group action: orbits, stabilisers, transitivity. *) @@ -23,20 +22,20 @@ Require Import fingroup morphism perm automorphism quotient. (* groupAction D R == the structure for group actions of D on R. This *) (* is a telescope on action D rT. *) (* gact_range to == the range R of to : groupAction D R. *) -(* GroupAction toAut == construct a groupAction for action to from *) +(* GroupAction toAut == constructs a groupAction for action to from *) (* toAut : actm to @* D \subset Aut R (actm to is *) (* the morphism to {perm rT} associated to 'to'). *) (* orbit to A x == the orbit of x under the action of A via to. *) (* orbit_transversal to A S == a transversal of the partition orbit to A @: S *) (* of S, provided A acts on S via to. *) -(* amove to A x y == the set of a in A whose action send x to y. *) +(* amove to A x y == the set of a in A whose action sends x to y. *) (* 'C_A[x | to] == the stabiliser of x : rT in A :&: D. *) -(* 'C_A(S | to) == the point-wise stabiliser of S : {set rT} in D :&: A. *) +(* 'C_A(S | to) == the pointwise stabiliser of S : {set rT} in D :&: A. *) (* 'N_A(S | to) == the global stabiliser of S : {set rT} in D :&: A. *) (* 'Fix_(S | to)[a] == the set of fixpoints of a in S. *) (* 'Fix_(S | to)(A) == the set of fixpoints of A in S. *) (* In the first three _A can be omitted and defaults to the domain D of to; *) -(* In the last two S can be omitted and defaults to [set: T], so 'Fix_to[a] *) +(* in the last two S can be omitted and defaults to [set: T], so 'Fix_to[a] *) (* is the set of all fixpoints of a. *) (* The domain restriction ensures that stabilisers have a canonical group *) (* structure, but note that 'Fix sets are generally not groups. Indeed, we *) @@ -44,7 +43,7 @@ Require Import fingroup morphism perm automorphism quotient. (* 'C_(G | to)(A) == the centraliser in R :&: G of the group action of *) (* D :&: A via to *) (* 'C_(G | to)[a] == the centraliser in R :&: G of a \in D, via to. *) -(* These sets are groups when G is. G can be omitted: 'C(|to)(A) is the *) +(* These sets are groups when G is; G can be omitted: 'C(|to)(A) is the *) (* centraliser in R of the action of D :&: A via to. *) (* [acts A, on S | to] == A \subset D acts on the set S via to. *) (* {acts A, on S | to} == A acts on the set S (Prop statement). *) @@ -66,13 +65,13 @@ Require Import fingroup morphism perm automorphism quotient. (* 'P == natural action of a permutation group via aperm. *) (* 'J == internal group action (conjugation) via conjg (_ ^ _). *) (* 'R == regular group action (right translation) via mulg (_ * _). *) -(* (but, to limit ambiguity, _ * _ is NOT a canonical action) *) +(* (However, to limit ambiguity, _ * _ is NOT a canonical action.) *) (* to^* == the action induced by to on {set rT} via to^* (== setact to). *) (* 'Js == the internal action on subsets via _ :^ _, equivalent to 'J^*. *) (* 'Rs == the regular action on subsets via rcoset, equivalent to 'R^*. *) (* 'JG == the conjugation action on {group rT} via (_ :^ _)%G. *) (* to / H == the action induced by to on coset_of H via qact to H, and *) -(* restricted to qact_dom to H == 'N(rcosets H 'N(H) | to^* ). *) +(* restricted to (qact_dom to H) == 'N(rcosets H 'N(H) | to^* ). *) (* 'Q == the action induced to cosets by conjugation; the domain is *) (* qact_dom 'J H, which is provably equal to 'N(H). *) (* to %% A == the action of coset_of A via modact to A, with domain D / A *) @@ -90,7 +89,7 @@ Require Import fingroup morphism perm automorphism quotient. (* be the actual morphism object (e.g., coset_morphism H), not *) (* the underlying function (e.g., coset H). *) (* The explicit application of an action to is usually written (to%act x a), *) -(* where the %act omitted if to is an abstract action or a set action to0^*. *) +(* but %act can be omitted if to is an abstract action or a set action to0^*. *) (* Note that this form will simplify and expose the acting function. *) (* There is a %gact scope for group actions; the notations above are *) (* recognised in %gact when they denote canonical group actions. *) @@ -100,7 +99,7 @@ Require Import fingroup morphism perm automorphism quotient. (* the identity function. If to is a group action with range R *) (* then actm to a is canonically a morphism on R. *) (* We also define here the restriction operation on permutations (the domain *) -(* of this operations is a stabiliser), and local automorphpism groups: *) +(* of this operations is a stabiliser), and local automorphism groups: *) (* restr_perm S p == if p acts on S, the permutation with support in S that *) (* coincides with p on S; else the identity. Note that *) (* restr_perm is a permutation group morphism that maps *) @@ -305,12 +304,12 @@ Lemma mem_setact S a x : x \in S -> to x a \in to^* S a. Proof. exact: mem_imset. Qed. Lemma card_setact S a : #|to^* S a| = #|S|. -Proof. by apply: card_imset; exact: act_inj. Qed. +Proof. by apply: card_imset; apply: act_inj. Qed. Lemma setact_is_action : is_action D to^*. Proof. split=> [a R S eqRS | a b Da Db S]; last first. - rewrite /setact /= -imset_comp; apply: eq_imset => x; exact: actMin. + by rewrite /setact /= -imset_comp; apply: eq_imset => x; apply: actMin. apply/setP=> x; apply/idP/idP=> /(mem_setact a). by rewrite eqRS => /imsetP[y Sy /act_inj->]. by rewrite -eqRS => /imsetP[y Sy /act_inj->]. @@ -335,13 +334,13 @@ by rewrite inE xfix. Qed. Lemma afixS A B : A \subset B -> 'Fix_to(B) \subset 'Fix_to(A). -Proof. by move=> sAB; apply/subsetP=> u; rewrite !inE; exact: subset_trans. Qed. +Proof. by move=> sAB; apply/subsetP=> u; rewrite !inE; apply: subset_trans. Qed. Lemma afixU A B : 'Fix_to(A :|: B) = 'Fix_to(A) :&: 'Fix_to(B). Proof. by apply/setP=> x; rewrite !inE subUset. Qed. Lemma afix1P a x : reflect (to x a = x) (x \in 'Fix_to[a]). -Proof. by rewrite inE sub1set inE; exact: eqP. Qed. +Proof. by rewrite inE sub1set inE; apply: eqP. Qed. Lemma astabIdom S : 'C_D(S | to) = 'C(S | to). Proof. by rewrite setIA setIid. Qed. @@ -423,7 +422,7 @@ Lemma astabCin A S : A \subset D -> (A \subset 'C(S | to)) = (S \subset 'Fix_to(A)). Proof. move=> sAD; apply/subsetP/subsetP=> [sAC x xS | sSF a aA]. - by apply/afixP=> a aA; exact: astab_act (sAC _ aA) xS. + by apply/afixP=> a aA; apply: astab_act (sAC _ aA) xS. rewrite !inE (subsetP sAD _ aA); apply/subsetP=> x xS. by move/afixP/(_ _ aA): (sSF _ xS); rewrite inE => ->. Qed. @@ -463,7 +462,7 @@ Qed. Lemma subset_faithful A B S : B \subset A -> [faithful A, on S | to] -> [faithful B, on S | to]. -Proof. by move=> sAB; apply: subset_trans; exact: setSI. Qed. +Proof. by move=> sAB; apply: subset_trans; apply: setSI. Qed. Section Reindex. @@ -478,7 +477,7 @@ Qed. Lemma reindex_acts A a F : [acts A, on S | to] -> a \in A -> \big[op/idx]_(i in S) F i = \big[op/idx]_(i in S) F (to i a). -Proof. by move=> nSA /(subsetP nSA); exact: reindex_astabs. Qed. +Proof. by move=> nSA /(subsetP nSA); apply: reindex_astabs. Qed. End Reindex. @@ -521,7 +520,7 @@ Proof. by move=> a Da /= x; rewrite -{2}(invgK a) actKin ?groupV. Qed. Lemma setactVin S a : a \in D -> to^* S a^-1 = to^~ a @^-1: S. Proof. -by move=> Da; apply: can2_imset_pre; [exact: actKVin | exact: actKin]. +by move=> Da; apply: can2_imset_pre; [apply: actKVin | apply: actKin]. Qed. Lemma actXin x a i : a \in D -> to x (a ^+ i) = iter i (to^~ a) x. @@ -539,10 +538,10 @@ Proof. by rewrite -{2}(setD1K (group1 G)) afixU afix1 setTI. Qed. Lemma orbit_refl G x : x \in orbit to G x. Proof. by rewrite -{1}[x]act1 mem_orbit. Qed. -Local Notation orbit_rel A := (fun x y => y \in orbit to A x). +Local Notation orbit_rel A := (fun x y => x \in orbit to A y). Lemma contra_orbit G x y : x \notin orbit to G y -> x != y. -Proof. by apply: contraNneq => ->; exact: orbit_refl. Qed. +Proof. by apply: contraNneq => ->; apply: orbit_refl. Qed. Lemma orbit_in_sym G : G \subset D -> symmetric (orbit_rel G). Proof. @@ -552,31 +551,31 @@ Qed. Lemma orbit_in_trans G : G \subset D -> transitive (orbit_rel G). Proof. -move=> sGD _ x _ /imsetP[a Ga ->] /imsetP[b Gb ->]. +move=> sGD _ _ z /imsetP[a Ga ->] /imsetP[b Gb ->]. by rewrite -actMin ?mem_orbit ?groupM // (subsetP sGD). Qed. -Lemma orbit_in_transl G x y : - G \subset D -> y \in orbit to G x -> orbit to G y = orbit to G x. +Lemma orbit_in_eqP G x y : + G \subset D -> reflect (orbit to G x = orbit to G y) (x \in orbit to G y). Proof. -move=> sGD Gxy; apply/setP=> z. -by apply/idP/idP; apply: orbit_in_trans; rewrite // orbit_in_sym. +move=> sGD; apply: (iffP idP) => [yGx|<-]; last exact: orbit_refl. +by apply/setP=> z; apply/idP/idP=> /orbit_in_trans-> //; rewrite orbit_in_sym. Qed. -Lemma orbit_in_transr G x y z : +Lemma orbit_in_transl G x y z : G \subset D -> y \in orbit to G x -> (y \in orbit to G z) = (x \in orbit to G z). Proof. -by move=> sGD Gxy; rewrite !(orbit_in_sym _ z) ?(orbit_in_transl _ Gxy). +by move=> sGD Gxy; rewrite !(orbit_in_sym sGD _ z) (orbit_in_eqP y x sGD Gxy). Qed. Lemma orbit_act_in x a G : G \subset D -> a \in G -> orbit to G (to x a) = orbit to G x. -Proof. by move=> sGD /mem_orbit/orbit_in_transl->. Qed. +Proof. by move=> sGD /mem_orbit/orbit_in_eqP->. Qed. Lemma orbit_actr_in x a G y : G \subset D -> a \in G -> (to y a \in orbit to G x) = (y \in orbit to G x). -Proof. by move=> sGD /mem_orbit/orbit_in_transr->. Qed. +Proof. by move=> sGD /mem_orbit/orbit_in_transl->. Qed. Lemma orbit_inv_in A x y : A \subset D -> (y \in orbit to A^-1 x) = (x \in orbit to A y). @@ -630,7 +629,7 @@ Lemma orbit_partition G S : Proof. move=> actsGS; have sGD := acts_dom actsGS. have eqiG: {in S & &, equivalence_rel [rel x y | y \in orbit to G x]}. - by move=> x y z * /=; rewrite orbit_refl; split=> // /orbit_in_transl->. + by move=> x y z * /=; rewrite orbit_refl; split=> // /orbit_in_eqP->. congr (partition _ _): (equivalence_partitionP eqiG). apply: eq_in_imset => x Sx; apply/setP=> y. by rewrite inE /= andb_idl // => /acts_in_orbit->. @@ -718,7 +717,7 @@ Lemma acts_sub_orbit G S x : [acts G, on S | to] -> (orbit to G x \subset S) = (x \in S). Proof. move/acts_act=> GactS. -apply/subsetP/idP=> [| Sx y]; first by apply; exact: orbit_refl. +apply/subsetP/idP=> [| Sx y]; first by apply; apply: orbit_refl. by case/orbitP=> a Ga <-{y}; rewrite GactS. Qed. @@ -738,7 +737,7 @@ by rewrite /= (afixP Cx) // memJ_norm // groupV (subsetP (normsGI _ _) _ nAa). Qed. Lemma atrans_orbit G x : [transitive G, on orbit to G x | to]. -Proof. by apply: mem_imset; exact: orbit_refl. Qed. +Proof. by apply: mem_imset; apply: orbit_refl. Qed. Section OrbitStabilizer. @@ -757,7 +756,7 @@ Lemma amove_orbit : amove to G x @: orbit to G x = rcosets 'C_G[x | to] G. Proof. apply/setP => Ha; apply/imsetP/rcosetsP=> [[y] | [a Ga ->]]. by case/imsetP=> b Gb -> ->{Ha y}; exists b => //; rewrite amove_act. -by rewrite -amove_act //; exists (to x a); first exact: mem_orbit. +by rewrite -amove_act //; exists (to x a); first apply: mem_orbit. Qed. Lemma amoveK : @@ -842,7 +841,7 @@ Qed. Lemma atransPin G S : G \subset D -> [transitive G, on S | to] -> forall x, x \in S -> orbit to G x = S. -Proof. by move=> sGD /imsetP[y _ ->] x; exact: orbit_in_transl. Qed. +Proof. by move=> sGD /imsetP[y _ ->] x; apply/orbit_in_eqP. Qed. Lemma atransP2in G S : G \subset D -> [transitive G, on S | to] -> @@ -869,7 +868,7 @@ apply: (iffP idP) => [trH | defG]. rewrite -(mulgK b a) mem_mulg ?groupV // !inE groupM //= sub1set inE. by rewrite actMin -?xab. apply/imsetP; exists x => //; apply/setP=> y; rewrite -(atransPin sGD trG Sx). -apply/imsetP/imsetP=> [] [a]; last by exists a; first exact: (subsetP sHG). +apply/imsetP/imsetP=> [] [a]; last by exists a; first apply: (subsetP sHG). rewrite -defG => /imset2P[c b /setIP[_ cxc] Hb ->] ->. exists b; rewrite ?actMin ?(astab_dom cxc) ?(subsetP sHD) //. by rewrite (astab_act cxc) ?inE. @@ -879,9 +878,10 @@ End PartialAction. Arguments Scope orbit_transversal [_ group_scope _ action_scope group_scope group_scope]. +Implicit Arguments orbit_in_eqP [aT D rT to G x y]. Implicit Arguments orbit1P [aT D rT to G x]. Implicit Arguments contra_orbit [aT D rT x y]. -Prenex Implicits orbit1P. +Prenex Implicits orbit_in_eqP orbit1P. Notation "''C' ( S | to )" := (astab_group to S) : Group_scope. Notation "''C_' A ( S | to )" := (setI_group A 'C(S | to)) : Group_scope. @@ -922,33 +922,31 @@ Proof. by rewrite !actM actK. Qed. Lemma actCJV a b x : to (to x a) b = to (to x (b ^ a^-1)) a. Proof. by rewrite (actCJ _ a) conjgKV. Qed. -Lemma orbit_sym G x y : (y \in orbit to G x) = (x \in orbit to G y). -Proof. by apply: orbit_in_sym; exact: subsetT. Qed. +Lemma orbit_sym G x y : (x \in orbit to G y) = (y \in orbit to G x). +Proof. exact/orbit_in_sym/subsetT. Qed. Lemma orbit_trans G x y z : - y \in orbit to G x -> z \in orbit to G y -> z \in orbit to G x. -Proof. by apply: orbit_in_trans; exact: subsetT. Qed. + x \in orbit to G y -> y \in orbit to G z -> x \in orbit to G z. +Proof. exact/orbit_in_trans/subsetT. Qed. -Lemma orbit_transl G x y : y \in orbit to G x -> orbit to G y = orbit to G x. -Proof. -move=> Gxy; apply/setP=> z; apply/idP/idP; apply: orbit_trans => //. -by rewrite orbit_sym. -Qed. +Lemma orbit_eqP G x y : + reflect (orbit to G x = orbit to G y) (x \in orbit to G y). +Proof. exact/orbit_in_eqP/subsetT. Qed. -Lemma orbit_transr G x y z : +Lemma orbit_transl G x y z : y \in orbit to G x -> (y \in orbit to G z) = (x \in orbit to G z). -Proof. by move=> Gxy; rewrite orbit_sym (orbit_transl Gxy) orbit_sym. Qed. +Proof. exact/orbit_in_transl/subsetT. Qed. Lemma orbit_act G a x: a \in G -> orbit to G (to x a) = orbit to G x. -Proof. by move/mem_orbit/orbit_transl; exact. Qed. +Proof. exact/orbit_act_in/subsetT. Qed. Lemma orbit_actr G a x y : a \in G -> (to y a \in orbit to G x) = (y \in orbit to G x). -Proof. by move/mem_orbit/orbit_transr; exact. Qed. +Proof. by move/mem_orbit/orbit_transl; apply. Qed. Lemma orbit_eq_mem G x y : (orbit to G x == orbit to G y) = (x \in orbit to G y). -Proof. by apply/eqP/idP=> [<-|]; [exact: orbit_refl | exact: orbit_transl]. Qed. +Proof. exact: sameP eqP (orbit_eqP G x y). Qed. Lemma orbit_inv A x y : (y \in orbit to A^-1 x) = (x \in orbit to A y). Proof. by rewrite orbit_inv_in ?subsetT. Qed. @@ -971,7 +969,7 @@ by rewrite !inE; apply/subsetP=> x Sx; rewrite inE cSa. Qed. Lemma astab1P x a : reflect (to x a = x) (a \in 'C[x | to]). -Proof. by rewrite !inE sub1set inE; exact: eqP. Qed. +Proof. by rewrite !inE sub1set inE; apply: eqP. Qed. Lemma sub_astab1 A x : (A \subset 'C[x | to]) = (x \in 'Fix_to(A)). Proof. by rewrite sub_astab1_in ?subsetT. Qed. @@ -1031,7 +1029,7 @@ Proof. by rewrite -astab_setact /setact imset_set1. Qed. Lemma atransP G S : [transitive G, on S | to] -> forall x, x \in S -> orbit to G x = S. -Proof. by case/imsetP=> x _ -> y; exact: orbit_transl. Qed. +Proof. by case/imsetP=> x _ -> y; apply/orbit_eqP. Qed. Lemma atransP2 G S : [transitive G, on S | to] -> {in S &, forall x y, exists2 a, a \in G & y = to x a}. @@ -1066,13 +1064,13 @@ rewrite (cardD1 X) {X}X_Gx mem_imset // ltnS leqn0 => /eqP GtrS. apply/imsetP; exists x => //; apply/eqP. rewrite eqEsubset acts_sub_orbit // Sx andbT. apply/subsetP=> y Sy; have:= card0_eq GtrS (orbit to G y). -rewrite !inE /= mem_imset // andbT => /eqP <-; exact: orbit_refl. +by rewrite !inE /= mem_imset // andbT => /eqP <-; apply: orbit_refl. Qed. Lemma atrans_dvd G S : [transitive G, on S | to] -> #|S| %| #|G|. -Proof. by case/imsetP=> x _ ->; exact: dvdn_orbit. Qed. +Proof. by case/imsetP=> x _ ->; apply: dvdn_orbit. Qed. -(* Aschbacher 5.2 *) +(* This is Aschbacher (5.2) *) Lemma acts_fix_norm A B : A \subset 'N(B) -> [acts A, on 'Fix_to(B) | to]. Proof. move=> nAB; have:= acts_subnorm_fix to B; rewrite !setTI. @@ -1084,8 +1082,8 @@ Lemma faithfulP A S : [faithful A, on S | to]. Proof. apply: (iffP subsetP) => [Cto1 a Aa Ca | Cto1 a]. - apply/set1P; rewrite Cto1 // inE Aa; exact/astabP. -case/setIP=> Aa /astabP Ca; apply/set1P; exact: Cto1. + by apply/set1P; rewrite Cto1 // inE Aa; apply/astabP. +by case/setIP=> Aa /astabP Ca; apply/set1P; apply: Cto1. Qed. (* This is the first part of Aschbacher (5.7) *) @@ -1100,13 +1098,13 @@ case/(atransP2 transG Su) => y Gy ->{uy}. by apply/astab1P; rewrite astab1_act (bigcapP cSx). Qed. -(* Aschbacher 5.20 *) +(* This is Aschbacher (5.20) *) Theorem subgroup_transitiveP G H S x : x \in S -> H \subset G -> [transitive G, on S | to] -> reflect ('C_G[x | to] * H = G) [transitive H, on S | to]. -Proof. by move=> Sx sHG; exact: subgroup_transitivePin (subsetT G). Qed. +Proof. by move=> Sx sHG; apply: subgroup_transitivePin (subsetT G). Qed. -(* Aschbacher 5.21 *) +(* This is Aschbacher (5.21) *) Lemma trans_subnorm_fixP x G H S : let C := 'C_G[x | to] in let T := 'Fix_(S | to)(H) in [transitive G, on S | to] -> x \in S -> H \subset C -> @@ -1139,12 +1137,13 @@ Qed. End TotalActions. Implicit Arguments astabP [aT rT to S a]. +Implicit Arguments orbit_eqP [aT rT to G x y]. Implicit Arguments astab1P [aT rT to x a]. Implicit Arguments astabsP [aT rT to S a]. Implicit Arguments atransP [aT rT to G S]. Implicit Arguments actsP [aT rT to A S]. Implicit Arguments faithfulP [aT rT to A S]. -Prenex Implicits astabP astab1P astabsP atransP actsP faithfulP. +Prenex Implicits astabP orbit_eqP astab1P astabsP atransP actsP faithfulP. Section Restrict. @@ -1158,7 +1157,7 @@ Variable sAD : A \subset D. Lemma ract_is_action : is_action A (ract sAD). Proof. rewrite /ract; case: to => f [injf fM]. -split=> // x; exact: (sub_in2 (subsetP sAD)). +by split=> // x; apply: (sub_in2 (subsetP sAD)). Qed. Canonical raction := Action ract_is_action. @@ -1334,14 +1333,14 @@ Qed. Lemma qactEcond x a : x \in 'N(H) -> - quotient_action (coset H x) a = - (if a \in qact_dom then coset H (to x a) else coset H x). + quotient_action (coset H x) a + = coset H (if a \in qact_dom then to x a else x). Proof. move=> Nx; apply: val_inj; rewrite val_subact //= qact_subdomE. have: H :* x \in rcosets H 'N(H) by rewrite -rcosetE mem_imset. case nNa: (a \in _); rewrite // -(astabs_act _ nNa). rewrite !val_coset ?(acts_act acts_qact_dom nNa) //=. -case/rcosetsP=> y Ny defHy; rewrite defHy; apply: rcoset_transl. +case/rcosetsP=> y Ny defHy; rewrite defHy; apply: rcoset_eqP. by rewrite rcoset_sym -defHy (mem_imset (_^~_)) ?rcoset_refl. Qed. @@ -1712,7 +1711,7 @@ Qed. Lemma restr_perm_Aut : restr_perm H @* Aut G \subset Aut H. Proof. -by apply/subsetP=> a'; case/morphimP=> a _ AutGa ->{a'}; exact: Aut_restr_perm. +by apply/subsetP=> a'; case/morphimP=> a _ AutGa ->{a'}; apply: Aut_restr_perm. Qed. Lemma Aut_in_isog : Aut_in (Aut G) H \isog restr_perm H @* Aut G. @@ -1730,11 +1729,11 @@ Proof. rewrite (isog_transl _ Aut_in_isog) /=; set rG := _ @* _. apply: (iffP idP) => [iso_rG h injh hH| AutHinG]. have: aut injh hH \in rG; last case/morphimP=> g nHg AutGg def_g. - suffices ->: rG = Aut H by exact: Aut_aut. + suffices ->: rG = Aut H by apply: Aut_aut. by apply/eqP; rewrite eqEcard restr_perm_Aut /= (card_isog iso_rG). exists (autm_morphism AutGg); rewrite injm_autm im_autm; split=> // x Hx. by rewrite -(autE injh hH Hx) def_g actpermE actbyE. -suffices ->: rG = Aut H by exact: isog_refl. +suffices ->: rG = Aut H by apply: isog_refl. apply/eqP; rewrite eqEsubset restr_perm_Aut /=. apply/subsetP=> h AutHh; have hH := im_autm AutHh. have [g [injg gG eq_gh]] := AutHinG _ (injm_autm AutHh) hH. @@ -1881,7 +1880,7 @@ Variable to : groupAction D R. Lemma actperm_Aut : is_groupAction R to. Proof. by case: to. Qed. Lemma im_actperm_Aut : actperm to @* D \subset Aut R. -Proof. by apply/subsetP=> _ /morphimP[a _ Da ->]; exact: actperm_Aut. Qed. +Proof. by apply/subsetP=> _ /morphimP[a _ Da ->]; apply: actperm_Aut. Qed. Lemma gact_out x a : a \in D -> x \notin R -> to x a = x. Proof. by move=> Da Rx; rewrite -actpermE (out_Aut _ Rx) ?actperm_Aut. Qed. @@ -1893,7 +1892,7 @@ by rewrite Aut_morphic ?actperm_Aut. Qed. Lemma actmM a : {in R &, {morph actm to a : x y / x * y}}. -Proof. rewrite /actm; case: ifP => //; exact: gactM. Qed. +Proof. by rewrite /actm; case: ifP => //; apply: gactM. Qed. Canonical act_morphism a := Morphism (actmM a). @@ -2155,7 +2154,7 @@ rewrite eqEcard (card_preimset _ (act_inj _ _)) leqnn andbT. apply/subsetP=> x Nx; rewrite inE; move/(astabs_act (H :* x)): HDa. rewrite mem_rcosets mulSGid ?normG // Nx => /rcosetsP[y Ny defHy]. suffices: to x a \in H :* y by apply: subsetP; rewrite mul_subG ?sub1set ?normG. -by rewrite -defHy; apply: mem_imset; exact: rcoset_refl. +by rewrite -defHy; apply: mem_imset; apply: rcoset_refl. Qed. Lemma qact_is_groupAction : is_groupAction (R / H) (to / H). @@ -2177,8 +2176,8 @@ Proof. move=> sHR; apply/setP=> a; apply/idP/idP=> nHa; have Da := astabs_dom nHa. rewrite !inE Da; apply/subsetP=> x Hx; rewrite inE -(rcoset1 H). have /rcosetsP[y Ny defHy]: to^~ a @: H \in rcosets H 'N(H). - by rewrite (astabs_act _ nHa) -{1}(mulg1 H) -rcosetE mem_imset ?group1. - by rewrite (@rcoset_transl _ H 1 y) -defHy -1?(gact1 Da) mem_setact. + by rewrite (astabs_act _ nHa); apply/rcosetsP; exists 1; rewrite ?mulg1. + by rewrite (rcoset_eqP (_ : 1 \in H :* y)) -defHy -1?(gact1 Da) mem_setact. rewrite !inE Da; apply/subsetP=> Hx; rewrite inE => /rcosetsP[x Nx ->{Hx}]. apply/imsetP; exists (to x a). case Rx: (x \in R); last by rewrite gact_out ?Rx. @@ -2201,14 +2200,14 @@ Variable H : {group aT}. Lemma modact_is_groupAction : is_groupAction 'C_(|to)(H) (to %% H). Proof. -move=> Ha /morphimP[a Na Da ->]; have NDa: a \in 'N_D(H) by exact/setIP. +move=> Ha /morphimP[a Na Da ->]; have NDa: a \in 'N_D(H) by apply/setIP. rewrite inE; apply/andP; split. apply/subsetP=> x; rewrite 2!inE andbC actpermE /= modactEcond //. by apply: contraR; case: ifP => // E Rx; rewrite gact_out. apply/morphicP=> x y /setIP[Rx cHx] /setIP[Ry cHy]. rewrite /= !actpermE /= !modactE ?gactM //. suffices: x * y \in 'C_(|to)(H) by case/setIP. -rewrite groupM //; exact/setIP. +by rewrite groupM //; apply/setIP. Qed. Canonical mod_groupAction := GroupAction modact_is_groupAction. @@ -2367,7 +2366,7 @@ Proof. have [[_ defD2] [injh _]] := (isomP iso_f, isomP iso_h). move=> sSR1; rewrite (morphimEsub _ sSR1). apply: (morph_astabs (gact_stable to1) (injmP injh)) => // u x. -by move/(subsetP sSR1); exact: hfJ. +by move/(subsetP sSR1); apply: hfJ. Qed. Lemma morph_gastab S : S \subset R1 -> f @* 'C(S | to1) = 'C(h @* S | to2). @@ -2375,7 +2374,7 @@ Proof. have [[_ defD2] [injh _]] := (isomP iso_f, isomP iso_h). move=> sSR1; rewrite (morphimEsub _ sSR1). apply: (morph_astab (gact_stable to1) (injmP injh)) => // u x. -by move/(subsetP sSR1); exact: hfJ. +by move/(subsetP sSR1); apply: hfJ. Qed. Lemma morph_gacent A : A \subset D1 -> h @* 'C_(|to1)(A) = 'C_(|to2)(f @* A). @@ -2653,10 +2652,10 @@ Lemma index_cent1 x : #|G : 'C_G[x]| = #|x ^: G|. Proof. by rewrite -astab1J -card_orbit. Qed. Lemma classes_partition : partition (classes G) G. -Proof. by apply: orbit_partition; apply/actsP=> x Gx y; exact: groupJr. Qed. +Proof. by apply: orbit_partition; apply/actsP=> x Gx y; apply: groupJr. Qed. Lemma sum_card_class : \sum_(C in classes G) #|C| = #|G|. -Proof. by apply: acts_sum_card_orbit; apply/actsP=> x Gx y; exact: groupJr. Qed. +Proof. by apply: acts_sum_card_orbit; apply/actsP=> x Gx y; apply: groupJr. Qed. Lemma class_formula : \sum_(C in classes G) #|G : 'C_G[repr C]| = #|G|. Proof. @@ -2668,7 +2667,7 @@ Qed. Lemma abelian_classP : reflect {in G, forall x, x ^: G = [set x]} (abelian G). Proof. rewrite /abelian -astabJ astabC. -by apply: (iffP subsetP) => cGG x Gx; apply/orbit1P; exact: cGG. +by apply: (iffP subsetP) => cGG x Gx; apply/orbit1P; apply: cGG. Qed. Lemma card_classes_abelian : abelian G = (#|classes G| == #|G|). diff --git a/mathcomp/fingroup/all.v b/mathcomp/fingroup/all_fingroup.v index 903da3c..903da3c 100644 --- a/mathcomp/fingroup/all.v +++ b/mathcomp/fingroup/all_fingroup.v diff --git a/mathcomp/fingroup/automorphism.v b/mathcomp/fingroup/automorphism.v index 8c9b0b9..c13a343 100644 --- a/mathcomp/fingroup/automorphism.v +++ b/mathcomp/fingroup/automorphism.v @@ -1,9 +1,8 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat. -From mathcomp.discrete -Require Import fintype finset. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat fintype finset. +From mathcomp Require Import fingroup perm morphism. (******************************************************************************) @@ -13,7 +12,7 @@ Require Import fingroup perm morphism. (* are permutations of type {perm gT} contained in Aut G : {set {perm gT}}. *) (* This lets us use the finGroupType of {perm gT}. Note also that while *) (* morphisms on G are undefined outside G, automorphisms have their support *) -(* in G, i.e., they are the identity ouside G. *) +(* in G, i.e., they are the identity outside G. *) (* Definitions: *) (* Aut G (or [Aut G]) == the automorphism group of G. *) (* [Aut G]%G == the group structure for Aut G. *) @@ -55,7 +54,7 @@ Lemma Aut_morphic A a : a \in Aut A -> morphic A a. Proof. by case/setIdP. Qed. Lemma out_Aut A a x : a \in Aut A -> x \notin A -> a x = x. -Proof. by case/setIdP=> Aa _; exact: out_perm. Qed. +Proof. by case/setIdP=> Aa _; apply: out_perm. Qed. Lemma eq_Aut A : {in Aut A &, forall a b, {in A, a =1 b} -> a = b}. Proof. @@ -91,7 +90,7 @@ Notation f := (autm AutGa). Notation fE := (autmE AutGa). Lemma injm_autm : 'injm f. -Proof. apply/injmP; apply: in2W; exact: perm_inj. Qed. +Proof. by apply/injmP; apply: in2W; apply: perm_inj. Qed. Lemma ker_autm : 'ker f = 1. Proof. by move/trivgP: injm_autm. Qed. @@ -102,7 +101,7 @@ by have:= AutGa; rewrite inE => /andP[/perm_closed <-]; rewrite permKV. Qed. Lemma Aut_closed x : x \in G -> a x \in G. -Proof. by move=> Gx; rewrite -im_autm; exact: mem_morphim. Qed. +Proof. by move=> Gx; rewrite -im_autm; apply: mem_morphim. Qed. End AutGroup. @@ -165,13 +164,13 @@ Lemma morphim_fixP A : A \subset G -> reflect (f @* A = A) (f @* A \subset A). Proof. rewrite /morphim => sAG; have:= eqEcard (f @: A) A. rewrite (setIidPr sAG) card_in_imset ?leqnn ?andbT => [<-|]; first exact: eqP. -move/injmP: injf; apply: sub_in2; exact/subsetP. +by move/injmP: injf; apply: sub_in2; apply/subsetP. Qed. Hypothesis Gf : f @* G = G. Lemma aut_closed : f @: G \subset G. -Proof. by rewrite -morphimEdom; exact/morphim_fixP. Qed. +Proof. by rewrite -morphimEdom; apply/morphim_fixP. Qed. Definition aut := perm_in (injmP injf) aut_closed. @@ -187,7 +186,7 @@ Proof. by rewrite inE morphic_aut perm_in_on. Qed. Lemma imset_autE A : A \subset G -> aut @: A = f @* A. Proof. move=> sAG; rewrite /morphim (setIidPr sAG). -apply: eq_in_imset; apply: sub_in1 autE; exact/subsetP. +by apply: eq_in_imset; apply: sub_in1 autE; apply/subsetP. Qed. Lemma preim_autE A : A \subset G -> aut @^-1: A = f @*^-1 A. @@ -259,7 +258,7 @@ Let domG := subsetP sGD. Lemma im_Aut_isom : Aut_isom injf sGD @* Aut G = Aut (f @* G). Proof. apply/eqP; rewrite eqEcard; apply/andP; split. - by apply/subsetP=> _ /morphimP[a _ AutGa ->]; exact: Aut_Aut_isom. + by apply/subsetP=> _ /morphimP[a _ AutGa ->]; apply: Aut_Aut_isom. have inj_isom' := injm_Aut_isom (injm_invm injf) (morphimS _ sGD). rewrite card_injm ?injm_Aut_isom // -(card_injm inj_isom') ?subset_leq_card //. apply/subsetP=> a /morphimP[a' _ AutfGa' def_a]. @@ -267,7 +266,7 @@ by rewrite -(morphim_invm injf sGD) def_a Aut_Aut_isom. Qed. Lemma Aut_isomP : isom (Aut G) (Aut (f @* G)) (Aut_isom injf sGD). -Proof. by apply/isomP; split; [exact: injm_Aut_isom | exact: im_Aut_isom]. Qed. +Proof. by apply/isomP; split; [apply: injm_Aut_isom | apply: im_Aut_isom]. Qed. Lemma injm_Aut : Aut (f @* G) \isog Aut G. Proof. by rewrite isog_sym (isom_isog _ _ Aut_isomP). Qed. @@ -294,7 +293,7 @@ Proof. by []. Qed. Variable G : {group gT}. Lemma injm_conj x : 'injm (conjgm G x). -Proof. by apply/injmP; apply: in2W; exact: conjg_inj. Qed. +Proof. by apply/injmP; apply: in2W; apply: conjg_inj. Qed. Lemma conj_isom x : isom G (G :^ x) (conjgm G x). Proof. by apply/isomP; rewrite morphim_conj setIid injm_conj. Qed. @@ -303,7 +302,7 @@ Lemma conj_isog x : G \isog G :^ x. Proof. exact: isom_isog (conj_isom x). Qed. Lemma norm_conjg_im x : x \in 'N(G) -> conjgm G x @* G = G. -Proof. by rewrite morphimEdom; exact: normP. Qed. +Proof. by rewrite morphimEdom; apply: normP. Qed. Lemma norm_conj_isom x : x \in 'N(G) -> isom G G (conjgm G x). Proof. by move/norm_conjg_im/restr_isom_to/(_ (conj_isom x))->. Qed. @@ -314,7 +313,7 @@ Lemma norm_conj_autE : {in 'N(G) & G, forall x y, conj_aut x y = y ^ x}. Proof. by move=> x y nGx Gy; rewrite /= autE //= subgK. Qed. Lemma conj_autE : {in G &, forall x y, conj_aut x y = y ^ x}. -Proof. by apply: sub_in11 norm_conj_autE => //; exact: subsetP (normG G). Qed. +Proof. by apply: sub_in11 norm_conj_autE => //; apply: subsetP (normG G). Qed. Lemma conj_aut_morphM : {in 'N(G) &, {morph conj_aut : x y / x * y}}. Proof. @@ -336,7 +335,7 @@ by rewrite perm1 norm_conj_autE // conjgE -cGx ?mulKg. Qed. Lemma Aut_conj_aut A : conj_aut @* A \subset Aut G. -Proof. by apply/subsetP=> _ /imsetP[x _ ->]; exact: Aut_aut. Qed. +Proof. by apply/subsetP=> _ /imsetP[x _ ->]; apply: Aut_aut. Qed. End ConjugationMorphism. @@ -359,17 +358,13 @@ Definition characteristic A B := Infix "\char" := characteristic. Lemma charP H G : - reflect [/\ H \subset G - & forall f : {morphism G >-> gT}, - 'injm f -> f @* G = G -> f @* H = H] - (H \char G). + let fixH (f : {morphism G >-> gT}) := 'injm f -> f @* G = G -> f @* H = H in + reflect [/\ H \subset G & forall f, fixH f] (H \char G). Proof. -apply: (iffP andP) => [] [sHG chHG]; split=> //. - move=> f injf Gf; apply/morphim_fixP=> //. - by have:= forallP chHG (aut injf Gf); rewrite Aut_aut imset_autE. -apply/forall_inP=> f Af; have injf := injm_autm Af. -move/(morphim_fixP injf _ sHG): (chHG _ injf (im_autm Af)). -by rewrite /morphim (setIidPr _). +do [apply: (iffP andP) => -[sHG chHG]; split] => // [f injf Gf|]. + by apply/morphim_fixP; rewrite // -imset_autE ?(forall_inP chHG) ?Aut_aut. +apply/forall_inP=> f Af; rewrite -(autmE Af) -morphimEsub //. +by rewrite chHG ?injm_autm ?im_autm. Qed. (* Characteristic subgroup properties : composition, relational properties *) @@ -385,13 +380,13 @@ Proof. case/charP=> sKH chKH; case/charP=> sHG chHG. apply/charP; split=> [|f injf Gf]; first exact: subset_trans sHG. rewrite -{1}(setIidPr sKH) -(morphim_restrm sHG) chKH //. - rewrite ker_restrm; move/trivgP: injf => ->; exact: subsetIr. + by rewrite ker_restrm; move/trivgP: injf => ->; apply: subsetIr. by rewrite morphim_restrm setIid chHG. Qed. Lemma char_norms H G : H \char G -> 'N(G) \subset 'N(H). Proof. -case/charP=> sHG chHG; apply/normsP=> x /normP Nx. +case/charP=> sHG chHG; apply/normsP=> x /normP-Nx. have:= chHG [morphism of conjgm G x] => /=. by rewrite !morphimEsub //=; apply; rewrite // injm_conj. Qed. @@ -400,7 +395,7 @@ Lemma char_sub A B : A \char B -> A \subset B. Proof. by case/andP. Qed. Lemma char_norm_trans H G A : H \char G -> A \subset 'N(G) -> A \subset 'N(H). -Proof. by move/char_norms=> nHnG nGA; exact: subset_trans nHnG. Qed. +Proof. by move/char_norms=> nHnG nGA; apply: subset_trans nHnG. Qed. Lemma char_normal_trans H G K : K \char H -> H <| G -> K <| G. Proof. @@ -418,10 +413,10 @@ Lemma charI G H K : H \char G -> K \char G -> H :&: K \char G. Proof. case/charP=> sHG chHG; case/charP=> _ chKG. apply/charP; split=> [|f injf Gf]; first by rewrite subIset // sHG. -rewrite morphimGI ?(chHG, chKG) //; exact: subset_trans (sub1G H). +by rewrite morphimGI ?(chHG, chKG) //; apply: subset_trans (sub1G H). Qed. -Lemma charMgen G H K : H \char G -> K \char G -> H <*> K \char G. +Lemma charY G H K : H \char G -> K \char G -> H <*> K \char G. Proof. case/charP=> sHG chHG; case/charP=> sKG chKG. apply/charP; split=> [|f injf Gf]; first by rewrite gen_subG subUset sHG. @@ -430,7 +425,7 @@ Qed. Lemma charM G H K : H \char G -> K \char G -> H * K \char G. Proof. -move=> chHG chKG; rewrite -norm_joinEl ?charMgen //. +move=> chHG chKG; rewrite -norm_joinEl ?charY //. exact: subset_trans (char_sub chHG) (char_norm chKG). Qed. @@ -440,7 +435,7 @@ Lemma lone_subgroup_char G H : Proof. move=> sHG Huniq; apply/charP; split=> // f injf Gf; apply/eqP. have{injf} injf: {in H &, injective f}. - by move/injmP: injf; apply: sub_in2; exact/subsetP. + by move/injmP: injf; apply: sub_in2; apply/subsetP. have fH: f @* H = f @: H by rewrite /morphim (setIidPr sHG). rewrite eqEcard {2}fH card_in_imset ?{}Huniq //=. by rewrite -{3}Gf morphimS. @@ -453,6 +448,7 @@ End Characteristicity. Arguments Scope characteristic [_ group_scope group_scope]. Notation "H \char G" := (characteristic H G) : group_scope. +Hint Resolve char_refl. Section InjmChar. diff --git a/mathcomp/fingroup/fingroup.v b/mathcomp/fingroup/fingroup.v index 9044cd0..40d25ba 100644 --- a/mathcomp/fingroup/fingroup.v +++ b/mathcomp/fingroup/fingroup.v @@ -1,9 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import choice fintype div path bigop prime finset. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq choice fintype. +From mathcomp +Require Import div path bigop prime finset. (******************************************************************************) (* This file defines the main interface for finite groups : *) @@ -27,7 +27,7 @@ Require Import choice fintype div path bigop prime finset. (* from proofs of the group axioms. *) (* [baseFinGroupType of T] == a clone of an existing baseFinGroupType *) (* structure on T, for T (the existing structure *) -(* might be for som delta-expansion of T). *) +(* might be for some delta-expansion of T). *) (* [finGroupType of T] == a clone of an existing finGroupType structure on *) (* T, for the canonical baseFinGroupType structure *) (* of T (the existing structure might be for the *) @@ -60,7 +60,9 @@ Require Import choice fintype div path bigop prime finset. (* x^-1 == the group inverse of x. *) (* x ^- n == the inverse of x ^+ n (notation for (x ^+ n)^-1). *) (* 1 == the unit element. *) -(* x ^ y == the conjugate of x by y. *) +(* x ^ y == the conjugate of x by y (i.e., y^-1 * (x * y)). *) +(* [~ x, y] == the commutator of x and y (i.e., x^-1 * x ^ y). *) +(* [~ x1, ..., xn] == the commutator of x1, ..., xn (associating left). *) (* \prod_(i ...) x i == the product of the x i (order-sensitive). *) (* commute x y <-> x and y commute. *) (* centralises x A <-> x centralises A. *) @@ -68,17 +70,19 @@ Require Import choice fintype div path bigop prime finset. (* 'C_G[x] == the set of elements of G that commute with x. *) (* <[x]> == the cyclic subgroup generated by the element x. *) (* #[x] == the order of the element x, i.e., #|<[x]>|. *) -(* [~ x1, ..., xn] == the commutator of x1, ..., xn. *) (* Operations on subsets/subgroups of a finite group: *) (* H * G == {xy | x \in H, y \in G}. *) (* 1 or [1] or [1 gT] == the unit group. *) (* [set: gT]%G == the group of all x : gT (in Group_scope). *) +(* group_set G == G contains 1 and is closed under binary product; *) +(* this is the characteristic property of the *) +(* {group gT} subtype of {set gT}. *) (* [subg G] == the subtype, set, or group of all x \in G: this *) (* notation is defined simultaneously in %type, %g *) (* and %G scopes, and G must denote a {group gT} *) (* structure (G is in the %G scope). *) (* subg, sgval == the projection into and injection from [subg G]. *) -(* H^# == the set H minus the unit element *) +(* H^# == the set H minus the unit element. *) (* repr H == some element of H if 1 \notin H != set0, else 1. *) (* (repr is defined over sets of a baseFinGroupType, *) (* so it can be used, e.g., to pick right cosets.) *) @@ -92,20 +96,23 @@ Require Import choice fintype div path bigop prime finset. (* classes G == the set of all conjugate classes of G. *) (* G :^: H == {G :^ x | x \in H}. *) (* class_support G H == {x ^ y | x \in G, y \in H}. *) -(* [~: H1, ..., Hn] == commutator subgroup of H1, ..., Hn. *) -(*{in G, centralised H} <-> G centralises H. *) +(* commg_set G H == {[~ x, y] | x \in G, y \in H}; NOT the commutator! *) +(* <<H>> == the subgroup generated by the set H. *) +(* [~: G, H] == the commmutator subgroup of G and H, i.e., *) +(* <<commg_set G H>>>. *) +(* [~: H1, ..., Hn] == commutator subgroup of H1, ..., Hn (left assoc.). *) +(* H <*> G == the subgroup generated by sets H and G (H join G). *) +(* (H * G)%G == the join of G H : {group gT} (convertible, but not *) +(* identical to (G <*> H)%G). *) +(* (\prod_(i ...) H i)%G == the group generated by the H i. *) +(* {in G, centralised H} <-> G centralises H. *) (* {in G, normalised H} <-> G normalises H. *) (* <-> forall x, x \in G -> H :^ x = H. *) (* 'N(H) == the normaliser of H. *) (* 'N_G(H) == the normaliser of H in G. *) -(* H <| G <=> H is normal in G. *) +(* H <| G <=> H is a normal subgroup of G. *) (* 'C(H) == the centraliser of H. *) (* 'C_G(H) == the centraliser of H in G. *) -(* <<H>> == the subgroup generated by the set H. *) -(* H <*> G == the subgroup generated by sets H and G (H join G). *) -(* (H * G)%G == the join of G H : {group gT} (convertible, but not *) -(* identical to (G <*> H)%G). *) -(* (\prod_(i ...) H i)%G == the group generated by the H i. *) (* gcore H G == the largest subgroup of H normalised by G. *) (* If H is a subgroup of G, this is the largest *) (* normal subgroup of G contained in H). *) @@ -142,7 +149,7 @@ Delimit Scope group_scope with g. Delimit Scope Group_scope with G. (* This module can be imported to open the scope for group element *) -(* operations locally to a file, without exporing the Open to *) +(* operations locally to a file, without exporting the Open to *) (* clients of that file (as Open would do). *) Module GroupScope. Open Scope group_scope. @@ -412,7 +419,7 @@ Lemma invMg x y : (x * y)^-1 = y^-1 * x^-1. Proof. by case: T x y => ? []. Qed. Lemma invg_inj : @injective T T invg. Proof. exact: can_inj invgK. Qed. Lemma eq_invg_sym x y : (x^-1 == y) = (x == y^-1). -Proof. by exact: (inv_eq invgK). Qed. +Proof. by apply: (inv_eq invgK). Qed. Lemma invg1 : 1^-1 = 1 :> T. Proof. by apply: invg_inj; rewrite -{1}[1^-1]mul1g invMg invgK mul1g. Qed. @@ -467,11 +474,11 @@ Proof. by move=> cxy cxz; rewrite /commute -mulgA -cxz !mulgA cxy. Qed. Lemma commuteX x y n : commute x y -> commute x (y ^+ n). Proof. -move=> cxy; elim: n => [|n]; [exact: commute1 | rewrite expgS; exact: commuteM]. +by move=> cxy; case: n; [apply: commute1 | elim=> // n; apply: commuteM]. Qed. Lemma commuteX2 x y m n : commute x y -> commute (x ^+ m) (y ^+ n). -Proof. by move=> cxy; exact/commuteX/commute_sym/commuteX. Qed. +Proof. by move=> cxy; apply/commuteX/commute_sym/commuteX. Qed. Lemma expgVn x n : x^-1 ^+ n = x ^- n. Proof. by elim: n => [|n IHn]; rewrite ?invg1 // expgSr expgS invMg IHn. Qed. @@ -507,7 +514,7 @@ Lemma mulKVg : rev_left_loop invg mulgT. Proof. by move=> x y; rewrite mulgA mulgV mul1g. Qed. Lemma mulgI : right_injective mulgT. -Proof. move=> x; exact: can_inj (mulKg x). Qed. +Proof. by move=> x; apply: can_inj (mulKg x). Qed. Lemma mulgK : right_loop invg mulgT. Proof. by move=> x y; rewrite -mulgA mulgV mulg1. Qed. @@ -516,7 +523,7 @@ Lemma mulgKV : rev_right_loop invg mulgT. Proof. by move=> x y; rewrite -mulgA mulVg mulg1. Qed. Lemma mulIg : left_injective mulgT. -Proof. move=> x; exact: can_inj (mulgK x). Qed. +Proof. by move=> x; apply: can_inj (mulgK x). Qed. Lemma eq_invg_mul x y : (x^-1 == y :> T) = (x * y == 1 :> T). Proof. by rewrite -(inj_eq (@mulgI x)) mulgV eq_sym. Qed. @@ -566,7 +573,7 @@ Lemma conjgKV : @rev_right_loop T T invg conjg. Proof. by move=> y x; rewrite -conjgM mulVg conjg1. Qed. Lemma conjg_inj : @left_injective T T T conjg. -Proof. move=> y; exact: can_inj (conjgK y). Qed. +Proof. by move=> y; apply: can_inj (conjgK y). Qed. Lemma conjg_eq1 x y : (x ^ y == 1) = (x == 1). Proof. by rewrite -(inj_eq (@conjg_inj y) x) conj1g. Qed. @@ -595,10 +602,10 @@ Lemma invg_comm x y : [~ x, y]^-1 = [~ y, x]. Proof. by rewrite commgEr conjVg invMg invgK. Qed. Lemma commgP x y : reflect (commute x y) ([~ x, y] == 1 :> T). -Proof. by rewrite [[~ x, y]]mulgA -invMg -eq_mulVg1 eq_sym; exact: eqP. Qed. +Proof. by rewrite [[~ x, y]]mulgA -invMg -eq_mulVg1 eq_sym; apply: eqP. Qed. Lemma conjg_fixP x y : reflect (x ^ y = x) ([~ x, y] == 1 :> T). -Proof. by rewrite -eq_mulVg1 eq_sym; exact: eqP. Qed. +Proof. by rewrite -eq_mulVg1 eq_sym; apply: eqP. Qed. Lemma commg1_sym x y : ([~ x, y] == 1 :> T) = ([~ y, x] == 1 :> T). Proof. by rewrite -invg_comm (inv_eq invgK) invg1. Qed. @@ -610,13 +617,13 @@ Lemma comm1g x : [~ 1, x] = 1. Proof. by rewrite -invg_comm commg1 invg1. Qed. Lemma commgg x : [~ x, x] = 1. -Proof. by exact/eqP/commgP. Qed. +Proof. exact/eqP/commgP. Qed. Lemma commgXg x n : [~ x, x ^+ n] = 1. Proof. exact/eqP/commgP/commuteX. Qed. Lemma commgVg x : [~ x, x^-1] = 1. -Proof. by exact/eqP/commgP/commuteV. Qed. +Proof. exact/eqP/commgP/commuteV. Qed. Lemma commgXVg x n : [~ x, x ^- n] = 1. Proof. exact/eqP/commgP/commuteV/commuteX. Qed. @@ -654,7 +661,7 @@ by rewrite /repr; case: ifP => // _; case: pickP => // A0; rewrite [x \in A]A0. Qed. Lemma card_mem_repr A : #|A| > 0 -> repr A \in A. -Proof. by rewrite lt0n => /existsP[x]; exact: mem_repr. Qed. +Proof. by rewrite lt0n => /existsP[x]; apply: mem_repr. Qed. Lemma repr_set1 x : repr [set x] = x. Proof. by apply/set1P/card_mem_repr; rewrite cards1. Qed. @@ -818,8 +825,8 @@ Notation "#| B : A |" := (indexg B A) : group_scope. (* No notation for lcoset and rcoset, which are to be used mostly *) (* in curried form; x *: B and A :* 1 denote singleton products, *) -(* so thus we can use mulgA, mulg1, etc, on, say, A :* 1 * B :* x. *) -(* No notation for the set commutator generator set set_commg. *) +(* so we can use mulgA, mulg1, etc, on, say, A :* 1 * B :* x. *) +(* No notation for the set commutator generator set commg_set. *) Notation "''N' ( A )" := (normaliser A) : group_scope. Notation "''N_' G ( A )" := (G%g :&: 'N(A)) : group_scope. @@ -856,24 +863,20 @@ Lemma prodsgP (I : finType) (P : pred I) (A : I -> {set gT}) x : (x \in \prod_(i | P i) A i). Proof. rewrite -big_filter filter_index_enum; set r := enum P. -pose inA c i := c i \in A i; set RHS := x \in _. -suffices IHr: reflect (exists2 c, all (inA c) r & x = \prod_(i <- r) c i) RHS. - apply: (iffP IHr) => [][c inAc ->]. - rewrite -[r]filter_index_enum big_filter; exists c => // i Pi. - by apply: (allP inAc); rewrite mem_enum. - rewrite -big_filter filter_index_enum; exists c => //; apply/allP=> i. - rewrite mem_enum; exact: inAc. -have: uniq r by [rewrite enum_uniq]; rewrite {}/RHS. -elim: {P}r x => /= [x _|i r IHr x /andP[r'i Ur]]. - by rewrite unlock; apply: (iffP set1P) => [-> | [] //]; exists (fun _ => 1). +pose inA c := all (fun i => c i \in A i); set piAx := x \in _. +suffices IHr: reflect (exists2 c, inA c r & x = \prod_(i <- r) c i) piAx. + apply: (iffP IHr) => -[c inAc ->]; do [exists c; last by rewrite big_filter]. + by move=> i Pi; rewrite (allP inAc) ?mem_enum. + by apply/allP=> i; rewrite mem_enum => /inAc. +have: uniq r by rewrite enum_uniq. +elim: {P}r x @piAx => /= [x _ | i r IHr x /andP[r'i /IHr{IHr}IHr]]. + by rewrite unlock; apply: (iffP set1P) => [-> | [] //]; exists (fun=> x). rewrite big_cons; apply: (iffP idP) => [|[c /andP[Aci Ac] ->]]; last first. by rewrite big_cons mem_mulg //; apply/IHr=> //; exists c. -case/mulsgP=> y _ Ai_y /IHr[//| c Ac ->] ->{x}. -exists [eta c with i |-> y] => /=. - rewrite /inA /= eqxx Ai_y; apply/allP=> j rj. - by case: eqP rj r'i => [-> -> // | _ rj _]; exact: (allP Ac). -rewrite big_cons eqxx !big_seq; congr (_ * _). -by apply: eq_bigr => j rj; case: eqP rj r'i => // -> ->. +case/mulsgP=> c_i _ Ac_i /IHr[c /allP-inAcr ->] ->{x}. +exists [eta c with i |-> c_i]; rewrite /= ?big_cons eqxx ?Ac_i. + by apply/allP=> j rj; rewrite /= ifN ?(memPn r'i) ?inAcr. +by congr (_ * _); apply: eq_big_seq => j rj; rewrite ifN ?(memPn r'i). Qed. Lemma mem_prodg (I : finType) (P : pred I) (A : I -> {set gT}) c : @@ -925,7 +928,7 @@ Lemma memV_invg x A : (x^-1 \in A^-1) = (x \in A). Proof. by rewrite inE invgK. Qed. Lemma card_invg A : #|A^-1| = #|A|. -Proof. by apply: card_preimset; exact: invg_inj. Qed. +Proof. exact/card_preimset/invg_inj. Qed. (* Product with singletons. *) @@ -938,7 +941,7 @@ Lemma mulg_set1 x y : [set x] :* y = [set x * y]. Proof. by rewrite [_ * _]imset2_set1l imset_set1. Qed. Lemma invg_set1 x : [set x]^-1 = [set x^-1]. -Proof. by apply/setP=> y; rewrite !inE inv_eq //; exact: invgK. Qed. +Proof. by apply/setP=> y; rewrite !inE inv_eq //; apply: invgK. Qed. End BaseSetMulProp. @@ -964,7 +967,7 @@ Lemma mem_lcoset A x y : (y \in x *: A) = (x^-1 * y \in A). Proof. by rewrite -lcosetE [_ x](can_imset_pre _ (mulKg _)) inE. Qed. Lemma lcosetP A x y : reflect (exists2 a, a \in A & y = x * a) (y \in x *: A). -Proof. rewrite -lcosetE; exact: imsetP. Qed. +Proof. by rewrite -lcosetE; apply: imsetP. Qed. Lemma lcosetsP A B C : reflect (exists2 x, x \in B & C = x *: A) (C \in lcosets A B). @@ -983,7 +986,7 @@ Lemma lcosetKV : rev_left_loop invg (fun x A => x *: A). Proof. by move=> x A; rewrite -lcosetM mulgV mul1g. Qed. Lemma lcoset_inj : right_injective (fun x A => x *: A). -Proof. by move=> x; exact: can_inj (lcosetK x). Qed. +Proof. by move=> x; apply: can_inj (lcosetK x). Qed. Lemma lcosetS x A B : (x *: A \subset x *: B) = (A \subset B). Proof. @@ -1009,7 +1012,7 @@ Lemma mem_rcoset A x y : (y \in A :* x) = (y * x^-1 \in A). Proof. by rewrite -rcosetE [_ x](can_imset_pre A (mulgK _)) inE. Qed. Lemma rcosetP A x y : reflect (exists2 a, a \in A & y = a * x) (y \in A :* x). -Proof. by rewrite -rcosetE; exact: imsetP. Qed. +Proof. by rewrite -rcosetE; apply: imsetP. Qed. Lemma rcosetsP A B C : reflect (exists2 x, x \in B & C = A :* x) (C \in rcosets A B). @@ -1028,7 +1031,7 @@ Lemma rcosetKV : rev_right_loop invg (fun A x => A :* x). Proof. by move=> x A; rewrite -rcosetM mulVg mulg1. Qed. Lemma rcoset_inj : left_injective (fun A x => A :* x). -Proof. by move=> x; exact: can_inj (rcosetK x). Qed. +Proof. by move=> x; apply: can_inj (rcosetK x). Qed. Lemma rcosetS x A B : (A :* x \subset B :* x) = (A \subset B). Proof. @@ -1042,14 +1045,11 @@ Proof. by rewrite -(rcosetS x^-1) rcosetK. Qed. Lemma sub_rcosetV x A B : (A \subset B :* x^-1) = (A :* x \subset B). Proof. by rewrite sub_rcoset invgK. Qed. -(* Inverse map lcosets to rcosets *) - -Lemma lcosets_invg A B : lcosets A^-1 B^-1 = invg @^-1: rcosets A B. +(* Inverse maps lcosets to rcosets *) +Lemma invg_lcosets A B : (lcosets A B)^-1 = rcosets A^-1 B^-1. Proof. -apply/setP=> C; rewrite inE. -apply/imsetP/imsetP=> [] [a]; rewrite -memV_invg ?invgK => Aa; - try move/(canRL invgK); move->; exists a^-1; - by rewrite // lcosetE rcosetE invMg invg_set1 ?invgK. +rewrite /A^-1/= - -[RHS]imset_comp -imset_comp. +by apply: eq_imset => x /=; rewrite lcosetE rcosetE invMg invg_set1. Qed. (* Conjugates. *) @@ -1082,7 +1082,7 @@ Lemma conjsgKV : @rev_right_loop _ gT invg conjugate. Proof. by move=> x A; rewrite -conjsgM mulVg conjsg1. Qed. Lemma conjsg_inj : @left_injective _ gT _ conjugate. -Proof. by move=> x; exact: can_inj (conjsgK x). Qed. +Proof. by move=> x; apply: can_inj (conjsgK x). Qed. Lemma cardJg A x : #|A :^ x| = #|A|. Proof. by rewrite (card_imset _ (conjg_inj x)). Qed. @@ -1232,7 +1232,7 @@ Lemma group_setP A : Proof. apply: (iffP andP) => [] [A1 AM]; split=> {A1}//. by move=> x y Ax Ay; apply: (subsetP AM); rewrite mem_mulg. -apply/subsetP=> _ /mulsgP[x y Ax Ay ->]; exact: AM. +by apply/subsetP=> _ /mulsgP[x y Ax Ay ->]; apply: AM. Qed. Structure group_type : Type := Group { @@ -1402,14 +1402,14 @@ Definition cardG_gt0_reduced : 0 < card (@mem gT (predPredType gT) G) Lemma indexg_gt0 A : 0 < #|G : A|. Proof. rewrite lt0n; apply/existsP; exists A. -rewrite -{2}[A]mulg1 -rcosetE; exact: mem_imset. +by rewrite -{2}[A]mulg1 -rcosetE; apply: mem_imset. Qed. Lemma trivgP : reflect (G :=: 1) (G \subset [1]). -Proof. by rewrite subG1; exact: eqP. Qed. +Proof. by rewrite subG1; apply: eqP. Qed. Lemma trivGP : reflect (G = 1%G) (G \subset [1]). -Proof. by rewrite subG1; exact: eqP. Qed. +Proof. by rewrite subG1; apply: eqP. Qed. Lemma proper1G : ([1] \proper G) = (G :!=: 1). Proof. by rewrite properEneq sub1G andbT eq_sym. Qed. @@ -1481,12 +1481,12 @@ Lemma groupVl x : x^-1 \in G -> x \in G. Proof. by move/groupVr; rewrite invgK. Qed. Lemma groupV x : (x^-1 \in G) = (x \in G). -Proof. by apply/idP/idP; [exact: groupVl | exact: groupVr]. Qed. +Proof. by apply/idP/idP; [apply: groupVl | apply: groupVr]. Qed. Lemma groupMl x y : x \in G -> (x * y \in G) = (y \in G). Proof. -move=> Gx; apply/idP/idP=> Gy; last exact: groupM. -rewrite -(mulKg x y); exact: groupM (groupVr _) _. +move=> Gx; apply/idP/idP=> [Gxy|]; last exact: groupM. +by rewrite -(mulKg x y) groupM ?groupVr. Qed. Lemma groupMr x y : x \in G -> (y * x \in G) = (y \in G). @@ -1505,7 +1505,7 @@ Proof. by move=> Gx Gy; rewrite !in_group. Qed. Lemma group_prod I r (P : pred I) F : (forall i, P i -> F i \in G) -> \prod_(i <- r | P i) F i \in G. -Proof. by move=> G_P; elim/big_ind: _ => //; exact: groupM. Qed. +Proof. by move=> G_P; elim/big_ind: _ => //; apply: groupM. Qed. (* Inverse is an anti-morphism. *) @@ -1549,20 +1549,20 @@ Proof. by rewrite mem_lcoset mulVg group1. Qed. Lemma lcoset_sym x y : (x \in y *: G) = (y \in x *: G). Proof. by rewrite !mem_lcoset -groupV invMg invgK. Qed. -Lemma lcoset_transl x y : x \in y *: G -> x *: G = y *: G. +Lemma lcoset_eqP {x y} : reflect (x *: G = y *: G) (x \in y *: G). Proof. -move=> Gyx; apply/setP=> u; rewrite !mem_lcoset in Gyx *. -by rewrite -{2}(mulKVg x u) mulgA (groupMl _ Gyx). +suffices <-: (x *: G == y *: G) = (x \in y *: G) by apply: eqP. +by rewrite eqEsubset !mulSG !sub1set lcoset_sym andbb. Qed. -Lemma lcoset_transr x y z : x \in y *: G -> (x \in z *: G) = (y \in z *: G). -Proof. by move=> Gyx; rewrite -2!(lcoset_sym z) (lcoset_transl Gyx). Qed. +Lemma lcoset_transl x y z : x \in y *: G -> (x \in z *: G) = (y \in z *: G). +Proof. by move=> Gyx; rewrite -2!(lcoset_sym z) (lcoset_eqP Gyx). Qed. Lemma lcoset_trans x y z : x \in y *: G -> y \in z *: G -> x \in z *: G. -Proof. by move/lcoset_transr->. Qed. +Proof. by move/lcoset_transl->. Qed. Lemma lcoset_id x : x \in G -> x *: G = G. -Proof. rewrite -{-2}(mul1g G); exact: lcoset_transl. Qed. +Proof. by move=> Gx; rewrite (lcoset_eqP (_ : x \in 1 *: G)) mul1g. Qed. (* Right cosets, with an elimination form for repr. *) @@ -1572,20 +1572,20 @@ Proof. by rewrite mem_rcoset mulgV group1. Qed. Lemma rcoset_sym x y : (x \in G :* y) = (y \in G :* x). Proof. by rewrite -!memV_lcosetV lcoset_sym. Qed. -Lemma rcoset_transl x y : x \in G :* y -> G :* x = G :* y. +Lemma rcoset_eqP {x y} : reflect (G :* x = G :* y) (x \in G :* y). Proof. -move=> Gyx; apply: invg_inj; rewrite !invg_rcoset. -by apply: lcoset_transl; rewrite memV_lcosetV. +suffices <-: (G :* x == G :* y) = (x \in G :* y) by apply: eqP. +by rewrite eqEsubset !mulGS !sub1set rcoset_sym andbb. Qed. -Lemma rcoset_transr x y z : x \in G :* y -> (x \in G :* z) = (y \in G :* z). -Proof. by move=> Gyx; rewrite -2!(rcoset_sym z) (rcoset_transl Gyx). Qed. +Lemma rcoset_transl x y z : x \in G :* y -> (x \in G :* z) = (y \in G :* z). +Proof. by move=> Gyx; rewrite -2!(rcoset_sym z) (rcoset_eqP Gyx). Qed. -Lemma rcoset_trans x y z : y \in G :* x -> z \in G :* y -> z \in G :* x. +Lemma rcoset_trans x y z : x \in G :* y -> y \in G :* z -> x \in G :* z. Proof. by move/rcoset_transl->. Qed. Lemma rcoset_id x : x \in G -> G :* x = G. -Proof. by rewrite -{-2}(mulg1 G); exact: rcoset_transl. Qed. +Proof. by move=> Gx; rewrite (rcoset_eqP (_ : x \in G :* 1)) mulg1. Qed. (* Elimination form. *) @@ -1603,22 +1603,20 @@ by rewrite -[repr _](mulgKV x); split; rewrite -mem_rcoset mem_repr_rcoset. Qed. Lemma rcoset_repr x : G :* (repr (G :* x)) = G :* x. -Proof. by apply: rcoset_transl; exact: mem_repr (rcoset_refl x). Qed. +Proof. exact/rcoset_eqP/mem_repr_rcoset. Qed. (* Coset spaces. *) -Lemma mem_lcosets A x : (x *: G \in lcosets G A) = (x \in A * G). +Lemma mem_rcosets A x : (G :* x \in rcosets G A) = (x \in G * A). Proof. -apply/imsetP/mulsgP=> [[a Aa eqxaG] | [a g Aa Gg ->{x}]]. - exists a (a^-1 * x); rewrite ?mulKVg //. - by rewrite -mem_lcoset -lcosetE -eqxaG lcoset_refl. -by exists a; rewrite // lcosetM lcosetE lcoset_id. +apply/rcosetsP/mulsgP=> [[a Aa /rcoset_eqP/rcosetP[g]] | ]; first by exists g a. +by case=> g a Gg Aa ->{x}; exists a; rewrite // rcosetM rcoset_id. Qed. -Lemma mem_rcosets A x : (G :* x \in rcosets G A) = (x \in G * A). +Lemma mem_lcosets A x : (x *: G \in lcosets G A) = (x \in A * G). Proof. -rewrite -memV_invg invMg invGid -mem_lcosets. -by rewrite -{4}invGid lcosets_invg inE invg_lcoset invgK. +rewrite -[LHS]memV_invg invg_lcoset invg_lcosets. +by rewrite -[RHS]memV_invg invMg invGid mem_rcosets. Qed. (* Conjugates. *) @@ -1650,17 +1648,19 @@ Lemma classGidr x : {in G, normalised (x ^: G)}. Proof. by move=> y Gy /=; rewrite -class_rcoset rcoset_id. Qed. Lemma class_refl x : x \in x ^: G. -Proof. by apply/imsetP; exists (1 : gT); rewrite ?conjg1. Qed. +Proof. by apply/imsetP; exists 1; rewrite ?conjg1. Qed. Hint Resolve class_refl. -Lemma class_transr x y : x \in y ^: G -> x ^: G = y ^: G. -Proof. by case/imsetP=> z Gz ->; rewrite classGidl. Qed. +Lemma class_eqP x y : reflect (x ^: G = y ^: G) (x \in y ^: G). +Proof. +by apply: (iffP idP) => [/imsetP[z Gz ->] | <-]; rewrite ?class_refl ?classGidl. +Qed. Lemma class_sym x y : (x \in y ^: G) = (y \in x ^: G). -Proof. by apply/idP/idP=> /class_transr->. Qed. +Proof. by apply/idP/idP=> /class_eqP->. Qed. Lemma class_transl x y z : x \in y ^: G -> (x \in z ^: G) = (y \in z ^: G). -Proof. by rewrite -!(class_sym z) => /class_transr->. Qed. +Proof. by rewrite -!(class_sym z) => /class_eqP->. Qed. Lemma class_trans x y z : x \in y ^: G -> y \in z ^: G -> x \in z ^: G. Proof. by move/class_transl->. Qed. @@ -1668,7 +1668,7 @@ Proof. by move/class_transl->. Qed. Lemma repr_class x : {y | y \in G & repr (x ^: G) = x ^ y}. Proof. set z := repr _; have: #|[set y in G | z == x ^ y]| > 0. - have: z \in x ^: G by exact: (mem_repr x). + have: z \in x ^: G by apply: (mem_repr x). by case/imsetP=> y Gy ->; rewrite (cardD1 y) inE Gy eqxx. by move/card_mem_repr; move: (repr _) => y /setIdP[Gy /eqP]; exists y. Qed. @@ -1693,7 +1693,7 @@ by have [y Gy ->] := repr_class x; rewrite classGidl ?groupJ. Qed. Lemma mem_repr_classes xG : xG \in classes G -> repr xG \in xG. -Proof. by case/repr_classesP=> _ {2}->; exact: class_refl. Qed. +Proof. by case/repr_classesP=> _ {2}->; apply: class_refl. Qed. Lemma classes_gt0 : 0 < #|classes G|. Proof. by rewrite (cardsD1 1) classes1. Qed. @@ -1703,7 +1703,7 @@ Proof. rewrite (cardsD1 1) classes1 ltnS lt0n cards_eq0. apply/set0Pn/trivgPn=> [[xG /setD1P[nt_xG]] | [x Gx ntx]]. by case/imsetP=> x Gx def_xG; rewrite def_xG classG_eq1 in nt_xG; exists x. -by exists (x ^: G); rewrite !inE classG_eq1 ntx; exact: mem_imset. +by exists (x ^: G); rewrite !inE classG_eq1 ntx; apply: mem_imset. Qed. Lemma mem_class_support A x : x \in A -> x \in class_support A G. @@ -1722,7 +1722,7 @@ Qed. Lemma class_support_subG A : A \subset G -> class_support A G \subset G. Proof. -by move=> sAG; rewrite class_supportEr; apply/bigcupsP=> x Gx; exact: conj_subG. +by move=> sAG; rewrite class_supportEr; apply/bigcupsP=> x Gx; apply: conj_subG. Qed. Lemma sub_class_support A : A \subset class_support A G. @@ -1768,12 +1768,12 @@ Definition subg_one := Subg group1. Definition subg_inv u := Subg (groupVr (subgP u)). Definition subg_mul u v := Subg (groupM (subgP u) (subgP v)). Lemma subg_oneP : left_id subg_one subg_mul. -Proof. move=> u; apply: val_inj; exact: mul1g. Qed. +Proof. by move=> u; apply: val_inj; apply: mul1g. Qed. Lemma subg_invP : left_inverse subg_one subg_inv subg_mul. -Proof. move=> u; apply: val_inj; exact: mulVg. Qed. +Proof. by move=> u; apply: val_inj; apply: mulVg. Qed. Lemma subg_mulP : associative subg_mul. -Proof. move=> u v w; apply: val_inj; exact: mulgA. Qed. +Proof. by move=> u v w; apply: val_inj; apply: mulgA. Qed. Definition subFinGroupMixin := FinGroup.Mixin subg_mulP subg_oneP subg_invP. Canonical subBaseFinGroupType := @@ -1788,7 +1788,7 @@ Definition subg : gT -> subg_of := insubd (1 : subg_of). Lemma subgK x : x \in G -> val (subg x) = x. Proof. by move=> Gx; rewrite insubdK. Qed. Lemma sgvalK : cancel sgval subg. -Proof. case=> x Gx; apply: val_inj; exact: subgK. Qed. +Proof. by case=> x Gx; apply: val_inj; apply: subgK. Qed. Lemma subg_default x : (x \in G) = false -> val (subg x) = 1. Proof. by move=> Gx; rewrite val_insubd Gx. Qed. Lemma subgM : {in G &, {morph subg : x y / x * y}}. @@ -1811,22 +1811,20 @@ Lemma mulGSid G H : H \subset G -> G * H = G. Proof. exact: mulGSgid (group1 H). Qed. Lemma mulGidPl G H : reflect (G * H = G) (H \subset G). -Proof. by apply: (iffP idP) => [|<-]; [exact: mulGSid | exact: mulG_subr]. Qed. +Proof. by apply: (iffP idP) => [|<-]; [apply: mulGSid | apply: mulG_subr]. Qed. Lemma mulGidPr G H : reflect (G * H = H) (G \subset H). -Proof. by apply: (iffP idP) => [|<-]; [exact: mulSGid | exact: mulG_subl]. Qed. +Proof. by apply: (iffP idP) => [|<-]; [apply: mulSGid | apply: mulG_subl]. Qed. Lemma comm_group_setP G H : reflect (commute G H) (group_set (G * H)). Proof. rewrite /group_set (subsetP (mulG_subl _ _)) ?group1 // andbC. have <-: #|G * H| <= #|H * G| by rewrite -invMG card_invg. -rewrite -mulgA mulGS mulgA mulSG -eqEcard eq_sym; exact: eqP. +by rewrite -mulgA mulGS mulgA mulSG -eqEcard eq_sym; apply: eqP. Qed. Lemma card_lcosets G H : #|lcosets H G| = #|G : H|. -Proof. -by rewrite -[#|G : H|](card_preimset _ invg_inj) -lcosets_invg !invGid. -Qed. +Proof. by rewrite -card_invg invg_lcosets !invGid. Qed. (* Group Modularity equations *) @@ -1861,11 +1859,14 @@ Bind Scope group_scope with subg_of. Implicit Arguments trivgP [gT G]. Implicit Arguments trivGP [gT G]. +Implicit Arguments lcoset_eqP [gT G x y]. +Implicit Arguments rcoset_eqP [gT G x y]. Implicit Arguments mulGidPl [gT G H]. Implicit Arguments mulGidPr [gT G H]. Implicit Arguments comm_group_setP [gT G H]. +Implicit Arguments class_eqP [gT G x y]. Implicit Arguments repr_classesP [gT G xG]. -Prenex Implicits trivgP trivGP comm_group_setP. +Prenex Implicits trivgP trivGP lcoset_eqP rcoset_eqP comm_group_setP class_eqP. Section GroupInter. @@ -1887,8 +1888,7 @@ Variables (I : finType) (P : pred I) (F : I -> {group gT}). Lemma group_set_bigcap : group_set (\bigcap_(i | P i) F i). Proof. -elim/big_rec: _ => [|i G _ gG]; first exact: groupP. -exact: group_setI (Group gG). +by elim/big_rec: _ => [|i G _ gG]; rewrite -1?(insubdK 1%G gG) groupP. Qed. Canonical bigcap_group := group group_set_bigcap. @@ -1901,6 +1901,10 @@ Canonical commutator_group A B : {group _} := Eval hnf in [group of [~: A, B]]. Canonical joing_group A B : {group _} := Eval hnf in [group of A <*> B]. Canonical cycle_group x : {group _} := Eval hnf in [group of <[x]>]. +Definition joinG G H := joing_group G H. + +Definition subgroups A := [set G : {group gT} | G \subset A]. + Lemma order_gt0 (x : gT) : 0 < #[x]. Proof. exact: cardG_gt0. Qed. @@ -1908,13 +1912,9 @@ End GroupInter. Hint Resolve order_gt0. -Definition joinG (gT : finGroupType) (G H : {group gT}) := joing_group G H. - -Definition subgroups (gT : finGroupType) (G : {set gT}) := - [set H : {group gT} | H \subset G]. - Arguments Scope generated_group [_ group_scope]. Arguments Scope joing_group [_ group_scope group_scope]. +Arguments Scope subgroups [_ group_scope]. Notation "G :&: H" := (setI_group G H) : Group_scope. Notation "<< A >>" := (generated_group A) : Group_scope. @@ -1923,7 +1923,7 @@ Notation "[ ~: A1 , A2 , .. , An ]" := (commutator_group .. (commutator_group A1 A2) .. An) : Group_scope. Notation "A <*> B" := (joing_group A B) : Group_scope. Notation "G * H" := (joinG G H) : Group_scope. -Prenex Implicits joinG. +Prenex Implicits joinG subgroups. Notation "\prod_ ( i <- r | P ) F" := (\big[joinG/1%G]_(i <- r | P%B) F%G) : Group_scope. @@ -1960,8 +1960,7 @@ Proof. rewrite -[#|G|]sum1_card (partition_big_imset (rcoset H)) /=. rewrite mulnC -sum_nat_const; apply: eq_bigr => _ /rcosetsP[x Gx ->]. rewrite -(card_rcoset _ x) -sum1_card; apply: eq_bigl => y. -rewrite rcosetE eqEcard mulGS !card_rcoset leqnn andbT. -by rewrite group_modr sub1set // inE. +by rewrite rcosetE (sameP eqP rcoset_eqP) group_modr (sub1set, inE). Qed. Lemma divgI G H : #|G| %/ #|G :&: H| = #|G : H|. @@ -1995,10 +1994,10 @@ Lemma divg_indexS G H : H \subset G -> #|G| %/ #|G : H| = #|H|. Proof. by move/Lagrange <-; rewrite mulnK. Qed. Lemma coprimeSg G H p : H \subset G -> coprime #|G| p -> coprime #|H| p. -Proof. by move=> sHG; exact: coprime_dvdl (cardSg sHG). Qed. +Proof. by move=> sHG; apply: coprime_dvdl (cardSg sHG). Qed. Lemma coprimegS G H p : H \subset G -> coprime p #|G| -> coprime p #|H|. -Proof. by move=> sHG; exact: coprime_dvdr (cardSg sHG). Qed. +Proof. by move=> sHG; apply: coprime_dvdr (cardSg sHG). Qed. Lemma indexJg G H x : #|G :^ x : H :^ x| = #|G : H|. Proof. by rewrite -!divgI -conjIg !cardJg. Qed. @@ -2051,26 +2050,23 @@ Proof. by rewrite -divgS ?sub1G // cards1 divn1. Qed. Lemma indexMg G A : #|G * A : G| = #|A : G|. Proof. -congr #|(_ : {set _})|; apply/eqP; rewrite eqEsubset andbC imsetS ?mulG_subr //. -by apply/subsetP=> _ /imsetP[x GAx ->]; rewrite rcosetE mem_rcosets. +apply/eq_card/setP/eqP; rewrite eqEsubset andbC imsetS ?mulG_subr //. +by apply/subsetP=> _ /rcosetsP[x GAx ->]; rewrite mem_rcosets. Qed. Lemma rcosets_partition_mul G H : partition (rcosets H G) (H * G). Proof. -have eqiR: {in H * G & &, equivalence_rel [rel x y | y \in rcoset H x]}. - by move=> *; rewrite /= !rcosetE rcoset_refl; split=> // /rcoset_transl->. -congr (partition _ _): (equivalence_partitionP eqiR); apply/setP=> Hx. -apply/imsetP/idP=> [[x HGx defHx] | /rcosetsP[x Gx ->]]. - suffices ->: Hx = H :* x by rewrite mem_rcosets. - apply/setP=> y; rewrite defHx inE /= rcosetE andb_idl //. - by apply: subsetP y; rewrite mulGS sub1set. -exists (1 * x); rewrite ?mem_mulg // mul1g. -apply/setP=> y; rewrite inE /= rcosetE andb_idl //. -by apply: subsetP y; rewrite mulgS ?sub1set. +set HG := H * G; have sGHG: {subset G <= HG} by apply/subsetP/mulG_subr. +have defHx x: x \in HG -> [set y in HG | rcoset H x == rcoset H y] = H :* x. + move=> HGx; apply/setP=> y; rewrite inE !rcosetE (sameP eqP rcoset_eqP). + by rewrite rcoset_sym; apply/andb_idl/subsetP; rewrite mulGS sub1set. +have:= preim_partitionP (rcoset H) HG; congr (partition _ _); apply/setP=> Hx. +apply/imsetP/idP=> [[x HGx ->] | ]; first by rewrite defHx // mem_rcosets. +by case/rcosetsP=> x /sGHG-HGx ->; exists x; rewrite ?defHx. Qed. Lemma rcosets_partition G H : H \subset G -> partition (rcosets H G) G. -Proof. by move/mulSGid=> {2}<-; exact: rcosets_partition_mul. Qed. +Proof. by move=> sHG; have:= rcosets_partition_mul G H; rewrite mulSGid. Qed. Lemma LagrangeMl G H : (#|G| * #|H : G|)%N = #|G * H|. Proof. @@ -2111,9 +2107,9 @@ Qed. Lemma prime_TIg G H : prime #|G| -> ~~ (G \subset H) -> G :&: H = 1. Proof. -case/primeP=> _; move/(_ _ (cardSg (subsetIl G H))). -rewrite (sameP setIidPl eqP) eqEcard subsetIl -ltnNge ltn_neqAle -trivg_card1. -by case/predU1P=> ->. +case/primeP=> _ /(_ _ (cardSg (subsetIl G H))). +rewrite (sameP setIidPl eqP) eqEcard subsetIl => /pred2P[/card1_trivg|] //= ->. +by case/negP. Qed. Lemma prime_meetG G H : prime #|G| -> G :&: H != 1 -> G \subset H. @@ -2145,7 +2141,7 @@ Lemma subset_gen A : A \subset <<A>>. Proof. exact/bigcapsP. Qed. Lemma sub_gen A B : A \subset B -> A \subset <<B>>. -Proof. by move/subset_trans=> -> //; exact: subset_gen. Qed. +Proof. by move/subset_trans=> -> //; apply: subset_gen. Qed. Lemma mem_gen x A : x \in A -> x \in <<A>>. Proof. exact: subsetP (subset_gen A) x. Qed. @@ -2163,10 +2159,10 @@ Lemma genGid G : <<G>> = G. Proof. by apply/eqP; rewrite eqEsubset gen_subG subset_gen andbT. Qed. Lemma genGidG G : <<G>>%G = G. -Proof. by apply: val_inj; exact: genGid. Qed. +Proof. by apply: val_inj; apply: genGid. Qed. Lemma gen_set_id A : group_set A -> <<A>> = A. -Proof. by move=> gA; exact: (genGid (group gA)). Qed. +Proof. by move=> gA; apply: (genGid (group gA)). Qed. Lemma genS A B : A \subset B -> <<A>> \subset <<B>>. Proof. by move=> sAB; rewrite gen_subG sub_gen. Qed. @@ -2279,19 +2275,19 @@ by rewrite joingE (setUidPl sHG) genGid. Qed. Lemma joing_idPr A G : reflect (A <*> G = G) (A \subset G). -Proof. by rewrite joingC; exact: joing_idPl. Qed. +Proof. by rewrite joingC; apply: joing_idPl. Qed. Lemma joing_subP A B G : reflect (A \subset G /\ B \subset G) (A <*> B \subset G). -Proof. by rewrite join_subG; exact: andP. Qed. +Proof. by rewrite join_subG; apply: andP. Qed. Lemma joing_sub A B C : A <*> B = C -> A \subset C /\ B \subset C. -Proof. by move <-; exact/joing_subP. Qed. +Proof. by move <-; apply/joing_subP. Qed. Lemma genDU A B C : A \subset C -> <<C :\: A>> = <<B>> -> <<A :|: B>> = <<C>>. Proof. move=> sAC; rewrite -joingE -joing_idr => <- {B}; rewrite joing_idr. -by congr <<_>>; rewrite setDE setUIr setUCr setIT; exact/setUidPr. +by congr <<_>>; rewrite setDE setUIr setUCr setIT; apply/setUidPr. Qed. Lemma joingA : associative joingT. @@ -2314,7 +2310,7 @@ Lemma mulG_subG G H K : (G * H \subset K) = (G \subset K) && (H \subset K). Proof. by rewrite -gen_subG genM_join join_subG. Qed. Lemma mulGsubP K H G : reflect (K \subset G /\ H \subset G) (K * H \subset G). -Proof. by rewrite mulG_subG; exact: andP. Qed. +Proof. by rewrite mulG_subG; apply: andP. Qed. Lemma mulG_sub K H A : K * H = A -> K \subset A /\ H \subset A. Proof. by move <-; rewrite mulG_subl mulG_subr. Qed. @@ -2326,20 +2322,20 @@ Qed. Lemma comm_joingE G H : commute G H -> G <*> H = G * H. Proof. -by move/comm_group_setP=> gGH; rewrite -genM_join; exact: (genGid (group gGH)). +by move/comm_group_setP=> gGH; rewrite -genM_join; apply: (genGid (group gGH)). Qed. Lemma joinGC : commutative joinGT. -Proof. by move=> G H; apply: val_inj; exact: joingC. Qed. +Proof. by move=> G H; apply: val_inj; apply: joingC. Qed. Lemma joinGA : associative joinGT. -Proof. by move=> G H K; apply: val_inj; exact: joingA. Qed. +Proof. by move=> G H K; apply: val_inj; apply: joingA. Qed. Lemma join1G : left_id 1%G joinGT. -Proof. by move=> G; apply: val_inj; exact: joing1G. Qed. +Proof. by move=> G; apply: val_inj; apply: joing1G. Qed. Lemma joinG1 : right_id 1%G joinGT. -Proof. by move=> G; apply: val_inj; exact: joingG1. Qed. +Proof. by move=> G; apply: val_inj; apply: joingG1. Qed. Canonical joinG_law := Monoid.Law joinGA join1G joinG1. Canonical joinG_abelaw := Monoid.ComLaw joinGC. @@ -2373,12 +2369,12 @@ Proof. by move=> sAB sCD; rewrite genS ?imset2S. Qed. Lemma der1_subG G : [~: G, G] \subset G. Proof. -by rewrite gen_subG; apply/subsetP=> _ /imset2P[x y Gx Gy ->]; exact: groupR. +by rewrite gen_subG; apply/subsetP=> _ /imset2P[x y Gx Gy ->]; apply: groupR. Qed. Lemma comm_subG A B G : A \subset G -> B \subset G -> [~: A, B] \subset G. Proof. -by move=> sAG sBG; apply: subset_trans (der1_subG G); exact: commgSS. +by move=> sAG sBG; apply: subset_trans (der1_subG G); apply: commgSS. Qed. Lemma commGC A B : [~: A, B] = [~: B, A]. @@ -2468,7 +2464,7 @@ Qed. Lemma cycleP x y : reflect (exists i, y = x ^+ i) (y \in <[x]>). Proof. -by apply: (iffP idP) => [/cyclePmin[i _]|[i ->]]; [exists i | exact: mem_cycle]. +by apply: (iffP idP) => [/cyclePmin[i _]|[i ->]]; [exists i | apply: mem_cycle]. Qed. Lemma expg_order x : x ^+ #[x] = 1. @@ -2497,7 +2493,7 @@ Lemma invg2id x : #[x] = 2 -> x^-1 = x. Proof. by move=> ox; rewrite invg_expg ox. Qed. Lemma cycleX x i : <[x ^+ i]> \subset <[x]>. -Proof. rewrite cycle_subG; exact: mem_cycle. Qed. +Proof. by rewrite cycle_subG; apply: mem_cycle. Qed. Lemma cycleV x : <[x^-1]> = <[x]>. Proof. @@ -2524,7 +2520,7 @@ Implicit Type G H K : {group gT}. Lemma normP x A : reflect (A :^ x = A) (x \in 'N(A)). Proof. -suffices ->: (x \in 'N(A)) = (A :^ x == A) by exact: eqP. +suffices ->: (x \in 'N(A)) = (A :^ x == A) by apply: eqP. by rewrite eqEcard cardJg leqnn andbT inE. Qed. Implicit Arguments normP [x A]. @@ -2540,7 +2536,7 @@ Canonical normaliser_group A := group (group_set_normaliser A). Lemma normsP A B : reflect {in A, normalised B} (A \subset 'N(B)). Proof. apply: (iffP subsetP) => nBA x Ax; last by rewrite inE nBA //. -by apply/normP; exact: nBA. +by apply/normP; apply: nBA. Qed. Implicit Arguments normsP [A B]. @@ -2560,13 +2556,13 @@ Lemma normCs A : 'N(~: A) = 'N(A). Proof. by apply/setP=> x; rewrite -groupV !inE conjCg setCS sub_conjg. Qed. Lemma normG G : G \subset 'N(G). -Proof. by apply/normsP; exact: conjGid. Qed. +Proof. by apply/normsP; apply: conjGid. Qed. Lemma normT : 'N([set: gT]) = [set: gT]. Proof. by apply/eqP; rewrite -subTset normG. Qed. Lemma normsG A G : A \subset G -> A \subset 'N(G). -Proof. move=> sAG; exact: subset_trans (normG G). Qed. +Proof. by move=> sAG; apply: subset_trans (normG G). Qed. Lemma normC A B : A \subset 'N(B) -> commute A B. Proof. @@ -2580,7 +2576,7 @@ Lemma norm_joinEl G H : G \subset 'N(H) -> G <*> H = G * H. Proof. by move/normC/comm_joingE. Qed. Lemma norm_joinEr G H : H \subset 'N(G) -> G <*> H = G * H. -Proof. by move/normC=> cHG; exact: comm_joingE. Qed. +Proof. by move/normC=> cHG; apply: comm_joingE. Qed. Lemma norm_rlcoset G x : x \in 'N(G) -> G :* x = x *: G. Proof. by rewrite -sub1set => /normC. Qed. @@ -2604,7 +2600,7 @@ Lemma norm_gen A : 'N(A) \subset 'N(<<A>>). Proof. by apply/normsP=> x Nx; rewrite -genJ (normP Nx). Qed. Lemma class_norm x G : G \subset 'N(x ^: G). -Proof. by apply/normsP=> y; exact: classGidr. Qed. +Proof. by apply/normsP=> y; apply: classGidr. Qed. Lemma class_normal x G : x \in G -> x ^: G <| G. Proof. by move=> Gx; rewrite /normal class_norm class_subG. Qed. @@ -2616,7 +2612,7 @@ by case/imsetP=> y Gy ->; rewrite memJ_norm ?(subsetP nAG). Qed. Lemma class_support_norm A G : G \subset 'N(class_support A G). -Proof. by apply/normsP; exact: class_supportGidr. Qed. +Proof. by apply/normsP; apply: class_supportGidr. Qed. Lemma class_support_sub_norm A B G : A \subset G -> B \subset 'N(G) -> class_support A B \subset G. @@ -2756,7 +2752,7 @@ Proof. by move=> nAG; rewrite /normal subsetIl normsI ?normG. Qed. Lemma normalGI G H A : H \subset G -> A <| G -> H :&: A <| H. Proof. -by move=> sHG /andP[_ nAG]; exact: norm_normalI (subset_trans sHG nAG). +by move=> sHG /andP[_ nAG]; apply: norm_normalI (subset_trans sHG nAG). Qed. Lemma normal_subnorm G H : (H <| 'N_G(H)) = (H \subset G). @@ -2860,16 +2856,16 @@ by apply/subsetP=> _ /imsetP[y Ay ->]; rewrite /conjg -cAx ?mulKg. Qed. Lemma cents_norm A B : A \subset 'C(B) -> A \subset 'N(B). -Proof. by move=> cAB; exact: subset_trans (cent_sub B). Qed. +Proof. by move=> cAB; apply: subset_trans (cent_sub B). Qed. Lemma centC A B : A \subset 'C(B) -> commute A B. -Proof. by move=> cAB; exact: normC (cents_norm cAB). Qed. +Proof. by move=> cAB; apply: normC (cents_norm cAB). Qed. Lemma cent_joinEl G H : G \subset 'C(H) -> G <*> H = G * H. -Proof. by move=> cGH; exact: norm_joinEl (cents_norm cGH). Qed. +Proof. by move=> cGH; apply: norm_joinEl (cents_norm cGH). Qed. Lemma cent_joinEr G H : H \subset 'C(G) -> G <*> H = G * H. -Proof. by move=> cGH; exact: norm_joinEr (cents_norm cGH). Qed. +Proof. by move=> cGH; apply: norm_joinEr (cents_norm cGH). Qed. Lemma centJ A x : 'C(A :^ x) = 'C(A) :^ x. Proof. @@ -2882,7 +2878,7 @@ Lemma cent_norm A : 'N(A) \subset 'N('C(A)). Proof. by apply/normsP=> x nCx; rewrite -centJ (normP nCx). Qed. Lemma norms_cent A B : A \subset 'N(B) -> A \subset 'N('C(B)). -Proof. move=> nBA; exact: subset_trans nBA (cent_norm B). Qed. +Proof. by move=> nBA; apply: subset_trans nBA (cent_norm B). Qed. Lemma cent_normal A : 'C(A) <| 'N(A). Proof. by rewrite /(_ <| _) cent_sub cent_norm. Qed. @@ -2891,11 +2887,11 @@ Lemma centS A B : B \subset A -> 'C(A) \subset 'C(B). Proof. by move=> sAB; rewrite centsC (subset_trans sAB) 1?centsC. Qed. Lemma centsS A B C : A \subset B -> C \subset 'C(B) -> C \subset 'C(A). -Proof. by move=> sAB cCB; exact: subset_trans cCB (centS sAB). Qed. +Proof. by move=> sAB cCB; apply: subset_trans cCB (centS sAB). Qed. Lemma centSS A B C D : A \subset C -> B \subset D -> C \subset 'C(D) -> A \subset 'C(B). -Proof. move=> sAC sBD cCD; exact: subset_trans (centsS sBD cCD). Qed. +Proof. by move=> sAC sBD cCD; apply: subset_trans (centsS sBD cCD). Qed. Lemma centI A B : 'C(A) <*> 'C(B) \subset 'C(A :&: B). Proof. by rewrite gen_subG subUset !centS ?(subsetIl, subsetIr). Qed. @@ -2916,7 +2912,7 @@ Lemma sub_cent1 A x : (A \subset 'C[x]) = (x \in 'C(A)). Proof. by rewrite -cent_cycle centsC cycle_subG. Qed. Lemma cents_cycle x y : commute x y -> <[x]> \subset 'C(<[y]>). -Proof. move=> cxy; rewrite cent_cycle cycle_subG; exact/cent1P. Qed. +Proof. by move=> cxy; rewrite cent_cycle cycle_subG; apply/cent1P. Qed. Lemma cycle_abelian x : abelian <[x]>. Proof. exact: cents_cycle. Qed. @@ -2939,7 +2935,7 @@ Lemma commG1P A B : reflect ([~: A, B] = 1) (A \subset 'C(B)). Proof. apply: (iffP (centsP A B)) => [cAB | cAB1 x Ax y By]. apply/trivgP; rewrite gen_subG; apply/subsetP=> _ /imset2P[x y Ax Ay ->]. - by rewrite inE; apply/commgP; exact: cAB. + by rewrite inE; apply/commgP; apply: cAB. by apply/commgP; rewrite -in_set1 -[[set 1]]cAB1 mem_commg. Qed. @@ -2948,7 +2944,7 @@ Lemma abelianE A : abelian A = (A \subset 'C(A)). Proof. by []. Qed. Lemma abelian1 : abelian [1 gT]. Proof. exact: sub1G. Qed. Lemma abelianS A B : A \subset B -> abelian B -> abelian A. -Proof. by move=> sAB; exact: centSS. Qed. +Proof. by move=> sAB; apply: centSS. Qed. Lemma abelianJ A x : abelian (A :^ x) = abelian A. Proof. by rewrite /abelian centJ conjSg. Qed. @@ -2976,7 +2972,7 @@ Lemma sub_abelian_cent : C \subset A -> A \subset 'C(C). Proof. by move=> sCA; rewrite centsC (subset_trans sCA). Qed. Lemma sub_abelian_cent2 : B \subset A -> C \subset A -> B \subset 'C(C). -Proof. by move=> sBA; move/sub_abelian_cent; exact: subset_trans. Qed. +Proof. by move=> sBA; move/sub_abelian_cent; apply: subset_trans. Qed. Lemma sub_abelian_norm : C \subset A -> A \subset 'N(C). Proof. by move=> sCA; rewrite cents_norm ?sub_abelian_cent. Qed. @@ -3048,7 +3044,7 @@ Lemma mingroupP : Proof. apply: (iffP minsetP); rewrite /= groupP genGidG /= => [] [-> minG]. by split=> // H gPH sGH; apply: minG; rewrite // groupP genGidG. -split=> // A; case/andP=> gA gPA; rewrite -(gen_set_id gA); exact: minG. +by split=> // A; case/andP=> gA gPA; rewrite -(gen_set_id gA); apply: minG. Qed. Lemma maxgroupP : @@ -3056,7 +3052,7 @@ Lemma maxgroupP : Proof. apply: (iffP maxsetP); rewrite /= groupP genGidG /= => [] [-> maxG]. by split=> // H gPH sGH; apply: maxG; rewrite // groupP genGidG. -split=> // A; case/andP=> gA gPA; rewrite -(gen_set_id gA); exact: maxG. +by split=> // A; case/andP=> gA gPA; rewrite -(gen_set_id gA); apply: maxG. Qed. Lemma maxgroupp : maxgroup G -> gP G. Proof. by case/maxgroupP. Qed. diff --git a/mathcomp/fingroup/gproduct.v b/mathcomp/fingroup/gproduct.v index 2706df6..39f19cc 100644 --- a/mathcomp/fingroup/gproduct.v +++ b/mathcomp/fingroup/gproduct.v @@ -1,10 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype bigop finset. -Require Import fingroup morphism quotient action. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div choice fintype. +From mathcomp +Require Import bigop finset fingroup morphism quotient action. (******************************************************************************) (* Partial, semidirect, central, and direct products. *) @@ -19,7 +18,7 @@ Require Import fingroup morphism quotient action. (* [splits G, over K] == [complements to K in G] is not empty. *) (* remgr A B x == the right remainder in B of x mod A, i.e., *) (* some element of (A :* x) :&: B. *) -(* divgr A B x == the "quotient" in B of x by A: for all x, *) +(* divgr A B x == the "division" in B of x by A: for all x, *) (* x = divgr A B x * remgr A B x. *) (* ++ External products : *) (* pairg1, pair1g == the isomorphisms aT1 -> aT1 * aT2, aT2 -> aT1 * aT2. *) @@ -179,7 +178,7 @@ Proof. by case/pprodP=> [[K H -> ->] <- /norm_joinEr]. Qed. Lemma pprodJ A B x : pprod A B :^ x = pprod (A :^ x) (B :^ x). Proof. rewrite /pprod !conjsg_eq1 !group_setJ normJ conjSg -conjsMg. -by do 3?case: ifP => // _; exact: conj0g. +by do 3?case: ifP => // _; apply: conj0g. Qed. (* Properties of the remainders *) @@ -223,7 +222,7 @@ Lemma remgr_id x : x \in H -> remgr K H x = x. Proof. move=> Hx; apply/eqP; rewrite eq_mulgV1 (sameP eqP set1gP) -tiKH inE. rewrite -mem_rcoset groupMr ?groupV // -in_setI remgrP. -by apply: subsetP Hx; exact: mulG_subr. +by apply: subsetP Hx; apply: mulG_subr. Qed. Lemma remgrMid x y : x \in K -> y \in H -> remgr K H (x * y) = y. @@ -326,7 +325,7 @@ Proof. by move/sdprodWpp/pprodWY. Qed. Lemma sdprodJ A B x : (A ><| B) :^ x = A :^ x ><| B :^ x. Proof. rewrite /sdprod -conjIg sub_conjg conjs1g -pprodJ. -by case: ifP => _ //; exact: imset0. +by case: ifP => _ //; apply: imset0. Qed. Lemma sdprod_context G K H : K ><| H = G -> @@ -337,7 +336,7 @@ by rewrite /normal mulG_subl mulG_subr mulG_subG normG. Qed. Lemma sdprod_compl G K H : K ><| H = G -> H \in [complements to K in G]. -Proof. by case/sdprodP=> _ mulKH _ tiKH; exact/complP. Qed. +Proof. by case/sdprodP=> _ mulKH _ tiKH; apply/complP. Qed. Lemma sdprod_normal_complP G K H : K <| G -> reflect (K ><| H = G) (K \in [complements to H in G]). @@ -504,7 +503,7 @@ Proof. by move/cprodWpp/pprodWY. Qed. Lemma cprodJ A B x : (A \* B) :^ x = A :^ x \* B :^ x. Proof. -by rewrite /cprod centJ conjSg -pprodJ; case: ifP => _ //; exact: imset0. +by rewrite /cprod centJ conjSg -pprodJ; case: ifP => _ //; apply: imset0. Qed. Lemma cprod_normal2 A B G : A \* B = G -> A <| G /\ B <| G. @@ -587,7 +586,7 @@ Qed. Lemma cprod_modr A B G H : A \* B = G -> B \subset H -> (H :&: A) \* B = H :&: G. -Proof. by rewrite -!(cprodC B) !(setIC H); exact: cprod_modl. Qed. +Proof. by rewrite -!(cprodC B) !(setIC H); apply: cprod_modl. Qed. Lemma bigcprodYP (I : finType) (P : pred I) (H : I -> {group gT}) : reflect (forall i j, P i -> P j -> i != j -> H i \subset 'C(H j)) @@ -692,12 +691,12 @@ Proof. by move/dprodWsd/sdprodWY. Qed. Lemma cprod_card_dprod G A B : A \* B = G -> #|A| * #|B| <= #|G| -> A \x B = G. -Proof. by case/cprodP=> [[K H -> ->] <- cKH] /cardMg_TI; exact: dprodE. Qed. +Proof. by case/cprodP=> [[K H -> ->] <- cKH] /cardMg_TI; apply: dprodE. Qed. Lemma dprodJ A B x : (A \x B) :^ x = A :^ x \x B :^ x. Proof. rewrite /dprod -conjIg sub_conjg conjs1g -cprodJ. -by case: ifP => _ //; exact: imset0. +by case: ifP => _ //; apply: imset0. Qed. Lemma dprod_normal2 A B G : A \x B = G -> A <| G /\ B <| G. @@ -754,11 +753,11 @@ Qed. Lemma bigdprodW I (r : seq I) P F G : \big[dprod/1]_(i <- r | P i) F i = G -> \prod_(i <- r | P i) F i = G. -Proof. by move/bigdprodWcp; exact: bigcprodW. Qed. +Proof. by move/bigdprodWcp; apply: bigcprodW. Qed. Lemma bigdprodWY I (r : seq I) P F G : \big[dprod/1]_(i <- r | P i) F i = G -> << \bigcup_(i <- r | P i) F i >> = G. -Proof. by move/bigdprodWcp; exact: bigcprodWY. Qed. +Proof. by move/bigdprodWcp; apply: bigcprodWY. Qed. Lemma bigdprodYP (I : finType) (P : pred I) (F : I -> {group gT}) : reflect (forall i, P i -> @@ -787,13 +786,13 @@ Qed. Lemma dprod_modr A B G H : A \x B = G -> B \subset H -> (H :&: A) \x B = H :&: G. -Proof. by rewrite -!(dprodC B) !(setIC H); exact: dprod_modl. Qed. +Proof. by rewrite -!(dprodC B) !(setIC H); apply: dprod_modl. Qed. Lemma subcent_dprod B C G A : B \x C = G -> A \subset 'N(B) :&: 'N(C) -> 'C_B(A) \x 'C_C(A) = 'C_G(A). Proof. move=> defG; have [_ _ cBC _] := dprodP defG; move: defG. -by rewrite !dprodEsd 1?(centSS _ _ cBC) ?subsetIl //; exact: subcent_sdprod. +by rewrite !dprodEsd 1?(centSS _ _ cBC) ?subsetIl //; apply: subcent_sdprod. Qed. Lemma dprod_card A B G : A \x B = G -> (#|A| * #|B|)%N = #|G|. @@ -895,7 +894,7 @@ Lemma morphim_coprime_sdprod : K ><| H = G -> coprime #|K| #|H| -> f @* K ><| f @* H = f @* G. Proof. rewrite /sdprod => defG coHK; move: defG. -by rewrite !coprime_TIg ?coprime_morph // !subxx; exact: morphim_pprod. +by rewrite !coprime_TIg ?coprime_morph // !subxx; apply: morphim_pprod. Qed. Lemma injm_sdprod : 'injm f -> K ><| H = G -> f @* K ><| f @* H = f @* G. @@ -920,7 +919,7 @@ Lemma morphim_coprime_dprod : K \x H = G -> coprime #|K| #|H| -> f @* K \x f @* H = f @* G. Proof. rewrite /dprod => defG coHK; move: defG. -by rewrite !coprime_TIg ?coprime_morph // !subxx; exact: morphim_cprod. +by rewrite !coprime_TIg ?coprime_morph // !subxx; apply: morphim_cprod. Qed. End OneProd. @@ -990,13 +989,13 @@ Definition extprod_mulg (x y : gT1 * gT2) := (x.1 * y.1, x.2 * y.2). Definition extprod_invg (x : gT1 * gT2) := (x.1^-1, x.2^-1). Lemma extprod_mul1g : left_id (1, 1) extprod_mulg. -Proof. case=> x1 x2; congr (_, _); exact: mul1g. Qed. +Proof. by case=> x1 x2; congr (_, _); apply: mul1g. Qed. Lemma extprod_mulVg : left_inverse (1, 1) extprod_invg extprod_mulg. -Proof. by move=> x; congr (_, _); exact: mulVg. Qed. +Proof. by move=> x; congr (_, _); apply: mulVg. Qed. Lemma extprod_mulgA : associative extprod_mulg. -Proof. by move=> x y z; congr (_, _); exact: mulgA. Qed. +Proof. by move=> x y z; congr (_, _); apply: mulgA. Qed. Definition extprod_groupMixin := Eval hnf in FinGroup.Mixin extprod_mulgA extprod_mul1g extprod_mulVg. @@ -1037,10 +1036,10 @@ Canonical fst_morphism := @Morphism _ _ setT _ (in2W fst_morphM). Canonical snd_morphism := @Morphism _ _ setT _ (in2W snd_morphM). Lemma injm_pair1g : 'injm pair1g. -Proof. by apply/subsetP=> x /morphpreP[_ /set1P[->]]; exact: set11. Qed. +Proof. by apply/subsetP=> x /morphpreP[_ /set1P[->]]; apply: set11. Qed. Lemma injm_pairg1 : 'injm pairg1. -Proof. by apply/subsetP=> x /morphpreP[_ /set1P[->]]; exact: set11. Qed. +Proof. by apply/subsetP=> x /morphpreP[_ /set1P[->]]; apply: set11. Qed. Lemma morphim_pairg1 (H1 : {set gT1}) : pairg1 @* H1 = setX H1 1. Proof. by rewrite -imset2_pair imset2_set1r morphimEsub ?subsetT. Qed. @@ -1453,7 +1452,7 @@ Lemma morphim_sdprodm A B : A \subset H -> B \subset K -> sdprodm @* (A * B) = fH @* A * fK @* B. Proof. move=> sAH sBK; rewrite morphim_restrm /= (setIidPr _) ?morphim_pprodm //. -case/sdprodP: eqHK_G => _ <- _ _; exact: mulgSS. +by case/sdprodP: eqHK_G => _ <- _ _; apply: mulgSS. Qed. Lemma im_sdprodm : sdprodm @* G = fH @* H * fK @* K. @@ -1523,7 +1522,7 @@ Lemma morphim_cprodm A B : A \subset H -> B \subset K -> cprodm @* (A * B) = fH @* A * fK @* B. Proof. move=> sAH sBK; rewrite morphim_restrm /= (setIidPr _) ?morphim_pprodm //. -case/cprodP: eqHK_G => _ <- _; exact: mulgSS. +by case/cprodP: eqHK_G => _ <- _; apply: mulgSS. Qed. Lemma im_cprodm : cprodm @* G = fH @* H * fK @* K. @@ -1690,7 +1689,7 @@ apply: (iffP misomP) => [[pM /isomP[injf /= <-]] | ]. have sH1H: setX H1 1 \subset setX H1 H2 by rewrite setXS ?sub1G. have sH2H: setX 1 H2 \subset setX H1 H2 by rewrite setXS ?sub1G. rewrite morphim1 injm_cent ?injmI //= subsetI => /andP[_]. - by rewrite !morphimEsub //= !imset_mulgm mulg1 mul1g; exact: dprodE. + by rewrite !morphimEsub //= !imset_mulgm mulg1 mul1g; apply: dprodE. case/dprodP=> _ defG cH12 trH12. have fM: morphic (setX H1 H2) mulgm. apply/morphicP=> [[x1 x2] [y1 y2] /setXP[_ Hx2] /setXP[Hy1 _]]. diff --git a/mathcomp/fingroup/morphism.v b/mathcomp/fingroup/morphism.v index 9649644..0c0ba5b 100644 --- a/mathcomp/fingroup/morphism.v +++ b/mathcomp/fingroup/morphism.v @@ -1,10 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import choice fintype finfun bigop finset. -Require Import fingroup. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq choice fintype finfun. +From mathcomp +Require Import bigop finset fingroup. (******************************************************************************) (* This file contains the definitions of: *) @@ -19,50 +18,50 @@ Require Import fingroup. (* idm D == the identity morphism with domain D, or more precisely *) (* the identity function, but with a canonical *) (* {morphism G -> gT} structure. *) -(* trivm D == the trivial morphism with domain D *) +(* trivm D == the trivial morphism with domain D. *) (* If f has a {morphism D >-> rT} structure *) -(* 'dom f == D *) -(* f @* A == the image of A by f, where f is defined *) +(* 'dom f == D, the domain of f. *) +(* f @* A == the image of A by f, where f is defined. *) (* := f @: (D :&: A) *) -(* f @*^-1 R == the pre-image of R by f, where f is defined *) +(* f @*^-1 R == the pre-image of R by f, where f is defined. *) (* := D :&: f @^-1: R *) -(* 'ker f == the kernel of f *) -(* := f @^-1: 1 *) -(* 'ker_G f == the kernel of f restricted to G *) +(* 'ker f == the kernel of f. *) +(* := f @*^-1 1 *) +(* 'ker_G f == the kernel of f restricted to G. *) (* := G :&: 'ker f (this is a pure notation) *) -(* 'injm f <=> f injective on D *) +(* 'injm f <=> f injective on D. *) (* <-> ker f \subset 1 (this is a pure notation) *) (* invm injf == the inverse morphism of f, with domain f @* D, when f *) -(* is injective (injf : 'injm f) *) +(* is injective (injf : 'injm f). *) (* restrm f sDom == the restriction of f to a subset A of D, given *) (* (sDom : A \subset D); restrm f sDom is transparently *) (* identical to f; the restrmP and domP lemmas provide *) (* opaque restrictions. *) (* invm f infj == the inverse morphism for an injective f, with domain *) -(* f @* D, given (injf : 'injm f) *) +(* f @* D, given (injf : 'injm f). *) (* *) -(* G \isog H <=> G and H are isomorphic as groups *) -(* H \homg G <=> H is a homomorphic image of G *) -(* isom G H f <=> f maps G isomorphically to H, provided D contains G *) -(* <-> f @: G^# == H^# *) +(* G \isog H <=> G and H are isomorphic as groups. *) +(* H \homg G <=> H is a homomorphic image of G. *) +(* isom G H f <=> f maps G isomorphically to H, provided D contains G. *) +(* := f @: G^# == H^# *) (* *) (* If, moreover, g : {morphism G >-> gT} with G : {group aT}, *) (* factm sKer sDom == the (natural) factor morphism mapping f @* G to g @* G *) -(* given sDom : G \subset D, sKer : 'ker f \subset 'ker g *) +(* with sDom : G \subset D, sKer : 'ker f \subset 'ker g. *) (* ifactm injf g == the (natural) factor morphism mapping f @* G to g @* G *) (* when f is injective (injf : 'injm f); here g must *) -(* be an actual morphism structure, not its function *) +(* denote an actual morphism structure, not its function *) (* projection. *) (* *) (* If g has a {morphism G >-> aT} structure for any G : {group gT}, then *) -(* f \o g has a canonical {morphism g @*^-1 D >-> rT} structure *) +(* f \o g has a canonical {morphism g @*^-1 D >-> rT} structure. *) (* *) (* Finally, for an arbitrary function f : aT -> rT *) (* morphic D f <=> f preserves group multiplication in D, i.e., *) -(* f (x * y) = (f x) * (f y) for all x, y in D *) +(* f (x * y) = (f x) * (f y) for all x, y in D. *) (* morphm fM == a function identical to f, but with a canonical *) -(* {morphism D >-> rT} structure, given fM : morphic D f *) -(* misom D C f <=> f maps D isomorphically to C *) +(* {morphism D >-> rT} structure, given fM : morphic D f. *) +(* misom D C f <=> f is a morphism that maps D isomorphically to C. *) (* := morphic D f && isom D C f *) (******************************************************************************) @@ -83,10 +82,10 @@ Structure morphism (D : {set aT}) : Type := Morphism { _ : {in D &, {morph mfun : x y / x * y}} }. -(* We give the most 'lightweight' possible specification to define morphisms:*) -(* local congruence with the group law of aT. We then provide the properties *) -(* for the 'textbook' notion of morphism, when the required structures are *) -(* available (e.g. its domain is a group). *) +(* We give the 'lightest' possible specification to define morphisms: local *) +(* congruence, in D, with the group law of aT. We then provide the properties *) +(* for the 'textbook' notion of morphism, when the required structures are *) +(* available (e.g. its domain is a group). *) Definition morphism_for D of phant rT := morphism D. @@ -106,7 +105,7 @@ by exists z; first apply/setIP. Qed. Lemma morphpreP : reflect (x \in D /\ f x \in R) (x \in D :&: f @^-1: R). -Proof. rewrite !inE; exact: andP. Qed. +Proof. by rewrite !inE; apply: andP. Qed. End MorphismStructure. @@ -122,7 +121,7 @@ Implicit Arguments morphimP [aT rT D A f y]. Implicit Arguments morphpreP [aT rT D R f x]. Prenex Implicits morphimP morphpreP. -(* domain, image, preimage, kernel, using phantom types to infer the domain *) +(* Domain, image, preimage, kernel, using phantom types to infer the domain. *) Section MorphismOps1. @@ -206,7 +205,7 @@ Qed. Lemma morphR : {in D &, {morph f : x y / [~ x, y]}}. Proof. by move=> * /=; rewrite morphM ?(groupV, groupJ) // morphJ ?morphV. Qed. -(* morphic image,preimage properties w.r.t. set-theoretic operations *) +(* Morphic image, preimage properties w.r.t. set-theoretic operations. *) Lemma morphimE A : f @* A = f @: (D :&: A). Proof. by []. Qed. Lemma morphpreE R : f @*^-1 R = D :&: f @^-1: R. Proof. by []. Qed. @@ -237,7 +236,7 @@ Lemma mem_morphim A x : x \in D -> x \in A -> f x \in f @* A. Proof. by move=> Dx Ax; apply/morphimP; exists x. Qed. Lemma mem_morphpre R x : x \in D -> f x \in R -> x \in f @*^-1 R. -Proof. by move=> Dx Rfx; exact/morphpreP. Qed. +Proof. by move=> Dx Rfx; apply/morphpreP. Qed. Lemma morphimS A B : A \subset B -> f @* A \subset f @* B. Proof. by move=> sAB; rewrite imsetS ?setIS. Qed. @@ -270,7 +269,7 @@ Lemma morphim_eq0 A : A \subset D -> (f @* A == set0) = (A == set0). Proof. by rewrite imset_eq0 => /setIidPr->. Qed. Lemma morphim_set1 x : x \in D -> f @* [set x] = [set f x]. -Proof. by rewrite /morphim -sub1set => /setIidPr->; exact: imset_set1. Qed. +Proof. by rewrite /morphim -sub1set => /setIidPr->; apply: imset_set1. Qed. Lemma morphim1 : f @* 1 = 1. Proof. by rewrite morphim_set1 ?morph1. Qed. @@ -373,13 +372,13 @@ Proof. by apply/setP=> x; rewrite !inE; case: (x \in D). Qed. (* kernel, domain properties *) Lemma kerP x : x \in D -> reflect (f x = 1) (x \in 'ker f). -Proof. move=> Dx; rewrite 2!inE Dx; exact: set1P. Qed. +Proof. by move=> Dx; rewrite 2!inE Dx; apply: set1P. Qed. Lemma dom_ker : {subset 'ker f <= D}. Proof. by move=> x /morphpreP[]. Qed. Lemma mker x : x \in 'ker f -> f x = 1. -Proof. by move=> Kx; apply/kerP=> //; exact: dom_ker. Qed. +Proof. by move=> Kx; apply/kerP=> //; apply: dom_ker. Qed. Lemma mkerl x y : x \in 'ker f -> y \in D -> f (x * y) = f y. Proof. by move=> Kx Dy; rewrite morphM // ?(dom_ker, mker Kx, mul1g). Qed. @@ -391,12 +390,12 @@ Lemma rcoset_kerP x y : x \in D -> y \in D -> reflect (f x = f y) (x \in 'ker f :* y). Proof. move=> Dx Dy; rewrite mem_rcoset !inE groupM ?morphM ?groupV //=. -rewrite morphV // -eq_mulgV1; exact: eqP. +by rewrite morphV // -eq_mulgV1; apply: eqP. Qed. Lemma ker_rcoset x y : x \in D -> y \in D -> f x = f y -> exists2 z, z \in 'ker f & x = z * y. -Proof. move=> Dx Dy eqfxy; apply/rcosetP; exact/rcoset_kerP. Qed. +Proof. by move=> Dx Dy eqfxy; apply/rcosetP; apply/rcoset_kerP. Qed. Lemma ker_norm : D \subset 'N('ker f). Proof. @@ -513,7 +512,7 @@ Lemma sub_morphim_pre A R : A \subset D -> (f @* A \subset R) = (A \subset f @*^-1 R). Proof. move=> sAD; rewrite -morphpreSK (morphimS, morphimK) //. -apply/idP/idP; first by apply: subset_trans; exact: mulG_subr. +apply/idP/idP; first by apply: subset_trans; apply: mulG_subr. by move/(mulgS ('ker f)); rewrite -morphpreMl ?(sub1G, mul1g). Qed. @@ -625,7 +624,7 @@ Qed. Lemma morphim_norms A B : A \subset 'N(B) -> f @* A \subset 'N(f @* B). Proof. -by move=> nBA; apply: subset_trans (morphim_norm B); exact: morphimS. +by move=> nBA; apply: subset_trans (morphim_norm B); apply: morphimS. Qed. Lemma morphim_subnorm A B : f @* 'N_A(B) \subset 'N_(f @* A)(f @* B). @@ -639,7 +638,7 @@ Proof. by move=> Dx; rewrite -(morphim_set1 Dx) morphim_norm. Qed. Lemma morphim_cent1s A x : x \in D -> A \subset 'C[x] -> f @* A \subset 'C[f x]. Proof. -by move=> Dx cAx; apply: subset_trans (morphim_cent1 Dx); exact: morphimS. +by move=> Dx cAx; apply: subset_trans (morphim_cent1 Dx); apply: morphimS. Qed. Lemma morphim_subcent1 A x : x \in D -> f @* 'C_A[x] \subset 'C_(f @* A)[f x]. @@ -648,12 +647,12 @@ Proof. by move=> Dx; rewrite -(morphim_set1 Dx) morphim_subnorm. Qed. Lemma morphim_cent A : f @* 'C(A) \subset 'C(f @* A). Proof. apply/bigcapsP=> fx; case/morphimP=> x Dx Ax ->{fx}. -by apply: subset_trans (morphim_cent1 Dx); apply: morphimS; exact: bigcap_inf. +by apply: subset_trans (morphim_cent1 Dx); apply: morphimS; apply: bigcap_inf. Qed. Lemma morphim_cents A B : A \subset 'C(B) -> f @* A \subset 'C(f @* B). Proof. -by move=> cBA; apply: subset_trans (morphim_cent B); exact: morphimS. +by move=> cBA; apply: subset_trans (morphim_cent B); apply: morphimS. Qed. Lemma morphim_subcent A B : f @* 'C_A(B) \subset 'C_(f @* A)(f @* B). @@ -670,7 +669,7 @@ Qed. Lemma morphpre_norms R S : R \subset 'N(S) -> f @*^-1 R \subset 'N(f @*^-1 S). Proof. -by move=> nSR; apply: subset_trans (morphpre_norm S); exact: morphpreS. +by move=> nSR; apply: subset_trans (morphpre_norm S); apply: morphpreS. Qed. Lemma morphpre_normal R S : @@ -728,7 +727,7 @@ Proof. by move=> sRfD; move/morphim_cents; rewrite morphpreK. Qed. Lemma morphpre_subcent R A : 'C_(f @*^-1 R)(A) \subset f @*^-1 'C_R(f @* A). Proof. -by rewrite -morphpreIdom -setIA setICA morphpreI setIS //; exact: morphpre_cent. +by rewrite -morphpreIdom -setIA setICA morphpreI setIS //; apply: morphpre_cent. Qed. (* local injectivity properties *) @@ -892,7 +891,7 @@ Lemma injm_idm G : 'injm (idm G). Proof. by apply/injmP=> x y _ _. Qed. Lemma ker_idm G : 'ker (idm G) = 1. -Proof. by apply/trivgP; exact: injm_idm. Qed. +Proof. by apply/trivgP; apply: injm_idm. Qed. Lemma morphim_idm A B : B \subset A -> idm A @* B = B. Proof. @@ -1194,7 +1193,7 @@ Proof. by rewrite ker_ifactm => /trivgP->; rewrite morphim1. Qed. End InjFactm. -(* Reflected (boolean) form of morphism and isomorphism properties *) +(* Reflected (boolean) form of morphism and isomorphism properties. *) Section ReflectProp. @@ -1204,7 +1203,7 @@ Section Defs. Variables (A : {set aT}) (B : {set rT}). -(* morphic is the morphM property of morphisms seen through morphicP *) +(* morphic is the morphM property of morphisms seen through morphicP. *) Definition morphic (f : aT -> rT) := [forall u in [predX A & A], f (u.1 * u.2) == f u.1 * f u.2]. @@ -1240,14 +1239,14 @@ Lemma misom_isog f : misom f -> isog. Proof. case/andP=> fM iso_f; apply/existsP; exists (finfun f). apply/andP; split; last by rewrite /misom /isom !(eq_imset _ (ffunE f)). -apply/forallP=> u; rewrite !ffunE; exact: forallP fM u. +by apply/forallP=> u; rewrite !ffunE; apply: forallP fM u. Qed. Lemma isom_isog (D : {group aT}) (f : {morphism D >-> rT}) : A \subset D -> isom f -> isog. Proof. move=> sAD isof; apply: (@misom_isog f); rewrite /misom isof andbT. -apply/morphicP; exact: (sub_in2 (subsetP sAD) (morphM f)). +by apply/morphicP; apply: (sub_in2 (subsetP sAD) (morphM f)). Qed. Lemma isog_isom : isog -> {f : {morphism A >-> rT} | isom f}. @@ -1313,7 +1312,7 @@ by rewrite subDset setUC subsetU ?sAG. Qed. Lemma sub_isog (A : {set aT}) : A \subset G -> 'injm f -> isog A (f @* A). -Proof. by move=> sAG injf; apply: (isom_isog f sAG); exact: sub_isom. Qed. +Proof. by move=> sAG injf; apply: (isom_isog f sAG); apply: sub_isom. Qed. Lemma restr_isom_to (A : {set aT}) (C R : {group rT}) (sAG : A \subset G) : f @* A = C -> isom G R f -> isom A C (restrm sAG f). @@ -1349,7 +1348,7 @@ Proof. exact: sub_isom (im_idm G) (injm_idm G). Qed. Lemma isog_refl : G \isog G. Proof. exact: isom_isog idm_isom. Qed. Lemma card_isog : G \isog H -> #|G| = #|H|. -Proof. case/isogP=> f injf <-; apply: isom_card (f) _; exact/isomP. Qed. +Proof. by case/isogP=> f injf <-; apply: isom_card (f) _; apply/isomP. Qed. Lemma isog_abelian : G \isog H -> abelian G = abelian H. Proof. by case/isogP=> f injf <-; rewrite injm_abelian. Qed. @@ -1358,7 +1357,7 @@ Lemma trivial_isog : G :=: 1 -> H :=: 1 -> G \isog H. Proof. move=> -> ->; apply/isogP. exists [morphism of @trivm gT hT 1]; rewrite /= ?morphim1 //. -rewrite ker_trivm; exact: subxx. +by rewrite ker_trivm; apply: subxx. Qed. Lemma isog_eq1 : G \isog H -> (G :==: 1) = (H :==: 1). @@ -1393,7 +1392,7 @@ Variables gT hT kT : finGroupType. Variables (G : {group gT}) (H : {group hT}) (K : {group kT}). Lemma isog_sym : (G \isog H) = (H \isog G). -Proof. apply/idP/idP; exact: isog_symr. Qed. +Proof. by apply/idP/idP; apply: isog_symr. Qed. Lemma isog_transl : G \isog H -> (G \isog K) = (H \isog K). Proof. @@ -1467,14 +1466,14 @@ Qed. Lemma eq_homgl gT aT rT (G : {group gT}) (H : {group aT}) (K : {group rT}) : G \isog H -> homg G K = homg H K. Proof. -by rewrite isogEhom => /andP[homGH homHG]; apply/idP/idP; exact: homg_trans. +by rewrite isogEhom => /andP[homGH homHG]; apply/idP/idP; apply: homg_trans. Qed. Lemma eq_homgr gT rT aT (G : {group gT}) (H : {group rT}) (K : {group aT}) : G \isog H -> homg K G = homg K H. Proof. rewrite isogEhom => /andP[homGH homHG]. -by apply/idP/idP=> homK; exact: homg_trans homK _. +by apply/idP/idP=> homK; apply: homg_trans homK _. Qed. End Homg. @@ -1495,10 +1494,10 @@ Canonical sgval_morphism := Morphism (@sgvalM _ G). Canonical subg_morphism := Morphism (@subgM _ G). Lemma injm_sgval : 'injm sgval. -Proof. apply/injmP; apply: in2W; exact: subg_inj. Qed. +Proof. by apply/injmP; apply: in2W; apply: subg_inj. Qed. Lemma injm_subg : 'injm (subg G). -Proof. apply/injmP; exact: can_in_inj (@subgK _ _). Qed. +Proof. by apply/injmP; apply: can_in_inj (@subgK _ _). Qed. Hint Resolve injm_sgval injm_subg. Lemma ker_sgval : 'ker sgval = 1. Proof. exact/trivgP. Qed. @@ -1511,7 +1510,7 @@ by apply/subsetP=> u _; rewrite -(sgvalK u) mem_imset ?subgP. Qed. Lemma sgval_sub A : sgval @* A \subset G. -Proof. apply/subsetP=> x; case/imsetP=> u _ ->; exact: subgP. Qed. +Proof. by apply/subsetP=> x; case/imsetP=> u _ ->; apply: subgP. Qed. Lemma sgvalmK A : subg G @* (sgval @* A) = A. Proof. diff --git a/mathcomp/fingroup/perm.v b/mathcomp/fingroup/perm.v index 3719835..664129b 100644 --- a/mathcomp/fingroup/perm.v +++ b/mathcomp/fingroup/perm.v @@ -1,10 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import path choice fintype tuple finfun bigop finset binomial. -Require Import fingroup. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq path choice fintype. +From mathcomp +Require Import tuple finfun bigop finset binomial fingroup. (******************************************************************************) (* This file contains the definition and properties associated to the group *) @@ -15,14 +14,14 @@ Require Import fingroup. (* 'S_n == the set of all permutations of 'I_n, i.e., of {0,.., n-1} *) (* perm_on A u == u is a permutation with support A, i.e., u only displaces *) (* elements of A (u x != x implies x \in A). *) -(* tperm x y == the transposition of x, y *) -(* aperm x s == the image of x under the action of the permutation s *) +(* tperm x y == the transposition of x, y. *) +(* aperm x s == the image of x under the action of the permutation s. *) (* := s x *) (* pcycle s x == the set of all elements that are in the same cycle of the *) -(* permutation s as x, i.e., {x, s x, (s ^+ 2) x, ...} *) -(* pcycles s == the set of all the cycles of the permutation s *) -(* (s : bool) == s is an odd permutation (the coercion is called odd_perm) *) -(* dpair u == u is a pair (x, y) of distinct objects (i.e., x != y) *) +(* permutation s as x, i.e., {x, s x, (s ^+ 2) x, ...}. *) +(* pcycles s == the set of all the cycles of the permutation s. *) +(* (s : bool) == s is an odd permutation (the coercion is called odd_perm). *) +(* dpair u == u is a pair (x, y) of distinct objects (i.e., x != y). *) (* lift_perm i j s == the permutation obtained by lifting s : 'S_n.-1 over *) (* (i |-> j), that maps i to j and lift i k to lift j (s k). *) (* Canonical structures are defined allowing permutations to be an eqType, *) @@ -117,7 +116,7 @@ Variable T : finType. Implicit Types (x y : T) (s t : {perm T}) (S : {set T}). Lemma permP s t : s =1 t <-> s = t. -Proof. by split=> [| -> //]; rewrite unlock => eq_sv; exact/val_inj/ffunP. Qed. +Proof. by split=> [| -> //]; rewrite unlock => eq_sv; apply/val_inj/ffunP. Qed. Lemma pvalE s : pval s = s :> (T -> T). Proof. by rewrite [@fun_of_perm]unlock. Qed. @@ -126,7 +125,7 @@ Lemma permE f f_inj : @perm T f f_inj =1 f. Proof. by move=> x; rewrite -pvalE [@perm]unlock ffunE. Qed. Lemma perm_inj s : injective s. -Proof. by rewrite -!pvalE; exact: (injectiveP _ (valP s)). Qed. +Proof. by rewrite -!pvalE; apply: (injectiveP _ (valP s)). Qed. Implicit Arguments perm_inj []. Hint Resolve perm_inj. @@ -204,7 +203,7 @@ by have [-> /tH | /sH] := eqVneq (s x) x. Qed. Lemma out_perm S u x : perm_on S u -> x \notin S -> u x = x. -Proof. by move=> uS; exact: contraNeq (subsetP uS x). Qed. +Proof. by move=> uS; apply: contraNeq (subsetP uS x). Qed. Lemma im_perm_on u S : perm_on S u -> u @: S = S. Proof. @@ -265,11 +264,11 @@ pose fA s : ffA := [ffun u => s (val u)]. rewrite -!sum1dep_card -sum1_card (reindex_onto fA pfT) => [|f]. apply: eq_bigl => p; rewrite andbC; apply/idP/and3P=> [onA | []]; first split. - apply/eqP; suffices fTAp: fT (fA p) = pval p. - by apply/permP=> x; rewrite -!pvalE insubdK fTAp //; exact: (valP p). + by apply/permP=> x; rewrite -!pvalE insubdK fTAp //; apply: (valP p). apply/ffunP=> x; rewrite ffunE pvalE. by case: insubP => [u _ <- | /out_perm->] //=; rewrite ffunE. - by apply/forallP=> [[x Ax]]; rewrite ffunE /= perm_closed. - - by apply/injectiveP=> u v; rewrite !ffunE => /perm_inj; exact: val_inj. + - by apply/injectiveP=> u v; rewrite !ffunE => /perm_inj; apply: val_inj. move/eqP=> <- _ _; apply/subsetP=> x; rewrite !inE -pvalE val_insubd fun_if. by rewrite if_arg ffunE; case: insubP; rewrite // pvalE perm1 if_same eqxx. case/andP=> /forallP-onA /injectiveP-f_inj. @@ -393,7 +392,7 @@ have lt_xf a b u n : n < xf a b u -> ~~ pred2 a b ((u ^+ n.+1) a). by rewrite permX iterSr nth_traject // (leq_trans lt_n). pose t a b u := tperm a b * u. have tC a b u : t a b u = t b a u by rewrite /t tpermC. -have tK a b: involutive (t a b) by move=> u; exact: tpermKg. +have tK a b: involutive (t a b) by move=> u; apply: tpermKg. have tXC a b u n: n <= xf a b u -> (t a b u ^+ n.+1) b = (u ^+ n.+1) a. elim: n => [|n IHn] lt_n_f; first by rewrite permM tpermR. rewrite !(expgSr _ n.+1) !permM {}IHn 1?ltnW //; congr (u _). @@ -408,7 +407,7 @@ have eq_xf a b u: pred2 a b ((u ^+ (xf a b u).+1) a). have xfC a b u: xf b a (t a b u) = xf a b u. without loss lt_a: a b u / xf b a (t a b u) < xf a b u. move=> IHab; set m := xf b a _; set n := xf a b u. - by case: (ltngtP m n) => // ltx; [exact: IHab | rewrite -[m]IHab tC tK]. + by case: (ltngtP m n) => // ltx; [apply: IHab | rewrite -[m]IHab tC tK]. by move/lt_xf: (lt_a); rewrite -(tXC a b) 1?ltnW //= orbC [_ || _]eq_xf. pose ts := t x y s; rewrite /= -[_ * s]/ts. pose dp u := #|pcycles u :\ pcycle u y :\ pcycle u x|. @@ -428,8 +427,8 @@ rewrite -/(dp s) !addnA !eq_pcycle_mem andbT; congr (_ + _); last first. by rewrite /aperm exp_id mem_pcycle. by rewrite /aperm -exp_id mem_pcycle. elim: n => // n IHn; rewrite !expgSr !permM {}IHn tpermD //. - apply: contraNneq sxz => ->; exact: mem_pcycle. - apply: contraNneq syz => ->; exact: mem_pcycle. + by apply: contraNneq sxz => ->; apply: mem_pcycle. + by apply: contraNneq syz => ->; apply: mem_pcycle. case: eqP {dp} => [<- | ne_xy]; first by rewrite /t tperm1 mul1g pcycle_id. suff ->: (x \in pcycle (t x y s) y) = (x \notin pcycle s y) by case: (x \in _). without loss xf_x: s x y ne_xy / (s ^+ (xf x y s).+1) x = x. @@ -556,13 +555,13 @@ congr (_ (+) _); last first. congr (_ (+) _); transitivity (tperm (lift j t.1) (lift j t.2)); last first. by rewrite odd_tperm (inj_eq (@lift_inj _ _)). congr odd_perm; apply/permP=> k; case: (unliftP j k) => [k'|] ->. - rewrite lift_perm_lift inj_tperm //; exact: lift_inj. + by rewrite lift_perm_lift inj_tperm //; apply: lift_inj. by rewrite lift_perm_id tpermD // eq_sym neq_lift. suff{i j s} odd_lift0 (k : 'I_n.+1): lift_perm ord0 k 1 = odd k :> bool. rewrite -!odd_lift0 -{2}invg1 -lift_permV odd_permV -odd_permM. by rewrite lift_permM mulg1. elim: {k}(k : nat) {1 3}k (erefl (k : nat)) => [|m IHm] k def_k. - rewrite (_ : k = ord0) ?lift_perm1 ?odd_perm1 //; exact: val_inj. + by rewrite (_ : k = ord0) ?lift_perm1 ?odd_perm1 //; apply: val_inj. have le_mn: m < n.+1 by [rewrite -def_k ltnW]; pose j := Ordinal le_mn. rewrite -(mulg1 1)%g -(lift_permM _ j) odd_permM {}IHm // addbC. rewrite (_ : _ 1 = tperm j k); first by rewrite odd_tperm neq_ltn def_k leqnn. diff --git a/mathcomp/fingroup/presentation.v b/mathcomp/fingroup/presentation.v index 8f6ddb3..13dd99a 100644 --- a/mathcomp/fingroup/presentation.v +++ b/mathcomp/fingroup/presentation.v @@ -1,9 +1,8 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import fintype finset. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq fintype finset. +From mathcomp Require Import fingroup morphism. (******************************************************************************) @@ -187,7 +186,7 @@ Implicit Types gT rT : finGroupType. Import Presentation. Lemma isoGrp_hom gT (G : {group gT}) p : G \isog Grp p -> G \homg Grp p. -Proof. by move <-; exact: homg_refl. Qed. +Proof. by move <-; apply: homg_refl. Qed. Lemma isoGrpP gT (G : {group gT}) p rT (H : {group rT}) : G \isog Grp p -> reflect (#|H| = #|G| /\ H \homg Grp p) (H \isog G). @@ -216,7 +215,7 @@ have and_relE xT x1 x2 r: @and_rel xT x1 x2 r = (x1 == x2) && r :> bool. by case: r => //=; rewrite andbT. have rsatG e f: all (mem G) e -> rel e f NoRel -> rel (map h e) f NoRel. move=> Ge; have: NoRel -> NoRel by []; move: NoRel {2 4}NoRel. - elim: f => [x1 x2 | f1 IH1 f2 IH2] r hr IHr; last by apply: IH1; exact: IH2. + elim: f => [x1 x2 | f1 IH1 f2 IH2] r hr IHr; last by apply: IH1; apply: IH2. by rewrite !and_relE !evalG //; case/andP; move/eqP->; rewrite eqxx. set s := env1; set vT := gT : finType in s *. set s' := env1; set vT' := rT : finType in s' *. @@ -239,19 +238,19 @@ Qed. Lemma eq_homGrp gT rT (G : {group gT}) (H : {group rT}) p : G \isog H -> (G \homg Grp p) = (H \homg Grp p). Proof. -by rewrite isogEhom => /andP[homGH homHG]; apply/idP/idP; exact: homGrp_trans. +by rewrite isogEhom => /andP[homGH homHG]; apply/idP/idP; apply: homGrp_trans. Qed. Lemma isoGrp_trans gT rT (G : {group gT}) (H : {group rT}) p : G \isog H -> H \isog Grp p -> G \isog Grp p. -Proof. by move=> isoGH isoHp kT K; rewrite -isoHp; exact: eq_homgr. Qed. +Proof. by move=> isoGH isoHp kT K; rewrite -isoHp; apply: eq_homgr. Qed. Lemma intro_isoGrp gT (G : {group gT}) p : G \homg Grp p -> (forall rT (H : {group rT}), H \homg Grp p -> H \homg G) -> G \isog Grp p. Proof. move=> homGp freeG rT H. -by apply/idP/idP=> [homHp|]; [exact: homGrp_trans homGp | exact: freeG]. +by apply/idP/idP=> [homHp|]; [apply: homGrp_trans homGp | apply: freeG]. Qed. End PresentationTheory. diff --git a/mathcomp/fingroup/quotient.v b/mathcomp/fingroup/quotient.v index 1983e59..aa5bc0a 100644 --- a/mathcomp/fingroup/quotient.v +++ b/mathcomp/fingroup/quotient.v @@ -1,20 +1,19 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype prime finset. -Require Import fingroup morphism automorphism. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq div choice. +From mathcomp +Require Import fintype prime finset fingroup morphism automorphism. (******************************************************************************) (* This file contains the definitions of: *) -(* coset_of H == the (sub)type of bilateral cosets of H (see below) *) -(* coset H == the canonical projection into coset_of H *) +(* coset_of H == the (sub)type of bilateral cosets of H (see below). *) +(* coset H == the canonical projection into coset_of H. *) (* A / H == the quotient of A by H, that is, the morphic image *) (* of A by coset H. We do not require H <| A, so in a *) (* textbook A / H would be written 'N_A(H) * H / H. *) (* quotm f (nHG : H <| G) == the quotient morphism induced by f, *) -(* mapping G / H onto f @* G / f @* H *) +(* mapping G / H onto f @* G / f @* H. *) (* qisom f (eqHG : H = G) == the identity isomorphism between *) (* [set: coset_of G] and [set: coset_of H]. *) (* We also prove the three isomorphism theorems, and counting lemmas for *) @@ -32,13 +31,13 @@ Section Cosets. Variables (gT : finGroupType) (Q A : {set gT}). (******************************************************************************) -(* Cosets are right cosets of elements in the normaliser *) +(* Cosets are right cosets of elements in the normaliser. *) (* We let cosets coerce to GroupSet.sort, so they inherit the group subset *) (* base group structure. Later we will define a proper group structure on *) (* cosets, which will then hide the inherited structure once coset_of unifies *) (* with FinGroup.sort; the coercion to GroupSet.sort will no longer be used. *) (* Note that for Hx Hy : coset_of H, Hx * Hy : {set gT} can mean either *) -(* set_of_coset (mulg Hx Hy) OR mulg (set_of_coset Hx) (set_of_coset Hy) *) +(* set_of_coset (mulg Hx Hy) OR mulg (set_of_coset Hx) (set_of_coset Hy). *) (* However, since the two terms are actually convertible, we can live with *) (* this ambiguity. *) (* We take great care that neither the type coset_of H, nor its Canonical *) @@ -46,12 +45,12 @@ Variables (gT : finGroupType) (Q A : {set gT}). (* group structure of H. Otherwise, rewriting would be extremely awkward *) (* because all our equalities are stated at the set level. *) (* The trick we use is to interpret coset_of A, when A is any set, as the *) -(* type of cosets of the group <A> generated by A, in the group A <*> N(A) *) +(* type of cosets of the group <<A>> generated by A, in the group A <*> N(A) *) (* generated by A and its normaliser. This coincides with the type of *) (* bilateral cosets of A when A is a group. We restrict the domain of coset A *) (* to 'N(A), so that we get almost all the same conversion equalities as if *) (* we had forced A to be a group in the first place; the only exception, that *) -(* 1 : coset_of A : set _ = <<A>> rather than A, is covered by genGid. *) +(* 1 : coset_of A : {set gT} = <<A>> rather than A, can be handled by genGid. *) (******************************************************************************) Notation H := <<A>>. @@ -120,7 +119,7 @@ Canonical coset_baseGroupType := Eval hnf in BaseFinGroupType coset_of coset_of_groupMixin. Canonical coset_groupType := FinGroupType coset_invP. -(* Projection of the initial group type over the cosets groupType *) +(* Projection of the initial group type over the cosets groupType. *) Definition coset x : coset_of := insubd (1 : coset_of) (H :* x). @@ -152,19 +151,19 @@ Lemma coset_mem y xbar : y \in xbar -> coset y = xbar. Proof. case: xbar => /= Hx NHx Hxy; apply: val_inj=> /=. case/rcosetsP: NHx (NHx) Hxy => x Nx -> NHx Hxy. -by rewrite val_insubd /= (rcoset_transl Hxy) NHx. +by rewrite val_insubd /= (rcoset_eqP Hxy) NHx. Qed. (* coset is an inverse to repr *) Lemma mem_repr_coset xbar : repr xbar \in xbar. -Proof. case: xbar => /= _ /rcosetsP[x _ ->]; exact: mem_repr_rcoset. Qed. +Proof. by case: xbar => /= _ /rcosetsP[x _ ->]; apply: mem_repr_rcoset. Qed. Lemma repr_coset1 : repr (1 : coset_of) = 1. Proof. exact: repr_group. Qed. Lemma coset_reprK : cancel (fun xbar => repr xbar) coset. -Proof. by move=> xbar; exact: coset_mem (mem_repr_coset xbar). Qed. +Proof. by move=> xbar; apply: coset_mem (mem_repr_coset xbar). Qed. (* cosetP is slightly stronger than using repr because we only *) (* guarantee repr xbar \in 'N(A) when A is a group. *) @@ -178,7 +177,7 @@ by exists x; last rewrite (coset_mem xbar_x). Qed. Lemma coset_id x : x \in A -> coset x = 1. -Proof. by move=> Ax; apply: coset_mem; exact: mem_gen. Qed. +Proof. by move=> Ax; apply: coset_mem; apply: mem_gen. Qed. Lemma im_coset : coset @* 'N(A) = setT. Proof. @@ -293,7 +292,7 @@ Lemma im_quotient : 'N(H) / H = setT. Proof. exact: im_coset. Qed. Lemma quotientT : setT / H = setT. -Proof. by rewrite -im_quotient; exact: morphimT. Qed. +Proof. by rewrite -im_quotient; apply: morphimT. Qed. (* Variant of morphimIdom. *) Lemma quotientInorm A : 'N_A(H) / H = A / H. @@ -359,18 +358,18 @@ Qed. Lemma rcoset_kercosetP x y : x \in 'N(H) -> y \in 'N(H) -> reflect (coset H x = coset H y) (x \in H :* y). -Proof. rewrite -{6}ker_coset; exact: rcoset_kerP. Qed. +Proof. by rewrite -{6}ker_coset; apply: rcoset_kerP. Qed. Lemma kercoset_rcoset x y : x \in 'N(H) -> y \in 'N(H) -> coset H x = coset H y -> exists2 z, z \in H & x = z * y. -Proof. by move=> Nx Ny eqfxy; rewrite -ker_coset; exact: ker_rcoset. Qed. +Proof. by move=> Nx Ny eqfxy; rewrite -ker_coset; apply: ker_rcoset. Qed. Lemma quotientGI G A : H \subset G -> (G :&: A) / H = G / H :&: A / H. -Proof. by rewrite -{1}ker_coset; exact: morphimGI. Qed. +Proof. by rewrite -{1}ker_coset; apply: morphimGI. Qed. Lemma quotientIG A G : H \subset G -> (A :&: G) / H = A / H :&: G / H. -Proof. by rewrite -{1}ker_coset; exact: morphimIG. Qed. +Proof. by rewrite -{1}ker_coset; apply: morphimIG. Qed. Lemma quotientD A B : A / H :\: B / H \subset (A :\: B) / H. Proof. exact: morphimD. Qed. @@ -379,16 +378,16 @@ Lemma quotientD1 A : (A / H)^# \subset A^# / H. Proof. exact: morphimD1. Qed. Lemma quotientDG A G : H \subset G -> (A :\: G) / H = A / H :\: G / H. -Proof. by rewrite -{1}ker_coset; exact: morphimDG. Qed. +Proof. by rewrite -{1}ker_coset; apply: morphimDG. Qed. Lemma quotientK A : A \subset 'N(H) -> coset H @*^-1 (A / H) = H * A. -Proof. by rewrite -{8}ker_coset; exact: morphimK. Qed. +Proof. by rewrite -{8}ker_coset; apply: morphimK. Qed. Lemma quotientYK G : G \subset 'N(H) -> coset H @*^-1 (G / H) = H <*> G. Proof. by move=> nHG; rewrite quotientK ?norm_joinEr. Qed. Lemma quotientGK G : H <| G -> coset H @*^-1 (G / H) = G. -Proof. by case/andP; rewrite -{1}ker_coset; exact: morphimGK. Qed. +Proof. by case/andP; rewrite -{1}ker_coset; apply: morphimGK. Qed. Lemma quotient_class x A : x \in 'N(H) -> A \subset 'N(H) -> x ^: A / H = coset H x ^: (A / H). @@ -400,7 +399,7 @@ Proof. exact: classes_morphim. Qed. Lemma cosetpre_set1 x : x \in 'N(H) -> coset H @*^-1 [set coset H x] = H :* x. -Proof. by rewrite -{9}ker_coset; exact: morphpre_set1. Qed. +Proof. by rewrite -{9}ker_coset; apply: morphpre_set1. Qed. Lemma cosetpre_set1_coset xbar : coset H @*^-1 [set xbar] = xbar. Proof. by case: (cosetP xbar) => x Nx ->; rewrite cosetpre_set1 ?val_coset. Qed. @@ -416,14 +415,14 @@ Lemma quotientS1 G : G \subset H -> G / H = 1. Proof. by move=> sGH; apply/trivgP; rewrite -trivg_quotient quotientS. Qed. Lemma sub_cosetpre M : H \subset coset H @*^-1 M. -Proof. by rewrite -{1}ker_coset; exact: ker_sub_pre. Qed. +Proof. by rewrite -{1}ker_coset; apply: ker_sub_pre. Qed. Lemma quotient_proper G K : H <| G -> H <| K -> (G / H \proper K / H) = (G \proper K). Proof. by move=> nHG nHK; rewrite -cosetpre_proper ?quotientGK. Qed. Lemma normal_cosetpre M : H <| coset H @*^-1 M. -Proof. rewrite -{1}ker_coset; exact: ker_normal_pre. Qed. +Proof. by rewrite -{1}ker_coset; apply: ker_normal_pre. Qed. Lemma cosetpreSK C D : (coset H @*^-1 C \subset coset H @*^-1 D) = (C \subset D). @@ -447,15 +446,15 @@ Proof. by move=> nHA; rewrite morphimSK ?ker_coset. Qed. Lemma quotientSGK A G : A \subset 'N(H) -> H \subset G -> (A / H \subset G / H) = (A \subset G). -Proof. by rewrite -{2}ker_coset; exact: morphimSGK. Qed. +Proof. by rewrite -{2}ker_coset; apply: morphimSGK. Qed. Lemma quotient_injG : {in [pred G : {group gT} | H <| G] &, injective (fun G => G / H)}. -Proof. by rewrite /normal -{1}ker_coset; exact: morphim_injG. Qed. +Proof. by rewrite /normal -{1}ker_coset; apply: morphim_injG. Qed. Lemma quotient_inj G1 G2 : H <| G1 -> H <| G2 -> G1 / H = G2 / H -> G1 :=: G2. -Proof. by rewrite /normal -{1 3}ker_coset; exact: morphim_inj. Qed. +Proof. by rewrite /normal -{1 3}ker_coset; apply: morphim_inj. Qed. Lemma quotient_neq1 A : H <| A -> (A / H != 1) = (H \proper A). Proof. @@ -494,7 +493,7 @@ Qed. Lemma quotient_cent1s A x : A \subset 'C[x] -> A / H \subset 'C[coset H x]. Proof. -by move=> sAC; exact: subset_trans (quotientS sAC) (quotient_cent1 x). +by move=> sAC; apply: subset_trans (quotientS sAC) (quotient_cent1 x). Qed. Lemma quotient_subcent1 A x : 'C_A[x] / H \subset 'C_(A / H)[coset H x]. @@ -515,7 +514,7 @@ Proof. exact: morphim_subcent. Qed. Lemma norm_quotient_pre A C : A \subset 'N(H) -> A / H \subset 'N(C) -> A \subset 'N(coset H @*^-1 C). Proof. -by move/sub_quotient_pre=> -> /subset_trans-> //; exact: morphpre_norm. +by move/sub_quotient_pre=> -> /subset_trans-> //; apply: morphpre_norm. Qed. Lemma cosetpre_normal C D : (coset H @*^-1 C <| coset H @*^-1 D) = (C <| D). @@ -638,7 +637,7 @@ Lemma quotient1_isom : isom A (A / 1) (coset 1). Proof. by apply: sub_isom coset1_injm; rewrite ?norms1. Qed. Lemma quotient1_isog : isog A (A / 1). -Proof. apply: isom_isog quotient1_isom; exact: norms1. Qed. +Proof. by apply: isom_isog quotient1_isom; apply: norms1. Qed. End Quotient1. @@ -805,10 +804,10 @@ exact: first_isom_loc. Qed. Lemma second_isog : H / (K :&: H) \isog H / K. -Proof. by rewrite setIC -{1 3}(ker_coset K); exact: first_isog_loc. Qed. +Proof. by rewrite setIC -{1 3}(ker_coset K); apply: first_isog_loc. Qed. Lemma weak_second_isog : H / (K :&: H) \isog H * K / K. -Proof. by rewrite quotientMidr; exact: second_isog. Qed. +Proof. by rewrite quotientMidr; apply: second_isog. Qed. End SecondIsomorphism. @@ -859,7 +858,7 @@ case/charP; rewrite quotientSGK // => sKG /= chKG. apply/charP; split=> // f injf Gf; apply/morphim_fixP => //. rewrite -(quotientSGK _ sHK); last by rewrite -morphimIim Gf subIset ?nHG. have{chHG} Hf: f @* H = H by case/charP: chHG => _; apply. -set q := quotm_morphism f nsHG; have{injf}: 'injm q by exact: injm_quotm. +set q := quotm_morphism f nsHG; have{injf}: 'injm q by apply: injm_quotm. have: q @* _ = _ := morphim_quotm _ _ _; move: q; rewrite Hf => q im_q injq. by rewrite -im_q chKG // im_q Gf. Qed. @@ -954,7 +953,7 @@ Proof. exact: dvdn_morphim. Qed. Lemma index_quotient_ker : K \subset G -> G \subset 'N(H) -> (#|G / H : K / H| * #|G :&: H : K|)%N = #|G : K|. -Proof. by rewrite -{5}(ker_coset H); exact: index_morphim_ker. Qed. +Proof. by rewrite -{5}(ker_coset H); apply: index_morphim_ker. Qed. Lemma index_quotient : G :&: K \subset 'N(H) -> #|G / H : K / H| %| #|G : K|. Proof. exact: index_morphim. Qed. diff --git a/mathcomp/odd_order/BGappendixAB.v b/mathcomp/odd_order/BGappendixAB.v index 4cbafd0..386429b 100644 --- a/mathcomp/odd_order/BGappendixAB.v +++ b/mathcomp/odd_order/BGappendixAB.v @@ -1,9 +1,16 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div. +From mathcomp Require Import fintype bigop prime finset ssralg fingroup morphism. +From mathcomp Require Import automorphism quotient gfunctor commutator zmodp center pgroup. +From mathcomp Require Import sylow gseries nilpotent abelian maximal. +From mathcomp Require Import matrix mxalgebra mxrepresentation mxabelem. +From mathcomp Require Import BGsection1 BGsection2. (******************************************************************************) @@ -72,7 +79,7 @@ have [sQG qQ]: Q \subset G /\ q.-group Q by case/and3P: sylQ. have{qQ p'q} p'Q: p^'.-group Q by apply: sub_in_pnat qQ => q' _ /eqnP->. have{q q_pr sylQ qGc} ncEQ: ~~ (Q \subset 'C(E)). apply: contraL qGc => cEQ; rewrite -p'natE // -partn_eq1 //. - have nCQ: Q \subset 'N('C(E)) by exact: subset_trans (normG _). + have nCQ: Q \subset 'N('C(E)) by apply: subset_trans (normG _). have sylQc: q.-Sylow(G / 'C(E)) (Q / 'C(E)) by rewrite morphim_pSylow. by rewrite -(card_Hall sylQc) -trivg_card1 (sameP eqP trivgP) quotient_sub1. have solE: solvable E := pgroup_sol pE. @@ -145,13 +152,13 @@ have cAG: centgmx rG A. rewrite 2!(mulmxDl _ 1 A) 2!(mulmxDr A _ 1) !mulmx1 !mul1mx. rewrite !(inj_eq (addIr A)) ![_ *m A]mulmxDr ![A *m _]mulmxDl. by rewrite -!mulmxA Ax2 Ay2 !mulmx0 !mulmxA Ax2 Ay2 !mul0mx !addr0 !add0r. -have irrG: mx_irreducible rG by exact/abelem_mx_irrP. +have irrG: mx_irreducible rG by apply/abelem_mx_irrP. pose rAG := gen_repr irrG cAG; pose inFA := in_gen irrG cAG. pose valFA := @val_gen _ _ _ _ _ _ irrG cAG. set dA := gen_dim A in rAG inFA valFA. rewrite -(rker_abelem abelE ntE nEG) -/rG -(rker_gen irrG cAG) -/rAG. have dA_gt0: dA > 0 by rewrite (gen_dim_gt0 irrG cAG). -have irrAG: mx_irreducible rAG by exact: gen_mx_irr. +have irrAG: mx_irreducible rAG by apply: gen_mx_irr. have: dA <= 2. case Ax0: (Ax == 0). by rewrite subr_eq0 in Ax0; case/eqP: ncxy; rewrite (eqP Ax0) mulmx1 mul1mx. @@ -226,14 +233,14 @@ Proof. set Q := 'O_p(G) => p'G1 sCQ_P. have sPQ: P \subset Q by rewrite pcore_max. have defQ: 'O_{p^', p}(G) = Q by rewrite pseries_pop2. -have pQ: p.-group Q by exact: pcore_pgroup. +have pQ: p.-group Q by apply: pcore_pgroup. have sCQ: 'C_G(Q) \subset Q. by rewrite -{2}defQ solvable_p_constrained //= defQ /pHall pQ indexgg subxx. have pC: p.-group C. apply/pgroupP=> q q_pr; case/Cauchy=> // u Cu q_u; apply/idPn=> p'q. suff cQu: u \in 'C_G(Q). case/negP: p'q; have{q_u}: q %| #[u] by rewrite q_u. - by apply: pnatP q q_pr => //; apply: mem_p_elt pQ _; exact: (subsetP sCQ). + by apply: pnatP q q_pr => //; apply: mem_p_elt pQ _; apply: (subsetP sCQ). have [Gu cPu] := setIP Cu; rewrite inE Gu /= -cycle_subG. rewrite coprime_nil_faithful_cent_stab ?(pgroup_nil pQ) //= -/C -/Q. - by rewrite cycle_subG; apply: subsetP Gu; rewrite normal_norm ?pcore_normal. @@ -262,7 +269,7 @@ Lemma Puig_char G : 'L(G) \char G. Proof. exact: gFchar. Qed. Lemma center_Puig_char G : 'Z('L(G)) \char G. -Proof. exact: char_trans (center_char _) (Puig_char _). Qed. +Proof. by rewrite !gFchar_trans. Qed. (* This is B & G, Lemma B.1(a). *) Lemma Puig_succS G D E : D \subset E -> 'L_[G](E) \subset 'L_[G](D). @@ -408,7 +415,7 @@ have sCT_L: 'C_T('L_{k.*2.+1}(T)) \subset 'L_{k.*2.+1}(T). have{sCT_L} sLT: 'L_{k.*2.+2}(S) \subset T. apply: odd_abelian_gen_constrained sCT_L => //. - exact: pgroupS (Puig_at_sub _ _) pT. - - by apply: char_normal_trans nsTG; apply: gFchar. + - exact: gFnormal_trans nsTG. - exact: sL_ sSG. by rewrite norm_abgen_pgroup // (pgroupS _ pS) ?Puig_at_sub. have sL2: 'L_{k.*2.+2}(S) \subset 'L_{k.*2.+2}(T) by apply: Puig_max. @@ -423,17 +430,17 @@ Let L := 'L(S). Theorem Puig_center_normal : 'Z(L) <| G. Proof. have [sLiST sLTS] := pcore_Sylow_Puig_sub. -have sLiLT: 'L_*(T) \subset 'L(T) by exact: Puig_sub_even_odd. +have sLiLT: 'L_*(T) \subset 'L(T) by apply: Puig_sub_even_odd. have sZY: 'Z(L) \subset Y. rewrite subsetI andbC subIset ?centS ?orbT //=. suffices: 'C_S('L_*(S)) \subset 'L(T). by apply: subset_trans; rewrite setISS ?Puig_sub ?centS ?Puig_sub_even_odd. apply: subset_trans (subset_trans sLiST sLiLT). by apply: sub_cent_Puig_at pS; rewrite double_gt0. -have chY: Y \char G := char_trans (center_Puig_char _) (pcore_char _ _). +have chY: Y \char G by rewrite !gFchar_trans. have nsCY_G: 'C_G(Y) <| G by rewrite char_normal 1?subcent_char ?char_refl. have [C defC sCY_C nsCG] := inv_quotientN nsCY_G (pcore_normal p _). -have sLG: L \subset G by rewrite (subset_trans _ (pHall_sub sylS)) ?Puig_sub. +have sLG: L \subset G by rewrite gFsub_trans ?(pHall_sub sylS). have nsL_nCS: L <| 'N_G(C :&: S). have sYLiS: Y \subset 'L_*(S). rewrite abelian_norm_Puig ?double_gt0 ?center_abelian //. @@ -446,8 +453,8 @@ have nsL_nCS: L <| 'N_G(C :&: S). rewrite odd_abelian_gen_stable ?char_normal ?norm_abgen_pgroup //. by rewrite (pgroupS _ pT) ?subIset // Puig_sub. by rewrite (pgroupS _ pS) ?Puig_sub. - rewrite -[L](sub_Puig_eq _ sLCS) ?subsetIr //. - by rewrite (char_normal_trans (Puig_char _)) ?normalSG // subIset // sSG orbT. + rewrite -[L](sub_Puig_eq _ sLCS) ?subsetIr // gFnormal_trans ?normalSG //. + by rewrite subIset // sSG orbT. have sylCS: p.-Sylow(C) (C :&: S) := Sylow_setI_normal nsCG sylS. have{defC} defC: 'C_G(Y) * (C :&: S) = C. apply/eqP; rewrite eqEsubset mulG_subG sCY_C subsetIl /=. @@ -461,7 +468,7 @@ have defG: 'C_G(Y) * 'N_G(C :&: S) = G. have sCS_N: C :&: S \subset 'N_G(C :&: S). by rewrite subsetI normG subIset // sSG orbT. by rewrite -(mulSGid sCS_N) mulgA defC (Frattini_arg _ sylCS). -have nsZ_N: 'Z(L) <| 'N_G(C :&: S) := char_normal_trans (center_char _) nsL_nCS. +have nsZ_N: 'Z(L) <| 'N_G(C :&: S) := gFnormal_trans _ nsL_nCS. rewrite /normal subIset ?sLG //= -{1}defG mulG_subG /=. rewrite cents_norm ?normal_norm // centsC. by rewrite (subset_trans sZY) // centsC subsetIr. @@ -487,7 +494,7 @@ have def_Zq: Z / D = 'Z('L(S / D)). rewrite -!(injmF _ injf) ?Puig_sub //= morphim_restrm. by rewrite (setIidPr _) // subIset ?Puig_sub. have{def_Zq} nZq: Z / D <| G / D. - have sylSq: p.-Sylow(G / D) (S / D) by exact: morphim_pHall. + have sylSq: p.-Sylow(G / D) (S / D) by apply: morphim_pHall. rewrite def_Zq (Puig_center_normal _ _ sylSq) ?quotient_odd ?quotient_sol //. exact: trivg_pcore_quotient. have sZS: Z \subset S by rewrite subIset ?Puig_sub. diff --git a/mathcomp/odd_order/BGappendixC.v b/mathcomp/odd_order/BGappendixC.v index 31bccfb..75de45b 100644 --- a/mathcomp/odd_order/BGappendixC.v +++ b/mathcomp/odd_order/BGappendixC.v @@ -1,9 +1,18 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype choice ssrnat seq div fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype choice ssrnat seq div fintype. +From mathcomp Require Import tuple finfun bigop ssralg finset prime binomial poly polydiv. +From mathcomp Require Import fingroup morphism quotient automorphism action finalg zmodp. -Require Import gproduct cyclic commutator pgroup abelian frobenius BGsection1. +From mathcomp +Require Import gfunctor gproduct cyclic commutator pgroup abelian frobenius. +From mathcomp +Require Import BGsection1. +From mathcomp Require Import matrix mxalgebra mxabelem vector falgebra fieldext galois. +From mathcomp Require Import finfield ssrnum algC classfun character integral_char inertia. Set Implicit Arguments. @@ -147,7 +156,7 @@ Proof. by rewrite cyclic_abelian // -(injm_cyclic inj_sigmaU) ?cycFU. Qed. (* This is B & G, Appendix C, Remark VII. *) Let im_psi (x : F) : (x \in psi @: U) = (Nm x == 1). Proof. -have /cyclicP[u0 defFU]: cyclic [set: {unit F}] by exact: cycFU. +have /cyclicP[u0 defFU]: cyclic [set: {unit F}] by apply: cycFU. have o_u0: #[u0] = (p ^ q).-1 by rewrite orderE -defFU card_finField_unit oF. have ->: psi @: U = uval @: (sigmaU @* U) by rewrite morphimEdom -imset_comp. have /set1P[->]: (sigmaU @* U)%G \in [set <[u0 ^+ (#[u0] %/ nU)]>%G]. @@ -528,7 +537,7 @@ have{nHt1} nHP1: P1 \subset 'N(H). by apply/trivgPn; exists t1; rewrite // inE P1t1. have{nHP1} nPP1: P1 \subset 'N(P). have /Hall_pi hallP: Hall H P by apply: Frobenius_ker_Hall frobH. - by rewrite -(normal_Hall_pcore hallP nsPH) (char_norm_trans (pcore_char _ _)). + by rewrite -(normal_Hall_pcore hallP nsPH) gFnorm_trans. have sylP0: p.-Sylow(Q <*> P0) P0. rewrite /pHall -divgS joing_subr ?(pgroupS sP0P) //=. by rewrite norm_joinEr // coprime_cardMg ?(coprimegS sP0P) ?mulnK. @@ -552,7 +561,7 @@ Fact BGappendixC3_Ediv : E = [set x^-1 | x in E]%R. Proof. suffices sEV_E: [set x^-1 | x in E]%R \subset E. by apply/esym/eqP; rewrite eqEcard sEV_E card_imset //=; apply: invr_inj. -have /mulG_sub[/(subset_trans sP0P)/subsetP sP0H /subsetP sUH] := sdprodW defH. +have /mulG_sub[/(subset_trans sP0P)/subsetP-sP0H /subsetP-sUH] := sdprodW defH. have Hs := sP0H s P0s; have P1t: t \in P1 by rewrite memJ_conjg. have nUP1 t1: t1 \in P1 -> U :^ t1 = U by move/(subsetP nUP0y)/normP. have nUtn n u: u \in U -> u ^ (t ^+ n) \in U. diff --git a/mathcomp/odd_order/BGsection1.v b/mathcomp/odd_order/BGsection1.v index 98a2d08..ebce90d 100644 --- a/mathcomp/odd_order/BGsection1.v +++ b/mathcomp/odd_order/BGsection1.v @@ -1,10 +1,16 @@ -Set Printing Width 50. (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div fintype. +From mathcomp Require Import bigop prime binomial finset fingroup morphism perm automorphism. +From mathcomp Require Import quotient action gproduct gfunctor commutator. +From mathcomp Require Import ssralg finalg zmodp cyclic center pgroup finmodule gseries. +From mathcomp Require Import nilpotent sylow abelian maximal hall extremal. +From mathcomp Require Import matrix mxalgebra mxrepresentation mxabelem. (******************************************************************************) @@ -135,13 +141,12 @@ Proof. by move=> minM solM; case: (minnormal_solvable minM (subxx _) solM). Qed. Lemma minnormal_solvable_Fitting_center gT (M G : {group gT}) : minnormal M G -> M \subset G -> solvable M -> M \subset 'Z('F(G)). Proof. -have nZG: 'Z('F(G)) <| G. - by apply: (char_normal_trans (center_char _)); exact: Fitting_normal. +have nZG: 'Z('F(G)) <| G by rewrite !gFnormal_trans. move=> minM sMG solM; have[/andP[ntM nMG] minM'] := mingroupP minM. apply/setIidPl/minM'; last exact: subsetIl. apply/andP; split; last by rewrite normsI // normal_norm. apply: meet_center_nil => //; first by apply: Fitting_nil. -apply/andP; split; last by apply: subset_trans nMG; exact: Fitting_sub. +apply/andP; split; last exact: gFsub_trans. apply: Fitting_max; rewrite // /normal ?sMG //; apply: abelian_nil. by move: (minnormal_solvable_abelem minM solM) => /abelem_abelian. Qed. @@ -160,10 +165,10 @@ Variables (gT : finGroupType) (G G' : {group gT}). Hypothesis solG : solvable G. Hypothesis nsG'G : G' <| G. -Let sG'G : G' \subset G := normal_sub nsG'G. -Let nG'G : G \subset 'N(G') := normal_norm nsG'G. -Let nsF'G : 'F(G') <| G := char_normal_trans (Fitting_char G') nsG'G. +Let sG'G : G' \subset G. Proof. exact: normal_sub. Qed. +Let nG'G : G \subset 'N(G'). Proof. exact: normal_norm. Qed. +Let nsF'G : 'F(G') <| G. Proof. exact: gFnormal_trans. Qed. Let Gchief (UV : {group gT} * {group gT}) := chief_factor G UV.2 UV.1. Let H := \bigcap_(UV | Gchief UV) 'C(UV.1 / UV.2 | 'Q). @@ -179,8 +184,8 @@ have{chiefUV} [/=/maxgroupp/andP[_ nVG] sUG nUG] := and3P chiefUV. have solUV: solvable (U / V) by rewrite quotient_sol // (solvableS sUG). have{solUV minUV}: U / V \subset 'Z('F(G / V)). exact: minnormal_solvable_Fitting_center minUV (quotientS V sUG) solUV. -rewrite sub_astabQ (subset_trans (normal_sub nsF'G) nVG) /=. -rewrite subsetI centsC => /andP[_]; apply: subset_trans. +rewrite sub_astabQ gFsub_trans ?(subset_trans sG'G) //=. +case/subsetIP=> _; rewrite centsC; apply: subset_trans. by rewrite Fitting_max ?quotient_normal ?quotient_nil ?Fitting_nil. Qed. @@ -237,14 +242,14 @@ move=> nGA coGA solG regAG; without loss cycA: A nGA coGA regAG / cyclic A. by apply/trivgP; rewrite -regAG setSI. pose X := G <*> A; pose F := 'F(X); pose pi := \pi(A); pose Q := 'O_pi(F). have pi'G: pi^'.-group G by rewrite /pgroup -coprime_pi' //= coprime_sym. -have piA: pi.-group A by exact: pgroup_pi. +have piA: pi.-group A by apply: pgroup_pi. have oX: #|X| = (#|G| * #|A|)%N by rewrite [X]norm_joinEr ?coprime_cardMg. have hallG: pi^'.-Hall(X) G. by rewrite /pHall -divgS joing_subl //= pi'G pnatNK oX mulKn. have nsGX: G <| X by rewrite /normal joing_subl join_subG normG. have{oX pi'G piA} hallA: pi.-Hall(X) A. by rewrite /pHall -divgS joing_subr //= piA oX mulnK. -have nsQX: Q <| X := char_normal_trans (pcore_char _ _) (Fitting_normal _). +have nsQX: Q <| X by rewrite !gFnormal_trans. have{solG cycA} solX: solvable X. rewrite (series_sol nsGX) {}solG /= norm_joinEr // quotientMidl //. by rewrite morphim_sol // abelian_sol // cyclic_abelian. @@ -260,7 +265,7 @@ have sFG: F \subset G. by rewrite (sub_normal_Hall hallG) ?gFsub //= -defF pi'F mul1g pcore_pgroup. have <-: F = 'F(G). apply/eqP; rewrite eqEsubset -{1}(setIidPr sFG) FittingS ?joing_subl //=. - by rewrite Fitting_max ?Fitting_nil // (char_normal_trans (Fitting_char _)). + by rewrite Fitting_max ?Fitting_nil // gFnormal_trans. apply/trivgP; rewrite /= -(coprime_TIg coGA) subsetI subsetIl andbT. apply: subset_trans (subset_trans (cent_sub_Fitting solX) sFG). by rewrite setSI ?joing_subr. @@ -442,20 +447,20 @@ without loss defR: G pG nGA coGA cAG1 / [~: G, A] = G. rewrite -(coprime_cent_prod nGA) ?mul_subG ?subsetIr //=. have sRG: [~: G, A] \subset G by rewrite commg_subl. rewrite IH ?coprime_commGid ?(pgroupS sRG) ?commg_normr ?(coprimeSg sRG) //. - by rewrite (subset_trans (OhmS 1 sRG)). + by apply: subset_trans cAG1; apply: OhmS. have [|[defPhi defG'] defC] := abelian_charsimple_special pG coGA defR. apply/bigcupsP=> H /andP[chH abH]; have sHG := char_sub chH. have nHA := char_norm_trans chH nGA. rewrite centsC coprime_abelian_faithful_Ohm1 ?(coprimeSg sHG) //. by rewrite centsC (subset_trans (OhmS 1 sHG)). -have abelZ: p.-abelem 'Z(G) by exact: center_special_abelem. +have abelZ: p.-abelem 'Z(G) by apply: center_special_abelem. have cAZ: {in 'Z(G), centralised A} by apply/centsP; rewrite -defC subsetIr. have cGZ: {in 'Z(G), centralised G} by apply/centsP; rewrite subsetIr. have defG1: 'Ohm_1(G) = 'Z(G). apply/eqP; rewrite eqEsubset -{1}defC subsetI Ohm_sub cAG1 /=. by rewrite -(Ohm1_id abelZ) OhmS ?center_sub. rewrite (subset_trans _ (subsetIr G _)) // defC -defG1 -{1}defR gen_subG /=. -apply/subsetP=> xa; case/imset2P=> x a Gx Aa ->{xa}; rewrite commgEl. +apply/subsetP=> _ /imset2P[x a Gx Aa ->]; rewrite commgEl. set u := x^-1; set v := x ^ a; pose w := [~ v, u]. have [Gu Gv]: u \in G /\ v \in G by rewrite groupV memJ_norm ?(subsetP nGA). have Zw: w \in 'Z(G) by rewrite -defG' mem_commg. @@ -526,7 +531,7 @@ have expH': {in H &, forall y z, [~ y, z] ^+ p = 1}. move=> y z Hy Hz; apply/eqP. have /setIP[_ cHyz]: [~ y, z] \in 'Z(H) by rewrite (subsetP clH) // mem_commg. rewrite -commXg; last exact/commute_sym/(centP cHyz). - suffices /setIP[_ cHyp]: y ^+ p \in 'Z(H) by exact/commgP/(centP cHyp). + suffices /setIP[_ cHyp]: y ^+ p \in 'Z(H) by apply/commgP/(centP cHyp). rewrite (subsetP sPhiZ) // (Phi_joing pH) mem_gen // inE orbC. by rewrite (Mho_p_elt 1) ?(mem_p_elt pH). have Hfx: f x \in H. @@ -687,7 +692,7 @@ Proof. case/andP=> sRG pR solG. without loss p'G1: gT G R sRG pR solG / 'O_p^'(G) = 1. have nOG_CR: 'C_G(R) \subset 'N('O_p^'(G)) by rewrite subIset ?gFnorm. - move=> IH; rewrite -quotient_sub1 ?(subset_trans (pcore_sub _ _)) //. + move=> IH; rewrite -quotient_sub1 ?gFsub_trans //. apply: subset_trans (morphimF _ _ nOG_CR) _; rewrite /= -quotientE. rewrite -(coprime_subcent_quotient_pgroup pR) ?pcore_sub //; first 1 last. - by rewrite (subset_trans sRG) ?gFnorm. @@ -695,7 +700,7 @@ without loss p'G1: gT G R sRG pR solG / 'O_p^'(G) = 1. have p'Gq1 : 'O_p^'(G / 'O_p^'(G)) = 1 := trivg_pcore_quotient p^' G. by rewrite -p'Gq1 IH ?morphimS ?morphim_pgroup ?morphim_sol. set M := 'O_p^'('C_G(R)); pose T := 'O_p(G). -have /subsetIP[sMG cMR]: M \subset 'C_G(R) by exact: pcore_sub. +have /subsetIP[sMG cMR]: M \subset 'C_G(R) by apply: pcore_sub. have [p'M pT]: p^'.-group M /\ p.-group T by rewrite !pcore_pgroup. have nRT: R \subset 'N(T) by rewrite (subset_trans sRG) ?gFnorm. have pRT: p.-group (R <*> T). @@ -731,7 +736,7 @@ move=> abelA nGA coGA; symmetry; move: {2}_.+1 (ltnSn #|G|) => n. elim: n gT => // n IHn gT in A G abelA nGA coGA *; rewrite ltnS => leGn. without loss nilG: G nGA coGA leGn / nilpotent G. move=> {IHn} IHn; apply/eqP; rewrite eqEsubset gen_subG. - apply/andP; split; last by apply/bigcupsP=> B _; exact: subsetIl. + apply/andP; split; last by apply/bigcupsP=> B _; apply: subsetIl. pose T := [set P : {group gT} | Sylow G P & A \subset 'N(P)]. rewrite -{1}(@Sylow_transversal_gen _ T G) => [|P | p _]; first 1 last. - by rewrite inE -!andbA; case/and4P. @@ -743,11 +748,10 @@ without loss nilG: G nGA coGA leGn / nilpotent G. by apply/bigcupsP=> B cycBq; rewrite (bigcup_max B) ?setSI. by rewrite (leq_trans (subset_leq_card sPG)). apply/eqP; rewrite eqEsubset gen_subG. -apply/andP; split; last by apply/bigcupsP=> B _; exact: subsetIl. +apply/andP; split; last by apply/bigcupsP=> B _; apply: subsetIl. have [Z1 | ntZ] := eqsVneq 'Z(G) 1. by rewrite (TI_center_nil _ (normal_refl G)) ?Z1 ?(setIidPr _) ?sub1G. -have nZA: A \subset 'N('Z(G)) := char_norm_trans (center_char G) nGA. -have{ntZ nZA} [M /= minM] := minnormal_exists ntZ nZA. +have{ntZ} [M /= minM] := minnormal_exists ntZ (gFnorm_trans _ nGA). rewrite subsetI centsC => /andP[sMG /cents_norm nMG]. have coMA := coprimeSg sMG coGA; have{nilG} solG := nilpotent_sol nilG. have [nMA ntM abelM] := minnormal_solvable minM sMG solG. @@ -780,7 +784,7 @@ Proposition coprime_abelian_gen_cent1 gT (A G : {group gT}) : Proof. move=> abelA ncycA nGA coGA. apply/eqP; rewrite eq_sym eqEsubset /= gen_subG. -apply/andP; split; last by apply/bigcupsP=> B _; exact: subsetIl. +apply/andP; split; last by apply/bigcupsP=> B _; apply: subsetIl. rewrite -{1}(coprime_abelian_gen_cent abelA nGA) ?genS //. apply/bigcupsP=> B; have [-> | /trivgPn[a Ba n1a]] := eqsVneq B 1. by rewrite injm_cyclic ?coset1_injm ?norms1 ?(negbTE ncycA). @@ -819,9 +823,9 @@ have abelSK : abelian (alpha @* S). set ker_trans := 'ker (transfer G abelSK). have G'ker : G' \subset ker_trans. rewrite gen_subG; apply/subsetP=> h; case/imset2P=> h1 h2 Gh1 Gh2 ->{h}. - by rewrite !inE groupR // morphR //; apply/commgP; exact: addrC. + by rewrite !inE groupR // morphR //; apply/commgP; apply: addrC. have transg0: transfer G abelSK g = 0%R. - by move/kerP: (subsetP G'ker g G'g); exact. + by move/kerP: (subsetP G'ker g G'g); apply. have partX := rcosets_cycle_partition sSG Gg. have trX := transversalP partX; set X := transversal _ _ in trX. have /and3P[_ sXG _] := trX. @@ -850,9 +854,9 @@ Theorem Burnside_normal_complement : 'N_G(S) \subset 'C(S) -> 'O_p^'(G) ><| S = G. Proof. move=> cSN; set K := 'O_p^'(G); have [sSG pS _] := and3P sylS. -have /andP[sKG nKG]: K <| G by exact: pcore_normal. +have /andP[sKG nKG]: K <| G by apply: pcore_normal. have{nKG} nKS := subset_trans sSG nKG. -have p'K: p^'.-group K by exact: pcore_pgroup. +have p'K: p^'.-group K by apply: pcore_pgroup. have{pS p'K} tiKS: K :&: S = 1 by rewrite setIC coprime_TIg ?(pnat_coprime pS). suffices{tiKS nKS} hallK: p^'.-Hall(G) K. rewrite sdprodE //= -/K; apply/eqP; rewrite eqEcard ?mul_subG //=. @@ -903,7 +907,7 @@ have /sdprodP[_ /= defG _ _] := Burnside_normal_complement cSN. set Q := 'O_p^'(G) in defG; have nQG: G \subset 'N(Q) := gFnorm _ _. left; rewrite coprime_TIg ?(pnat_coprime pS) //. apply: pgroupS (pcore_pgroup _ G); rewrite /= -/Q. -rewrite -quotient_sub1 ?(subset_trans (der_sub _ _)) ?quotientR //= -/Q. +rewrite -quotient_sub1 ?gFsub_trans ?quotientR //= -/Q. rewrite -defG quotientMidl (sameP trivgP commG1P) -abelianE. by rewrite morphim_abelian ?cyclic_abelian. Qed. @@ -935,7 +939,7 @@ Lemma cyclic_pdiv_normal_complement gT (S G : {group gT}) : Proof. set p := pdiv _ => sylS cycS; have cSS := cyclic_abelian cycS. exists 'O_p^'(G)%G; apply: Burnside_normal_complement => //. -have [-> | ntS] := eqsVneq S 1; first exact: cents1. +have [-> | ntS] := eqsVneq S 1; first apply: cents1. have [sSG pS p'iSG] := and3P sylS; have [pr_p _ _] := pgroup_pdiv pS ntS. rewrite -['C(S)]mulg1 -ker_conj_aut -morphimSK ?subsetIr // setIC morphimIdom. set A_G := _ @* _; pose A := Aut S. @@ -955,7 +959,7 @@ Lemma Zgroup_metacyclic gT (G : {group gT}) : Zgroup G -> metacyclic G. Proof. elim: {G}_.+1 {-2}G (ltnSn #|G|) => // n IHn G; rewrite ltnS => leGn ZgG. have{n IHn leGn} solG: solvable G. - have [-> | ntG] := eqsVneq G 1; first exact: solvable1. + have [-> | ntG] := eqsVneq G 1; first apply: solvable1. have [S sylS] := Sylow_exists (pdiv #|G|) G. have cycS: cyclic S := forall_inP ZgG S (p_Sylow sylS). have [H defG] := cyclic_pdiv_normal_complement sylS cycS. @@ -1074,7 +1078,7 @@ Lemma p'quo_plength1 G H : Proof. rewrite /plength_1 => nHG p'H; apply/idP/idP; last exact: plength1_quo. move=> pGH1; rewrite eqEsubset pseries_sub. -have nOG: 'O_{p^'}(G) <| G by exact: pseries_normal. +have nOG: 'O_{p^'}(G) <| G by apply: pseries_normal. rewrite -(quotientSGK (normal_norm nOG)) ?(pseries_sub_catl [:: _]) //. have [|f f_inj im_f] := third_isom _ nHG nOG. by rewrite /= pseries1 pcore_max. @@ -1092,11 +1096,11 @@ rewrite /plength_1 => nHG pH trO; apply/idP/idP; last exact: plength1_quo. rewrite (pseries_pop _ trO) => pGH1; rewrite eqEsubset pseries_sub /=. rewrite pseries_pop //; last first. apply/eqP; rewrite -subG1; have <-: H :&: 'O_p^'(G) = 1. - by apply: coprime_TIg; exact: pnat_coprime (pcore_pgroup _ _). + by apply: coprime_TIg; apply: pnat_coprime (pcore_pgroup _ _). rewrite setIC subsetI subxx -quotient_sub1. by rewrite -trO morphim_pcore. - apply: subset_trans (normal_norm nHG); exact: pcore_sub. -have nOG: 'O_{p}(G) <| G by exact: pseries_normal. + exact/gFsub_trans/normal_norm. +have nOG: 'O_{p}(G) <| G by apply: pseries_normal. rewrite -(quotientSGK (normal_norm nOG)) ?(pseries_sub_catl [:: _]) //. have [|f f_inj im_f] := third_isom _ nHG nOG. by rewrite /= pseries1 pcore_max. @@ -1137,26 +1141,24 @@ apply/idP/idP=> [p1G | pU]. rewrite [#[_]](@pnat_1 p) //; first exact: morph_p_elt. apply: mem_p_elt (pcore_pgroup _ (G / _)) _. by rewrite /= -quotient_pseries /= (eqP p1G); apply/morphimP; exists x. -have nOG: 'O_{p^', p}(G) <| G by exact: pseries_normal. +have nOG: 'O_{p^', p}(G) <| G by apply: pseries_normal. rewrite eqEsubset pseries_sub. rewrite -(quotientSGK (normal_norm nOG)) ?(pseries_sub_catl [:: _; _]) //=. rewrite (quotient_pseries [::_;_]) pcore_max //. rewrite /pgroup card_quotient ?normal_norm //. -apply: (@pnat_dvd _ #|G : p_elt_gen p G|); last first. +apply: pnat_dvd (indexgS G (_ : p_elt_gen p G \subset _)) _; last first. case p_pr: (prime p); last by rewrite p'natEpi // mem_primes p_pr. rewrite -card_quotient // p'natE //; apply/negP=> /Cauchy[] // Ux. case/morphimP=> x Nx Gx -> /= oUx_p; have:= prime_gt1 p_pr. rewrite -(part_pnat_id (pnat_id p_pr)) -{1}oUx_p {oUx_p} -order_constt. rewrite -morph_constt //= coset_id ?order1 //. by rewrite mem_gen // inE groupX // p_elt_constt. -apply: indexgS. have nOU: p_elt_gen p G \subset 'N('O_{p^'}(G)). by rewrite (subset_trans sUG) // normal_norm ?pseries_normal. rewrite -(quotientSGK nOU) ?(pseries_sub_catl [:: _]) //=. rewrite (quotient_pseries [::_]) pcore_max ?morphim_normal //. rewrite /pgroup card_quotient //= pseries1; apply: pnat_dvd pU. -apply: indexgS; rewrite pcore_max ?pcore_pgroup //. -apply: char_normal_trans nUG; exact: pcore_char. +by apply: indexgS; rewrite pcore_max ?pcore_pgroup // gFnormal_trans. Qed. End Plength1. @@ -1169,7 +1171,7 @@ Proof. move=> nHG nKG trHK. have [p_pr | p_nonpr] := boolP (prime p); last by rewrite !plength1_nonprime. apply/andP/idP=> [[pH1 pK1] | pG1]; last by rewrite !plength1_quo. -pose U := p_elt_gen p G; have nU : U <| G by exact: p_elt_gen_normal. +pose U := p_elt_gen p G; have nU : U <| G by apply: p_elt_gen_normal. have exB (N : {group gT}) : N <| G -> p.-length_1 (G / N) -> exists B : {group gT}, @@ -1183,7 +1185,7 @@ have exB (N : {group gT}) : congr <<_>>; apply/setP=> Nx; rewrite inE setIdE quotientGI // inE. apply: andb_id2l => /morphimP[x NNx Gx ->{Nx}] /=. apply/idP/idP=> [pNx | /morphimP[y NNy]]; last first. - by rewrite inE => p_y ->; exact: morph_p_elt. + by rewrite inE => p_y ->; apply: morph_p_elt. rewrite -(constt_p_elt pNx) -morph_constt // mem_morphim ?groupX //. by rewrite inE p_elt_constt. have nNU: U \subset 'N(N) := subset_trans (normal_sub nU) nNG. @@ -1242,7 +1244,7 @@ Implicit Type gT : finGroupType. Canonical Puig_succ_group gT (D E : {set gT}) := [group of 'L_[D](E)]. Fact Puig_at_group_set n gT D : @group_set gT 'L_{n}(D). -Proof. case: n => [|n]; exact: groupP. Qed. +Proof. by case: n => [|n]; apply: groupP. Qed. Canonical Puig_at_group n gT D := Group (@Puig_at_group_set n gT D). Canonical Puig_inf_group gT (D : {set gT}) := [group of 'L_*(D)]. diff --git a/mathcomp/odd_order/BGsection10.v b/mathcomp/odd_order/BGsection10.v index 88b4c39..f05aeb4 100644 --- a/mathcomp/odd_order/BGsection10.v +++ b/mathcomp/odd_order/BGsection10.v @@ -1,9 +1,16 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div path fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div path fintype. +From mathcomp Require Import bigop finset prime fingroup morphism perm automorphism quotient. +From mathcomp Require Import action gproduct gfunctor pgroup cyclic center commutator. +From mathcomp Require Import gseries nilpotent sylow abelian maximal hall. +From mathcomp Require Import BGsection1 BGsection3 BGsection4 BGsection5 BGsection6. +From mathcomp Require Import BGsection7 BGsection9. (******************************************************************************) @@ -143,7 +150,7 @@ exact: def_uniq_mmax (rank3_Uniqueness (mFT_pgroup_proper pP) rP) maxM sPM. Qed. Remark beta_sub_sigma : {subset \beta(M) <= \sigma(M)}. -Proof. by move=> p; move/beta_sub_alpha; exact: alpha_sub_sigma. Qed. +Proof. by move=> p; move/beta_sub_alpha; apply: alpha_sub_sigma. Qed. Remark Mbeta_sub_Malpha : M`_\beta \subset M`_\alpha. Proof. exact: sub_pcore beta_sub_alpha. Qed. @@ -291,7 +298,7 @@ have M0tC_M2: M2 \in orbit 'Js 'C(X) (M0 :^ t). rewrite (atransP (IHX _ ltXX2 pX2)) inE; first by case/setIdP: O_M2 => ->. rewrite (acts_act (acts_orbit _ _ _)) ?inE ?MG_M0 //. by rewrite (subset_trans sX2Pt) ?conjSg. -rewrite (orbit_transl M0C_M1) (orbit_transr _ M0tC_M2). +rewrite (orbit_eqP M0C_M1) (orbit_transl _ M0tC_M2). have maxM0 := maxOM _ _ O_M0; have ltMG := mmax_proper maxM0. have [rPgt2 | rPle2] := ltnP 2 'r(P). have uP: P \in 'U by rewrite rank3_Uniqueness ?(mFT_pgroup_proper pP). @@ -302,7 +309,7 @@ have pl1L: p.-length_1 L. by case/rank2_der1_complement; rewrite ?mFT_sol ?plength1_pseries2_quo. have [|u v nLPu Lp'_v ->] := imset2P (_ : t \in 'N_L(P) * 'O_p^'(L)). by rewrite normC ?plength1_Frattini // subIset ?gFnorm. -rewrite actM (orbit_transr _ (mem_orbit _ _ _)); last first. +rewrite actM (orbit_transl _ (mem_orbit _ _ _)); last first. have coLp'X: coprime #|'O_p^'(L)| #|X| := p'nat_coprime (pcore_pgroup _ _) pX. apply: subsetP Lp'_v; have [sLp'L nLp'L] := andP (pcore_normal p^' L). rewrite -subsetIidl -coprime_norm_cent ?subsetIidl //. @@ -493,7 +500,7 @@ rewrite leq_eqVlt; case: ltngtP => // rCPB _. have: 2 < 'r('C(B)) by rewrite (leq_trans rCPB) ?rankS ?subsetIr. by apply: cent_rank3_Uniqueness; rewrite -dimB -rank_abelem. have cPX: P \subset 'C(X). - have EpPB: B \in 'E_p(P) by exact/pElemP. + have EpPB: B \in 'E_p(P) by apply/pElemP. have coPX: coprime #|P| #|X| := coprimeSg sPMa coMaX. rewrite centsC (coprime_odd_faithful_cent_abelem EpPB) ?mFT_odd //. rewrite -(setIid 'C(B)) setIA (pmaxElem_LdivP p_pr _) 1?centsC //. @@ -513,7 +520,7 @@ apply: contraL => /= s_p; have piMp := sigma_sub_pi maxM s_p. have p_pr: prime p by move: piMp; rewrite mem_primes; case/andP. rewrite -p'natE ?(pi'_p'nat _ s_p) // -pgroupE -partG_eq1. rewrite -(card_Hall (quotient_pHall _ Msigma_Hall)) /=; last first. - exact: subset_trans (pcore_sub _ _) (der_norm _ _). + exact/gFsub_trans/gFnorm. by rewrite quotientS1 ?cards1 // Msigma_der1. Qed. @@ -527,8 +534,8 @@ Lemma cent1_sigma'_Zgroup P : exists x, [/\ x \in 'Ohm_1('Z(P))^#, 'M('C[x]) != [set M] & Zgroup 'C_(M`_\alpha)[x]]. Proof. -move=> sylP ntP; set T := 'Ohm_1(_); have [sPM pP _] := and3P sylP. -have [charT nilP] := (char_trans (Ohm_char 1 _) (center_char P), pgroup_nil pP). +move=> sylP ntP; have [sPM pP _] := and3P sylP; have nilP := pgroup_nil pP. +set T := 'Ohm_1('Z(P)); have charT: T \char P by rewrite !gFchar_trans. suffices [x Tx not_uCx]: exists2 x, x \in T^# & 'M('C[x]) != [set M]. exists x; split=> //; rewrite odd_rank1_Zgroup ?mFT_odd //= leqNgt. apply: contra not_uCx; rewrite -cent_cycle; set X := <[x]> => rCMaX. @@ -596,7 +603,7 @@ have [[_ abelX dimX] p_pr] := (pnElemP EpX, pnElem_prime EpX). have pX := abelem_pgroup abelX; have rpM2 := sigma'_norm_mmax_rank2 pX sNXM. have [P sylP sXP] := Sylow_superset sXM pX; have [sPM pP _] := and3P sylP. pose T := 'Ohm_1('Z(P)); pose A := X <*> T; have nilP := pgroup_nil pP. -have charT: T \char P := char_trans (Ohm_char 1 _) (center_char P). +have charT: T \char P by apply/gFchar_trans/gFchar. have neqTX: T != X. apply: contraNneq s'p => defX; apply/exists_inP; exists P => //. by rewrite (subset_trans _ sNXM) // -defX char_norms. @@ -669,7 +676,7 @@ Corollary mFT_Sylow_der1 : P \subset 'N(P)^`(1). Proof. have [-> | ntP] := eqsVneq P 1; first exact: sub1G. have ltNG: 'N(P) \proper G := mFT_norm_proper ntP (mFT_pgroup_proper pP). -have [M] := mmax_exists ltNG; case/setIdP=> /= maxM sNM. +have [M /setIdP[/= maxM sNM]] := mmax_exists ltNG. have [ltMG solM] := (mmax_proper maxM, mmax_sol maxM). have [pl1M sPM] := (mFT_proper_plength1 p ltMG, subset_trans (normG P) sNM). have sylP := pHall_subl sPM (subsetT M) sylP_G. @@ -833,7 +840,7 @@ apply/bigcapsP=> [[p /= _] q'p]; have [b_p | b'p] := boolP (p \in \beta(M)). by rewrite pcore_pgroup_id ?(pi'_p'group _ b_p) // /pgroup card_quotient. have p'Mb: p^'.-group M`_\beta := pi_p'group bMb b'p. rewrite sub_Hall_pcore ?(pi_p'group qQ) {Q qQ sQM'}//. -rewrite pquotient_pcore ?quotient_pHall ?(subset_trans (pcore_sub _ _)) //. +rewrite pquotient_pcore ?quotient_pHall 1?gFsub_trans //. by have [-> _ _] := beta_max_pdiv b'p. Qed. @@ -866,7 +873,7 @@ have pqX: pq.-group X by rewrite (pi_pgroup qX) ?inE ?eqxx ?orbT. have{solXM' pqX} [W /= hallW sXW] := Hall_superset solXM' (joing_subl _ _) pqX. have [sWXM' pqW _] := and3P hallW; have sWM := subset_trans sWXM' sXM'M. have{b'q} b'W: \beta(M)^'.-group W. (* GG -- Coq diverges on b'p <> b'q *) - by apply: sub_pgroup pqW => r /pred2P[]->; [exact: b'p | exact: b'q]. + by apply: sub_pgroup pqW => r /pred2P[]->; [apply: b'p | apply: b'q]. have nilM'W: nilpotent (M^`(1) :&: W). by rewrite beta'_der1_nil ?subsetIl ?(pgroupS (subsetIr _ _)). have{nilM'W} nilW: nilpotent W. @@ -881,7 +888,7 @@ have{nilM'W} nilW: nilpotent W. rewrite [Wq](sub_pHall sylQ _ _ (subsetIr _ W)) //= -/Wq. apply/pgroupP=> r r_pr r_dv_Wp'. have:= pgroupP (pgroupS sWqMp' (pcore_pgroup _ _)) r r_pr r_dv_Wp'. - by apply/implyP; rewrite implyNb; exact: (pgroupP (pgroupS sWp'W pqW)). + by apply/implyP; rewrite implyNb; apply: (pgroupP (pgroupS sWp'W pqW)). have [[_ _ max_p] sQM] := (beta_max_pdiv b'p, subset_trans sQW sWM). rewrite subsetI sQW -quotient_sub1 ?(subset_trans sQM nMp'M) //. apply: contraLR lt_pq; rewrite -leqNgt andbT subG1 -rank_gt0. @@ -892,7 +899,7 @@ have{nilM'W} nilW: nilpotent W. rewrite (isog_pgroup _ (second_isog _)) ?(pgroupS (quotientS _ sWXM')) //=. by rewrite (quotientYidr (subset_trans sXW nM'W)) quotient_pgroup. have{qWWM'} sylWp: p.-Sylow(W) Wp. - rewrite /pHall pcore_pgroup (subset_trans (pcore_sub _ _)) ?subsetIr //=. + rewrite /pHall pcore_pgroup gFsub_trans ?subsetIr //=. rewrite -(Lagrange_index (subsetIr _ _) (pcore_sub _ _)) pnat_mul //. rewrite -(divgS (pcore_sub _ _)) -card_quotient ?normsI ?normG //= -pgroupE. rewrite (pi_p'group qWWM') //= -(dprod_card (nilpotent_pcoreC p nilM'W)). @@ -906,7 +913,7 @@ have{nilM'W} nilW: nilpotent W. by rewrite !inE andb_orl andbN orbF; apply: andb_idr; move/eqP->. apply: nilpotentS (mul_subG _ _) (Fitting_nil W). rewrite Fitting_max ?(pgroup_nil pWp) //. - by rewrite (char_normal_trans (pcore_char _ _)) //= setIC norm_normalI. + by rewrite gFnormal_trans //= setIC norm_normalI. by rewrite Fitting_max ?(pgroup_nil qWq) //= setIC norm_normalI. have part1: exists2 P : {group gT}, p.-Sylow(M`_\sigma) P & X \subset 'C(P). have sMsXM' := subset_trans sMsM' (joing_subr X _). @@ -929,14 +936,14 @@ have nsMbXM : M`_\beta <*> X <| M. rewrite -{2}(quotientGK nsMbM) -quotientYK ?cosetpre_normal //=. rewrite (eq_Hall_pcore _ (quotient_pHall nMbX sylX)); last first. exact: nilpotent_pcore_Hall Mbeta_quo_nil. - by rewrite (char_normal_trans (pcore_char _ _)) ?quotient_normal ?der_normal. + by rewrite gFnormal_trans ?quotient_normal ?gFnormal. pose U := 'N_M(X); have defM: M`_\beta * U = M. have sXU : X \subset U by rewrite subsetI sXM normG. rewrite -[U](mulSGid sXU) /= -/U mulgA -norm_joinEr //. apply: Frattini_arg nsMbXM (pHall_subl (joing_subr _ X) _ sylX). by rewrite join_subG Mbeta_der1 (pHall_sub sylX). have sWpU: 'O_p(W) \subset U. - rewrite (subset_trans (pcore_sub _ _)) // subsetI sWM normal_norm //=. + rewrite gFsub_trans // subsetI sWM normal_norm //=. have sylX_W: q.-Sylow(W) X := pHall_subl sXW sWM' sylX. by rewrite (eq_Hall_pcore (nilpotent_pcore_Hall q nilW) sylX_W) pcore_normal. have sylWp: p.-Sylow(M^`(1)) 'O_p(W). @@ -967,7 +974,7 @@ have nMbSM: M`_\beta <*> S <| M. have sylS_M' := pHall_subl sSM' sM'M sylS. rewrite (eq_Hall_pcore _ (quotient_pHall nMbS sylS_M')); last first. exact: nilpotent_pcore_Hall Mbeta_quo_nil. - by rewrite (char_normal_trans (pcore_char _ _)) ?quotient_normal ?der_normal. + by rewrite gFnormal_trans ?quotient_normal ?gFnormal. have defM: M`_\beta * 'N_M(S) = M. have sSNM: S \subset 'N_M(S) by rewrite subsetI sSM normG. rewrite -(mulSGid sSNM) /= mulgA -norm_joinEr //. @@ -1124,8 +1131,8 @@ have{max2A maxQ neq_pq q_dv_CA} [P [sylP sAP] sPNQ']: have{sNQM} defP: 'O_p(M^`(1)) = P. rewrite (nilpotent_Hall_pcore nilM' (pHall_subl _ _ sylP)) ?subsetT //. by rewrite (subset_trans sPNQ') ?dergS. -have charP: P \char M by rewrite -defP (char_trans (pcore_char p _)) ?der_char. -have [sPM nsPM] := (char_sub charP, char_normal charP). +have nsPM: P <| M by rewrite -defP !gFnormal_trans. +have sPM := normal_sub nsPM. case/exists_inP: sg'p; exists P; first exact: pHall_subl (subsetT M) sylP. by rewrite (mmax_normal maxM) // -rank_gt0 ltnW // -dimA -rank_abelem ?rankS. Qed. @@ -1135,15 +1142,15 @@ Proposition sub'cent_sigma_cyclic K (Y := 'C_K(M`_\sigma) :&: M^`(1)) : K \subset M -> sigma'.-group K -> cyclic Y /\ Y <| M. Proof. move=> sKM sg'K; pose Z := 'O_sigma'('F(M)). -have nsZM: Z <| M := char_normal_trans (pcore_char _ _) (Fitting_normal M). +have nsZM: Z <| M by rewrite !gFnormal_trans. have [sZM nZM] := andP nsZM; have Fnil := Fitting_nil M. have rZle1: 'r(Z) <= 1. apply: leq_trans (rankS _) (sub'cent_sigma_rank1 sZM (pcore_pgroup _ _)). rewrite subsetI subxx (sameP commG1P trivgP) /=. rewrite -(TI_pcoreC \sigma(M) M 'F(M)) subsetI commg_subl commg_subr. - by rewrite (subset_trans sZM) ?gFnorm ?(subset_trans (pcore_sub _ _)). + by rewrite (subset_trans sZM) ?gFnorm ?gFsub_trans. have{rZle1} cycZ: cyclic Z. - have nilZ: nilpotent Z := nilpotentS (pcore_sub _ _) Fnil. + have nilZ: nilpotent Z := nilpotentS (gFsub _ _) Fnil. by rewrite nil_Zgroup_cyclic // odd_rank1_Zgroup // mFT_odd. have cZM': M^`(1) \subset 'C_M(Z). rewrite der1_min ?normsI ?normG ?norms_cent //= -ker_conj_aut. @@ -1237,12 +1244,11 @@ do [split; apply: contra neqHgM] => [|nilMs]. have defUS: 'M(S) = [set M] := def_uniq_mmax uniqS maxM sSM. by rewrite (eq_uniq_mmax defUS _ sSHg) ?mmaxJ. have nsSM: S <| M. - have nsMsM: M`_\sigma <| M by exact: pcore_normal. + have nsMsM: M`_\sigma <| M by apply: pcore_normal. have{sylS} sylS: p.-Sylow(M`_\sigma) S. apply: pHall_subl (pcore_sub _ _) sylS => //. by rewrite (sub_Hall_pcore (Msigma_Hall maxM)) ?(pi_pgroup pS). - rewrite (nilpotent_Hall_pcore nilMs sylS). - by rewrite (char_normal_trans (pcore_char _ _)). + by rewrite (nilpotent_Hall_pcore nilMs sylS) gFnormal_trans. have sNS_Hg: 'N(S) \subset H :^ g. rewrite -sub_conjgV -normJ (norm_sigma_Sylow sHp) //. by rewrite (pHall_subl _ (subsetT _)) ?sub_conjgV // pHallJ ?in_setT. @@ -1275,7 +1281,7 @@ have [pS sAS] := (pHall_pgroup sylS, subset_trans sAP sPS). have [maxAS defC1] := maxA S pS sAS; set C := 'C_S(A) in defC1. have sZ0A: Z0 \subset A by rewrite -defC1 OhmS // setISS // centS. have sZ1A: Z1 \subset A by rewrite -defC1 OhmS // setIS // centS. -have [pZ0 pZ1]: p.-group Z0 /\ p.-group Z1 by split; exact: pgroupS pA. +have [pZ0 pZ1]: p.-group Z0 /\ p.-group Z1 by split; apply: pgroupS pA. have sZ10: Z1 \subset Z0. rewrite -[gval Z1]Ohm_id OhmS // subsetI (subset_trans sZ1A) //=. by rewrite (subset_trans sZ1Z) // subIset // centS ?orbT. @@ -1286,7 +1292,7 @@ have ntZ1: Z1 :!=: 1. have EpZ01: abelian C -> Z1 = Z0 /\ Z0 \in E1A. move=> cCC; have [eqZ0A | ltZ0A] := eqVproper sZ0A. rewrite (abelianS _ cCC) // in not_cPP. - by rewrite subsetI sPS centsC -eqZ0A (subset_trans (Ohm_sub _ _)) ?subsetIr. + by rewrite subsetI sPS centsC -eqZ0A gFsub_trans ?subsetIr. have leZ0p: #|Z0| <= p ^ 1. by rewrite (card_pgroup pZ0) leq_exp2l // -ltnS -dimA properG_ltn_log. have [_ _ [e oZ1]] := pgroup_pdiv pZ1 ntZ1. @@ -1299,7 +1305,7 @@ have [A1 neqA1Z EpA1]: exists2 A1, A1 != Z1 & #|Z1| = p -> A1 \in E1A. have [A1 defA]:= abelem_split_dprod abelA sZ1A. have{defA} [_ defA _ tiA1Z1] := dprodP defA. have EpZ1: Z1 \in E1A by rewrite [E1A]p1ElemE // !inE sZ1A /= oZ1. - suffices: A1 \in E1A by exists A1; rewrite // eq_sym; exact/(TIp1ElemP EpZ1). + suffices: A1 \in E1A by exists A1; rewrite // eq_sym; apply/(TIp1ElemP EpZ1). rewrite [E1A]p1ElemE // !inE -defA mulG_subr /=. by rewrite -(mulKn #|A1| p_gt0) -{1}oZ1 -TI_cardMg // defA oA mulKn. pose cplA1C Y := [/\ cyclic Y, Z0 \subset Y, A1 \x Y = C & abelian C]. @@ -1325,7 +1331,7 @@ have [Y [{cplA1C} cycY sZ0Y defC cCC]]: exists Y, cplA1C Y. rewrite (center_idP (abelem_abelian abelA1)). by rewrite (center_idP (cyclic_abelian cycY)). have{EpZ01} [<- _] := EpZ01 cCC; rewrite subsetI (subset_trans sZ1Z) //. - by rewrite setIS ?centS // (subset_trans (Ohm_sub 1 _)) ?ucn_sub. + by rewrite setIS ?centS ?gFsub_trans. have not_cSS := contra (abelianS sPS) not_cPP. have:= mFT_rank2_Sylow_cprod sylS rSle2 not_cSS. case=> E [_ dimE3 eE] [Y cycY [defS defY1]]. @@ -1389,8 +1395,7 @@ apply/imsetP; exists A1; first by rewrite 2!inE neqA1Z. apply/eqP; rewrite eq_sym eqEcard; apply/andP; split. apply/subsetP=> _ /imsetP[x /setIP[Px nAx] ->]. rewrite 2!inE /E1A -(normP nAx) pnElemJ EpA1 andbT -val_eqE /=. - have nZ0P: P \subset 'N(Z0). - by rewrite (char_norm_trans (Ohm_char 1 _)) // gFnorm. + have nZ0P: P \subset 'N(Z0) by rewrite !gFnorm_trans. by rewrite -(normsP nZ0P x Px) (inj_eq (@conjsg_inj _ x)). have pN: p.-group 'N_P(_) := pgroupS (subsetIl P _) pP. have defCPA: 'N_('N_P(A))(A1) = 'C_P(A). @@ -1398,7 +1403,7 @@ have defCPA: 'N_('N_P(A))(A1) = 'C_P(A). rewrite subIset /=; last by rewrite orbC cents_norm ?centS. rewrite setIAC (subset_trans (subsetIl _ _)) //= subsetI subsetIl /=. rewrite -defA centM subsetI andbC subIset /=; last first. - by rewrite centsC (subset_trans (Ohm_sub 1 _)) ?subsetIr. + by rewrite centsC gFsub_trans ?subsetIr. have nC_NP: 'N_P(A1) \subset 'N('C(A1)) by rewrite norms_cent ?subsetIr. rewrite -quotient_sub1 // subG1 trivg_card1. rewrite (pnat_1 (quotient_pgroup _ (pN _))) //. @@ -1407,7 +1412,7 @@ have defCPA: 'N_('N_P(A))(A1) = 'C_P(A). have sCN: 'C_P(A) \subset 'N_P(A) by rewrite setIS ?cent_sub. have nA_NCPA: 'N_P('C_P(A)) \subset 'N_P(A). have [_ defCPA1] := maxA P pP sAP. - by rewrite -{2}defCPA1 setIS // (char_norm_trans (Ohm_char 1 _)). + by rewrite -[in 'N(A)]defCPA1 setIS // gFnorm_trans. rewrite card_orbit astab1JG /= {}defCPA. rewrite -(leq_add2l (Z0 \in E1A)) -cardsD1 EpZ0 (card_p1Elem_p2Elem Ep2A) ltnS. rewrite dvdn_leq ?(pfactor_dvdn 1) ?indexg_gt0 // -divgS // logn_div ?cardSg //. @@ -1481,7 +1486,7 @@ have ntX: X != 1. by rewrite /= p_core_Fitting -/X X1 cards1. have bMq: q \in \beta(M) by apply: (pgroupP (pgroupS (Fitting_sub Y) bY)). have b_q: q \in \beta(G) by move: bMq; rewrite -predI_sigma_beta //; case/andP. -have sXM: X \subset M := subset_trans (pcore_sub q Y) sYM. +have sXM: X \subset M := gFsub_trans _ sYM. have [P sylP sXP] := Sylow_superset sXM qX; have [sPM qP _] := and3P sylP. have sylPG: q.-Sylow(G) P by rewrite (sigma_Sylow_G maxM) ?beta_sub_sigma. have uniqNX: 'M('N_P(X)) = [set M]. diff --git a/mathcomp/odd_order/BGsection11.v b/mathcomp/odd_order/BGsection11.v index 6fccf96..be228fc 100644 --- a/mathcomp/odd_order/BGsection11.v +++ b/mathcomp/odd_order/BGsection11.v @@ -1,9 +1,16 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div path fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div path fintype. +From mathcomp Require Import bigop finset prime fingroup morphism perm automorphism quotient. +From mathcomp Require Import action gproduct gfunctor pgroup cyclic center commutator. +From mathcomp Require Import gseries nilpotent sylow abelian maximal hall. +From mathcomp Require Import BGsection1 BGsection3 BGsection4 BGsection5 BGsection6. +From mathcomp Require Import BGsection7 BGsection10. (******************************************************************************) @@ -256,8 +263,7 @@ have regX12: 'C_Q2(X1) = 1 by case: (regA X1) nregX11 => // ->; rewrite eqxx. pose X := 'Ohm_1('Z(P))%G. have eqCQ12_X: ('C_Q1(X) == 1) = ('C_Q2(X) == 1). rewrite -(inj_eq (@conjsg_inj _ g)) conjs1g conjIg -/Q2 -centJ (normP _) //. - rewrite (subsetP (char_norm_trans (Ohm_char 1 _) _) g nPg) //. - by rewrite char_norms ?center_char. + by rewrite (subsetP (gFnorm_trans _ _) g nPg) ?gFnorms. have{EpX1} EpX1: X1 \in 'E_p^1(A) :\ X. rewrite 2!inE EpX1 andbT; apply: contraNneq nregX11 => defX1. by rewrite defX1 eqCQ12_X -defX1 regX12. @@ -308,7 +314,7 @@ have EpA2: A2 \in 'E_p^1(A) by rewrite -(normP nAg2) pnElemJ. have{neMg12} neqA12: A1 :!=: A2. rewrite -(canF_eq (conjsgKV g2)) -conjsgM (sameP eqP normP). rewrite (contra (subsetP sNA0_M _)) // -mem_rcoset. - by apply: contra neMg12 => g1Mg2; rewrite defMg1 defMg2 (rcoset_transl g1Mg2). + by apply: contra neMg12 => g1Mg2; rewrite defMg1 defMg2 (rcoset_eqP g1Mg2). have{notMg1 nAg1} regA1: 'C_Ms(A1) = 1. by case/exceptional_TI_MsigmaJ: notMg1; rewrite // -(normP nAg1) conjSg. have{notMg2 nAg2} regA2: 'C_Ms(A2) = 1. @@ -364,26 +370,25 @@ have [cKA | not_cKA]:= boolP (A \subset 'C(K)). apply/eqP; rewrite eqEcard mulG_subG pcore_sub sEM /= TI_cardMg //. by rewrite (card_Hall hallE) (card_Hall (Msigma_Hall maxM)) ?partnC. rewrite norm_joinEr -?quotientK ?(subset_trans sAE) //= cosetpre_normal. - rewrite quotient_normal // -defA (char_normal_trans (Ohm_char _ _)) //. - by rewrite (char_normal_trans (pcore_char p _)). + by rewrite quotient_normal // -defA !gFnormal_trans. pose q := pdiv #|K : 'C_K(A)|. have q_pr: prime q by rewrite pdiv_prime // indexg_gt1 subsetI subxx centsC. have [nKA coKA] := (subset_trans sAP nKP, coprimegS sAP coKP). have [Q sylQ nQA]: exists2 Q : {group gT}, q.-Sylow(K) Q & A \subset 'N(Q). - by apply: sol_coprime_Sylow_exists => //; exact: (pgroup_sol pA). + by apply: sol_coprime_Sylow_exists => //; apply: (pgroup_sol pA). have [sQK qQ q'iQK] := and3P sylQ; have [sKE tauK _]:= and3P hallK. have{q'iQK} not_cQA: ~~ (A \subset 'C(Q)). apply: contraL q'iQK => cQA; rewrite p'natE // negbK. rewrite -(Lagrange_index (subsetIl K 'C(A))) ?dvdn_mulr ?pdiv_dvd //. by rewrite subsetI sQK centsC. -have ntQ: Q :!=: 1 by apply: contraNneq not_cQA => ->; exact: cents1. +have ntQ: Q :!=: 1 by apply: contraNneq not_cQA => ->; apply: cents1. have q_dv_K: q %| #|K| := dvdn_trans (pdiv_dvd _) (dvdn_indexg _ _). have sM'q: q \in (\sigma(M))^' := pgroupP (pgroupS sKE s'E) q q_pr q_dv_K. have{q_dv_K} tau_q: q \in tau := pgroupP tauK q q_pr q_dv_K. have sylQ_E: q.-Sylow(E) Q := subHall_Sylow hallK tau_q sylQ. have sylQ_M: q.-Sylow(M) Q := subHall_Sylow hallE sM'q sylQ_E. have q'p: p != q by rewrite neq_ltn [p < q]tau_q. -have [regQ | nregQ] := eqVneq 'C_Q(A) 1; last first. +suffices nregQ: 'C_Q(A) != 1. have ncycQ: ~~ cyclic Q. apply: contra not_cQA => cycQ. rewrite (coprime_odd_faithful_Ohm1 qQ) ?mFT_odd ?(coprimeSg sQK) //. @@ -408,31 +413,30 @@ have [regQ | nregQ] := eqVneq 'C_Q(A) 1; last first. have{Eq2B} Eq2B := subsetP (pnElemS q 2 sQ_Qstar) B Eq2B. rewrite inE Eq2B (subsetP (pmaxElemS q (subsetT _))) // inE maxB inE. by have [? _ _] := pnElemP Eq2B. -pose Q0 := 'Z(Q); have charQ0: Q0 \char Q := center_char Q. -have nQ0A: A \subset 'N(Q0) := char_norm_trans charQ0 nQA. +pose Q0 := 'Z(Q); have sQ0Q: Q0 \subset Q by apply: gFsub. +have nQ0A: A \subset 'N(Q0) by apply: gFnorm_trans. +have ntQ0: Q0 != 1 by apply: contraNneq ntQ => /(trivg_center_pgroup qQ)->. +apply: contraNneq (sM'q) => regQ; apply/exists_inP; exists Q => //. +suffices nsQ0M: Q0 <| M by rewrite -(mmax_normal _ nsQ0M) ?gFnorms. +have sQ0M: Q0 \subset M := subset_trans sQ0Q (pHall_sub sylQ_M). +have qQ0: q.-group Q0 := pgroupS sQ0Q qQ. +have p'Q0: p^'.-group Q0 by apply: (pi_pnat qQ0); rewrite eq_sym in q'p. +have sM'Q0: \sigma(M)^'.-group Q0 := pi_pnat qQ0 sM'q. +have cQ0Q0: abelian Q0 := center_abelian Q. have defQ0: [~: A, Q0] = Q0. - rewrite -{2}[Q0](coprime_abelian_cent_dprod nQ0A) ?center_abelian //. - by rewrite setIAC regQ (setIidPl (sub1G _)) dprodg1 commGC. - by rewrite (coprimeSg (subset_trans (center_sub Q) sQK)). + rewrite -{2}[Q0](coprime_abelian_cent_dprod nQ0A) //. + by rewrite setIAC regQ setI1g dprodg1 commGC. + by rewrite (coprimeSg (subset_trans sQ0Q sQK)). have [_ _ [A1 EpA1 [A2 EpA2 [neqA12 regA1 regA2]]]] := exceptional_structure. have defA: A1 \x A2 = A by apply/(p2Elem_dprodP Ep2A EpA1 EpA2). have{defQ0} defQ0: [~: A1, Q0] * [~: A2, Q0] = Q0. have{defA} [[_ defA cA12 _] [sA2A _ _]] := (dprodP defA, pnElemP EpA2). by rewrite -commMG ?defA // normsR ?(cents_norm cA12) // (subset_trans sA2A). -have nsQ0M: Q0 <| M. - have sQ0M: Q0 \subset M := subset_trans (center_sub Q) (pHall_sub sylQ_M). - have qQ0: q.-group Q0 := pgroupS (center_sub Q) qQ. - have p'Q0: p^'.-group Q0 by apply: (pi_pnat qQ0); rewrite eq_sym in q'p. - have sM'Q0: \sigma(M)^'.-group Q0 := pi_pnat qQ0 sM'q. - have cQ0Q0: abelian Q0 := center_abelian Q. - have sA_NQ0: A \subset 'N_M(Q0) by rewrite subsetI sAM. - have sEpA_EpN := subsetP (pnElemS p 1 sA_NQ0). - have nsRQ0 := commG_sigma'_1Elem_cyclic maxM sQ0M sM'Q0 sM'p (sEpA_EpN _ _). - rewrite -defQ0 -!(commGC Q0). - by apply: normalM; [case/nsRQ0: EpA1 | case/nsRQ0: EpA2]. -case/exists_inP: sM'q; exists Q => //. -rewrite (subset_trans (char_norms charQ0)) ?(mmax_normal maxM nsQ0M) //= -/Q0. -by apply: contraNneq ntQ; move/(trivg_center_pgroup qQ)->. +have sA_NQ0: A \subset 'N_M(Q0) by rewrite subsetI sAM. +have sEpA_EpN := subsetP (pnElemS p 1 sA_NQ0). +have nsRQ0 := commG_sigma'_1Elem_cyclic maxM sQ0M sM'Q0 sM'p (sEpA_EpN _ _). +rewrite -defQ0 -!(commGC Q0). +by apply: normalM; [case/nsRQ0: EpA1 | case/nsRQ0: EpA2]. Qed. End Section11. diff --git a/mathcomp/odd_order/BGsection12.v b/mathcomp/odd_order/BGsection12.v index b831ebc..e392285 100644 --- a/mathcomp/odd_order/BGsection12.v +++ b/mathcomp/odd_order/BGsection12.v @@ -1,9 +1,16 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq choice div fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq choice div fintype. +From mathcomp Require Import path bigop finset prime fingroup morphism perm automorphism. +From mathcomp Require Import quotient action gproduct gfunctor pgroup cyclic commutator. +From mathcomp Require Import center gseries nilpotent sylow abelian maximal hall frobenius. +From mathcomp Require Import BGsection1 BGsection3 BGsection4 BGsection5 BGsection6. +From mathcomp Require Import BGsection7 BGsection9 BGsection10 BGsection11. (******************************************************************************) @@ -97,7 +104,7 @@ Let solE := sigma_compl_sol. Let exHallE pi := exists Ei : {group gT}, pi.-Hall(E) Ei. Lemma ex_tau13_compl : exHallE \tau1(M) /\ exHallE \tau3(M). -Proof. by split; exact: Hall_exists. Qed. +Proof. by split; apply: Hall_exists. Qed. Lemma ex_tau2_compl E1 E3 : \tau1(M).-Hall(E) E1 -> \tau3(M).-Hall(E) E3 -> @@ -228,7 +235,7 @@ Lemma tau2_not_beta p : p \in \tau2(M) -> p \notin \beta(G) /\ {subset 'E_p^2(M) <= 'E*_p(G)}. Proof. case/andP=> s'p /eqP rpM; split; first exact: sigma'_rank2_beta' rpM. -by apply/subsetP; exact: sigma'_rank2_max. +by apply/subsetP; apply: sigma'_rank2_max. Qed. End Introduction. @@ -262,8 +269,7 @@ have cycE1: cyclic E1. by rewrite subsetI (der_sub 1) (dergS 1). have solE: solvable E := solvableS sEM solM. have nilE': nilpotent E^`(1) := der1_sigma_compl_nil maxM hallE. -have nsE'piE pi: 'O_pi(E^`(1)) <| E. - exact: char_normal_trans (pcore_char _ _) (der_normal _ _). +have nsE'piE pi: 'O_pi(E^`(1)) <| E by rewrite !gFnormal_trans. have SylowE3 P: Sylow E3 P -> [/\ cyclic P, P \subset E^`(1) & 'C_P(E) = 1]. - case/SylowP=> p p_pr sylP; have [sPE3 pP _] := and3P sylP. have [-> | ntP] := eqsVneq P 1. @@ -276,18 +282,17 @@ have SylowE3 P: Sylow E3 P -> [/\ cyclic P, P \subset E^`(1) & 'C_P(E) = 1]. rewrite (odd_pgroup_rank1_cyclic pP) ?mFT_odd //. rewrite (tau3E maxM hallE) in t3p. by case/andP: t3p => _ /eqP <-; rewrite p_rankS. - have nEp'E: E \subset 'N('O_p^'(E)) by exact: gFnorm. + have nEp'E: E \subset 'N('O_p^'(E)) by apply: gFnorm. have nEp'P := subset_trans sPE nEp'E. have sylP_E := subHall_Sylow hallE3 t3p sylP. have nsEp'P_E: 'O_p^'(E) <*> P <| E. rewrite sub_der1_normal ?join_subG ?pcore_sub //=. - rewrite norm_joinEr // -quotientSK //=; last first. - by rewrite (subset_trans (der_sub 1 E)). + rewrite norm_joinEr // -quotientSK ?gFsub_trans //=. have [_ /= <- _ _] := dprodP (nilpotent_pcoreC p nilE'). rewrite -quotientMidr -mulgA (mulSGid (pcore_max _ _)) ?pcore_pgroup //=. rewrite quotientMidr quotientS //. apply: subset_trans (pcore_sub_Hall sylP_E). - by rewrite pcore_max ?pcore_pgroup /=. + by rewrite pcore_max ?pcore_pgroup ?nsE'piE. have nEP_sol: solvable 'N_E(P) by rewrite (solvableS _ solE) ?subsetIl. have [K hallK] := Hall_exists p^' nEP_sol; have [sKNEP p'K _] := and3P hallK. have coPK: coprime #|P| #|K| := pnat_coprime pP p'K. @@ -295,8 +300,7 @@ have SylowE3 P: Sylow E3 P -> [/\ cyclic P, P \subset E^`(1) & 'C_P(E) = 1]. have mulPK: P * K = 'N_E(P). apply/eqP; rewrite eqEcard mul_subG //= coprime_cardMg // (card_Hall hallK). by rewrite (card_Hall (pHall_subl sP_NEP (subsetIl E _) sylP_E)) partnC. - rewrite subsetI in sKNEP; case/andP: sKNEP => sKE nPK. - have nEp'K := subset_trans sKE nEp'E. + have{sKNEP} [sKE nPK] := subsetIP sKNEP; have nEp'K := subset_trans sKE nEp'E. have defE: 'O_p^'(E) <*> K * P = E. have sP_Ep'P: P \subset 'O_p^'(E) <*> P := joing_subr _ _. have sylP_Ep'P := pHall_subl sP_Ep'P (normal_sub nsEp'P_E) sylP_E. @@ -412,8 +416,8 @@ have p_pr := pnElem_prime Ep2A. have [sAM abelA dimA] := pnElemP Ep2A; have [pA cAA _] := and3P abelA. have [sA0A _ _] := pnElemP EpA0; have pA0 := pgroupS sA0A pA. have sAH: A \subset H. - by apply: subset_trans (cents_norm _) sNH; exact: subset_trans (centS sA0A). -have nsHsH: H`_\sigma <| H by exact: pcore_normal. + by apply: subset_trans (cents_norm _) sNH; apply: subset_trans (centS sA0A). +have nsHsH: H`_\sigma <| H by apply: pcore_normal. have [sHsH nHsH] := andP nsHsH; have nHsA := subset_trans sAH nHsH. have nsHsA_H: H`_\sigma <*> A <| H. have [sHp | sH'p] := boolP (p \in \sigma(H)). @@ -439,14 +443,14 @@ have [sMp | sM'p] := boolP (p \in \sigma(M)). apply: subset_trans (cAp' _ _ _ (subsetIr _ _)) _. - exact: pi_p'group (pgroupS (subsetIl _ _) (pcore_pgroup _ _)) aM'p. - by rewrite (normsI _ (normsG sAH)) // (subset_trans sAM) ?gFnorm. - by rewrite setIAC; case/sigma_disjoint: notMGH => // -> _ _; exact: subsetIl. + by rewrite setIAC; case/sigma_disjoint: notMGH => // -> _ _; apply: subsetIl. suffices cMaA: A \subset 'C(M`_\sigma :&: H). by rewrite !{1}(subset_trans cMaA) ?centS ?setSI // Malpha_sub_Msigma. have [sHp | sH'p] := boolP (p \in \sigma(H)); last first. apply/commG1P; apply: contraNeq neqMH => ntA_MsH. have [P sylP sAP] := Sylow_superset sAH pA. - have excH: exceptional_FTmaximal p H A0 A by split=> //; exact/pnElemP. - have maxAM: M \in 'M(A) by exact/setIdP. + have excH: exceptional_FTmaximal p H A0 A by split=> //; apply/pnElemP. + have maxAM: M \in 'M(A) by apply/setIdP. rewrite (exceptional_sigma_uniq maxH excH sylP sAP maxAM) //. apply: contraNneq ntA_MsH => tiMsHs; rewrite -subG1. have [sHsA_H nHsA_H] := andP nsHsA_H. @@ -520,9 +524,9 @@ have{EpAnonuniq} sCMkApCA y: y \in A^# -> rewrite (eq_uniq_mmax defU maxH) ?subIset //. by rewrite orbC (subset_trans (cent_sub Y)). have [cAMs cAMa] := nonuniq_p2Elem_cent_sigma maxM maxH neqHM Ep2A EpY sNYH. - rewrite !{1}subsetI !{1}(subset_trans (subsetIl _ _) (pcore_sub _ _)). - by split=> // [/cAMs | /cAMa]; rewrite centsC; apply: subset_trans; - rewrite setIS ?(subset_trans (cent_sub Y)). + do 2!rewrite {1}subsetI {1}(subset_trans (subsetIl _ _) (pcore_sub _ _)). + have sCYH: 'C(Y) \subset H := subset_trans (cent_sub Y) sNYH. + by split=> // [/cAMs | /cAMa]; rewrite centsC; apply/subset_trans/setIS. have ntA: A :!=: 1 by rewrite -rank_gt0 (rank_abelem abelA) dimA. have ncycA: ~~ cyclic A by rewrite (abelem_cyclic abelA) dimA. have rCMAle2: 'r('C_M(A)) <= 2. @@ -556,7 +560,7 @@ have rPle2: 'r(P) <= 2. by apply/subsetPn; rewrite subG1. have [|rCMz _ _] := sCMkApCA z; first by rewrite inE ntz (subsetP sZA). rewrite (leq_trans _ rCMz) ?rankS // subsetI sPM centsC cycle_subG. - by rewrite (subsetP _ z Zz) // (subset_trans (Ohm_sub 1 _)) ?subsetIr. + by rewrite (subsetP _ z Zz) // gFsub_trans ?subsetIr. have aM'p: p \in \alpha(M)^'. by rewrite !inE -leqNgt -(p_rank_Sylow sylP) -rank_pgroup. have sMaCMA: M`_\alpha \subset 'C_M(A). @@ -575,12 +579,11 @@ have Ma1: M`_\alpha = 1. have nilMs: nilpotent M`_\sigma. rewrite (nilpotentS (Msigma_der1 maxM)) // (isog_nil (quotient1_isog _)). by rewrite -Ma1 Malpha_quo_nil. -rewrite (subset_trans (cents_norm (centS sZA))) ?(mmax_normal maxM) //. -apply: char_normal_trans (char_trans (Ohm_char 1 _) (center_char P)) _. +rewrite (subset_trans (cents_norm (centS sZA))) ?(mmax_normal maxM) //=. have{sylP} sylP: p.-Sylow(M`_\sigma) P. apply: pHall_subl _ (pcore_sub _ _) sylP. by rewrite (sub_Hall_pcore (Msigma_Hall maxM)) // (pi_pgroup pP). -by rewrite (nilpotent_Hall_pcore _ sylP) ?(char_normal_trans (pcore_char _ _)). +by rewrite (nilpotent_Hall_pcore _ sylP) ?gFnormal_trans. Qed. (* This is B & G, Theorem 12.5(a) -- this part does not mention a specific *) @@ -634,8 +637,7 @@ have sylP_Hs: p.-Sylow(H`_\sigma) P. rewrite (pHall_subl _ (pcore_sub _ _) sylP) //. by rewrite (sub_Hall_pcore (Msigma_Hall maxH)) // (pi_pgroup pP). have nPH: H \subset 'N(P). - rewrite (nilpotent_Hall_pcore nilHs sylP_Hs). - by rewrite !(char_norm_trans (pcore_char _ _)) ?normG. + by rewrite (nilpotent_Hall_pcore nilHs sylP_Hs) !gFnorm_trans ?normG. have coMsP: coprime #|M`_\sigma| #|P|. exact: pnat_coprime (pcore_pgroup _ _) (pi_pnat pP _). rewrite (sameP commG1P trivgP) -(coprime_TIg coMsP) commg_subI ?setIS //. @@ -670,7 +672,7 @@ have sAsylE P: p.-Sylow(E) P -> 'Ohm_1(P) = A /\ ~~ ('N(P) \subset M). exact: subset_trans (pcore_max pA nsAE) (pcore_sub_Hall sylP). have not_sNA_M: ~~ ('N(A) \subset M). have [P sylP] := Sylow_exists p E; have [<-]:= sAsylE P sylP. - exact: contra (subset_trans (char_norms (Ohm_char 1 P))). + exact/contra/subset_trans/gFnorms. have{sAsylE syl_p_M} defEpE: 'E_p^1(E) = 'E_p^1(A). apply/eqP; rewrite eqEsubset andbC pnElemS //. apply/subsetP=> X /pnElemP[sXE abelX dimX]; apply/pnElemP; split=> //. @@ -798,7 +800,7 @@ have def_t2: \tau2(M) =i (p : nat_pred). have{EqB} Eq2B: B \in 'E_q^2(E). by move: t2Mq; rewrite (tau2E hallE) => /andP[_ /eqP <-]. have [sBE abelB dimB]:= pnElemP Eq2B; have [qB _] := andP abelB. - have coBA: coprime #|B| #|A| by exact: pnat_coprime qB (pi_pnat pA _). + have coBA: coprime #|B| #|A| by apply: pnat_coprime qB (pi_pnat pA _). have [[nsBE _] [sCBE _ _] _ _] := tau2_compl_context maxM hallE t2Mq Eq2B. have nBA: A \subset 'N(B) by rewrite (subset_trans sAE) ?normal_norm. have cAB: B \subset 'C(A). @@ -819,7 +821,7 @@ have def_t2: \tau2(M) =i (p : nat_pred). rewrite (card_Hall sylS) -(card_Hall sylP1). by rewrite (card_Hall sylP_E) -(card_Hall sylP1_E). have coMsA: coprime #|Ms| #|A|. - by exact: pnat_coprime (pcore_pgroup _ _) (pi_pnat pA _). + by apply: pnat_coprime (pcore_pgroup _ _) (pi_pnat pA _). have defMs: <<\bigcup_(X in 'E_p^1(A)) 'C_Ms(X)>> = Ms. have ncycA: ~~ cyclic A by rewrite (abelem_cyclic abelA) dimA. have [sAM _ _] := pnElemP Ep2A_M. @@ -843,7 +845,7 @@ have{ltPS} not_sSM: ~~ (S \subset M). by rewrite (sameP setIidPl eqP) defSM proper_neq. have not_sA0Z: ~~ (A0 \subset 'Z(S)). apply: contra not_sSM; rewrite subsetI centsC; case/andP=> _ sS_CA0. - by case/mem_uniq_mmax: uniqCA0 => _; exact: subset_trans sS_CA0. + by case/mem_uniq_mmax: uniqCA0 => _; apply: subset_trans sS_CA0. have [EpZ0 dxCSA transNSA] := basic_p2maxElem_structure max2A pS sAS not_cSS. do [set Z0 := 'Ohm_1('Z(S))%G; set EpA' := _ :\ Z0] in EpZ0 dxCSA transNSA. have sZ0Z: Z0 \subset 'Z(S) := Ohm_sub 1 _. @@ -906,8 +908,7 @@ have defFM: Ms \x A0 = 'F(M). have [_ /= defFM cFpp' _] := dprodP (nilpotent_pcoreC p nilF). have defFp': 'O_p^'('F(M)) = Ms. apply/eqP; rewrite eqEsubset. - rewrite (sub_Hall_pcore (Msigma_Hall maxM)); last first. - exact: subset_trans (pcore_sub _ _) (Fitting_sub _). + rewrite (sub_Hall_pcore (Msigma_Hall maxM)); last by rewrite !gFsub_trans. rewrite /pgroup (sub_in_pnat _ (pcore_pgroup _ _)) => [|q piFq]; last first. have [Q sylQ] := Sylow_exists q 'F(M); have [sQF qQ _] := and3P sylQ. have ntQ: Q :!=: 1. @@ -919,8 +920,8 @@ have defFM: Ms \x A0 = 'F(M). apply/implyP; rewrite implyNb /= -def_t2 orbC. by rewrite (prime_class_mmax_norm maxM qQ). rewrite pcore_max ?(pi_p'group (pcore_pgroup _ _)) //. - rewrite /normal (subset_trans (Fitting_sub M)) ?gFnorm //. - rewrite Fitting_max ?pcore_normal ?(tau2_Msigma_nil _ t2Mp) //. + rewrite [_ <| _]andbC gFsub_trans ?gFnorm //. + rewrite Fitting_max ?gFnormal ?(tau2_Msigma_nil _ t2Mp) //. rewrite p_core_Fitting defFp' centsC in defFM cFpp'. rewrite -defFM (centC cFpp'); congr (Ms * _). apply/eqP; rewrite eqEsubset pcore_max //. @@ -1021,7 +1022,7 @@ have Ep2A_M := subsetP (pnElemS p 2 sEM) A Ep2A. have eqFC H: A <| H -> 'C(A) \subset H -> 'F(H) = 'F('C(A)). move=> nsAH sCH; have [_ nAH] := andP nsAH. apply/eqP; rewrite eqEsubset !Fitting_max ?Fitting_nil //. - by rewrite (char_normal_trans (Fitting_char _)) // /normal sCH norms_cent. + by rewrite gFnormal_trans // /normal sCH norms_cent. apply: normalS sCH (Fitting_normal H). have [_ defF cFpFp' _] := dprodP (nilpotent_pcoreC p (Fitting_nil H)). have sAFp: A \subset 'O_p('F(H)) by rewrite p_core_Fitting pcore_max. @@ -1043,7 +1044,7 @@ have sSE: S \subset E by rewrite (subset_trans _ sCAE) // (centSS _ _ cSS). have nA_NS: 'N(S) \subset 'N(A). have [ ] := tau2_context maxM t2Mp Ep2A_M; have sSM := subset_trans sSE sEM. have sylS_M: p.-Sylow(M) S := pHall_subl sSM (subsetT M) sylS. - by case/(_ S) => // _ [// |<- _] _ _ _ _; exact: char_norms (Ohm_char 1 _). + by case/(_ S) => // _ [// |<- _] _ _ _ _; apply: char_norms (Ohm_char 1 _). have sS_NS': S \subset 'N(S)^`(1) := mFT_Sylow_der1 sylS. have sNS'_FE: 'N(S)^`(1) \subset 'F(E). by rewrite -eqFN_FE (subset_trans (dergS 1 nA_NS)). @@ -1125,8 +1126,7 @@ have cE2E2: abelian E2 := abelianS sE2_ZFE (center_abelian _). have sE2FE: E2 \subset 'F(E) := subset_trans sE2_ZFE (center_sub _). have nsE2E: E2 <| E. have hallE2_F := pHall_subl sE2FE (Fitting_sub E) hallE2. - rewrite (nilpotent_Hall_pcore nilF hallE2_F). - exact: char_normal_trans (pcore_char _ _) (Fitting_normal E). + by rewrite (nilpotent_Hall_pcore nilF hallE2_F) !gFnormal_trans. have [_ _ [cycE1 cycE3] [_ defEl] _] := sigma_compl_context maxM complEi. have [[K _ defK _] _ _ _] := sdprodP defEl; rewrite defK in defEl. have [nsKE _ mulKE1 nKE1 _] := sdprod_context defEl; have [sKE _] := andP nsKE. @@ -1227,7 +1227,7 @@ rewrite /normal subIset ?comm_subG ?normG //=; split. have ->: 'C_S(X) = 'C_S('C(S) * X). by rewrite centM setIA; congr (_ :&: _); rewrite (setIidPl _) // centsC. by rewrite normsI ?norms_cent. -have CS_S_1: [~: 'C(S), S] = 1 by exact/commG1P. +have CS_S_1: [~: 'C(S), S] = 1 by apply/commG1P. by rewrite commGC -[[~: X, S]]mul1g -CS_S_1 -commMG ?CS_S_1 ?norms1 ?normsR. Qed. @@ -1379,7 +1379,7 @@ have [cSS | not_cSS] := boolP (abelian S). have pS := pHall_pgroup sylS. have [def_t2 _ _ _] := nonabelian_tau2 maxM hallE t2p Ep2A pS not_cSS. apply: sigma'_nil_abelian (subset_trans _ sEM) (pgroupS _ s'E) _ => //. -by rewrite (eq_pgroup _ def_t2) in t2E2; exact: pgroup_nil t2E2. +by rewrite (eq_pgroup _ def_t2) in t2E2; apply: pgroup_nil t2E2. Qed. (* This is B & G, Corollary 12.10(c). *) @@ -1477,7 +1477,7 @@ Lemma primes_norm_tau2Elem M E p A Mstar : Proof. move=> maxM hallE t2Mp Ep2A; move: Mstar. have [sAE abelA dimA] := pnElemP Ep2A; have [pA cAA _] := and3P abelA. -have ntA: A :!=: 1 by exact: (nt_pnElem Ep2A). +have ntA: A :!=: 1 by apply: (nt_pnElem Ep2A). have [sEM solE] := (pHall_sub hallE, sigma_compl_sol hallE). have [_ nsCA_E t1CEAb] := tau1_cent_tau2Elem_factor maxM hallE t2Mp Ep2A. have [sAM nCA_E] := (subset_trans sAE sEM, normal_norm nsCA_E). @@ -1553,7 +1553,7 @@ have part_b H: have hallHb: \beta(H).-Hall(H) H`_\beta := Mbeta_Hall maxH. have nilH'b: nilpotent (H^`(1) / H`_\beta) := Mbeta_quo_nil maxH. have{nilH'b} sAQ_Hb: [~: A, Q] \subset H`_\beta. - rewrite -quotient_cents2 ?(subset_trans _ (gFnorm _ _)) // centsC. + rewrite commGC -quotient_cents2 ?gFnorm_trans ?normsG //=. rewrite (sub_nilpotent_cent2 nilH'b) ?quotientS ?coprime_morph //. rewrite (pnat_coprime (pi_pnat pA t2Mp) (pi_pnat qQ _)) ?tau2'1 //. by rewrite (pnatPpi t1CEAb) // mem_primes q_pr cardG_gt0. @@ -1848,7 +1848,7 @@ have not_sNS_M: ~~ ('N(S) \subset M). have regNNS Z (Z1 := 'Ohm_1(Z)%G): Z \subset S -> cyclic Z -> Z :!=: 1 -> 'N(S) \subset 'N(Z1) -> 'C_Ms(Z1) = 1. - move=> sZS cycZ ntZ nZ1_NS; apply: contraNeq not_sNS_M => nregZ1. - have sZ1S: Z1 \subset S := subset_trans (Ohm_sub 1 Z) sZS. + have sZ1S: Z1 \subset S by apply: gFsub_trans. have EpZ1: Z1 \in 'E_p^1(E). rewrite p1ElemE // !inE (subset_trans sZ1S) //=. by rewrite (Ohm1_cyclic_pgroup_prime _ (pgroupS sZS pS)). @@ -1867,7 +1867,7 @@ have [cSU | not_cSU] := boolP (U \subset 'C(S)). by rewrite !inE groupX //= -expgM -expS expg_exponent. have sSZ: S \subset 'Z(U) by rewrite subsetI sSU centsC. have{sSZ} nsZU z: z \in S -> <[z]> <| U. - by move/(subsetP sSZ)=> ZUz; rewrite sub_center_normal ?cycle_subG. + by move=> Sz; rewrite sub_center_normal ?cycle_subG ?(subsetP sSZ). have [homoS | ltSnS1] := eqVproper sSnS1. have Ep2A_M := subsetP (pnElemS p 2 sEM) A Ep2A. have [_ _ _ _ [A1 EpA1 regA1]] := tau2_context maxM t2p Ep2A_M. @@ -1937,7 +1937,7 @@ have ltQ01: Q0 \proper Q1. have [X]: exists2 X, X \in subgroups Q & ('C_S(X) != 1) && ([~: S, X] != 1). apply/exists_inP; apply: contraFT (ltnn 1); rewrite negb_exists_in => irrS. have [sQ01 not_sQ10] := andP ltQ01. - have qQb: q.-group (Q / Q0) by exact: quotient_pgroup. + have qQb: q.-group (Q / Q0) by apply: quotient_pgroup. have ntQ1b: Q1 / Q0 != 1 by rewrite -subG1 quotient_sub1. have ntQb: Q / Q0 != 1 := subG1_contra (quotientS _ sQ1Q) ntQ1b. have{irrS} regQ: semiregular (S / Q0) (Q / Q0). @@ -1951,7 +1951,7 @@ have [X]: exists2 X, X \in subgroups Q & ('C_S(X) != 1) && ([~: S, X] != 1). apply: contraNneq ntXb; move/commG1P => cXS. by rewrite quotientS1 // subsetI sXQ centsC. have{regQ} cycQb: cyclic (Q / Q0). - have nSQb: Q / Q0 \subset 'N(S / Q0) by exact: quotient_norms. + have nSQb: Q / Q0 \subset 'N(S / Q0) by apply: quotient_norms. apply: odd_regular_pgroup_cyclic qQb (mFT_quo_odd _ _) _ nSQb regQ. rewrite -(isog_eq1 (quotient_isog _ _)) ?coprime_TIg 1?coprime_sym //. by rewrite cents_norm // centsC subsetIr. @@ -1959,7 +1959,7 @@ have [X]: exists2 X, X \in subgroups Q & ('C_S(X) != 1) && ([~: S, X] != 1). apply/eqP; rewrite (rank_Sylow sylQ1). by rewrite (tau1E maxM hallE) in t1q; case/and3P: t1q. have: 'r(Q) <= 1; last apply: leq_trans. - have nQ0_Ohm1Q := subset_trans (Ohm_sub 1 Q) nQ0Q. + have nQ0_Ohm1Q: 'Ohm_1(Q) \subset 'N(Q0) by apply: gFsub_trans. rewrite -rQ1 -rank_Ohm1 rankS // -(quotientSGK _ sQ01) //. rewrite (subset_trans (morphim_Ohm _ _ nQ0Q)) //= -quotientE -/Q0. rewrite -(cardSg_cyclic cycQb) ?Ohm_sub ?quotientS //. @@ -1977,17 +1977,16 @@ have [X]: exists2 X, X \in subgroups Q & ('C_S(X) != 1) && ([~: S, X] != 1). have: 'Z(E) \subset 'C_E(A); last apply: subset_trans. by rewrite setIS ?centS // normal_sub. have [x Ex sQ1xE1] := Hall_pJsub hallE1 t1q sQ1E qQ1. - rewrite -(conjSg _ _ x) (normsP (normal_norm (center_normal E))) //. + rewrite -(conjSg _ _ x) ['Z(E) :^ x](normsP _ x Ex) ?gFnorm //. set Q11x := _ :^ x; have oQ11x: #|Q11x| = q. by rewrite cardJg (Ohm1_cyclic_pgroup_prime _ qQ1) // -rank_gt0 rQ1. - apply: regE1subZ. + apply: regE1subZ; rewrite /= -/Q11x. apply/nElemP; exists q; rewrite p1ElemE // !inE oQ11x. by rewrite (subset_trans _ sQ1xE1) //= conjSg Ohm_sub. - have: cyclic Q11x by rewrite prime_cyclic ?oQ11x. - case/cyclicP=> y defQ11x; rewrite /= -/Q11x defQ11x cent_cycle regU13 //. + have /cyclicP[y defQ11x]: cyclic Q11x by rewrite prime_cyclic ?oQ11x. + rewrite defQ11x cent_cycle regU13 //. rewrite !inE -order_gt1 -cycle_subG /order -defQ11x oQ11x prime_gt1 //. - rewrite sub_conjg (subset_trans (Ohm_sub 1 Q1)) //. - by rewrite (normsP (normal_norm nsUE)) ?groupV. + by rewrite -(normsP (normal_norm nsUE) x Ex) conjSg gFsub_trans. by rewrite /p_elt /order -defQ11x oQ11x pnatE //; apply/orP; left. rewrite inE in sylS2; have [sS2H _]:= andP nsS2H. have sylS2_H := pHall_subl sS2H (subsetT H) sylS2. @@ -2067,7 +2066,7 @@ have defZ: 'Ohm_1 ('Z(P)) = Z. have ncycQ: ~~ cyclic Q := contra (@cyclic_abelian _ Q) not_cQQ. have rQgt1: 'r_p(Q) > 1. by rewrite ltnNge -(odd_pgroup_rank1_cyclic pQ) ?mFT_odd. -have [A Ep2A]: exists A, A \in 'E_p^2(Q) by exact/p_rank_geP. +have [A Ep2A]: exists A, A \in 'E_p^2(Q) by apply/p_rank_geP. wlog uniqNEpA: M H maxM maxH sP_MH sNMH sPM sPH sylP_M sylP_H / ~~ [exists A0 in 'E_p^1(A) :\ Z, 'M('N(A0)) == [set M]]. - move=> IH; case: exists_inP (IH M H) => [[A0 EpA0 defMA0] _ | _ -> //]. @@ -2153,7 +2152,7 @@ have sMp: p \in \sigma(M). have [bMp | sXMs'] := orP bMp_sXMs'; first by rewrite beta_sub_sigma. rewrite -pnatE // -[p]oX; apply: pgroupS (subset_trans sXMs' (der_sub 1 _)) _. exact: pcore_pgroup. -have hallMs: \sigma(M).-Hall(M) Ms by exact: Msigma_Hall. +have hallMs: \sigma(M).-Hall(M) Ms by apply: Msigma_Hall. have sXMs: X \subset Ms by rewrite (sub_Hall_pcore hallMs) // /pgroup oX pnatE. have [P sylP sXP]:= Sylow_superset sXMs pX. have sylP_M: p.-Sylow(M) P := subHall_Sylow hallMs sMp sylP. @@ -2178,7 +2177,7 @@ have{bMp_sXMs'} [bM'p sXMs']: p \notin \beta(M) /\ X \subset Ms^`(1). move: bMp_sXMs'; rewrite !inE -negb_exists_in. by case: exists_inP => // [[]]; exists P. have defMs: 'O_p^'(Ms) ><| P = Ms. - by have [_ hallMp' _] := beta_max_pdiv maxM bM'p; exact/sdprod_Hall_p'coreP. + by have [_ hallMp' _] := beta_max_pdiv maxM bM'p; apply/sdprod_Hall_p'coreP. have{defMs} sXP': X \subset P^`(1). have{defMs} [_ defMs nMp'P tiMp'P] := sdprodP defMs. have [injMp'P imMp'P] := isomP (quotient_isom nMp'P tiMp'P). @@ -2198,7 +2197,7 @@ have sXZ: X \subset 'Z(P). rewrite (subset_trans sXP') // -(der_cprod 1 defP) (derG1P cRR) cprodg1. have{dimQ} dimQ: logn p #|Q| <= 3 by rewrite dimQ. have [[_ ->] _] := p3group_extraspecial (pgroupS sQP pP) not_cQQ dimQ. - by case/cprodP: (center_cprod defP) => _ <- _; exact: mulG_subl. + by case/cprodP: (center_cprod defP) => _ <- _; apply: mulG_subl. have uniqP: 'M(P) = [set M]. exact: def_uniq_mmax (nonabelian_Uniqueness pP not_cPP) maxM sPM. rewrite (def_uniq_mmaxS _ ltCXG uniqP) //. @@ -2376,7 +2375,7 @@ without loss sXMs: M maxM sM_Y sMq / X \subset M`_\sigma. by exists (y * x^-1); rewrite conjsgM sub_conjgV -MsigmaJ. have:= parts_ab (E :^ x)%G p H; rewrite tau1J /= cardJg pHallJ2. rewrite (eq_pHall _ _ (eq_negn (sigmaJ _ _))). - by rewrite 2!orbit_sym (orbit_transl (mem_orbit _ _ _)) //; apply. + by rewrite 2!orbit_sym (orbit_eqP (mem_orbit _ _ _)) //; apply. have pre_part_a E p H: \sigma(M)^'.-Hall(M) E -> p \in \pi(E) -> H \in 'M(Y) -> gval H \notin M :^: G -> 'r_p(H :&: M) <= 1. @@ -2456,10 +2455,9 @@ split=> // t1Mp; rewrite (contra ((piSg (dergS 1 sNHY_L)) p)) // -p'groupEpi. have nsKL: K <| L by rewrite /K; case: ifP => _; apply: pcore_normal. have [sKL nKL] := andP nsKL; have nKML := subset_trans (subsetIr M L) nKL. suffices: p^'.-group (K * (M :&: L)^`(1)). - have sder := subset_trans (der_sub 1 _). - rewrite -norm_joinEr ?sder //; apply: pgroupS => /=. - rewrite norm_joinEr -?quotientSK ?sder //= !quotient_der //. - by rewrite -{1}defL quotientMidl. + rewrite -norm_joinEr ?gFsub_trans //; apply: pgroupS => /=. + rewrite norm_joinEr -?quotientSK ?gFsub_trans //= !quotient_der //. + by rewrite -[in L / K]defL quotientMidl. rewrite pgroupM p'K (pgroupS (dergS 1 (subsetIl M L))) // p'groupEpi. by rewrite mem_primes andbA andbC negb_and; case/and3P: t1Mp => _ _ ->. Qed. @@ -2641,9 +2639,8 @@ have sQM': Q \subset M^`(1). have ntMa: Ma != 1. apply: contraNneq nonuniqNQ => Ma1. rewrite (mmax_normal maxM _ ntQ) ?mmax_sup_id //. - apply: char_normal_trans (der_normal 1 M). have sylQ_M': q.-Sylow(M^`(1)) Q := pHall_subl sQM' (der_sub 1 M) sylQ. - rewrite (nilpotent_Hall_pcore _ sylQ_M') ?pcore_char //. + rewrite (nilpotent_Hall_pcore _ sylQ_M') ?gFnormal_trans //. by rewrite (isog_nil (quotient1_isog _)) -Ma1 Malpha_quo_nil. have a'q: q \notin \alpha(M). apply: contra nonuniqNQ => a_q. diff --git a/mathcomp/odd_order/BGsection13.v b/mathcomp/odd_order/BGsection13.v index be68f9d..f5e3a9c 100644 --- a/mathcomp/odd_order/BGsection13.v +++ b/mathcomp/odd_order/BGsection13.v @@ -1,9 +1,16 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div path fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div path fintype. +From mathcomp Require Import bigop finset prime fingroup morphism perm automorphism quotient. +From mathcomp Require Import action gproduct gfunctor pgroup cyclic center commutator. +From mathcomp Require Import gseries nilpotent sylow abelian maximal hall frobenius. +From mathcomp Require Import BGsection1 BGsection3 BGsection4 BGsection5 BGsection6. +From mathcomp Require Import BGsection7 BGsection9 BGsection10 BGsection12. (******************************************************************************) @@ -54,9 +61,9 @@ have nsHbH: H`_\beta <| H := pcore_normal _ _; have [_ nHbH] := andP nsHbH. have sYH := subset_trans sYH' (der_sub 1 H); have nHbY := subset_trans sYH nHbH. have nsHbY_H: H`_\beta <*> Y <| H. rewrite -{2}(quotientGK nsHbH) -quotientYK ?cosetpre_normal //. - rewrite (char_normal_trans _ (der_normal 1 _)) //= -quotient_der //. - rewrite (nilpotent_Hall_pcore _ (quotient_pHall nHbY sylY)) ?pcore_char //. - exact: Mbeta_quo_nil. + have ->: Y / H`_\beta = 'O_q(H^`(1) / H`_\beta). + by apply: nilpotent_Hall_pcore; rewrite ?Mbeta_quo_nil ?quotient_pHall. + by rewrite quotient_der ?gFnormal_trans. have sYNY: Y \subset 'N_H(Y) by rewrite subsetI sYH normG. have{nsHbY_H} defH: H`_\beta * 'N_H(Y) = H. rewrite -(mulSGid sYNY) mulgA -(norm_joinEr nHbY). @@ -102,10 +109,10 @@ split=> // [P sPMH pP | t1Mp]; last first. have nsHaH: H`_\alpha <| H := pcore_normal _ _; have [_ nHaH] := andP nsHaH. have [sPM sPH] := subsetIP sPMH; have nHaS := subset_trans sSH nHaH. have nsHaS_H: H`_\alpha <*> S <| H. - rewrite -{2}(quotientGK nsHaH) -quotientYK ?cosetpre_normal //. - rewrite (char_normal_trans _ (der_normal 1 _)) //= -quotient_der //. - rewrite (nilpotent_Hall_pcore _ (quotient_pHall nHaS sylS_H')) ?pcore_char //. - exact: Malpha_quo_nil. + rewrite -[H in _ <| H](quotientGK nsHaH) -quotientYK ?cosetpre_normal //. + have ->: S / H`_\alpha = 'O_p(H^`(1) / H`_\alpha). + by apply: nilpotent_Hall_pcore; rewrite ?Malpha_quo_nil ?quotient_pHall. + by rewrite quotient_der ?gFnormal_trans. have [sHaS_H nHaS_H] := andP nsHaS_H. have sP_HaS: P \subset H`_\alpha <*> S. have [x Hx sPSx] := Sylow_subJ sylS sPH pP; apply: subset_trans sPSx _. @@ -246,7 +253,7 @@ have [[sPE abelP dimP] [sRE abelR dimR]] := (pnElemP EpP, pnElemP ErR). have [sPM sRM] := (subset_trans sPE sEM, subset_trans sRE sEM). have [[pP cPP _] [rR _]] := (and3P abelP, andP abelR). have coCR: coprime #|C| #|R| := coprimeSg sCMs (coprimegS sRE coMsE). -have ntP: P :!=: 1 by exact: nt_pnElem EpP _. +have ntP: P :!=: 1 by apply: nt_pnElem EpP _. pose ST := [set S | Sylow C (gval S) & R \subset 'N(S)]. have sST_CP S: S \in ST -> S \subset C by case/setIdP=> /SylowP[q _ /andP[]]. rewrite -{sST_CP}[C](Sylow_transversal_gen sST_CP) => [|q _]; last first. @@ -349,7 +356,7 @@ have{abelR dimR} ErR: R \in 'E_r^1('C_E(P)). rewrite !inE abelR dimR (subset_trans sRE1) // subsetI sE1E. by rewrite sub_abelian_cent ?cyclic_abelian. rewrite centsC (subset_trans (cent_tau1Elem_Msigma t1p EpP ErR)) //. -have [y defR]: exists y, R :=: <[y]> by apply/cyclicP; exact: cyclicS cycS. +have [y defR]: exists y, R :=: <[y]> by apply/cyclicP; apply: cyclicS cycS. have sylS_E: r.-Sylow(E) S. apply: subHall_Sylow hallE1 (pnatPpi t1E1 _) (sylS). by rewrite -p_rank_gt0 -(rank_Sylow sylS) rank_gt0. @@ -676,9 +683,8 @@ suffices sXL: X \subset L. rewrite coprime_TIg // coprime_morphl // (coprimeSg (subsetIl _ _)) //. exact: pnat_coprime (pcore_pgroup _ _) (pi_pnat (pcore_pgroup _ _) _). rewrite commg_subI // subsetI. - rewrite quotientS; last by rewrite subsetI sXMb. - rewrite (char_norm_trans (pcore_char _ _)) ?quotient_norms //. - by rewrite (subset_trans sXL) ?der_norm. + rewrite quotientS /=; last by rewrite subsetI sXMb. + by rewrite quotient_der ?gFnorm_trans ?normsG ?quotientS. rewrite (sub_Hall_pcore (nilpotent_pcore_Hall _ (Mbeta_quo_nil _))) //. rewrite quotient_pgroup ?quotient_norms //. by rewrite normsI ?(subset_trans sQM nMbM) ?normsG. @@ -696,10 +702,10 @@ have defM: M`_\beta * 'N_M(Q) = M. have nMbQ := subset_trans sQM nMbM. have nsMbQ_M: M`_\beta <*> Q <| M. rewrite -{2}(quotientGK nsMbM) -quotientYK ?cosetpre_normal //. - rewrite (eq_Hall_pcore (nilpotent_pcore_Hall q (Mbeta_quo_nil _))) //. - apply: char_normal_trans (pcore_char _ _) (quotient_normal _ _). - exact: der_normal. - rewrite quotient_pHall // (pHall_subl _ (der_sub 1 M) sylQ) //. + suffices ->: Q / M`_\beta = 'O_q(M^`(1) / M`_\beta). + by rewrite quotient_der ?nMbM ?gFnormal_trans. + apply: nilpotent_Hall_pcore; first exact: Mbeta_quo_nil. + rewrite quotient_pHall // (pHall_subl _ _ sylQ) ?gFsub //. by rewrite -defQ commgSS // (subset_trans nUP). have sylQ_MbQ := pHall_subl (joing_subr _ Q) (normal_sub nsMbQ_M) sylQ. rewrite -{3}(Frattini_arg nsMbQ_M sylQ_MbQ) /= norm_joinEr // -mulgA. @@ -727,7 +733,7 @@ have{not_pM'} [R ErR nQR]: exists2 R, R \in 'E_r^1('C_M(P)) & R \subset 'N(Q). rewrite pHallE sSK /= -/K -(setIidPr sKM) -defM -group_modl // setIAC. rewrite (setIidPr sKM) -LagrangeMr partnM // -(card_Hall sylS). rewrite part_p'nat ?mul1n 1?(pnat_dvd (dvdn_indexg _ _)) //. - by apply: (pi_p'nat bMb); apply: contra sM'r; exact: beta_sub_sigma. + by apply: (pi_p'nat bMb); apply: contra sM'r; apply: beta_sub_sigma. have rC: 'r_r('C_M(P)) > 0 by rewrite p_rank_gt0 (piSg _ piHr) // subsetI sHM. have{rC} [R ErR] := p_rank_geP rC; have [sRcMP abelR _] := pnElemP ErR. have{sRcMP abelR} [[sRM cPR] [rR _]] := (subsetIP sRcMP, andP abelR). @@ -774,7 +780,7 @@ without loss sylS_L: L maxL sLq notMGL / q.-Sylow(L) S. have [T sylT] := Sylow_exists q L; have sylT_G := sigma_Sylow_G maxL sLq sylT. have [x Gx ->] := Sylow_trans sylT_G (sigma_Sylow_G maxM sMq sylS_M). case/(_ (L :^ x)%G); rewrite ?mmaxJ ?sigmaJ ?pHallJ2 //. - by rewrite (orbit_transr _ (mem_orbit 'Js L Gx)). + by rewrite (orbit_transl _ (mem_orbit 'Js L Gx)). have [[sSL _] [[E1 hallE1] [E3 hallE3]]] := (andP sylS_L, ex_tau13_compl hallE). have [E2 hallE2 complEi] := ex_tau2_compl hallE hallE1 hallE3. have E2_1: E2 :==: 1. @@ -1104,7 +1110,7 @@ have defP: 'C_A(Q) = P. rewrite (card_pgroup (pgroupS _ pA)) ?subsetIl // (card_pgroup pP) dimP. rewrite leq_exp2l ?prime_gt1 ?(pnElem_prime EpP) //. by rewrite -ltnS -dimA properG_ltn_log // /proper subsetIl subsetIidl. -have EqFQ: Q \in 'E_q^1(F) by exact/pnElemP. +have EqFQ: Q \in 'E_q^1(F) by apply/pnElemP. have regQLs: 'C_(L`_\sigma)(Q) = 1. by rewrite (tau12_regular maxL hallF t1Lq EqFQ t2Lp Ep2A) // defP. have ntAQ: [~: A, Q] != 1 by rewrite (sameP eqP commG1P). diff --git a/mathcomp/odd_order/BGsection14.v b/mathcomp/odd_order/BGsection14.v index 493f634..3414af3 100644 --- a/mathcomp/odd_order/BGsection14.v +++ b/mathcomp/odd_order/BGsection14.v @@ -1,10 +1,18 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div path fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div path fintype. +From mathcomp Require Import bigop finset prime fingroup morphism perm automorphism quotient. +From mathcomp Require Import action gproduct gfunctor pgroup cyclic center commutator. +From mathcomp Require Import gseries nilpotent sylow abelian maximal hall frobenius. +From mathcomp Require Import ssralg ssrnum ssrint rat. +From mathcomp Require Import BGsection1 BGsection3 BGsection4 BGsection5 BGsection6. +From mathcomp Require Import BGsection7 BGsection9 BGsection10 BGsection12 BGsection13. (******************************************************************************) @@ -428,7 +436,7 @@ have [ntS cycS]: S :!=: 1 /\ cyclic S. by rewrite -2!andb_orr orNb andbT inE /= sM'p in t13p. have [p_pr _ _] := pgroup_pdiv pS ntS. have oA: #|A| = p by rewrite (Ohm1_cyclic_pgroup_prime cycS pS). -have sAM: A \subset M := subset_trans (Ohm_sub 1 S) sSM. +have sAM: A \subset M by apply: gFsub_trans. have regA: 'C_Ms(A) = 1. apply: contraNeq kM'p => nregA; rewrite unlock; apply/andP; split=> //. by apply/exists_inP; exists [group of A]; rewrite ?p1ElemE // !inE sAM oA /=. @@ -448,7 +456,7 @@ move=> notP1maxM; suffices [maxM]: M \in 'M /\ ~~ \sigma_kappa(M).-group M. have [/setIdP[maxM k'M] | /setDP[PmaxM]] := orP notP1maxM; last first. by rewrite inE PmaxM; case/setDP: PmaxM. split=> //; apply: contra (not_sigma_mmax maxM). -by apply: sub_in_pnat => p piMp /orP[] // /idPn[]; exact: (pnatPpi k'M). +by apply: sub_in_pnat => p piMp /orP[] // /idPn[]; apply: (pnatPpi k'M). Qed. (* This is B & G, Proposition 14.2. *) @@ -592,7 +600,7 @@ have [have_a nK1K ntE1 sE1K]: [/\ part_a, b1_hyp, E1 :!=: 1 & E1 \subset K]. by rewrite setIS ?centS. have defM := sdprod_sigma maxM hallE. rewrite /b1_hyp -defK; split=> //; exists 1%G; last first. - by split; [exact: abelian1 | rewrite -defK | exact: semiregular1l]. + by split; [apply: abelian1 | rewrite -defK | apply: semiregular1l]. rewrite sdprod1g; do 2?split=> //; rewrite ?mul1g ?groupP -?defK //. rewrite pHallE sub1G cards1 eq_sym partG_eq1 pgroupNK /=. have{defM} [_ defM _ _] := sdprodP defM; rewrite -{2}defM defK pgroupM. @@ -706,9 +714,9 @@ have [regMsU nilMs]: 'C_Ms(U) = 1 /\ nilpotent Ms. have sk'p := pnatPpi sk'U piUp. have [S sylS] := Sylow_exists p U; have [sSU _] := andP sylS. have sylS_M := subHall_Sylow hallU sk'p sylS. - rewrite -(setIidPr (centS (subset_trans (Ohm_sub 1 S) sSU))) setIA. - have [|_ _ -> ->] := sigma'_kappa'_facts maxM sylS_M; last by rewrite setI1g. - by rewrite -negb_or (piSg sUM). + have [|_ _ regMsS1 nilMs] := sigma'_kappa'_facts maxM sylS_M. + by rewrite -negb_or (piSg sUM). + by split=> //; apply/trivgP; rewrite -regMsS1 setIS ?centS ?gFsub_trans. have [[_ F _ defF] _ _ _] := sdprodP defM; rewrite defF in defM. have hallMs: \sigma(M).-Hall(M) Ms by apply: Msigma_Hall. have hallF: \sigma(M)^'.-Hall(M) F by apply/(sdprod_Hall_pcoreP hallMs). @@ -834,7 +842,7 @@ have kp: p \in \kappa(M). have [sXM abelX dimX] := pnElemP EpX; have [pX _] := andP abelX. have [K hallK sXK] := Hall_superset (mmax_sol maxM) sXM (pi_pgroup pX kp). have PmaxM: M \in 'M_'P. - by rewrite 2!inE maxM andbT; apply: contraL kp => k'M; exact: (pnatPpi k'M). + by rewrite 2!inE maxM andbT; apply: contraL kp => k'M; apply: (pnatPpi k'M). have [_ [defNK defNX] [_ uniqCKs] _ _] := Ptype_structure PmaxM hallK. have{defNX} sCMy_nMK: 'C_M[y] \subset 'N_M(K). have [|<- _] := defNX X. @@ -931,7 +939,7 @@ have neqNM: N :!=: M by apply: contraNneq not_sCX_M => <-. have maxNX'_N: N \in 'M('N(X)) :\ M by rewrite 2!inE neqNM. have [notMGN _] := sigma_subgroup_embedding maxM sMq sXM qX ntX maxNX'_N. have sN'q: q \notin \sigma(N). - by apply: contraFN (sigma_partition maxM maxN notMGN q) => sNq; exact/andP. + by apply: contraFN (sigma_partition maxM maxN notMGN q) => sNq; apply/andP. rewrite (negPf sN'q) => [[t2Nq s_piM_bN hallMN]]. have defN: N`_\sigma ><| (M :&: N) = N. exact/(sdprod_Hall_pcoreP (Msigma_Hall maxN)). @@ -1155,7 +1163,7 @@ have tiP: trivIset P. have->: class_support M^~~ G = cover P. apply/setP=> az; apply/imset2P/bigcupP=> [[a z] | [xRz]]. case/bigcupP=> x Ms_x xRa Gz ->; exists (x ^ z *: 'R[x ^ z]). - by apply: mem_imset; exact: mem_imset2. + by apply: mem_imset; apply: mem_imset2. by rewrite sigma_coverJ memJ_conjg. case/imsetP=> _ /imset2P[x z Ms_x Gz ->] ->; rewrite sigma_coverJ. by case/imsetP=> a xRa ->; exists a z => //; apply/bigcupP; exists x. @@ -1646,7 +1654,7 @@ have [Mi MXi P2maxMi]: exists2 Mi, Mi \in MX & Mi \in 'M_'P2. have [[PmaxMi _] [PmaxMj _]] := (PmaxMX _ MXi, PmaxMX _ MXj). have [[maxMi _] [maxMj _]] := (setDP PmaxMi, setDP PmaxMj). apply: sigma_support_disjoint; rewrite ?mmaxJ //. - rewrite (orbit_transr _ (mem_orbit _ _ _)) ?inE //=. + rewrite (orbit_transl _ (mem_orbit _ _ _)) ?inE //=. apply: contra (ntKsX _ MXi); case/imsetP=> y _ /= defMj; rewrite -/(Ks_ _). have sKisKj: Ks_ Mi \subset K_ Mj by rewrite sKsKX // eq_sym. rewrite -(setIidPl sKisKj) coprime_TIg //. @@ -1884,7 +1892,7 @@ rewrite {}/S {}/Ls in Sx; without loss a1: a H L PmaxH hallL Sx / a = 1. move/(_ 1 (H :^ a)%G (L :^ a)%G); rewrite conjsg1 PtypeJ PmaxH pHallJ2. rewrite (eq_pHall _ _ (kappaJ H a)) hallL MsigmaJ centJ. rewrite -conjIg -conjYg -conjUg -conjDg Sx !inE. - by rewrite !(orbit_transr _ (mem_orbit _ _ _)) ?inE //; exact. + by rewrite !(orbit_transl _ (mem_orbit _ _ _)) ?inE //; apply. have [_ [defNL _] [_ uniqH] _ _] := Ptype_structure PmaxH hallL. do [rewrite {a}a1 conjsg1; set Ls := 'C_(_)(L)] in Sx defNL. have{x Sx Tx} [Mk MXk ntLsMks]: exists2 Mk, Mk \in MX & Ls :&: Ks_ Mk != 1. @@ -1962,8 +1970,8 @@ have tiPcover: trivIset Pcover. rewrite -setI_eq0 !{1}class_supportEr big_distrr big1 //= => a Ga. rewrite big_distrl big1 //= => b Gb; apply/eqP. rewrite -!{1}sigma_supportJ setI_eq0 sigma_support_disjoint ?mmaxJ //. - apply: contra notMGH; rewrite {a Ga}(orbit_transr _ (mem_orbit _ _ Ga)). - rewrite {b Gb}(orbit_transl (mem_orbit _ _ Gb))=> /imsetP[c Gc ->] /=. + apply: contra notMGH; rewrite {a Ga}(orbit_transl _ (mem_orbit _ _ Ga)). + rewrite {b Gb}(orbit_eqP (mem_orbit _ _ Gb))=> /imsetP[c Gc ->] /=. by rewrite sigma_supportJ class_supportGidl. have ntPcover: cover Pcover \subset G^#. apply/bigcupsP=> _ /imsetP[M maxM ->]; rewrite class_supportEr. @@ -2174,11 +2182,11 @@ have piKp: p \in \pi(K) by rewrite pi_pdiv cardG_gt1. have t2Mp: p \in \tau2(M). have s'p := pnatPpi s'K piKp. have sylKp: p.-Sylow(K) 'O_p(K) := nilpotent_pcore_Hall p (abelian_nil cKK). + have ntKp: 'O_p(K) != 1 by rewrite -rank_gt0 (rank_Sylow sylKp) p_rank_gt0. rewrite inE /= s'p ?(sigma'_norm_mmax_rank2 maxM s'p (pHall_pgroup sylKp)) //. - rewrite (mmax_normal maxM) ?(char_normal_trans (pcore_char _ _)) //. - by rewrite -rank_gt0 (rank_Sylow sylKp) p_rank_gt0. + by rewrite (mmax_normal maxM) ?gFnormal_trans. have [A EpA _] := ex_tau2Elem hallE t2Mp. -have [sAE] := pnElemP EpA; case/andP=> pA _ dimA. +have [sAE /andP[pA _] dimA] := pnElemP EpA. have [[nsAE _] _ _ _] := tau2_compl_context maxM hallE t2Mp EpA. have nAQ := subset_trans sQE (normal_norm nsAE). have [S sylS sAS]:= Sylow_superset (subsetT A) pA. @@ -2288,8 +2296,8 @@ have snK_sMst L: K <|<| L -> L \subset Mst. apply/subsetP=> a Aa; rewrite -groupV sK_uniqMst // (subset_trans sKL) //. by rewrite -sub_conjg (normsP (normal_norm nsLA)). have sEH: E \subset H. - apply: subset_trans (char_norm_trans _ (normal_norm nsUE)) sNRH. - by rewrite (nilpotent_Hall_pcore (abelian_nil cUU) sylR) pcore_char. + have defR: R :=: 'O_r(U) := nilpotent_Hall_pcore (abelian_nil cUU) sylR. + by apply: subset_trans sNRH; rewrite defR gFnorm_trans ?normal_norm. have [sUH sKH]: U \subset H /\ K \subset H by apply/mulGsubP; rewrite mulUK. have notMstGH: gval H \notin Mst :^: G. apply: contra ntR => /imsetP[a _ defH]. @@ -2332,7 +2340,7 @@ have defUK: [~: U, K] = U. have qK: q.-group K := pnat_id q_pr. have sUHs: U \subset H`_\sigma. have [nsHsH _ mulHsD nHsD _] := sdprod_context (sdprod_sigma maxH hallD). - have nHsDq := subset_trans (pcore_sub q D) nHsD. + have nHsDq: 'O_q(D) \subset 'N(H`_\sigma) by apply: gFsub_trans. pose HsDq := H`_\sigma <*> 'O_q(D). have defHsDq: H`_\sigma * 'O_q(D) = HsDq by rewrite -norm_joinEr. have hallHs_HsDq: q^'.-Hall(HsDq) H`_\sigma. @@ -2347,16 +2355,14 @@ have sUHs: U \subset H`_\sigma. by rewrite -{3}mulHsD quotientMidl quotient_normal // pcore_normal. have sU_HsDq: U \subset HsDq. by rewrite -defUK (subset_trans (commgSS sUH sK_HsDq)) // commg_subr. - rewrite (sub_normal_Hall hallHs_HsDq) //. - rewrite p'groupEpi; apply: (contraL (pnatPpi sk'M_U)) => /=. - by rewrite inE /= orbC (pnatPpi kK). - exact: normalS (joing_subl _ _) _ (pcore_normal _ _). + rewrite (sub_normal_Hall hallHs_HsDq) ?normalYl // p'groupEpi. + by apply: contraL (pnatPpi sk'M_U) _; rewrite !inE /= orbC (pnatPpi kK). have defNMU: 'N_M(U) = E. have [_ mulHsE nHsE _] := sdprodP (sdprod_sigma maxM hallE). have [sUE nUE] := andP nsUE; rewrite -mulHsE -normC // -group_modl //=. rewrite coprime_norm_cent ?(subset_trans sUE) //; last first. exact: coprimegS sUE (coprime_sigma_compl hallE). - have sR1U: 'Ohm_1(R) \subset U := subset_trans (Ohm_sub 1 R) (pHall_sub sylR). + have sR1U: 'Ohm_1(R) \subset U := gFsub_trans _ (pHall_sub sylR). rewrite (trivgP (subset_trans (setIS _ (centS sR1U)) _)) ?mulg1 //. have [|_ _ -> //] := sigma'_kappa'_facts maxM sylR_M. by rewrite s'Mr (piSg sUM). @@ -2371,7 +2377,7 @@ suffices ->: H :&: Mst = D. have [sUFH nilFH] := (subset_trans sUHs sHsFH, Fitting_nil H). have hallFu: sk'.-Hall('F(H)) Fu := nilpotent_pcore_Hall sk' nilFH. have sUFu: U \subset Fu by rewrite (sub_Hall_pcore hallFu). - have nsFuH: Fu <| H := char_normal_trans (pcore_char _ _) (Fitting_normal _). + have nsFuH: Fu <| H by rewrite !gFnormal_trans. have [[sFuFH sk'Fu _] [sFuH nFuH]] := (and3P hallFu, andP nsFuH). have defU: M :&: Fu = U. have sk'MFu: sk'.-group(M :&: Fu) := pgroupS (subsetIr M _) sk'Fu. diff --git a/mathcomp/odd_order/BGsection15.v b/mathcomp/odd_order/BGsection15.v index 2238534..c3e8004 100644 --- a/mathcomp/odd_order/BGsection15.v +++ b/mathcomp/odd_order/BGsection15.v @@ -1,10 +1,18 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq choice div fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq choice div fintype. +From mathcomp Require Import path bigop finset prime fingroup morphism perm automorphism. +From mathcomp Require Import quotient action gproduct gfunctor pgroup cyclic commutator. +From mathcomp Require Import center gseries nilpotent sylow abelian maximal hall frobenius. +From mathcomp Require Import BGsection1 BGsection2 BGsection3 BGsection4 BGsection5. +From mathcomp Require Import BGsection6 BGsection7 BGsection9 BGsection10 BGsection12. +From mathcomp Require Import BGsection13 BGsection14. (******************************************************************************) @@ -71,7 +79,7 @@ rewrite -(nilpotent_Fitting nilH) FittingEgen genS //. apply/bigcupsP=> [[p /= _] piHp]; rewrite (bigcup_max 'O_p(H)%G) //. have sylHp := nilpotent_pcore_Hall p nilH. have sylHp_M := subHall_Sylow hallH (pnatPpi pi_H piHp) sylHp. -by rewrite (p_Sylow sylHp_M) (char_normal_trans (pcore_char _ _)). +by rewrite (p_Sylow sylHp_M) gFnormal_trans. Qed. Lemma Fcore_dprod : \big[dprod/1]_(P | Sylow M (gval P) && (P <| M)) P = M`_\F. @@ -384,7 +392,7 @@ have{sMsM'} sKsQ: Ks \subset Q. apply/eqP; rewrite eqEsubset Fitting_sub /= -{1}defMsK. rewrite (odd_sdprod_primact_commg_sub_Fitting defM) ?mFT_odd //. apply/trivgP; rewrite -tiKsFM setIAC setSI //= -/Ms subsetI Fitting_sub /=. - by rewrite Fitting_max ?Fitting_nil // (char_normal_trans (Fitting_char _)). + by rewrite Fitting_max ?Fitting_nil // gFnormal_trans. have nilMs_Q: nilpotent (Ms / Q). have [nMsK1 tiQK1] := (subset_trans sK1K nMsK, coprime_TIg coQK1). have prK1b: prime #|K1 / Q| by rewrite -(card_isog (quotient_isog _ _)). @@ -436,7 +444,7 @@ have{K1 sK1M sK1K coMsK1 coQK1 prK1 defCMsK1 nQK1 solMs} Qi_rec Qi: case/subsetIP=> sLM nQij nLN; exists L. have{sLbQ} sLQ: L \subset Q by rewrite -(quotientSGK nQij sQiQ) -defLb. rewrite inE /psubgroup /normal sLQ sQij nQij (pgroupS sLQ qQ) -defLb. - have nLDK: D <*> K \subset 'N(L) by apply: subset_trans nLN; exact/subsetIP. + have nLDK: D <*> K \subset 'N(L) by apply: subset_trans nLN; apply/subsetIP. have sLD_Ms: L <*> D \subset Ms by rewrite join_subG (subset_trans sLQ). have coLD_K1: coprime #|L <*> D| #|K1| := coprimeSg sLD_Ms coMsK1. have [[nQiD nQiK] [nLD nLK]] := (joing_subP nQiDK, joing_subP nLDK). @@ -490,7 +498,7 @@ have nsQ0M: Q0 <| M by rewrite /normal subIset ?sQM. have sFM_QCQ: 'F(M) \subset Q <*> 'C_M(Q). have [_ /= mulQQ' cQQ' _] := dprodP (nilpotent_pcoreC q (Fitting_nil M)). rewrite -{3}mulQQ' p_core_Fitting cent_joinEr ?subsetIr //= -/Q in cQQ' *. - by rewrite mulgS // subsetI (subset_trans (pcore_sub _ _) (Fitting_sub M)). + by rewrite mulgS // subsetI gFsub_trans ?gFsub. have sQCQ_CMsQb: Q <*> 'C_M(Q) \subset 'C_Ms(Qb | 'Q). rewrite join_subG !(subsetI _ Ms) sQMs /= !sub_astabQ nQ0Q /= -/Q0 -/Qb. rewrite -abelianE cQbQb quotient_cents ?subsetIr //= andbC subIset ?nQ0M //=. @@ -688,10 +696,8 @@ have nsMsM: M`_\sigma <| M := pcore_normal _ _; have [sMsM _] := andP nsMsM. have [nsHM | not_nsHM] := boolP (H <| M). by exists b; rewrite // (mmax_normal maxM) ?setIid. have neqMFs: M`_\F != M`_\sigma. - apply: contraNneq not_nsHM => eq_MF_Ms. - have nilMs: nilpotent M`_\sigma by apply/Fcore_eq_Msigma. - rewrite (eq_Hall_pcore (nilpotent_pcore_Hall _ nilMs) hallH). - exact: char_normal_trans (pcore_char _ _) nsMsM. + apply: contraNneq not_nsHM => /(Fcore_eq_Msigma maxM)nilMs. + by rewrite (nilpotent_Hall_pcore nilMs hallH) gFnormal_trans. have [K hallK] := Hall_exists \kappa(M) (mmax_sol maxM). pose q := #|'C_(M`_\sigma)(K)|. have [D hallD] := Hall_exists q^' (solvableS sMsM (mmax_sol maxM)). @@ -707,11 +713,10 @@ have [_ mulQD nQD tiQD] := sdprodP defMs; rewrite -/Q in mulQD nQD tiQD. have nQH := subset_trans sHMs (normal_norm nsQ_Ms). have nsQHM: Q <*> H <| M. rewrite -(quotientGK nsQM) -quotientYK // cosetpre_normal //= -/Q. - have nilMsQ: nilpotent (M`_\sigma / Q). - by rewrite -mulQD quotientMidl -(isog_nil (quotient_isog _ _)). - have hallMsQpi := nilpotent_pcore_Hall pi nilMsQ. - rewrite (eq_Hall_pcore hallMsQpi (quotient_pHall nQH hallH)). - by rewrite (char_normal_trans (pcore_char _ _)) ?quotient_normal. + suffices ->: H / Q = 'O_pi(M`_\sigma / Q). + by rewrite gFnormal_trans ?quotient_normal. + apply: nilpotent_Hall_pcore; last exact: quotient_pHall. + by rewrite /= -mulQD quotientMidl -(isog_nil (quotient_isog _ _)). have tiQH: Q :&: H = 1. apply: coprime_TIg (p'nat_coprime (pi_pgroup qQ _) piH). apply: contra not_nsHM => pi_q; rewrite (joing_idPr _) // in nsQHM. @@ -720,9 +725,9 @@ have defM: Q * 'N_M(H) = M. have hallH_QH: pi.-Hall(Q <*> H) H. by rewrite (pHall_subl (joing_subr _ _) _ hallH) // join_subG sQMs. have solQH := solvableS (normal_sub nsQHM) (mmax_sol maxM). - rewrite -{2}(Hall_Frattini_arg solQH nsQHM hallH_QH) /= norm_joinEr //. + rewrite -[RHS](Hall_Frattini_arg solQH nsQHM hallH_QH) /= norm_joinEr //. by rewrite -mulgA [H * _]mulSGid // subsetI (subset_trans sHMs sMsM) normG. -move: Mb; rewrite -{1}defM; case/mulsgP=> z n Qz Nn defb; exists n => //. +rewrite -defM in Mb; case/mulsgP: Mb => z n Qz Nn defb; exists n => //. rewrite def_xa defb conjgM [x ^ z](conjg_fixP _) // -in_set1 -set1gE -tiQH. rewrite inE {1}commgEr groupMr // -mem_conjgV groupV /=. rewrite (normsP (normal_norm nsQM)) ?Qz; last first. @@ -805,7 +810,7 @@ have defF: 'F(M`_\sigma) \x Y = 'F(M). rewrite Fitting_max ?(nilpotentS (pcore_sub _ _)) //=. by rewrite -(pcore_setI_normal _ nsFM) norm_normalI ?(subset_trans sMsM). rewrite /normal (char_norm_trans (Fitting_char _)) ?(subset_trans sFM) //. - by rewrite Fitting_max ?Fitting_nil // (char_normal_trans (Fitting_char _)). + by rewrite Fitting_max ?Fitting_nil ?gFnormal_trans. have [[ntH sHMs sMsM' _] nnil_struct] := Fcore_structure maxM. have hallH: \pi(H).-Hall(M`_\sigma) H := pHall_subl sHMs sMsM (Fcore_Hall M). have [X [_ cycX t2X defCH]] := cent_Hall_sigma_sdprod maxM hallH ntH. @@ -815,7 +820,7 @@ have hallX: \sigma(M)^'.-Hall('C_M(H)) X. by apply: sub_pgroup t2X => p /andP[]. have sYX: Y \subset X. rewrite (normal_sub_max_pgroup (Hall_max hallX)) ?pcore_pgroup //. - rewrite /normal (char_norm_trans (pcore_char _ _)) ?subIset ?nFM //= -/Y. + rewrite /normal gFnorm_trans ?subIset ?nFM //= -/Y andbT. have [_ _ cFsY _] := dprodP defF. rewrite subsetI sYM (sub_nilpotent_cent2 nilF) //= -/Y -/H. exact: pnat_coprime (pgroupS sHMs sMs) (pcore_pgroup _ _). @@ -826,25 +831,25 @@ have [U complU] := ex_kappa_compl maxM hallK. have [[defM _ cM'M'b] defM' _ _ _] := kappa_structure maxM complU. have d_holds: M \in 'M_'P -> 'F(M) \subset M^`(1). rewrite inE maxM andbT -(trivg_kappa maxM hallK) => ntK. - rewrite -(dprodW defF) mulG_subG (subset_trans (Fitting_sub _)) //= -/Y. + rewrite -(dprodW defF) mulG_subG gFsub_trans //= -/Y. have{defM'} [[defM' _] nsM'M] := (defM' ntK, der_normal 1 M). have hallM': \kappa(M)^'.-Hall(M) M^`(1). by apply/(sdprod_normal_pHallP nsM'M hallK); rewrite /= -defM'. - rewrite (sub_normal_Hall hallM') ?(sub_pgroup _ t2Y) // => p. - by case/andP=> _; apply: contraL => /rank_kappa->. + rewrite (sub_normal_Hall hallM') ?(sub_pgroup _ t2Y) // => p /andP[_]. + by apply: contraL => /rank_kappa->. have defF_H: 'C_M(H) \subset 'F(M) -> H \* 'C_M(H) = 'F(M). move=> sCHF; apply/eqP; rewrite cprodE ?subsetIr // eqEsubset ?mul_subG //=. have hallH_F := pHall_subl sHF sFM (Fcore_Hall M). have nsHF := normalS sHF sFM (Fcore_normal M). have /dprodP[_] := nilpotent_pcoreC \pi(H) nilF. rewrite (normal_Hall_pcore hallH_F nsHF) /= -/H => defF_H cHFH' _. - by rewrite -defF_H mulgS // subsetI (subset_trans (pcore_sub _ _)). + by rewrite -defF_H mulgS // subsetI gFsub_trans. have [eqHMs | neqHMs] := eqVneq H M`_\sigma. split=> //; [split=> // | by rewrite eqHMs abelian_nil]. by rewrite (subset_trans _ sHF) //= eqHMs der1_min ?comm_subG. rewrite defF_H // -(sdprodW defCH) -eqHMs mulG_subG subIset ?sHF //=. rewrite Fitting_max ?abelian_nil ?cyclic_abelian //. - rewrite -(normal_Hall_pcore hallX) ?(char_normal_trans (pcore_char _ _)) //. + rewrite -(normal_Hall_pcore hallX) ?gFnormal_trans //. by rewrite norm_normalI // eqHMs norms_cent. move: defCH; rewrite -dprodEsd; first by case/dprod_normal2. by rewrite -eqHMs (centsS (subsetIl _ _)); case/subsetIP: (pHall_sub hallX). @@ -1138,9 +1143,7 @@ have cycHp': cyclic 'O_p^'(H). by rewrite subsetI pcore_sub. rewrite {1}defX oX1 /= -[M`_\F]/(gval H) -/P; split=> //. pose Z q := 'Ohm_1('Z('O_q(H)))%G. -have charZ q: Z q \char H. - have:= char_trans (center_char _) (pcore_char q H). - exact: char_trans (Ohm_char 1 _). +have charZ q: Z q \char H by rewrite 3?gFchar_trans. have{cycHp'} oZ: {in \pi(H), forall q, #|Z q| = q}. move=> q piHp; have [-> // | p'q] := eqVneq q p. have qHq: q.-group 'O_q(H) := pcore_pgroup q H. @@ -1190,10 +1193,10 @@ have [K_dv_p1 | {regZq_dv_q1}] := altP (@implyP (Ks :==: Z0) (#|K| %| p.-1)). apply: contra neqZqKs => sKsZq; rewrite eqEsubset sKsZq /=. by rewrite prime_meetG ?oZ // (setIidPr sKsZq). rewrite {Z oZ charZ}negb_imply; case/andP; move/eqP=> defKs not_K_dv_p1. -have nPK: K \subset 'N(P) := char_norm_trans (pcore_char p H) nHK. +have nPK: K \subset 'N(P) by apply: gFnorm_trans. have defZP: 'Z(P) = Ks. apply/eqP; rewrite eqEsubset andbC {1}defKs Ohm_sub subsetI subIset ?sPH //. - have nZPK: K \subset 'N('Z(P)) := char_norm_trans (center_char _) nPK. + have nZPK: K \subset 'N('Z(P)) by apply: gFnorm_trans. have coZPK: coprime #|'Z(P)| #|K| := coprimeSg (center_sub _) coPK. rewrite centsC (coprime_odd_faithful_Ohm1 pZP) ?mFT_odd //. by rewrite /= -/Z0 -defKs centsC subsetIr. @@ -1217,7 +1220,7 @@ have rPle2: 'r(P) <= 2. by rewrite inE Ep2B (subsetP (pmaxElemS p (subsetT P))) // inE pmaxB inE. have [x defK] := cyclicP cycK; have Kx: x \in K by rewrite defK cycle_id. have nPx := subsetP nPK x Kx; rewrite /A defK morphim_cycle //. - have Axj: conj_aut [group of P] x \in A by exact: mem_morphim. + have Axj: conj_aut [group of P] x \in A by apply: mem_morphim. have [_ _ -> //] := Aut_narrow pP (mFT_odd _) nnP solA AutA oddA. by rewrite morph_p_elt // (mem_p_elt p'K). have{rPle2} dimP: logn p #|P| = 3. @@ -1279,9 +1282,9 @@ have{not_t2'H} [q1 t2Hq neq_q]: exists2 q1, q1 \in \tau2(H) & q1 = q -> uniq_q. by rewrite (partition_pi_mmax maxH) t2q1 !orbT. have [all_q | ] := altP (@allP _ (pred1 q) s); last first. by case/allPn=> q1; rewrite mem_s=> t2q1; move/eqnP=> ne_q1q; exists q1. - have s_q1: head q s \in s by case: (s) nts => // q1 s' _; exact: predU1l. + have s_q1: head q s \in s by case: (s) nts => // q1 s' _; apply: predU1l. exists (head q s) => [|def_q q1]; rewrite -mem_s //. - by apply/idP/idP; [exact: all_q | move/eqnP->; rewrite -def_q]. + by apply/idP/idP; [apply: all_q | move/eqnP->; rewrite -def_q]. have [A /= Eq2A Eq2A_H] := ex_tau2Elem hallD t2Hq; rewrite -/D in Eq2A. have [b'q qmaxA]: q1 \notin \beta(G) /\ A \in 'E*_q1(G). by have [-> ->] := tau2_not_beta maxH t2Hq. @@ -1305,9 +1308,9 @@ have sLq: q \in \sigma(L). by rewrite -pnatE // -pgroupE (pgroupS sKLs) ?pcore_pgroup. have sLq1: q1 \in \sigma(L). apply: contraLR sLq => s'Lq1; rewrite -negnK negbK /= -pnatE // -pgroupE. - have s'LA: \sigma(L)^'.-group A by exact: pi_pgroup qA _. + have s'LA: \sigma(L)^'.-group A by apply: pi_pgroup qA _. have [E hallE sAE] := Hall_superset (mmax_sol maxL) sAL s'LA. - have EqA_E: A \in 'E_q1^2(E) by exact/pnElemP. + have EqA_E: A \in 'E_q1^2(E) by apply/pnElemP. have [[sEL s'E _] t2Lq1] := (and3P hallE, sigma'2Elem_tau2 maxL hallE EqA_E). have [_ [sCAE _ _] _ _] := tau2_compl_context maxL hallE t2Lq1 EqA_E. by apply: pgroupS (subset_trans _ sCAE) s'E; rewrite centsC. @@ -1455,7 +1458,7 @@ have FmaxM: M \in 'M_'F; last split=> //. have nilMs: nilpotent M`_\sigma by rewrite notP1type_Msigma_nil ?FmaxM. have sMsF: M`_\sigma \subset 'F(M) by rewrite Fitting_max ?pcore_normal. have defR: R :=: 'O_r(U) := nilpotent_Hall_pcore (abelian_nil cUU) sylR_U. -have nRK: K \subset 'N(R) by rewrite defR (char_norm_trans (pcore_char r U)). +have nRK: K \subset 'N(R) by rewrite defR gFnorm_trans. have ntR: R :!=: 1. rewrite -rank_gt0 (rank_Sylow sylR_N) p_rank_gt0. by rewrite (partition_pi_mmax maxN) t2Nr !orbT. diff --git a/mathcomp/odd_order/BGsection16.v b/mathcomp/odd_order/BGsection16.v index a37edba..8aed3eb 100644 --- a/mathcomp/odd_order/BGsection16.v +++ b/mathcomp/odd_order/BGsection16.v @@ -1,10 +1,18 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div path fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div path fintype. +From mathcomp Require Import bigop finset prime fingroup morphism perm automorphism quotient. +From mathcomp Require Import action gproduct gfunctor pgroup cyclic center commutator. +From mathcomp Require Import gseries nilpotent sylow abelian maximal hall frobenius. +From mathcomp Require Import BGsection1 BGsection2 BGsection3 BGsection4 BGsection5. +From mathcomp Require Import BGsection6 BGsection7 BGsection9 BGsection10 BGsection12. +From mathcomp Require Import BGsection13 BGsection14 BGsection15. (******************************************************************************) @@ -175,7 +183,7 @@ Proof. by rewrite /FTtype; do 4!case: ifP => // _. Qed. Definition FTcore M := if 0 < FTtype M <= 2 then M`_\F else M^`(1). Fact FTcore_is_group M : group_set (FTcore M). -Proof. rewrite /FTcore; case: ifP => _; exact: groupP. Qed. +Proof. by rewrite [group_set _]fun_if !groupP if_same. Qed. Canonical Structure FTcore_group M := Group (FTcore_is_group M). Definition FTsupport1 M := (FTcore M)^#. @@ -617,7 +625,7 @@ split. have{sylS} sylS := subHall_Sylow hallU sk'p sylS. have [[sSM pS _] [/= s'p _]] := (and3P sylS, norP sk'p). rewrite (sigma'_nil_abelian maxM) ?(pi_pgroup pS) ?(pgroup_nil pS) //. - rewrite (rank_Sylow sylS) leqNgt (contra _ s'p) //; exact: alpha_sub_sigma. + by rewrite (rank_Sylow sylS) leqNgt (contra _ s'p) //; apply: alpha_sub_sigma. - have [_ _ _ cUA_UA _] := kappa_structure maxM complU. apply: abelianS cUA_UA; rewrite genS // -big_distrr ?setIS -?def_FTcore //=. by apply/bigcupsP=> x A1x; rewrite (bigcup_max x) // setDE setIAC subsetIr. @@ -770,7 +778,7 @@ split; last 1 first. rewrite -[Z](defCK _ K1ya) inE groupJ // cent1C -consttJ groupX ?cent1id //. by rewrite (contra (mem_p_elt su'K)) ?(contra (mem_p_elt suKs)) ?p_eltJ. rewrite (trivg_kappa_compl maxM complU) => notP1maxM. -have P2maxM: M \in 'M_'P2 by exact/setDP. +have P2maxM: M \in 'M_'P2 by apply/setDP. split; first by have [_ _ _ _ []] := Ptype_structure PmaxM hallK. apply: contraR notP1maxM; case/nonTI_Fitting_facts=> //. by case/setUP=> //; case/idPn; case/setDP: PmaxM. @@ -1070,7 +1078,7 @@ have [K1 | ntK] := eqsVneq K 1. by have [_] := kappa_compl_context maxM complU; rewrite defH K1 sdprodg1. exists U; split. have [_ _ _ cU1U1 exU0] := kappa_structure maxM complU. - split=> //; last by rewrite -/Ms -defH in exU0; exact: exU0. + split=> //; last by rewrite -/Ms -defH in exU0; apply: exU0. exists [group of <<\bigcup_(x in (M`_\sigma)^#) 'C_U[x]>>]. split=> //= [|x Hx]; last by rewrite sub_gen //= -/Ms -defH (bigcup_max x). rewrite -big_distrr /= /normal gen_subG subsetIl. @@ -1093,7 +1101,7 @@ have [K1 | ntK] := eqsVneq K 1. have [_ <- nHU tiHU] := sdprodP defM. by rewrite quotientMidl -(exponent_isog (quotient_isog _ _)). have sylP: p.-Sylow(H) P := nilpotent_pcore_Hall _ (Fcore_nil M). - have ntP: P != 1 by apply: contraNneq not_cPP => ->; exact: abelian1. + have ntP: P != 1 by apply: contraNneq not_cPP => ->; apply: abelian1. by exists p; rewrite // -p_rank_gt0 -(rank_Sylow sylP) rank_gt0. have PmaxM: M \in 'M_'P by rewrite inE -(trivg_kappa maxM hallK) ntK. have [Mstar _ [_ _ _ [cycW _ _ _ _]]] := Ptype_embedding PmaxM hallK. @@ -1131,7 +1139,7 @@ have [Ueq1 | ntU] := eqsVneq U 1; last first. apply/eqP; rewrite eqEsubset (Fcore_max hallH_M') ?Fcore_nil // andbT. rewrite (Fcore_max (subHall_Hall hallM' _ (Fcore_Hall _))) ?Fcore_nil //. by move=> p piM'Fp; apply: pnatPpi k'M' (piSg (Fcore_sub _) piM'Fp). - exact: char_normal_trans (Fcore_char _) nsM'M. + exact: gFnormal_trans nsM'M. exists U _ K _ defW; split=> //; split; first by rewrite defM'F. by exists U; split=> // x _; apply: subsetIl. have [_ _ _ _ /(_ ntU)] := kappa_structure maxM complU. @@ -1165,7 +1173,7 @@ have [_ _ _ _] := nonTI_Fitting_structure maxM notMy ntX. case=> [[] | [_]]; first by case/idPn; case/setDP: PmaxM. move: #|_| => p; set P := 'O_p(H); rewrite /= -/H => not_cPP cycHp'. have sylP: p.-Sylow(H) P := nilpotent_pcore_Hall _ (Fcore_nil M). -have ntP: P != 1 by apply: contraNneq not_cPP => ->; exact: abelian1. +have ntP: P != 1 by apply: contraNneq not_cPP => ->; apply: abelian1. have piHp: p \in \pi(H) by rewrite -p_rank_gt0 -(rank_Sylow sylP) rank_gt0. have defH: H = Ms by apply/eqP; rewrite defY1 Y1. rewrite -defMs -defH in defM; have [_ <- nHU tiHU] := sdprodP defM. @@ -1200,7 +1208,7 @@ split=> [H x a hallH nilH Hx|]. have [allFM | ] := boolP (('M : {set {group gT}}) \subset 'M_'F). by left=> M maxM; rewrite -FTtype_Fmax // (subsetP allFM). case/subsetPn=> S maxS notFmaxS; right. -have PmaxS: S \in 'M_'P by exact/setDP. +have PmaxS: S \in 'M_'P by apply/setDP. have [[U W1] /= complU] := kappa_witness maxS; have [_ hallW1 _] := complU. have ntW1: W1 :!=: 1 by rewrite (trivg_kappa maxS). have [[_ [_]]] := BGsummaryC maxS complU ntW1; set W2 := 'C_(_)(W1) => ntW2 _. @@ -1263,7 +1271,7 @@ without loss {defX} ->: X / X = 'A0(M). set D0 := finset _ => [[sD0A1 tameA0 signD0]] D. have sDD0: D \subset D0 by rewrite /D /D0 !setIdE setSI. split=> [|x Ax a Axa|x Dx]; first exact: subset_trans sDD0 sD0A1. - by apply: tameA0; exact: (subsetP sAA0). + by apply: tameA0; apply: (subsetP sAA0). have [/= -> -> [-> coA0L -> -> frobL]] := signD0 x (subsetP sDD0 x Dx). by do 2![split=> //] => y Ay; rewrite coA0L // (subsetP sAA0). move=> {X} D; pose Ms := M`_\sigma. @@ -1291,7 +1299,7 @@ have tiA0 x a: x \in 'A0(M) :\: 'A1(M) -> x ^ a \in 'A0(M) -> a \in M. have sDA1: D \subset 'A1(M). apply/subsetPn=> [[x /setIdP[A0x not_sCxM] notA1x]]. case/subsetP: not_sCxM => a cxa. - by apply: (tiA0 x); [exact/setDP | rewrite /conjg -(cent1P cxa) mulKg]. + by apply: (tiA0 x); [apply/setDP | rewrite /conjg -(cent1P cxa) mulKg]. have sDMs1: D \subset Ms^# by rewrite /Ms -def_FTcore. have [tameMs _ signM PsignM] := BGsummaryD maxM. split=> // [x A0x a A0xa|x Dx]. @@ -1318,7 +1326,7 @@ have MSx_gt1: #|'M_\sigma[x]| > 1. by apply: contra not_sCxM; move/cent1_sub_uniq_sigma_mmax->. have [FmaxM t2'M]: M \in 'M_'F /\ \tau2(M)^'.-group M. apply: (non_disjoint_signalizer_Frobenius ell1x MSx_gt1 SMxM). - by apply: contra not_sNx'CMy; exact: pgroupS (subsetIl _ _). + by apply: contra not_sNx'CMy; apply: pgroupS (subsetIl _ _). have defA0: 'A0(M) = Ms^#. rewrite FTsupp0_type1; last by rewrite -FTtype_Fmax. rewrite /'A(M) /'A1(M) -FTtype_Fmax // FmaxM def_FTcore //= -/Ms. diff --git a/mathcomp/odd_order/BGsection2.v b/mathcomp/odd_order/BGsection2.v index fc5f489..5faf8ca 100644 --- a/mathcomp/odd_order/BGsection2.v +++ b/mathcomp/odd_order/BGsection2.v @@ -1,11 +1,19 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div fintype. +From mathcomp Require Import bigop prime binomial finset fingroup morphism perm automorphism. +From mathcomp Require Import quotient action gfunctor commutator gproduct. +From mathcomp Require Import ssralg finalg zmodp cyclic center pgroup gseries nilpotent. +From mathcomp Require Import sylow abelian maximal hall. Require poly ssrint. +From mathcomp Require Import matrix mxalgebra mxrepresentation mxabelem. +From mathcomp Require Import BGsection1. (******************************************************************************) @@ -62,8 +70,7 @@ have absM f: (M *m f <= M)%MS -> {a | (a \in E_H)%MS & M *m a = M *m f}. rewrite /gring_mx /= !mulmx_sum_row !linear_sum; apply: eq_bigr => i /= _. by rewrite !linearZ /= !rowK !mxvecK -in_submodJ. rewrite /gring_mx /= mulmxKpV ?submx_full ?mxvecK //; last first. - suffices: mx_absolutely_irreducible rM by case/andP. - by apply: closedF; exact/submod_mx_irr. + by have/andP[]: mx_absolutely_irreducible rM by apply/closedF/submod_mx_irr. rewrite {1}[in_submod]lock in_submodE -mulmxA mulmxA -val_submodE -lock. by rewrite mulmxA -in_submodE in_submodK. have /morphimP[x nHx Gx defHx]: Hx \in (G / H)%g by rewrite defGH cycle_id. @@ -108,7 +115,7 @@ have{cHtau_x} cGtau_x: centgmx rG (tau *m rG x). rewrite !(mulmx_suml, mulmx_sumr); apply: eq_bigr => y Hy. by rewrite -!(scalemxAl, scalemxAr) (centgmxP cHtau_x) ?mulmxA. have{cGtau_x} [a def_tau_x]: exists a, tau *m rG x = a%:M. - by apply/is_scalar_mxP; apply: mx_abs_irr_cent_scalar cGtau_x; exact: closedF. + by apply/is_scalar_mxP; apply: mx_abs_irr_cent_scalar cGtau_x; apply: closedF. apply: mx_iso_simple (eqmx_iso _ _) simM; apply/eqmxP; rewrite submx1 sub1mx. case/mx_irrP: (irrG) => _ -> //; rewrite /mxmodule {1}defG join_subG /=. rewrite cycle_subG inE Gx andbC (subset_trans modM) ?rstabs_subg ?subsetIr //=. @@ -137,41 +144,40 @@ have ltHG: H \proper G. by rewrite properEcard sHG -(Lagrange sHG) ltn_Pmulr // prime_gt1. have dvLH: \rank L %| #|H|. have absL: mx_absolutely_irreducible (submod_repr (mxsimple_module simL)). - by apply: closF; exact/submod_mx_irr. + exact/closF/submod_mx_irr. apply: IHm absL (solvableS (normal_sub nsHG) solG). by rewrite (leq_trans (proper_card ltHG)). have [_ [x Gx H'x]] := properP ltHG. have prGH: prime #|G / H|%g by rewrite card_quotient ?normal_norm. -wlog sH: / socleType rH by exact: socle_exists. +wlog sH: / socleType rH by apply: socle_exists. pose W := PackSocle (component_socle sH simL). have card_sH: #|sH| = #|G : 'C_G[W | 'Cl]|. rewrite -cardsT; have ->: setT = orbit 'Cl G W. apply/eqP; rewrite eqEsubset subsetT. have /imsetP[W' _ defW'] := Clifford_atrans irrG sH. have WW': W' \in orbit 'Cl G W by rewrite orbit_in_sym // -defW' inE. - by rewrite defW' andbT; apply/subsetP=> W''; exact: orbit_in_trans. + by rewrite defW' andbT; apply/subsetP=> W'' /orbit_in_trans->. rewrite orbit_stabilizer // card_in_imset //. exact: can_in_inj (act_reprK _). have sHcW: H \subset 'C_G[W | 'Cl]. apply: subset_trans (subset_trans (joing_subl _ _) (Clifford_astab sH)) _. - apply/subsetP=> z; rewrite !inE => /andP[->]; apply: subset_trans. - exact: subsetT. + by rewrite subsetI subsetIl astabS ?subsetT. have [|] := prime_subgroupVti ('C_G[W | 'Cl] / H)%G prGH. rewrite quotientSGK ?normal_norm // => cClG. have def_sH: setT = [set W]. apply/eqP; rewrite eq_sym eqEcard subsetT cards1 cardsT card_sH. by rewrite -indexgI (setIidPl cClG) indexgg. suffices L1: (L :=: 1%:M)%MS. - by rewrite L1 mxrank1 in dvLH; exact: dvdn_trans (cardSg sHG). + by rewrite L1 mxrank1 in dvLH; apply: dvdn_trans (cardSg sHG). apply/eqmxP; rewrite submx1. have cycH: cyclic (G / H)%g by rewrite prime_cyclic. have [y Gy|_ _] := mx_irr_prime_index closF irrG cycH simL; last first. by apply; rewrite ?submx1 //; case simL. - have simLy: mxsimple rH (L *m rG y) by exact: Clifford_simple. + have simLy: mxsimple rH (L *m rG y) by apply: Clifford_simple. pose Wy := PackSocle (component_socle sH simLy). have: (L *m rG y <= Wy)%MS by rewrite PackSocleK component_mx_id. have ->: Wy = W by apply/set1P; rewrite -def_sH inE. - by rewrite PackSocleK; exact: component_mx_iso. + by rewrite PackSocleK; apply: component_mx_iso. rewrite (setIidPl _) ?quotientS ?subsetIl // => /trivgP. rewrite quotient_sub1 //; last by rewrite subIset // normal_norm. move/setIidPl; rewrite (setIidPr sHcW) /= => defH. @@ -238,7 +244,7 @@ Let g_mod i := expr_mod i gh1. Let EiP i e : reflect (e^g = eps ^+ i *: e) (e \in 'E_i)%MS. Proof. rewrite (sameP eigenspaceP eqP) mul_vec_lin -linearZ /=. -by rewrite (can_eq mxvecK); exact: eqP. +by rewrite (can_eq mxvecK); apply: eqP. Qed. Let E2iP i t e : @@ -260,7 +266,7 @@ Proposition mxdirect_sum_eigenspace_cycle : (sumV :=: 1%:M)%MS /\ mxdirect sumV. Proof. have splitF: group_splitting_field F (Zp_group h). move: prim_eps (abelianS (subsetT (Zp h)) (Zp_abelian _)). - by rewrite -{1}(card_Zp h_gt0); exact: primitive_root_splitting_abelian. + by rewrite -{1}(card_Zp h_gt0); apply: primitive_root_splitting_abelian. have F'Zh: [char F]^'.-group (Zp h). apply/pgroupP=> p p_pr; rewrite card_Zp // => /dvdnP[d def_h]. apply/negP=> /= charFp. @@ -290,7 +296,7 @@ split=> //; apply/eqmxP; rewrite submx1. wlog [I M /= simM <- _]: / mxsemisimple rZ 1. exact: mx_reducible_semisimple (mxmodule1 _) (mx_Maschke rZ F'Zh) _. apply/sumsmx_subP=> i _; have simMi := simM i; have [modMi _ _] := simMi. -set v := nz_row (M i); have nz_v: v != 0 by exact: nz_row_mxsimple simMi. +set v := nz_row (M i); have nz_v: v != 0 by apply: nz_row_mxsimple simMi. have rankMi: \rank (M i) = 1%N. by rewrite (mxsimple_abelian_linear splitF _ simMi) //= ZhT Zp_abelian. have defMi: (M i :=: v)%MS. @@ -302,7 +308,7 @@ have: a ^+ h - 1 == 0. apply: contraR nz_v => nz_pZa; rewrite -(eqmx_eq0 (eqmx_scale _ nz_pZa)). by rewrite scalerBl scale1r -lin_rZ // subr_eq0 char_Zp ?mulmx1. rewrite subr_eq0; move/eqP; case/(prim_rootP prim_eps) => k def_a. -by rewrite defMi (sumsmx_sup k) // /V_ -def_a; exact/eigenspaceP. +by rewrite defMi (sumsmx_sup k) // /V_ -def_a; apply/eigenspaceP. Qed. (* This is B & G, Proposition 2.4(b). *) @@ -431,7 +437,7 @@ have /mxdirect_sumsE[/= dx_diag rank_diag]: mxdirect sum_diagE. symmetry; rewrite /p -val_eqE /= -(eqn_modDl (h - i)). by rewrite addnA subnK 1?ltnW // modnDl modn_small. have dx_sumE1: mxdirect (\sum_(i < h) 'E_i). - by apply: mxdirect_sum_eigenspace => i j _ _; exact: inj_eps. + by apply: mxdirect_sum_eigenspace => i j _ _; apply: inj_eps. have diag_mod n: diagE (n %% h) = diagE n. by apply: eq_bigl=> it; rewrite modnDmr. split; last first. @@ -457,7 +463,7 @@ rewrite (mxdirectP dx_diag) /= (reindex (fun i : 'I_h => (i, inh (i + m)))) /=. apply: eq_big => [i | i _]; first by rewrite modn_mod eqxx. by rewrite rank_proj_eigenspace_cycle /n_ Vi_mod. exists (@fst _ _) => // [] [i t] /=. -by rewrite !inE /= (modn_small (valP t)) => def_t; exact/eqP/andP. +by rewrite !inE /= (modn_small (valP t)) => def_t; apply/eqP/andP. Qed. (* This is B & G, Proposition 2.4(h). *) @@ -610,7 +616,7 @@ have oZp := card_center_extraspecial pP esP; have[_ prZ] := esP. have{sdPH_G} [nsPG sHG defG nPH tiPH] := sdprod_context sdPH_G. have sPG := normal_sub nsPG. have coPH: coprime #|P| #|H| by rewrite oPpn coprime_pexpl. -have nsZG: 'Z(P) <| G := char_normal_trans (center_char _) nsPG. +have nsZG: 'Z(P) <| G := gFnormal_trans _ nsPG. have defCP: 'C_G(P) = 'Z(P). apply/eqP; rewrite eqEsubset andbC setSI //=. rewrite -(coprime_mulG_setI_norm defG) ?norms_cent ?normal_norm //=. @@ -618,8 +624,8 @@ have defCP: 'C_G(P) = 'Z(P). apply/subsetP=> x; case/setIP; case/setU1P=> [-> // | H'x]. rewrite -sub_cent1; move/setIidPl; rewrite primeHP // => defP. by have:= min_card_extraspecial pP esP; rewrite -defP oZp (leq_exp2l 3 1). -have F'P: [char F]^'.-group P by exact: pgroupS sPG F'G. -have F'H: [char F]^'.-group H by exact: pgroupS sHG F'G. +have F'P: [char F]^'.-group P by apply: pgroupS sPG F'G. +have F'H: [char F]^'.-group H by apply: pgroupS sHG F'G. wlog{ffulG F'G} [irrG regZ]: q rG / mx_irreducible rG /\ rfix_mx rG 'Z(P) = 0. move=> IH; wlog [I W /= simW defV _]: / mxsemisimple rG 1%:M. exact: (mx_reducible_semisimple (mxmodule1 _) (mx_Maschke rG F'G)). @@ -633,7 +639,7 @@ wlog{ffulG F'G} [irrG regZ]: q rG / mx_irreducible rG /\ rfix_mx rG 'Z(P) = 0. by move/negbFE: (z1 i) => /rstab_act-> //; rewrite submxMl. have [modW _ _] := simW i; pose rW := submod_repr modW. rewrite -(eqmx_rstab _ (val_submod1 (W i))) -(rstab_submod modW) in ffZ. - have irrW: mx_irreducible rW by exact/submod_mx_irr. + have irrW: mx_irreducible rW by apply/submod_mx_irr. have regZ: rfix_mx rW 'Z(P)%g = 0. apply/eqP; apply: contraR ffZ; case/mx_irrP: irrW => _ minW /minW. by rewrite normal_rfix_mx_module // -sub1mx inE Gz /= => /implyP/rfix_mxP->. @@ -662,7 +668,7 @@ have{M simM irrG regZ F'P} [irrP def_q]: mx_irreducible rP /\ q = (p ^ n)%N. rewrite -submx0 -regZ; apply/rfix_mxP=> z; move/(subsetP cMZ)=> cMz. by rewrite (rstab_act cMz). suffices irrP: mx_irreducible rP. - by split=> //; apply/eqP; rewrite eq_sym; case/mx_irrP: irrP => _; exact. + by split=> //; apply/eqP; rewrite eq_sym; case/mx_irrP: irrP => _; apply. apply: (@mx_irr_prime_index F _ G P _ M nsPG) => // [|x Gx]. by rewrite -defG quotientMidl quotient_cyclic. rewrite (bool_irrelevance (normal_sub nsPG) sPG). @@ -742,11 +748,11 @@ have mulBg x: x \in P -> B x *m gE = yr *m B x. apply/row_matrixP=> i; have Hi := enum_valP i; have Gi := subsetP sHG _ Hi. rewrite 2!row_mul !rowK mul_vec_lin /= -rowE rowK gring_indexK ?groupM //. by rewrite conjgM -repr_mxV // -!repr_mxM // ?(groupJ, groupM, groupV). -wlog sH: / irrType F H by exact: socle_exists. +wlog sH: / irrType F H by apply: socle_exists. have{cycH} linH: irr_degree (_ : sH) = 1%N. exact: irr_degree_abelian (cyclic_abelian cycH). have baseH := linear_irr_comp F'H (closF H) (linH _). -have{linH} linH (W : sH): \rank W = 1%N by rewrite baseH; exact: linH. +have{linH} linH (W : sH): \rank W = 1%N by rewrite baseH; apply: linH. have [w] := cycle_repr_structure sH defH F'H (closF H). rewrite -/h => prim_w [Wi [bijWi _ _ Wi_yg]]. have{Wi_yg baseH} Wi_yr i: Wi i *m yr = w ^+ i *: (Wi i : 'M_h). @@ -767,7 +773,7 @@ have{yr Wi_yr Pb mulBg} sB1E i: (B1 i <= E_ i)%MS. have{bijWi sumB cl1 F'H} defSB: (SB :=: 1%:M)%MS. apply/eqmxP; rewrite submx1 -sumB (big_setD1 _ cl1) addsmxS //=. rewrite exchange_big sumsmxS // => ZxH _; rewrite genmxE /= -sumsmxMr_gen. - rewrite -((reindex Wi) xpredT val) /=; last by exact: onW_bij. + rewrite -((reindex Wi) xpredT val) /=; last by apply: onW_bij. by rewrite -/(Socle _) (reducible_Socle1 sH (mx_Maschke _ F'H)) mul1mx. rewrite mxdirect_addsE /= in dxB; case/and3P: dxB => _ dxB dxB1. have{linH Bfree dxB} rankB1 i: \rank (B1 i) = #|clPqH^#|. @@ -833,7 +839,7 @@ move=> oddG ffulG. without loss closF: F rG ffulG / group_closure_field F gT. move=> IH; apply: (@group_closure_field_exists gT F) => [[Fc f closFc]]. rewrite -(eq_pgroup _ (fmorph_char f)). - by rewrite -(map_mx_faithful f) in ffulG; exact: IH ffulG closFc. + by rewrite -(map_mx_faithful f) in ffulG; apply: IH ffulG closFc. elim: {G}_.+1 {-2}G (ltnSn #|G|) => // m IHm G le_g_m in rG oddG ffulG *. apply/pgroupP=> p p_pr pG'; rewrite !inE p_pr /=; apply: wlog_neg => p_nz. have [P sylP] := Sylow_exists p G. @@ -852,7 +858,7 @@ have nPG: G \subset 'N(P). rewrite (sameP commG1P trivgP) -tiPN' subsetI commgS //. by rewrite commg_subr subsetIr. have /sdprodP[_ /= defG nKP _] := Burnside_normal_complement sylP cPN. - set K := 'O_p^'(G) in defG nKP; have nKG: G \subset 'N(K) by exact: gFnorm. + set K := 'O_p^'(G) in defG nKP; have nKG: G \subset 'N(K) by apply: gFnorm. suffices p'G': p^'.-group G^`(1)%g by case/eqnP: (pgroupP p'G' p p_pr pG'). apply: pgroupS (pcore_pgroup p^' G); rewrite -quotient_cents2 //= -/K. by rewrite -defG quotientMidl /= -/K quotient_cents ?(subset_trans sPN). @@ -881,7 +887,7 @@ have{IHm} abelQ: abelian Q. pose rQ := subg_repr rG sQG. wlog [U simU sU1]: / exists2 U, mxsimple rQ U & (U <= 1%:M)%MS. by apply: mxsimple_exists; rewrite ?mxmodule1 ?oner_eq0. -have Uscal: \rank U = 1%N by exact: (mxsimple_abelian_linear (closF _)) simU. +have Uscal: \rank U = 1%N by apply: (mxsimple_abelian_linear (closF _)) simU. have{simU} [Umod _ _] := simU. have{sU1} [|V Vmod sumUV dxUV] := mx_Maschke _ _ Umod sU1. have: p.-group Q by apply: pgroupS (pHall_pgroup sylP); rewrite subsetIr. @@ -1024,7 +1030,7 @@ have [v defUc]: exists u : 'rV_2, (u :=: U^C)%MS. pose B := col_mx u v; have uB: B \in unitmx. rewrite -row_full_unit -sub1mx -(eqmxMfull _ (addsmx_compl_full U)). by rewrite mulmx1 -addsmxE addsmxS ?defU ?defUc. -have Umod: mxmodule rP U by exact: rfix_mx_module. +have Umod: mxmodule rP U by apply: rfix_mx_module. pose W := rfix_mx (factmod_repr Umod) P. have ntW: W != 0. apply: (rfix_pgroup_char charFp) => //. @@ -1064,7 +1070,7 @@ pose rQ := abelem_repr abelP ntP nPQ. have [|P1 simP1 _] := dec_mxsimple_exists (mxmodule1 rQ). by rewrite oner_eq0. have [modP1 nzP1 _] := simP1. -have ffulQ: mx_faithful rQ by exact: abelem_mx_faithful. +have ffulQ: mx_faithful rQ by apply: abelem_mx_faithful. have linP1: \rank P1 = 1%N. apply/eqP; have:= abelem_cyclic abelQ; rewrite logQ; apply: contraFT. rewrite neq_ltn ltnNge lt0n mxrank_eq0 nzP1 => P1full. diff --git a/mathcomp/odd_order/BGsection3.v b/mathcomp/odd_order/BGsection3.v index 25879a6..aa4db05 100644 --- a/mathcomp/odd_order/BGsection3.v +++ b/mathcomp/odd_order/BGsection3.v @@ -1,10 +1,18 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div. +From mathcomp Require Import fintype tuple bigop prime binomial finset ssralg fingroup finalg. +From mathcomp Require Import morphism perm automorphism quotient action commutator gproduct. +From mathcomp Require Import zmodp cyclic gfunctor center pgroup gseries nilpotent sylow. +From mathcomp Require Import finmodule abelian frobenius maximal extremal hall. +From mathcomp Require Import matrix mxalgebra mxrepresentation mxabelem wielandt_fixpoint. +From mathcomp Require Import BGsection1 BGsection2. (******************************************************************************) @@ -241,8 +249,8 @@ without loss{m IHm leGm} [ffulG cycZ]: / rker rG = 1 /\ cyclic 'Z(G). have [N1 _ | ntN] := eqVneq N 1. apply: IH; split. by apply/trivgP; rewrite -N1 /N rker_submod rstabS ?submx1. - have: mx_irreducible (submod_repr modM) by exact/submod_mx_irr. - by apply: mx_faithful_irr_center_cyclic; exact/trivgP. + have: mx_irreducible (submod_repr modM) by apply/submod_mx_irr. + by apply: mx_faithful_irr_center_cyclic; apply/trivgP. have tiRN: R :&: N = 1. by apply: prime_TIg; rewrite //= rker_submod rfix_mx_rstabC // regR submx0. have nsNG: N <| G := rker_normal _; have [sNG nNG] := andP nsNG. @@ -261,7 +269,7 @@ without loss{m IHm leGm} [ffulG cycZ]: / rker rG = 1 /\ cyclic 'Z(G). rewrite (sameP commG1P trivgP). apply: subset_trans (kquo_mx_faithful (submod_repr modM)). rewrite IHm ?quotient_sol ?coprime_morph ?morphim_odd ?quotient_pgroup //. - - apply: leq_trans leGm; exact: ltn_quotient. + - by apply: leq_trans leGm; apply: ltn_quotient. - by rewrite card_quotient // -indexgI tiRN indexg1. apply/eqP; rewrite -submx0 rfix_quo // rfix_submod //. by rewrite regR capmx0 linear0 sub0mx. @@ -280,7 +288,7 @@ case cKK: (abelian K). by apply/Frobenius_semiregularP; rewrite // -cardG_gt1 prime_gt1. have [spK defZK]: special K /\ 'C_K(R) = 'Z(K). apply: (abelian_charsimple_special qK) => //. - apply/bigcupsP=> H; case/andP=> chHK cHH. + apply/bigcupsP=> H /andP[chHK cHH]. have:= char_sub chHK; rewrite subEproper. case/predU1P=> [eqHK | ltHK]; first by rewrite eqHK cKK in cHH. have nHR: R \subset 'N(H) := char_norm_trans chHK nKR. @@ -352,7 +360,7 @@ case: (eqVneq (rker rG) 1) => [ffulG | ntC]; last first. set C := rker rG in ntC *; have nsCG: C <| G := rker_normal rG. have [sCG nCG] := andP nsCG. have nCK := subset_trans sKG nCG; have nCR := subset_trans sRG nCG. - case sKC: (K \subset C); first exact: subset_trans (der_sub _ _) _. + case sKC: (K \subset C); first exact: gFsub_trans. have sCK: C \subset K. by rewrite proper_sub // (Frobenius_normal_proper_ker frobG) ?sKC. have frobGq: [Frobenius G / C = (K / C) ><| (R / C)]. @@ -368,7 +376,7 @@ case: (eqVneq (rker rG) 1) => [ffulG | ntC]; last first. by rewrite cycle_id -cycle_eq1 -defRq ntRq. - move=> Hq; rewrite -(group_inj (cosetpreK Hq)). by apply: quotient_splitting_field; rewrite ?subsetIl. - by apply: leq_trans leKm; exact: ltn_quotient. + by apply: leq_trans leKm; apply: ltn_quotient. have ltK_abelian (N : {group gT}): R \subset 'N(N) -> N \proper K -> abelian N. move=> nNR ltNK; have [sNK _] := andP ltNK; apply/commG1P/trivgP. rewrite -(setIidPr (sub1G (N <*> R))) /= -ffulG; set G1 := N <*> R. @@ -378,10 +386,9 @@ have ltK_abelian (N : {group gT}): R \subset 'N(N) -> N \proper K -> abelian N. rewrite -(rker_subg _ sG1). apply: IHm defG1 _; rewrite ?(solvableS sG1) ?(pgroupS sG1) //. by apply/trivgP; rewrite -regR setSI. - by apply: leq_trans leKm; exact: proper_card. + by apply: leq_trans leKm; apply: proper_card. have cK'K': abelian K^`(1). - apply: ltK_abelian; first exact: char_norm_trans (der_char _ _) nKR. - exact: (sol_der1_proper solK). + exact: ltK_abelian (gFnorm_trans _ nKR) (sol_der1_proper solK _ ntK). pose fixG := rfix_mx rG; pose NRmod N (U : 'M_n) := N <*> R \subset rstabs rG U. have dx_modK_rfix (N : {group gT}) U V: N \subset K -> R \subset 'N(N) -> NRmod N U -> NRmod N V -> @@ -408,17 +415,17 @@ have redG := mx_Maschke rG F'G. wlog [U simU nfixU]: / exists2 U, mxsimple rG U & ~~ (U <= fixG K)%MS. move=> IH; wlog [I U /= simU sumU _]: / mxsemisimple rG 1%:M. exact: (mx_reducible_semisimple (mxmodule1 _) redG). - case: (pickP (fun i => ~~ (U i <= fixG K))%MS) => [i nfixU | fixK]. + have [i nfixU | fixK] := pickP (fun i => ~~ (U i <= fixG K)%MS). by apply: IH; exists (U i). - apply: (subset_trans (der_sub _ _)); rewrite rfix_mx_rstabC // -sumU. + rewrite gFsub_trans // rfix_mx_rstabC // -sumU. by apply/sumsmx_subP=> i _; apply/idPn; rewrite fixK. have [modU ntU minU] := simU; pose rU := submod_repr modU. -have irrU: mx_irreducible rU by exact/submod_mx_irr. +have irrU: mx_irreducible rU by apply/submod_mx_irr. have [W modW sumUW dxUW] := redG U modU (submx1 U). have cWK: (W <= fixG K)%MS. have:= dx_modK_rfix _ _ _ (subxx _) nKR _ _ dxUW. - by rewrite /NRmod /= norm_joinEr // defKR (negPf nfixU); exact. -have nsK'G: K^`(1) <| G by exact: char_normal_trans (der_char _ _) nsKG. + by rewrite /NRmod /= norm_joinEr // defKR (negPf nfixU); apply. +have nsK'G: K^`(1) <| G by rewrite gFnormal_trans. have [sK'G nK'G] := andP nsK'G. suffices nregK'U: (rfix_mx rU K^`(1))%MS != 0. rewrite rfix_mx_rstabC ?normal_sub // -sumUW addsmx_sub andbC. @@ -427,27 +434,26 @@ suffices nregK'U: (rfix_mx rU K^`(1))%MS != 0. apply: contra nregK'U => cUK'; rewrite (eqmx_eq0 (rfix_submod _ _)) //. by rewrite (eqP cUK') linear0. pose rK := subg_repr rU (normal_sub nsKG); set p := #|R| in p_pr. -wlog sK: / socleType rK by exact: socle_exists. +wlog sK: / socleType rK by apply: socle_exists. have [i _ def_sK]: exists2 i, i \in setT & [set: sK] = orbit 'Cl G i. - by apply/imsetP; exact: Clifford_atrans. + exact/imsetP/Clifford_atrans. have card_sK: #|[set: sK]| = #|G : 'C[i | 'Cl]|. by rewrite def_sK card_orbit_in ?indexgI. have ciK: K \subset 'C[i | 'Cl]. apply: subset_trans (astabS _ (subsetT _)). - by apply: subset_trans (Clifford_astab _); exact: joing_subl. + by apply: subset_trans (Clifford_astab _); apply: joing_subl. pose M := socle_base i; have simM: mxsimple rK M := socle_simple i. have [sKp | sK1 {ciK card_sK}]: #|[set: sK]| = p \/ #|[set: sK]| = 1%N. - apply/pred2P; rewrite orbC card_sK; case/primeP: p_pr => _; apply. by rewrite (_ : p = #|G : K|) ?indexgS // -divgS // -(sdprod_card defG) mulKn. - have{def_sK} def_sK: [set: sK] = orbit 'Cl R i. - apply/eqP; rewrite eq_sym -subTset def_sK. - apply/subsetP=> i_yz; case/imsetP=> yz; rewrite -{1}defKR. - case/imset2P=> y z; move/(subsetP ciK); rewrite !inE sub1set inE. - case/andP=> Gy; move/eqP=> ciy Rz -> ->{yz i_yz}. + apply/eqP; rewrite eq_sym -subTset def_sK -[G in orbit _ G i]defKR. + apply/subsetP=> _ /imsetP[_ /imset2P[y z /(subsetP ciK)ciy Rz ->] ->]. + rewrite !(inE, sub1set) in ciy; have{ciy}[Gy /eqP-ciy]:= andP ciy. by rewrite actMin ?(subsetP sRG z Rz) // ciy mem_orbit. have inj_i: {in R &, injective ('Cl%act i)}. - apply/dinjectiveP; apply/card_uniqP; rewrite size_map -cardE -/p. - by rewrite -sKp def_sK /orbit Imset.imsetE cardsE. + apply/dinjectiveP/card_uniqP; rewrite size_map -cardE -/p. + by rewrite -sKp def_sK /orbit [in _ @: _]unlock cardsE. pose sM := (\sum_(y in R) M *m rU y)%MS. have dxM: mxdirect sM. apply/mxdirect_sumsP=> y Ry; have Gy := subsetP sRG y Ry. @@ -482,9 +488,9 @@ have{sK i M simM sK1 def_sK} irrK: mx_irreducible rK. apply: (mx_irr_prime_index closF irrU cycGq simM) => x Gx /=. apply: (component_mx_iso simM); first exact: Clifford_simple. have jP: component_mx rK (M *m rU x) \in socle_enum sK. - by apply: component_socle; exact: Clifford_simple. + exact/component_socle/Clifford_simple. pose j := PackSocle jP; apply: submx_trans (_ : j <= _)%MS. - by rewrite PackSocleK component_mx_id //; exact: Clifford_simple. + by rewrite PackSocleK component_mx_id //; apply: Clifford_simple. have def_i: [set i] == [set: sK] by rewrite eqEcard subsetT cards1 sK1. by rewrite ((j =P i) _) // -in_set1 (eqP def_i) inE. pose G' := K^`(1) <*> R. @@ -518,7 +524,7 @@ have irrK'K: mx_absolutely_irreducible rK'K. exact: atrans_dvd_in (Clifford_atrans _ _). have nsK'G': K^`(1) <| G' := normalS (joing_subl _ _) sG'G nsK'G. pose rK'G' := subg_repr rG' (normal_sub nsK'G'). - wlog sK'G': / socleType rK'G' by exact: socle_exists. + wlog sK'G': / socleType rK'G' by apply: socle_exists. have coKp: coprime #|K| p := Frobenius_coprime frobG. have nK'R := subset_trans sRG nK'G. have sK'_dv_p: #|[set: sK'G']| %| p. @@ -549,9 +555,9 @@ have irrK'K: mx_absolutely_irreducible rK'K. apply closF; apply: (mx_irr_prime_index closF irrG' cycGq simM) => x K'x /=. apply: (component_mx_iso simM); first exact: Clifford_simple. have jP: component_mx rK'G' (M *m rG' x) \in socle_enum sK'G'. - by apply: component_socle; exact: Clifford_simple. + exact/component_socle/Clifford_simple. pose j := PackSocle jP; apply: submx_trans (_ : j <= _)%MS. - by rewrite PackSocleK component_mx_id //; exact: Clifford_simple. + by rewrite PackSocleK component_mx_id //; apply: Clifford_simple. by rewrite ((j =P i) _) // -in_set1 -def_i inE. have linU: \rank U = 1%N by apply/eqP; rewrite abelian_abs_irr in irrK'K. case: irrU => _ nz1 _; apply: contra nz1; move/eqP=> fix0. @@ -620,12 +626,12 @@ have{n leGn IHn tiHR} IHquo (X : {group gT}): rewrite -defHR quotientE morphimR // -!quotientE. have ltGbn: #|G / X| < n. exact: leq_trans (ltn_quotient ntX (subset_trans sXH sHG)) _. - have defGb: (H / X) ><| (R / X) = G / X by exact: quotient_coprime_sdprod. + have defGb: (H / X) ><| (R / X) = G / X by apply: quotient_coprime_sdprod. have pr_R0b: prime #|R0 / X|. have tiXR0: X :&: R0 = 1 by apply/trivgP; rewrite -tiHR setISS. by rewrite card_quotient // -indexgI setIC tiXR0 indexg1 oR0. - have solGb: solvable (G / X) by exact: quotient_sol. - have coHRb: coprime #|H / X| #|R / X| by exact: coprime_morph. + have solGb: solvable (G / X) by apply: quotient_sol. + have coHRb: coprime #|H / X| #|R / X| by apply: coprime_morph. apply: IHn defGb coHRb _ pr_R0b _; rewrite ?quotientS ?quotient_odd //. by rewrite -coprime_quotient_cent ?(coprimegS sR0R) // morphim_Zgroup. without loss Op'H: / 'O_p^'(H) = 1. @@ -633,10 +639,10 @@ without loss Op'H: / 'O_p^'(H) = 1. suffices: p.-length_1 (H / 'O_p^'(H)). by rewrite p'quo_plength1 ?pcore_normal ?pcore_pgroup. apply: IHquo => //; first by rewrite normal_sub ?pcore_normal. - by rewrite normal_norm // (char_normal_trans (pcore_char _ _)). + by rewrite normal_norm // gFnormal_trans. move defV: 'F(H)%G => V. have charV: V \char H by rewrite -defV Fitting_char. -have [sVH nVH]: V \subset H /\ H \subset 'N(V) := andP (char_normal charV). +have /andP[sVH nVH]: V <| H := char_normal charV. have nsVG: V <| G := char_normal_trans charV nsHG. have [_ nVG] := andP nsVG; have nVR: R \subset 'N(V) := subset_trans sRG nVG. without loss ntV: / V :!=: 1. @@ -650,7 +656,7 @@ have p'r: r != p. by rewrite -oR0 (coprimegS sR0R) // (coprimeSg sVH). without loss{charV} abelV: / p.-abelem V; last have [_ cVV eV] := and3P abelV. move/implyP; rewrite implybE -trivg_Phi //; case/orP=> // ntPhi. - have charPhi: 'Phi(V) \char H := char_trans (Phi_char _) charV. + have charPhi: 'Phi(V) \char H := gFchar_trans _ charV. have nsPhiH := char_normal charPhi; have [sPhiH nPhiH] := andP nsPhiH. have{charPhi} nPhiG: G \subset 'N('Phi(V)):= char_norm_trans charPhi nHG. rewrite -(pquo_plength1 nsPhiH) 1?IHquo ?(pgroupS (Phi_sub _)) //. @@ -683,8 +689,7 @@ without loss{IHquo} indecomposableV: / forall U W, have [nsUH nsWH]: U <| H /\ W <| H. by rewrite /normal !(subset_trans sHG) ?andbT. by rewrite -(quo2_plength1 _ nsUH nsWH) ?tiUW ?IHquo. -have nsFb: 'F(H / V) <| G / V. - exact: char_normal_trans (Fitting_char _) (morphim_normal _ _). +have nsFb: 'F(H / V) <| G / V by rewrite gFnormal_trans ?quotient_normal. have{nsVG nsFb} [/= U defU sVU nsUG] := inv_quotientN nsVG nsFb. have{nsUG} [sUG nUG] := andP nsUG. have [solU nVU] := (solvableS sUG solG, subset_trans sUG nVG). @@ -717,23 +722,23 @@ have coKP: coprime #|K| #|P| by rewrite coprime_sym (pnat_coprime pP). have{sylP} sylVP: p.-Sylow(H) (V <*> P). rewrite pHallE /= norm_joinEr ?mul_subG //= -defH -!LagrangeMl. rewrite partnM // part_pnat_id // -!card_quotient //. - by apply/eqP; congr (_ * _)%N; apply: card_Hall; exact: quotient_pHall. + by apply/eqP; congr (_ * _)%N; apply: card_Hall; apply: quotient_pHall. have [trKP | {sylV sVU nVU}ntKP] := eqVneq [~: K, P] 1. suffices sylVH: p.-Sylow(H) V. rewrite p_elt_gen_length1 // (_ : p_elt_gen p H = V). rewrite /pHall pcore_sub pcore_pgroup /= pnatNK. - by apply: pnat_dvd pV; exact: dvdn_indexg. + by apply: pnat_dvd pV; apply: dvdn_indexg. rewrite -(genGid V) -(setIidPr sVH); congr <<_>>; apply/setP=> x. rewrite !inE; apply: andb_id2l => Hx. by rewrite (mem_normal_Hall sylVH) /normal ?sVH. suffices sPV: P \subset V by rewrite -(joing_idPl sPV). - suffices sPU: P \subset U by rewrite (sub_normal_Hall sylV) //; exact/andP. + suffices sPU: P \subset U by rewrite (sub_normal_Hall sylV) //; apply/andP. have cUPb: P / V \subset 'C_(H / V)(U / V). rewrite subsetI morphimS // -mulVK quotientMidl quotient_cents2r //. by rewrite commGC trKP sub1G. rewrite -(quotientSGK nVP sVU) (subset_trans cUPb) //. by rewrite -defU cent_sub_Fitting ?quotient_sol. -have{sylVP} dxV: [~: V, K] \x 'C_V(K) = V by exact: coprime_abelian_cent_dprod. +have{sylVP} dxV: [~: V, K] \x 'C_V(K) = V by apply: coprime_abelian_cent_dprod. have tiVsub_VcK: 'C_V(K) = 1 \/ 'C_V(K) = V. apply: (indecomposableV _ [~: V, K]); first by rewrite dprodC. rewrite -mulHR -defH -mulgA mul_subG // subsetI. @@ -777,7 +782,7 @@ have nKR0: R0 \subset 'N(K) := subset_trans sR0R nKR. have mulKR0: K * R0 = K <*> R0 by rewrite norm_joinEr. have sKR0_G : K <*> R0 \subset G by rewrite -mulKR0 -mulHR mulgSS. have nV_KR0: K <*> R0 \subset 'N(V) := subset_trans sKR0_G nVG. -have solKR0: solvable (K <*> R0) by exact: solvableS solG. +have solKR0: solvable (K <*> R0) by apply: solvableS solG. have coKR0: coprime #|K| #|R0| by rewrite (coprimeSg sKH) ?(coprimegS sR0R). have r'K: r^'.-group K. by rewrite /pgroup p'natE -?prime_coprime // coprime_sym -oR0. @@ -790,7 +795,7 @@ have tiKR0cV: 'C_(K <*> R0)(V) = 1. suff defC: C == R0 by rewrite -(eqP defC) (subset_trans (joing_subl K R0)). have sC_R0: C \subset R0. rewrite -[C](coprime_mulG_setI_norm mulKR0) ?norms_cent //= tiKcV mul1g. - exact: subsetIl. + apply: subsetIl. rewrite eqEsubset sC_R0; apply: contraR ntC => not_sR0C. by rewrite -(setIidPr sC_R0) prime_TIg ?oR0. have{nKR0 mulKR0 sKR0_G solKR0 nV_KR0} oCVR0: #|'C_V(R0)| = p. @@ -801,7 +806,7 @@ have{nKR0 mulKR0 sKR0_G solKR0 nV_KR0} oCVR0: #|'C_V(R0)| = p. apply: odd_prime_sdprod_abelem_cent1 abelV nV_KR0 _ _; rewrite // ?oR0 //=. by rewrite -mulKR0 pgroupM p'K /pgroup oR0 pnatE. have [x defC]: exists x, 'C_V(R0) = <[x]>. - have ZgrC: Zgroup 'C_V(R0) by apply: ZgroupS ZgrCHR0; exact: setSI. + have ZgrC: Zgroup 'C_V(R0) by apply: ZgroupS ZgrCHR0; apply: setSI. apply/cyclicP; apply: (forall_inP ZgrC); apply/SylowP; exists p => //. by rewrite /pHall subxx indexgg (pgroupS (subsetIl V _)). rewrite defC; apply: nt_prime_order => //; last by rewrite -cycle_eq1 -defC. @@ -835,15 +840,14 @@ have{IHsub nVH} IHsub: forall X : {group gT}, by rewrite 2?normsY ?nXR0 ?(subset_trans sR0R) // (subset_trans sRG). have Op'H0: 'O_p^'(H0) = 1. have [sOp' nOp'] := andP (pcore_normal _ _ : 'O_p^'(H0) <| H0). - have p'Op': p^'.-group 'O_p^'(H0) by exact: pcore_pgroup. + have p'Op': p^'.-group 'O_p^'(H0) by apply: pcore_pgroup. apply: card1_trivg (pnat_1 (pgroupS _ pV) p'Op'). rewrite -scVH subsetI (subset_trans sOp') //= centsC; apply/setIidPl. rewrite -coprime_norm_cent ?(pnat_coprime pV p'Op') //. by rewrite (setIidPl (subset_trans _ nOp')) // /H0 -joingA joing_subl. exact: subset_trans (subset_trans sH0H nVH). have Op'HR0: 'O_p^'([~: H0, R0]) = 1. - apply/trivgP; rewrite -Op'H0 pcore_max ?pcore_pgroup //. - apply: char_normal_trans (pcore_char _ _) _. + apply/trivgP; rewrite -Op'H0 pcore_max ?pcore_pgroup // gFnormal_trans //. by rewrite /(_ <| _) commg_norml andbT commg_subl. have{ltG0G IHsub} p1_HR0: p.-length_1 [~: H0, R0]. by apply: IHsub ltG0G => //=; rewrite mul_subG ?normG. @@ -854,8 +858,7 @@ have{IHsub nVH} IHsub: forall X : {group gT}, by rewrite -{1}((_ :=P: _) p1_HR0) (quotient_pseries [::_;_]) pcore_pgroup. apply/trivgP; have <-: K :&: 'O_p([~: H0, R0]) = 1. by rewrite setIC coprime_TIg // (pnat_coprime (pcore_pgroup p _)). - rewrite commg_subI // subsetI ?sPOpHR0 ?sXK //=. - by rewrite (char_norm_trans (pcore_char _ _)) // normsRl. + by rewrite commg_subI // subsetI ?sPOpHR0 ?sXK //= gFnorm_trans // normsRl. have{defH sR0R} [defH defR0]: V * K * P = H /\ R0 :=: R. suffices: (V * K * P == H) && (R0 :==: R) by do 2!case: eqP => // ->. apply: contraR ntKP; rewrite -subG1 !eqEcard sR0R ?mul_subG //= negb_and. @@ -901,23 +904,23 @@ have{scKH} tiPRcK: 'C_(P <*> R)(K) = 1. have [K1 | ntK]:= eqsVneq K 1; first by rewrite K1 comm1G eqxx in ntKR. have [K1 | [q q_pr q_dv_K]] := trivgVpdiv K; first by case/eqP: ntK. have q_gt1 := prime_gt1 q_pr. -have p'q: q != p by exact: (pgroupP p'K). -have{r'K} q'r: r != q by rewrite eq_sym; exact: (pgroupP r'K). +have p'q: q != p by apply: (pgroupP p'K). +have{r'K} q'r: r != q by rewrite eq_sym; apply: (pgroupP r'K). have{defK} qK: q.-group K. have{defK} nilK: nilpotent K by rewrite -defK Fitting_nil. have{nilK} [_ defK _ _] := dprodP (nilpotent_pcoreC q nilK). have{IHsub} IHpi: forall pi, 'O_pi(K) = K \/ 'O_pi(K) \subset 'C(P). move=> pi; apply: IHsub (pcore_sub _ _). - by apply: char_norm_trans (pcore_char _ _) _; rewrite join_subG nKP. + by rewrite gFnorm_trans // join_subG nKP. case: (IHpi q) => [<-| cPKq]; first exact: pcore_pgroup. case/eqP: ntKP; apply/commG1P; rewrite -{}defK mul_subG //. case: (IHpi q^') => // defK; case/idPn: q_dv_K. - rewrite -p'natE // -defK; exact: pcore_pgroup. + by rewrite -p'natE // -defK; apply: pcore_pgroup. pose K' := K^`(1); have charK': K' \char K := der_char 1 K. have nsK'K: K' <| K := der_normal 1 K; have [sK'K nK'K] := andP nsK'K. have nK'PR: P <*> R \subset 'N(K') := char_norm_trans charK' nKPR. have iK'K: 'C_(P <*> R / K')(K / K') = 1 -> #|K / K'| > q ^ 2. - have qKb: q.-group (K / K') by exact: morphim_pgroup qK. + have qKb: q.-group (K / K') by apply: morphim_pgroup qK. rewrite ltnNge => trCK'; apply: contra ntKP => Kq_le_q2. suffices sPR_K': [~: P, R] \subset K'. rewrite -defP -(setIidPl sPR_K') coprime_TIg ?commG1 //. @@ -925,13 +928,13 @@ have iK'K: 'C_(P <*> R / K')(K / K') = 1 -> #|K / K'| > q ^ 2. rewrite -quotient_cents2 ?(char_norm_trans charK') //. suffices cPRbPrb: abelian (P <*> R / K'). by rewrite (sub_abelian_cent2 cPRbPrb) ?quotientS ?joing_subl ?joing_subr. - have nKbPR: P <*> R / K' \subset 'N(K / K') by exact: quotient_norms. + have nKbPR: P <*> R / K' \subset 'N(K / K') by apply: quotient_norms. case cycK: (cyclic (K / K')). rewrite (isog_abelian (quotient1_isog _)) -trCK' -ker_conj_aut. rewrite (isog_abelian (first_isog_loc _ _)) //. by rewrite (abelianS (Aut_conj_aut _ _)) ?Aut_cyclic_abelian. have{cycK} [oKb abelKb]: #|K / K'| = (q ^ 2)%N /\ q.-abelem (K / K'). - have sKb1: 'Ohm_1(K / K') \subset K / K' by exact: Ohm_sub. + have sKb1: 'Ohm_1(K / K') \subset K / K' by apply: Ohm_sub. have cKbKb: abelian (K / K') by rewrite sub_der1_abelian. have: #|'Ohm_1(K / K')| >= q ^ 2. rewrite (card_pgroup (pgroupS sKb1 qKb)) leq_exp2l // ltnNge. @@ -949,8 +952,8 @@ have iK'K: 'C_(P <*> R / K')(K / K') = 1 -> #|K / K'| > q ^ 2. by rewrite /pgroup (pi_pnat rR) // (pi_pnat pP) // !inE eq_sym. case cKK: (abelian K); last first. have [|[dPhiK dK'] dCKP] := abelian_charsimple_special qK coKP defKP. - apply/bigcupsP=> L; case/andP=> charL; have sLK := char_sub charL. - by case/IHsub: sLK cKK => // [|-> -> //]; exact: (char_norm_trans charL). + apply/bigcupsP=> L /andP[charL]; have sLK := char_sub charL. + by case/IHsub: sLK cKK => // [|-> -> //]; apply: char_norm_trans charL _. have eK: exponent K %| q. have oddK: odd #|K| := oddSg sKG oddG. have [Q [charQ _ _ eQ qCKQ]] := critical_odd qK oddK ntK; rewrite -eQ. @@ -1002,7 +1005,7 @@ case cKK: (abelian K); last first. have [x defCKR]: exists x, 'C_K(R) = <[x]>. have ZgrCKR: Zgroup 'C_K(R) := ZgroupS (setSI _ sKH) ZgrCHR. have qCKR: q.-group 'C_K(R) by rewrite (pgroupS (subsetIl K _)). - by apply/cyclicP; exact: nil_Zgroup_cyclic (pgroup_nil qCKR). + by apply/cyclicP; apply: nil_Zgroup_cyclic (pgroup_nil qCKR). have Kx: x \in K by rewrite -cycle_subG -defCKR subsetIl. rewrite defCKR cycle_subG in not_sCKR_K' *. exact: nt_prime_order (exponentP eK x Kx) (group1_contra not_sCKR_K'). @@ -1044,7 +1047,7 @@ case cKK: (abelian K); last first. rewrite -divgS // -{1}(coprime_cent_prod nKR) // TI_cardMg ?mulKn //. by rewrite setIA (setIidPl sKR_K). have sKRx_K: [~: K, R] :^ x \subset K by rewrite -{2}(normsP nKP x Px) conjSg. - have nKR_K: K \subset 'N([~: K, R]) by exact: commg_norml. + have nKR_K: K \subset 'N([~: K, R]) by apply: commg_norml. have mulKR_Krx: [~: K, R] * [~: K, R] :^ x = K. have maxKR: maximal [~: K, R] K by rewrite p_index_maximal ?iKR. apply: mulg_normal_maximal; rewrite ?(p_maximal_normal qK) //. @@ -1059,12 +1062,11 @@ case cKK: (abelian K); last first. by rewrite mulKR_Krx mulnAC leq_pmul2l ?muln_gt0 ?cardG_gt0 ?subset_leq_card. have tiKcP: 'C_K(P) = 1 by rewrite -defKP coprime_abel_cent_TI. have{IHsub} abelK: q.-abelem K. - have [|cPK1] := IHsub _ (char_norm_trans (Ohm_char 1 K) nKPR) (Ohm_sub 1 K). - by move/abelem_Ohm1P->. + have [/abelem_Ohm1P->//|cPK1] := IHsub _ (gFnorm_trans _ nKPR) (Ohm_sub 1 K). rewrite -(setIid K) TI_Ohm1 ?eqxx // in ntK. by apply/eqP; rewrite -subG1 -tiKcP setIS. have{K' iK'K charK' nsK'K sK'K nK'K nK'PR} oKq2: q ^ 2 < #|K|. - have K'1: K' :=: 1 by exact/commG1P. + have K'1: K' :=: 1 by apply/commG1P. rewrite -indexg1 -K'1 -card_quotient ?normal_norm // iK'K // K'1. by rewrite -injm_subcent ?coset1_injm ?norms1 //= tiPRcK morphim1. pose S := [set Vi : {group gT} | 'C_V('C_K(Vi)) == Vi & maximal 'C_K(Vi) K]. @@ -1174,7 +1176,7 @@ have [cSR | not_cSR] := boolP (R \subset 'C(S | 'JG)). rewrite join_subG (subset_trans sRnSV) /= ?bigcap_inf // andbT -defP. apply: (subset_trans (commgS P sRnSV)). have:= subset_trans (joing_subl P R) sPRnSV; rewrite -commg_subr /=. - move/subset_trans; apply; exact: bigcap_inf. + by move/subset_trans; apply; apply: bigcap_inf. rewrite -afixJG; move/orbit1P => -> allV1. have defV1: V1 = V by apply: group_inj; rewrite /= -defV allV1 big_set1. case/idPn: oKq2; rewrite -(Lagrange (subsetIl K 'C(V1))). @@ -1232,7 +1234,7 @@ have nVjR Vj: Vj \in S :\: D -> 'C_K(Vj) = [~: K, R]. have iKj: #|K : Kj| = q by rewrite (p_maximal_index qK (maxSK Vj Sj)). have dxKR: [~: K, R] \x 'C_K(R) = K by rewrite coprime_abelian_cent_dprod. have{dxKR} [_ defKR _ tiKRcR] := dprodP dxKR. - have Z_CK: Zgroup 'C_K(R) by apply: ZgroupS ZgrCHR; exact: setSI. + have Z_CK: Zgroup 'C_K(R) by apply: ZgroupS ZgrCHR; apply: setSI. have abelCKR: q.-abelem 'C_K(R) := abelemS (subsetIl _ _) abelK. have [qCKR _] := andP abelCKR. apply/eqP; rewrite eq_sym eqEcard sKRVj -(leq_pmul2r (ltnW q_gt1)). @@ -1250,9 +1252,9 @@ have nVjR Vj: Vj \in S :\: D -> 'C_K(Vj) = [~: K, R]. by rewrite (actsP actsR) //= S1 mem_imset. rewrite (subset_trans sCVj) // class_supportEr -(bigdprodWY defU) genS //. apply/bigcupsP=> x Rx; rewrite (bigcup_max (Vj :^ x)%G) // inE. - by rewrite (actsP actsR) // Sj andbT (orbit_transr _ (mem_orbit 'JG Vj Rx)). + by rewrite (actsP actsR) // Sj andbT (orbit_transl _ (mem_orbit 'JG Vj Rx)). have sDS: D \subset S. - by rewrite acts_sub_orbit //; apply: subset_trans actsPR; exact: joing_subr. + by rewrite acts_sub_orbit //; apply: subset_trans actsPR; apply: joing_subr. have [eqDS | ltDS] := eqVproper sDS. have [fix0 | [Vj cVjP]] := set_0Vmem 'Fix_(S | 'JG)(P). case/negP: p'r; rewrite eq_sym -dvdn_prime2 // -oD eqDS /dvdn. @@ -1282,7 +1284,7 @@ move=> defG solG R_pr regR. elim: {K}_.+1 {-2}K (ltnSn #|K|) => // m IHm K leKm in G defG solG regR *. have [nsKG sRG defKR nKR tiKR] := sdprod_context defG. have [sKG nKG] := andP nsKG. -wlog ntK: / K :!=: 1 by case: eqP => [-> _ | _ ->] //; exact: nilpotent1. +wlog ntK: / K :!=: 1 by case: eqP => [-> _ | _ ->] //; apply: nilpotent1. have [L maxL _]: {L : {group gT} | maxnormal L K G & [1] \subset L}. by apply: maxgroup_exists; rewrite proper1G ntK norms1. have [ltLK nLG]:= andP (maxgroupp maxL); have [sLK not_sKL]:= andP ltLK. @@ -1393,7 +1395,7 @@ have nsFK := char_normal chF; have [sFK nFK] := andP nsFK. pose KqF := K / 'F(K); have solK := solvableS sKG solG. without loss [p p_pr pKqF]: / exists2 p, prime p & p.-group KqF. move=> IHp; apply: wlog_neg => IH_KR; rewrite -quotient_cents2 //= -/KqF. - set Rq := R / 'F(K); have nKRq: Rq \subset 'N(KqF) by exact: quotient_norms. + set Rq := R / 'F(K); have nKRq: Rq \subset 'N(KqF) by apply: quotient_norms. rewrite centsC. apply: subset_trans (coprime_cent_Fitting nKRq _ _); last first. - exact: quotient_sol. @@ -1409,7 +1411,7 @@ without loss [p p_pr pKqF]: / exists2 p, prime p & p.-group KqF. by apply: char_from_quotient nsFP (Fitting_char _) _; rewrite -defPq. have defFP: 'F(P) = 'F(K). apply/eqP; rewrite eqEsubset !Fitting_max ?Fitting_nil //. - by rewrite char_normal ?(char_trans (Fitting_char _)). + by rewrite char_normal ?gFchar_trans. have coPR := coprimeSg sPK coKR. have nPR: R \subset 'N(P) := char_norm_trans chP nKR. pose G1 := P <*> R. @@ -1504,7 +1506,7 @@ rewrite quotient_cents2 ?norms_cent ?nW_ //= commGC. pose Hv := (P / V) <*> (R / V). have sHGv: Hv \subset G / V by rewrite join_subG !quotientS. have solHv: solvable Hv := solvableS sHGv (quotient_sol V solG). -have sPHv: P / V \subset Hv by exact: joing_subl. +have sPHv: P / V \subset Hv by apply: joing_subl. have nPRv: R / V \subset 'N(P / V) := quotient_norms _ nPR. have coPRv: coprime #|P / V| #|R / V| := coprime_morph _ (coprimeSg sPK coKR). apply: subset_trans (subsetIr (P / V) _). @@ -1639,7 +1641,7 @@ have nRqP: P \subset 'N('O_q(R)) by rewrite (subset_trans sPR) ?gFnorm. rewrite centsC (coprime_odd_faithful_Ohm1 qRq) ?(oddSg sRqR) //. apply: sub_abelian_cent2 (joing_subl _ _) (joing_subr _ _) => /=. set PQ := P <*> _; have oPQ: #|PQ| = (p * q)%N. - rewrite /PQ norm_joinEl ?(char_norm_trans (Ohm_char 1 _)) //. + rewrite /PQ norm_joinEl 1?gFnorm_trans //. rewrite coprime_cardMg 1?coprime_sym ?(coprimeSg (Ohm_sub 1 _)) // -/p. by congr (p * _)%N; apply: Ohm1_cyclic_pgroup_prime => /=. have sPQ_R: PQ \subset R by rewrite join_subG sPR (subset_trans (Ohm_sub 1 _)). @@ -1813,18 +1815,16 @@ have cMK': K^`(1) / V \subset 'C_(K / V)(M / V). - move=> Vx; case/setD1P=> ntVx; case/morphimP=> x nVx Rx defVx. rewrite defVx /= -cent_cycle -quotient_cycle //; congr 'C__(_ / V). apply/eqP; rewrite eqEsubset cycle_subG Rx /=. - apply: contraR ntVx; move/(prime_TIg Rpr); move/trivgP. - rewrite defVx /= (setIidPr _) cycle_subG //; move/set1P->. - by rewrite morph1. + apply: contraR ntVx => /(prime_TIg Rpr)/trivgP. + by rewrite defVx /= (setIidPr _) cycle_subG // => /set1P->; rewrite morph1. - rewrite -coprime_norm_quotient_cent ?(coprimegS sKG) ?(subset_trans sKG) //. by rewrite tcKM quotient1. move=> _ _ -> //; rewrite -coprime_quotient_cent ?quotient_cyclic //. by rewrite (coprimegS sRG). -rewrite !subsetI in cVK' cMK' *. -case/andP: cVK' => sK'K cVK'; case/andP: cMK' => _ cMVK'; rewrite sK'K. +have{cVK' cMK'} [[sK'K cVK'] [_ cMVK']] := (subsetIP cVK', subsetIP cMK'). have sK'G: K^`(1) \subset G by rewrite (subset_trans sK'K) ?proper_sub. have coMK': coprime #|M| #|K^`(1)| := coprimegS sK'G coMG. -rewrite (stable_factor_cent cVK') // /stable_factor /normal sVM nVM !andbT. +rewrite subsetI sK'K (stable_factor_cent cVK') //; apply/and3P; split=> //. by rewrite commGC -quotient_cents2 // (subset_trans sK'G). Qed. diff --git a/mathcomp/odd_order/BGsection4.v b/mathcomp/odd_order/BGsection4.v index c33bd2f..65be7a3 100644 --- a/mathcomp/odd_order/BGsection4.v +++ b/mathcomp/odd_order/BGsection4.v @@ -1,10 +1,18 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div. +From mathcomp Require Import fintype finfun bigop ssralg finset prime binomial. +From mathcomp Require Import fingroup morphism automorphism perm quotient action gproduct. +From mathcomp Require Import gfunctor commutator zmodp cyclic center pgroup gseries nilpotent. +From mathcomp Require Import sylow abelian maximal extremal hall. +From mathcomp Require Import matrix mxalgebra mxrepresentation mxabelem. +From mathcomp Require Import BGsection1 BGsection2. (******************************************************************************) @@ -60,7 +68,7 @@ have p3_L21: p <= 3 -> {in R & &, forall u v w, [~ u, v, w] = 1}. have{fS gS} expMR_fg: {in R &, forall u v n (r := [~ v, u]), (u * v) ^+ n = u ^+ n * v ^+ n * r ^+ 'C(n, 2) * [~ r, u] ^+ f n * [~ r, v] ^+ g n}. -- move=> u v Ru Rv n r; have Rr: r \in R by exact: groupR. +- move=> u v Ru Rv n r; have Rr: r \in R by apply: groupR. have cRr: {in R &, forall x y, commute x [~ r, y]}. move=> x y Rx Ry /=; red; rewrite (centerC Rx) //. have: [~ r, y] \in 'L_3(R) by rewrite !mem_commg. @@ -94,12 +102,11 @@ have expR1p: exponent 'Ohm_1(R) %| p. by apply: (centsP (cycle_abelian x)); rewrite ?defQ. have:= maxgroupp maxS; rewrite properEcard => /andP[sSQ ltSQ]. have pQ := pgroupS sQR pR; have pS := pgroupS sSQ pQ. - have{ltSQ leQn} ltSn: #|S| < n by exact: leq_trans ltSQ _. + have{ltSQ leQn} ltSn: #|S| < n by apply: leq_trans ltSQ _. have expS1p := IHn _ ltSn (subset_trans sSQ sQR). have defS1 := Ohm1Eexponent p_pr expS1p; move/exp_dv_p: expS1p => expS1p. have nS1Q: [~: Q, 'Ohm_1(S)] \subset 'Ohm_1(S). - rewrite commg_subr (char_norm_trans (Ohm_char 1 S)) ?normal_norm //. - exact: p_maximal_normal pQ maxS. + by rewrite commg_subr gFnorm_trans ?normal_norm // (p_maximal_normal pQ). have S1x : x \in 'Ohm_1(S) by rewrite defS1 !inE -cycle_subG sxS xp1 /=. have S1yx : [~ y, x] \in 'Ohm_1(S) by rewrite (subsetP nS1Q) ?mem_commg. have S1yxx : [~ y, x, x] \in 'Ohm_1(S) by rewrite groupR. @@ -129,8 +136,7 @@ have CiP_eq : C :&: R = A by rewrite -CRA_A setIC setIA (setIidPl sRG). have sylA: p.-Sylow(C) A. rewrite -CiP_eq; apply: (Sylow_setI_normal (subcent_normal _ _)). by apply: pHall_subl sylR; rewrite ?subsetIl // subsetI sRG normal_norm. -rewrite dprodEsd; last first. - by rewrite centsC (subset_trans (pcore_sub _ _)) ?subsetIr. +rewrite dprodEsd; last by rewrite centsC gFsub_trans ?subsetIr. by apply: Burnside_normal_complement; rewrite // subIset ?subsetIr. Qed. @@ -140,7 +146,7 @@ Lemma Ohm1_extremal_odd gT (R : {group gT}) p x : ('Ohm_1(R))%G \in 'E_p^2(R). Proof. move=> pR oddR ncycR Rx ixR; rewrite -cycle_subG in Rx. -have ntR: R :!=: 1 by apply: contra ncycR; move/eqP->; exact: cyclic1. +have ntR: R :!=: 1 by apply: contra ncycR; move/eqP->; apply: cyclic1. have [p_pr _ [e oR]]:= pgroup_pdiv pR ntR. case p2: (p == 2); first by rewrite oR odd_exp (eqP p2) in oddR. have [cRR | not_cRR] := orP (orbN (abelian R)). @@ -175,7 +181,7 @@ have [M minM]: {M | [min M | M <| R & ~~ cyclic M]}. by apply: ex_mingroup; exists R; rewrite normal_refl. have{minM} [[nsMR ncycM] [_ minM]] := (andP (mingroupp minM), mingroupP minM). have [sMR _] := andP nsMR; have pM := pgroupS sMR pR. -exists ('Ohm_1(M))%G; first exact: char_normal_trans (Ohm_char 1 M) nsMR. +exists ('Ohm_1(M))%G; first exact: gFnormal_trans. apply: (subsetP (pnElemS _ _ sMR)). have [M1 | ntM] := eqsVneq M 1; first by rewrite M1 cyclic1 in ncycM. have{ntM} [p_pr _ [e oM]] := pgroup_pdiv pM ntM. @@ -219,12 +225,11 @@ have sS_Z2R: S \subset 'Z_2(R). rewrite ucnSnR; apply/subsetP=> s Ss; rewrite inE (subsetP sSR) //= ucn1. by rewrite (subset_trans _ sSR_ZR) ?commSg ?sub1set. have sZ2R_R := ucn_sub 2 R; have pZ2R := pgroupS sZ2R_R pR. -have pZ: p.-group Z. - apply: pgroupS pR; apply: subset_trans (Ohm_sub _ _) (ucn_sub 2 R). +have pZ: p.-group Z by apply: pgroupS pR; apply/gFsub_trans/gFsub. have sSZ: S \subset Z. by rewrite /Z (OhmE 1 pZ2R) sub_gen // subsetI sS_Z2R sub_LdivT. have ncycX: ~~ cyclic S by rewrite (abelem_cyclic abelS) dimS. -split; first by apply: contra ncycX; exact: cyclicS. +split; first by apply: contra ncycX; apply: cyclicS. have nclZ2R : nil_class 'Z_2(R) <= 2 + _ := leq_trans (nil_class_ucn _ _) _. by have [] := exponent_odd_nil23 pZ2R (oddSg sZ2R_R oddR) (nclZ2R _ _). Qed. @@ -256,7 +261,7 @@ Proof. move=> pR oddR nsSR ncycS; have sSR := normal_sub nsSR. have{sSR ncycS} []:= Ohm1_odd_ucn2 (pgroupS sSR pR) (oddSg sSR oddR) ncycS. set Z := 'Ohm_1(_) => ncycZ expZp. -have chZS: Z \char S := char_trans (Ohm_char 1 _) (ucn_char 2 S). +have chZS: Z \char S by rewrite !gFchar_trans. have{nsSR} nsZR: Z <| R := char_normal_trans chZS nsSR. have [sZR _] := andP nsZR; have pZ: p.-group Z := pgroupS sZR pR. have geZ2: 2 <= logn p #|Z|. @@ -268,7 +273,7 @@ have{geZ2} p_pr: prime p by move: geZ2; rewrite lognE; case: (prime p). have{oE p_pr} dimE2: logn p #|E| = 2 by rewrite oE pfactorK. exists E; split; rewrite ?(subset_trans _ (char_sub chZS)) {chZS nsZR}//. rewrite !inE /abelem sER pE (p2group_abelian pE) dimE2 //= andbT. -by apply: (dvdn_trans _ expZp); apply: exponentS. +exact/(dvdn_trans _ expZp)/exponentS. Qed. (* This is B & G, Lemma 4.7, and (except for the trivial converse) Gorenstein *) @@ -351,7 +356,7 @@ apply/group_setP; rewrite !inE group1 expg1n. split=> //= x y; case/LdivP=> Rx xp1; case/LdivP=> Ry yp1. rewrite !inE groupM //=; apply: contraR not_clR_le3 => nt_xyp. pose XY := <[x]> <*> <[y]>. -have [XYx XYy]: x \in XY /\ y \in XY by rewrite -!cycle_subG; exact/joing_subP. +have [XYx XYy]: x \in XY /\ y \in XY by rewrite -!cycle_subG; apply/joing_subP. have{nt_xyp} defR: XY = R. have sXY_R : XY \subset R by rewrite join_subG !cycle_subG Rx Ry. have pXY := pgroupS sXY_R pR; have [// | ltXY_R] := eqVproper sXY_R. @@ -363,7 +368,7 @@ have [<- | ltXR] := eqVproper sXR. have [S maxS sXS]: {S : {group gT} | maximal S R & <[x]> \subset S}. exact: maxgroup_exists. have nsSR: S <| R := p_maximal_normal pR maxS; have [sSR _] := andP nsSR. -have{nsSR} nsS1R: 'Ohm_1(S) <| R := char_normal_trans (Ohm_char 1 S) nsSR. +have{nsSR} nsS1R: 'Ohm_1(S) <| R := gFnormal_trans _ nsSR. have [sS1R nS1R] := andP nsS1R; have pS1 := pgroupS sS1R pR. have expS1p: exponent 'Ohm_1(S) %| p := minR S (maxgroupp maxS). have{expS1p} dimS1: logn p #|'Ohm_1(S)| <= 3. @@ -469,8 +474,7 @@ have iTs: #|T : <[s]>| = p. by rewrite quotient_sub1 ?(contra (fun sRS => cyclicS sRS cycS)). have defR1: 'Ohm_1(R) = 'Ohm_1(T). apply/eqP; rewrite eqEsubset (OhmS _ sTR) andbT -Ohm_id OhmS //. - rewrite -(quotientSGK _ sST); last by rewrite (subset_trans _ nSR) ?Ohm_sub. - by rewrite -defTb morphim_Ohm. + by rewrite -(quotientSGK _ sST) ?gFsub_trans // -defTb morphim_Ohm. rewrite (subsetP (pnElemS _ _ sTR)) // (group_inj defR1). apply: Ohm1_extremal_odd iTs => //; apply: contra ncycR. by rewrite !(@odd_pgroup_rank1_cyclic _ p) // -p_rank_Ohm1 -defR1 p_rank_Ohm1. @@ -497,8 +501,7 @@ have nsR'R: R' <| R := der_normal 1 R; have [sR'R nR'R] := andP nsR'R. have [T EpT]: exists T, T \in 'E_p^1('Mho^e(R') :&: 'Z(R)). apply/p_rank_geP; rewrite -rank_pgroup; last first. by rewrite (pgroupS _ pR) //= setIC subIset ?center_sub. - rewrite rank_gt0 (meet_center_nil (pgroup_nil pR)) //. - exact: char_normal_trans (Mho_char e _) nsR'R. + rewrite rank_gt0 (meet_center_nil (pgroup_nil pR)) ?gFnormal_trans //. by case ntR'1: e; rewrite //= Mho0 (sameP eqP derG1P). have [p_gt1 p_pr] := (ltnW (ltnW p_gt3), pnElem_prime EpT). have p_odd: odd p by case/even_prime: p_pr p_gt3 => ->. @@ -530,7 +533,7 @@ have{yb defRb def_yb} defR: TX <*> <[y]> = R. rewrite -defTX -joingA -quotientYK ?join_subG ?quotientY ?cycle_subG ?nTx //. by rewrite !quotient_cycle // -def_xb -def_yb -defXb -defRb quotientGK. have sXYR: <[x]> <*> <[y]> \subset R by rewrite -defR -defTX -joingA joing_subr. -have [Rx Ry]: x \in R /\ y \in R by rewrite -!cycle_subG; exact/joing_subP. +have [Rx Ry]: x \in R /\ y \in R by rewrite -!cycle_subG; apply/joing_subP. have cTXY := subset_trans sXYR cTR; have [cTX cTY] := joing_subP cTXY. have [R'1_1 {e sTR'e} | ntR'1] := eqVneq 'Mho^1(R') 1; last first. have sR'TX: R' \subset TX. @@ -550,11 +553,11 @@ have{R'1_1} eR': exponent R' %| p. have <-: 'Ohm_1(R') = R' by apply/eqP; rewrite trivg_Mho ?R'1_1. rewrite -sub_LdivT (OhmEabelian (pgroupS sR'R pR)) ?subsetIr //. by rewrite (abelianS (OhmS 1 sR'R)) // (p2group_abelian pR1). -pose r := [~ x, y]; have Rr: r \in R by exact: groupR. +pose r := [~ x, y]; have Rr: r \in R by apply: groupR. have{defXb ntXb nsXbR} [i def_rb]: exists i, coset T r = (xb ^+ p) ^+ i. have p_xb: p.-elt xb by rewrite def_xb morph_p_elt ?(mem_p_elt pR). have pRbb: p.-group (R / T / 'Mho^1(Xb)) by rewrite !quotient_pgroup. - have [_ nXb1R] := andP (char_normal_trans (Mho_char 1 Xb) nsXbR). + have /andP[_ nXb1R]: 'Mho^1(Xb) <| R / T by apply: gFnormal_trans. apply/cycleP; rewrite -(Mho_p_cycle 1 p_xb) -defXb. apply: coset_idr; first by rewrite (subsetP nXb1R) ?mem_quotient. apply/eqP; rewrite !morphR ?(subsetP nXb1R) ?mem_quotient //=; apply/commgP. @@ -571,8 +574,8 @@ have{xb def_xb def_rb} [t Tt def_r]: exists2 t, t \in T & r = t * x ^+ (p * i). apply/rcosetP; rewrite -val_coset ?groupX ?morphX //= -def_xb. by rewrite expgM -def_rb val_coset ?groupR // rcoset_refl. have{eR' def_r cTT} defR': R' = <[r]>. - have R'r : r \in R' by exact: mem_commg. - have cxt: t \in 'C[x] by apply/cent1P; exact: (centsP cRT). + have R'r : r \in R' by apply: mem_commg. + have cxt: t \in 'C[x] by apply/cent1P; apply: (centsP cRT). have crx: x \in 'C[r] by rewrite cent1C def_r groupM ?groupX ?cent1id. have def_xy: x ^ y = t * x ^+ (p * i).+1. by rewrite conjg_mulR -/r def_r expgS !mulgA (cent1P cxt). @@ -615,7 +618,7 @@ have defU1: 'Ohm_1(U) = 'Ohm_1(R). by rewrite (card_pgroup pR1) leq_exp2l. apply/eqP; rewrite eqEcard oUs defUs -{1}defU1 quotientS ?Ohm_sub //. rewrite dvdn_leq ?cardG_gt0 //; case/pgroup_pdiv: (quotient_pgroup S pR1) => //. -rewrite -subG1 quotient_sub1 ?(subset_trans (Ohm_sub 1 R) nSR) //. +rewrite -subG1 quotient_sub1 ?(gFsub_trans _ nSR) //. apply: contraL (cycS) => sR1S; rewrite abelian_rank1_cyclic ?cyclic_abelian //. rewrite -ltnNge (rank_pgroup (pgroupS sSR pR)); apply/p_rank_geP. by exists 'Ohm_1(U)%G; rewrite -(setIidPr sSU) pnElemI inE EpU1 inE /= defU1. @@ -659,7 +662,7 @@ suffices{C T} cTT: abelian [~: R, A]. have ncycR: ~~ cyclic R by apply: contra not_cRR; apply: cyclic_abelian. have: 'Ohm_1(R)%G \in 'E_p^2(R) by apply: Ohm1_metacyclic_p2Elem. have nT1C1: 'Ohm_1(C) \subset 'N('Ohm_1(T)). - by rewrite (subset_trans (Ohm_sub 1 _)) ?(char_norm_trans (Ohm_char 1 _)). + by rewrite gFsub_trans ?gFnorm_trans. by case/pnElemP=> _ _ <-; rewrite -norm_joinEr ?lognSg // join_subG !OhmS. without loss defR: R pR oddR metaR nRA / [~: R, A] = R. set T := [~: R, A] => IH; have sTR: T \subset R by rewrite commg_subl. @@ -671,14 +674,13 @@ rewrite defR; apply: wlog_neg => not_cRR. have ncycR: ~~ cyclic R := contra (@cyclic_abelian _ R) not_cRR. pose cycR_nA S := [&& cyclic S, S \subset R & A \subset 'N(S)]. have [S maxS sR'S] : {S | [max S | cycR_nA S] & R^`(1) \subset S}. - apply: maxgroup_exists; rewrite {}/cycR_nA der_sub /=. - rewrite (char_norm_trans (der_char 1 _)) // andbT. + apply: maxgroup_exists; rewrite {}/cycR_nA der_sub /= gFnorm_trans // andbT. have [K [cycK nsKR cycKR]] := metacyclicP metaR. by rewrite (cyclicS _ cycK) // der1_min ?normal_norm // cyclic_abelian. -case/maxgroupP: maxS; case/and3P=> cycS sSR nSA maxS. +have{maxS} [/and3P[cycS sSR nSA] maxS] := maxgroupP maxS. have ntS: S :!=: 1 by rewrite (subG1_contra sR'S) // (sameP eqP derG1P). have nSR: R \subset 'N(S) := sub_der1_norm sR'S sSR. -have nsSR: S <| R by exact/andP. +have nsSR: S <| R by apply/andP. have sSZ: S \subset 'Z(R). have sR_NS': R \subset 'N(S)^`(1) by rewrite -{1}defR commgSS. rewrite subsetI sSR centsC (subset_trans sR_NS') // der1_min ?cent_norm //=. @@ -689,8 +691,8 @@ have pRb: p.-group (R / S) := quotient_pgroup S pR. pose R1 := 'Ohm_1(R); pose Rb1 := 'Ohm_1(R / S). have [Xb]: exists2 Xb, R1 / S \x gval Xb = Rb1 & A / S \subset 'N(Xb). have MaschkeRb1 := Maschke_abelem (Ohm1_abelem pRb cRbRb). - pose normOhm1 := (char_norm_trans (Ohm_char 1 _), quotient_norms S). - by apply: MaschkeRb1; rewrite ?quotient_pgroup ?morphim_Ohm ?normOhm1. + pose normOhm1 := (morphim_Ohm, gFnorm_trans, quotient_norms S). + by apply: MaschkeRb1; rewrite ?quotient_pgroup ?normOhm1. case/dprodP=> _ defRb1 _ tiR1bX nXbA. have sXbR: Xb \subset R / S. by apply: subset_trans (Ohm_sub 1 _); rewrite -defRb1 mulG_subr. @@ -708,7 +710,7 @@ have{tiR1bX} cycX: cyclic X. rewrite (cyclic_factor_abelian sSZ) // abelian_rank1_cyclic //. rewrite (rank_abelian_pgroup pRb cRbRb) -defRb1 defXb. rewrite (maxS X) ?trivg_quotient ?mulg1 //; last exact/and3P. -have EpR1: 'Ohm_1(R)%G \in 'E_p^2(R) by exact: Ohm1_metacyclic_p2Elem. +have EpR1: 'Ohm_1(R)%G \in 'E_p^2(R) by apply: Ohm1_metacyclic_p2Elem. have [sR1R _ dimR1] := pnElemP EpR1; have pR1 := pgroupS sR1R pR. rewrite -(card_isog (second_isog _)) ?(subset_trans sR1R) // -ltnS -dimR1. by rewrite (ltn_log_quotient pR1) ?subsetIr //= meet_Ohm1 // (setIidPl sSR). @@ -761,7 +763,7 @@ have{recR} IH: forall S, gval S \proper R -> A \subset 'N(S) -> A \subset 'C(S). have defR1: 'Ohm_1(R) = R. apply: contraNeq not_cRA; rewrite eqEproper Ohm_sub negbK => ltR1R. rewrite (coprime_odd_faithful_Ohm1 pR) ?IH ?(odd_pgroup_odd p_odd) //. - by rewrite (char_norm_trans (Ohm_char 1 R)). + exact: gFnorm_trans. have defRA: [~: R, A] = R. apply: contraNeq not_cRA; rewrite eqEproper commg_subl nRA negbK => ltRAR. rewrite centsC; apply/setIidPl. @@ -781,13 +783,13 @@ have [cRR | not_cRR] := boolP (abelian R). by rewrite (subn_sqr p 1) mulnA !Euclid_dvdM ?orbb. have [[defPhi defR'] _]: special R /\ 'C_R(A) = 'Z(R). apply: (abelian_charsimple_special pR) => //. - apply/bigcupsP=> S; case/andP=> charS cSS. + apply/bigcupsP=> S /andP[charS cSS]. rewrite centsC IH ?(char_norm_trans charS) // properEneq char_sub // andbT. by apply: contraNneq not_cRR => <-. have ntZ: 'Z(R) != 1 by rewrite -defR' (sameP eqP derG1P). have ltRbR: #|R / 'Z(R)| < #|R| by rewrite ltn_quotient ?center_sub. have pRb: p.-group (R / 'Z(R)) by apply: quotient_pgroup. -have nAZ: A \subset 'N('Z(R)) by rewrite (char_norm_trans (center_char R)). +have nAZ: A \subset 'N('Z(R)) by apply: gFnorm_trans. have defAb: A / 'Z(R) = <[coset _ a]> by rewrite quotient_cycle -?cycle_subG. have oab: #[coset 'Z(R) a] = q. rewrite orderE -defAb -(card_isog (quotient_isog _ _)) //. @@ -834,8 +836,8 @@ have p_gt3: p > 3; last split => //. by case/(pi_Aut_rank2_pgroup pR): (pgroupP (pgroup_pi _) q q_pr q_dv_A). pose S := 'Ohm_1(R); pose S' := S^`(1); pose C := 'C_R(S). have pS: p.-group S := pgroupS (Ohm_sub 1 _) pR. -have nsSR: S <| R := Ohm_normal 1 R. -have nsS'R: S' <| R := char_normal_trans (der_char 1 _) nsSR. +have nsSR: S <| R := gFnormal _ R. +have nsS'R: S' <| R := gFnormal_trans _ nsSR. have [sSR nSR] := andP nsSR; have [_ nS'R] := andP nsS'R. have [Sle2 | Sgt2] := leqP (logn p #|S|) 2. have metaR: metacyclic R := p2_Ohm1_metacyclic pR p_gt3 Sle2. @@ -894,7 +896,7 @@ have{cBbBb} abelBb: p.-abelem (B / C). have nsCB: C <| B by rewrite (normalS _ _ nsCR) ?setIS ?subsetIl // centS. have p'Ab: p^'.-group (A / C) by apply: quotient_pgroup. have sTbB: T / C \subset B / C by rewrite quotientS. -have nSA: A \subset 'N(S) := char_norm_trans (Ohm_char 1 _) nRA. +have nSA: A \subset 'N(S) := gFnorm_trans _ nRA. have nTA: A \subset 'N(T) := normsR nSA nRA. have nTbA: A / C \subset 'N(T / C) := quotient_norms _ nTA. have nBbA: A / C \subset 'N(B / C). @@ -914,7 +916,7 @@ have{nXbA} nXA: A \subset 'N(X). by rewrite -(quotientSGK nCA) ?normsG // quotient_normG -?defXb. have{abelT} defB1: 'Ohm_1(B) = T. apply/eqP; rewrite eq_sym eqEcard -{1}[T](Ohm1_id abelT) OhmS //. - have pB1: p.-group 'Ohm_1(B) := pgroupS (subset_trans (Ohm_sub 1 _) sBR) pR. + have pB1: p.-group 'Ohm_1(B) by apply: pgroupS pR; apply: gFsub_trans. rewrite (card_pgroup pT) (card_pgroup pB1) leq_exp2l //= -/T -/B. rewrite dimT -ltnS -dimS properG_ltn_log // properEneq OhmS ?subsetIl //= -/S. by case: eqP not_sSB => // <-; rewrite Ohm_sub. @@ -950,7 +952,7 @@ have [Ta AbTa not_cSbTa]: exists2 Ta, Ta \in A / T & Ta \notin 'C(S / T). by rewrite cents_norm ?subIset ?centS ?orbT. by rewrite mul_subG ?commg_subl ?normsI ?norms_cent // (subset_trans sSA_T). have [a nTa Aa defTa] := morphimP AbTa. -have nS'a: a \in 'N(S') := subsetP (char_norm_trans (der_char 1 _) nSA) a Aa. +have nS'a: a \in 'N(S') := subsetP (gFnorm_trans _ nSA) a Aa. have [i xa]: exists i, x ^ a = x ^+ i. by apply/cycleP; rewrite -cycle_subG cycleJ /= -defX (normsP nXA). have [j Tya]: exists j, Ty ^ Ta = Ty ^+ j. @@ -1077,14 +1079,14 @@ have s_p'C_B X: gval X \subset C -> p^'.-group X -> X \subset B. have [x Nx Hx def_v] := morphimP Vv; rewrite {1}def_v qactE //=. by rewrite -qactE ?(astab_dom cVa) ?(astab_act cVa) -?def_v. have{B pB s_p'C_B} pC : p.-group C. - apply/pgroupP=> q q_pr; case/Cauchy=> // a Ca oa; apply: wlog_neg => p'q. + apply/pgroupP=> q q_pr /Cauchy[] // a Ca oa; apply: wlog_neg => p'q. apply: (pgroupP pB) => //; rewrite -oa cardSg // s_p'C_B ?cycle_subG //. by rewrite /pgroup -orderE oa pnatE. have nVA: A \subset qact_dom <[nHA]> 'Phi(H) by rewrite qact_domE // acts_char. have nCA: A \subset 'N(C). by rewrite (subset_trans _ (astab_norm _ _)) // astabs_range. suffices{pC nCA}: p.-group (A / C)^`(1). - by rewrite -quotient_der ?pquotient_pgroup // (subset_trans (der_sub 1 A)). + by rewrite -quotient_der ?pquotient_pgroup // gFsub_trans. pose toAV := ((<[nHA]> / 'Phi(H)) \ nVA)%gact. have defC: C = 'C(V | toAV). by symmetry; rewrite astab_ract; apply/setIidPr; rewrite subIset ?subsetIl. @@ -1151,14 +1153,14 @@ without loss Gp'1: gT G solG oddG rG / 'O_p^'(G) = 1. have nsGp': 'O_p^'(G) <| G := pcore_normal p^' G; have [_ nGp'] := andP nsGp'. move/(_ _ (G / 'O_p^'(G))%G); rewrite quotient_sol // quotient_odd //=. have Gp'1 := trivg_pcore_quotient p^' G. - rewrite p_rank_p'quotient ?pcore_pgroup // Gp'1 indexg1; case=> //=. - rewrite -quotient_der // card_quotient ?(subset_trans (der_sub 1 G)) // => ->. + rewrite p_rank_p'quotient ?pcore_pgroup // Gp'1 indexg1 => -[] //=. + rewrite -quotient_der // card_quotient ?gFsub_trans // => ->. rewrite (pseries_pop2 _ Gp'1) /= -pseries1 -quotient_pseries /= /pgroup. pose isos := (isog_abelian (third_isog _ _ _), card_isog (third_isog _ _ _)). by rewrite !{}isos ?pseries_normal ?pseries_sub_catl. rewrite pseries_pop2 // Gp'1 indexg1 -pgroupE /=. set R := 'O_p(G); pose C := 'C_G(R). -have [sRG nRG] := andP (pcore_normal p G : R <| G). +have /andP[sRG nRG]: R <| G by apply: gFnormal. have sCR: C \subset R by rewrite /C /R -(Fitting_eq_pcore _) ?cent_sub_Fitting. have pR: p.-group R := pcore_pgroup p G; have pC: p.-group C := pgroupS sCR pR. have nCG: G \subset 'N(C) by rewrite normsI ?normG ?norms_cent. @@ -1208,7 +1210,7 @@ wlog Gs_p'_1: gT G Gs U V oddG solG nsGsG rGs chiefUf pUf sUGs / 'O_p^'(Gs) = 1. rewrite trivg_pcore_quotient quotient_odd ?quotient_sol ?quotientS //. have p'K: p^'.-group K := pcore_pgroup p^' Gs. have tiUfK := coprime_TIg (pnat_coprime pUf (quotient_pgroup V p'K)). - have nsKG: K <| G := char_normal_trans (pcore_char p^' Gs) nsGsG. + have nsKG: K <| G by apply: gFnormal_trans. have [[sG'G sGsG] nKG] := (der_sub 1 G, normal_sub nsGsG, normal_norm nsKG). have{sGsG} [nKG' nKGs] := (subset_trans sG'G nKG, subset_trans sGsG nKG). case/andP: chiefUf; case/maxgroupP; case/andP=> ltVU nVG maxV nsUG. @@ -1235,7 +1237,7 @@ wlog Gs_p'_1: gT G Gs U V oddG solG nsGsG rGs chiefUf pUf sUGs / 'O_p^'(Gs) = 1. by rewrite subsetI sVU andbT -(quotientSGK nKV sKW) -defWb. by rewrite andbT /proper subsetIr subsetIidr -(quotientSGK nKU sKW) -defWb. pose R := 'O_p(Gs); have pR: p.-group R := pcore_pgroup p Gs. -have nsRG: R <| G := char_normal_trans (pcore_char p Gs) nsGsG. +have nsRG: R <| G by apply: gFnormal_trans. have [[sGsG nGsG] [sRG nRG]] := (andP nsGsG, andP nsRG). have nsRGs: R <| Gs := pcore_normal p Gs; have [sRGs nRGs] := andP nsRGs. have sylR: p.-Sylow(Gs) R. @@ -1358,9 +1360,9 @@ have hallHp': p^'.-Hall(H) 'O_p^'(H). rewrite -(p_rank_Sylow sylP) (leq_trans (p_rank_le_rank _ _)) //. exact: leq_trans (rankS sPF) Fle2. rewrite -(card_Hall hallHp') part_p'nat ?pnatNK ?muln1 // subset_leqif_card. - by rewrite pcore_max ?pcore_pgroup ?(char_normal_trans (pcore_char _ _)). + by rewrite pcore_max ?pcore_pgroup ?gFnormal_trans. rewrite pcore_max ?pcore_pgroup // (normalS _ _ (pcore_normal _ _)) //. -rewrite -quotient_sub1 ?(subset_trans (pcore_sub _ _)) //. +rewrite -quotient_sub1 ?gFsub_trans //. rewrite -(setIidPr (quotientS _ (pcore_sub _ _))) coprime_TIg //. by rewrite coprime_morphr // (pnat_coprime pGq (pcore_pgroup _ _)). Qed. @@ -1403,10 +1405,10 @@ move=> oddG solG Fle2; pose pi := [pred q | p <= q]. rewrite pHallE pcore_sub eqn_leq -{1}(part_pnat_id (pcore_pgroup _ _)). rewrite dvdn_leq ?partn_dvd ?cardSg ?pcore_sub // /=. rewrite (@eq_in_partn _ pi) => [|q piGq]; last first. - by rewrite !inE eqn_leq; apply: andb_idl => le_q_p; exact: max_pdiv_max. + by rewrite !inE eqn_leq; apply: andb_idl => le_q_p; apply: max_pdiv_max. rewrite -(card_Hall (rank2_ge_pcore_Hall p oddG solG Fle2)) -/pi. rewrite subset_leq_card // pcore_max ?pcore_normal //. -apply: sub_in_pnat (pcore_pgroup _ _) => q; move/(piSg (pcore_sub _ _)) => piGq. +apply: sub_in_pnat (pcore_pgroup _ _) => q /(piSg (pcore_sub _ _))-piGq. by rewrite !inE eqn_leq max_pdiv_max. Qed. diff --git a/mathcomp/odd_order/BGsection5.v b/mathcomp/odd_order/BGsection5.v index ab5a14a..cfb8133 100644 --- a/mathcomp/odd_order/BGsection5.v +++ b/mathcomp/odd_order/BGsection5.v @@ -1,8 +1,14 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div. +From mathcomp Require Import fintype finset prime fingroup morphism perm automorphism action. +From mathcomp Require Import quotient cyclic gfunctor pgroup gproduct center commutator. +From mathcomp Require Import gseries nilpotent sylow abelian maximal hall. +From mathcomp Require Import BGsection1 BGsection4. (******************************************************************************) @@ -121,7 +127,7 @@ have [B Ep3B nBR]: exists2 B, B \in 'E_p^3(R) & R \subset 'N(B). have [sCR _] := andP nsCR; have pC: p.-group C := pgroupS sCR pR. have{pC cCC} abelC1: p.-abelem 'Ohm_1(C) := Ohm1_abelem pC cCC. have dimC1: 3 <= logn p #|'Ohm_1(C)| by rewrite -rank_abelem // rank_Ohm1. - have nsC1R: 'Ohm_1(C) <| R := char_normal_trans (Ohm_char 1 _) nsCR. + have nsC1R: 'Ohm_1(C) <| R := gFnormal_trans _ nsCR. have [B [sBC1 nsBR oB]] := normal_pgroup pR nsC1R dimC1. have [sBR nBR] := andP nsBR; exists B => //; apply/pnElemP. by rewrite oB pfactorK // (abelemS sBC1). @@ -157,26 +163,18 @@ Let T := 'C_R(W). Let ntZ : Z != 1. Proof. by rewrite Ohm1_eq1 (center_nil_eq1 (pgroup_nil pR)). Qed. -Let sZR : Z \subset R. -Proof. exact: subset_trans (Ohm_sub 1 _) (center_sub R). Qed. +Let sZR : Z \subset R. Proof. by rewrite !gFsub_trans. Qed. Let abelZ : p.-abelem (Z). Proof. by rewrite (Ohm1_abelem (pgroupS _ pR)) ?center_sub ?center_abelian. Qed. -Let pZ : p.-group Z. -Proof. exact: abelem_pgroup abelZ. Qed. +Let pZ : p.-group Z. Proof. exact: abelem_pgroup abelZ. Qed. Let defCRZ : 'C_R(Z) = R. -Proof. -apply/eqP; rewrite eqEsubset subsetIl subsetIidl centsC. -by rewrite (subset_trans (Ohm_sub 1 _)) ?subsetIr. -Qed. - -Let sWR : W \subset R. -Proof. exact: subset_trans (Ohm_sub 1 _) (ucn_sub 2 R). Qed. +Proof. by apply/setIidPl; rewrite centsC gFsub_trans ?subsetIr. Qed. -Let nWR : R \subset 'N(W). -Proof. exact: char_norm_trans (Ohm_char 1 _) (char_norm (ucn_char 2 R)). Qed. +Let sWR : W \subset R. Proof. exact/gFsub_trans/gFsub. Qed. +Let nWR : R \subset 'N(W). Proof. exact/gFnorm_trans/gFnorm. Qed. (* This is B & G, Lemma 5.2. *) Lemma Ohm1_ucn_p2maxElem E : @@ -188,7 +186,7 @@ Proof. case/setIP=> Ep2E maxE; have defCRE1 := Ohm1_cent_max maxE pR. have [[sER abelE dimE] oE] := (pnElemP Ep2E, card_pnElem Ep2E). have [[sZR_R nZR_R] [pE _ eE]] := (andP (center_normal R), and3P abelE). -have{nZR_R} nZR: R \subset 'N(Z) := char_norm_trans (Ohm_char 1 _) nZR_R. +have{nZR_R} nZR: R \subset 'N(Z) := gFnorm_trans _ nZR_R. have{sZR_R} [pZR pW] := (pgroupS sZR_R pR, pgroupS sWR pR). have sZE: Z \subset E by rewrite -defCRE1 OhmS ?setIS // centS. have rCRE : 'r_p('C_R(E)) = 2 by rewrite -p_rank_Ohm1 defCRE1 p_rank_abelem. @@ -228,8 +226,7 @@ have dimW: logn p #|W| = 2. by rewrite -divgS // logn_div ?cardSg // subn_gt0 properG_ltn_log. have abelW: p.-abelem W. by rewrite (abelem_Ohm1 (pgroupS _ pR)) ?(p2group_abelian pW) ?dimW ?ucn_sub. -have charT: T \char R. - by rewrite subcent_char ?char_refl //= (char_trans (Ohm_char 1 _)) ?ucn_char. +have charT: T \char R by rewrite subcent_char ?char_refl ?gFchar_trans. rewrite 2!inE sWR abelW dimW; do 2?split => //. by apply: contra (proper_subn ltZW); rewrite -defZ !subsetI subxx sER centsC. apply/prime_nt_dvdP=> //. @@ -281,7 +278,7 @@ have defST: S * T = R. have cRRb: abelian (R / T) by rewrite -defST quotientMidr quotient_abelian. have sR'T: R^`(1) \subset T by rewrite der1_min ?char_norm. have TI_SR': S :&: R^`(1) :=: 1. - by rewrite prime_TIg ?oS // (contra _ not_sST) //; move/subset_trans->. + by rewrite prime_TIg ?oS // (contra _ not_sST) // => /subset_trans->. have defCRS : S \x 'C_T(S) = 'C_R(S). rewrite (dprodE _ _) ?subsetIr //= -/T; last by rewrite setIA tiST setI1g. rewrite -{1}(center_idP cSS) subcent_TImulg ?defST //. @@ -459,7 +456,7 @@ have cycLb: cyclic (L / K) by rewrite prime_cyclic ?oLb. rewrite -(quotientSGK _ sKCX) // quotientGI // subsetI quotientS //= -/K. have actsXK: [acts X, on K | toX] by rewrite acts_ract subxx acts_char. rewrite ext_coprime_quotient_cent ?(pnat_coprime pK p'X) ?(pgroup_sol pK) //. -have actsAL : {acts A, on group L | [Aut R]} by exact: gacts_char. +have actsAL : {acts A, on group L | [Aut R]} by apply: gacts_char. have sAD: A \subset qact_dom <[actsAL]> [~: L, R]. by rewrite qact_domE // acts_actby subxx (setIidPr sKL) acts_char. suffices cLbX: X \subset 'C(L / K | <[actsAL]> / _). @@ -505,7 +502,7 @@ wlog Gp'1: gT G S oddG nnS solG sylS rS pl1G / 'O_p^'(G) = 1. have isoS := isog_symr (quotient_isog nKS tiKS). rewrite (isog_narrow p isoS) {isoS}(isog_rank isoS) quotient_pHall //. rewrite plength1_quo // trivg_pcore_quotient indexg1 /= -quotient_der //. - by rewrite card_quotient //= -/K -(card_isog (quotient1_isog _)); exact. + by rewrite card_quotient //= -/K -(card_isog (quotient1_isog _)); apply. rewrite Gp'1 indexg1 -(card_isog (quotient1_isog _)) -pgroupE. have [sSG pS _] := and3P sylS; have oddS: odd #|S| := oddSg sSG oddG. have ntS: S :!=: 1 by rewrite -rank_gt0 (leq_trans _ rS). @@ -522,13 +519,13 @@ have{defS} pKfA: p.-group ('ker fA). by rewrite -defS -Fitting_eq_pcore ?cent_sub_Fitting. split=> [|q]. rewrite -(pmorphim_pgroup pKfA) ?der_sub // morphim_der //. - by rewrite (pgroupS (der1_min (char_norm _) cAbAb)) ?pcore_pgroup ?pcore_char. -rewrite mem_primes; case/and3P=> q_pr _; case/Cauchy=> // x Gx ox. + by rewrite (pgroupS (der1_min _ cAbAb)) ?pcore_pgroup ?gFnorm. +rewrite mem_primes => /and3P[q_pr _ /Cauchy[] // x Gx ox]. rewrite leq_eqVlt -implyNb; apply/implyP=> p'q; rewrite -(ltn_predK p_gt1) ltnS. have ofAx: #[fA x] = q. apply/prime_nt_dvdP=> //; last by rewrite -ox morph_order. rewrite order_eq1; apply: contraNneq p'q => fAx1. - by apply: (pgroupP pKfA); rewrite // -ox order_dvdG //; exact/kerP. + by apply: (pgroupP pKfA); rewrite // -ox order_dvdG //; apply/kerP. have p'fAx: p^'.-elt (fA x) by rewrite /p_elt ofAx pnatE. by rewrite -ofAx dvdn_leq ?p'A_dv_p1 ?mem_morphim // -(subnKC p_gt1). Qed. diff --git a/mathcomp/odd_order/BGsection6.v b/mathcomp/odd_order/BGsection6.v index 234313c..f5323a2 100644 --- a/mathcomp/odd_order/BGsection6.v +++ b/mathcomp/odd_order/BGsection6.v @@ -1,8 +1,14 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq div fintype finset. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq div fintype finset. +From mathcomp Require Import prime fingroup morphism automorphism quotient gproduct gfunctor. +From mathcomp Require Import cyclic center commutator pgroup nilpotent sylow abelian hall. +From mathcomp Require Import maximal. +From mathcomp Require Import BGsection1 BGappendixAB. (******************************************************************************) @@ -51,8 +57,8 @@ Theorem Puig_center_p'core_normal p G S : odd #|G| -> solvable G -> p.-Sylow(G) S -> 'O_p^'(G) * 'Z('L(S)) <| G. Proof. move=> oddG solG sylS; rewrite -{2}(Puig_factorisation _ _ sylS) //. -have sZL_G := subset_trans (char_sub (center_Puig_char S)) (pHall_sub sylS). -rewrite -!quotientK ?(subset_trans _ (gFnorm _ _)) ?subsetIl //. +have sZL_G: 'Z('L(S)) \subset G by rewrite !gFsub_trans ?(pHall_sub sylS). +rewrite -!quotientK ?(subset_trans sZL_G) ?subIset ?gFnorm //=. by rewrite cosetpre_normal quotient_normal // normalSG. Qed. @@ -68,7 +74,7 @@ Lemma coprime_der1_sdprod K H G : Proof. case/sdprodP=> _ defG nKH tiKH coKH solK sKG'. set K' := K^`(1); have [sK'K nK'K] := andP (der_normal 1 K : K' <| K). -have nK'H: H \subset 'N(K') := char_norm_trans (der_char 1 K) nKH. +have nK'H: H \subset 'N(K') := gFnorm_trans _ nKH. set R := [~: K, H]; have sRK: R \subset K by rewrite commg_subl. have [nRK nRH] := joing_subP (commg_norm K H : K <*> H \subset 'N(R)). have sKbK'H': K / R \subset (K / R)^`(1) * (H / R)^`(1). @@ -98,7 +104,7 @@ Lemma prime_nil_der1_factor G : Proof. move=> nilG' /=; set G' := G^`(1); set p := #|G / G'| => p_pr. have nsG'G: G' <| G := der_normal 1 G; have [sG'G nG'G] := andP nsG'G. -have nsG'p'G: 'O_p^'(G') <| G := char_normal_trans (pcore_char _ _) nsG'G. +have nsG'p'G: 'O_p^'(G') <| G := gFnormal_trans _ nsG'G. have nG'p'G := normal_norm nsG'p'G; have solG' := nilpotent_sol nilG'. have{nilG'} pGb: p.-group (G / 'O_p^'(G')). rewrite /pgroup card_quotient -?(Lagrange_index sG'G (pcore_sub _ _)) //=. @@ -183,7 +189,7 @@ apply/eqP; rewrite eqEsubset mul_subG ?setISS ?cent_sub //=. apply/subsetP=> g /setIP[Gg /normP nHg]. have [|c Cc [u Uu defg]] := pprod_trans_coprime Gg; first by rewrite nHg. rewrite defg mem_mulg // !inE Uu -{2}nHg defg conjsgM conjSg (normP _) //=. -by case/setIP: Cc => _; exact: (subsetP (cent_sub H)). +by case/setIP: Cc => _; apply: (subsetP (cent_sub H)). Qed. End PprodSubCoprime. @@ -194,10 +200,10 @@ Variables (p : nat) (G S : {group gT}). Hypotheses (sylS : p.-Sylow(G) S) (pl1G : p.-length_1 G). Let K := 'O_p^'(G). Let sSG : S \subset G. Proof. by case/andP: sylS. Qed. -Let nsKG : K <| G. Proof. exact: pcore_normal. Qed. +Let nsKG : K <| G. Proof. apply: pcore_normal. Qed. Let sKG : K \subset G. Proof. by case/andP: nsKG. Qed. Let nKG : G \subset 'N(K). Proof. by case/andP: nsKG. Qed. -Let nKS : S \subset 'N(K). Proof. exact: subset_trans sSG nKG. Qed. +Let nKS : S \subset 'N(K). Proof. apply: subset_trans sSG nKG. Qed. Let coKS : coprime #|K| #|S|. Proof. exact: p'nat_coprime (pcore_pgroup _ G) (pHall_pgroup sylS). Qed. Let sSN : S \subset 'N_G(S). Proof. by rewrite subsetI sSG normG. Qed. diff --git a/mathcomp/odd_order/BGsection7.v b/mathcomp/odd_order/BGsection7.v index 9982283..f9db9f7 100644 --- a/mathcomp/odd_order/BGsection7.v +++ b/mathcomp/odd_order/BGsection7.v @@ -1,8 +1,14 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div fintype bigop. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div fintype bigop. +From mathcomp Require Import finset prime fingroup morphism automorphism action quotient. +From mathcomp Require Import gfunctor cyclic pgroup center commutator gseries nilpotent. +From mathcomp Require Import sylow abelian maximal hall. +From mathcomp Require Import BGsection1 BGsection6. (******************************************************************************) @@ -119,10 +125,10 @@ Lemma mFT_sol M : M \proper G -> solvable M. Proof. by case: gT M => ? []. Qed. Lemma mFT_nonAbelian : ~~ abelian G. -Proof. apply: contra mFT_nonSolvable; exact: abelian_sol. Qed. +Proof. by apply: contra mFT_nonSolvable; apply: abelian_sol. Qed. Lemma mFT_neq1 : G != 1. -Proof. by apply: contraNneq mFT_nonAbelian => ->; exact: abelian1. Qed. +Proof. by apply: contraNneq mFT_nonAbelian => ->; apply: abelian1. Qed. Lemma mFT_gt1 : [1] \proper G. Proof. by rewrite proper1G mFT_neq1. Qed. @@ -132,7 +138,7 @@ Proof. by rewrite quotient_odd ?mFT_odd. Qed. Lemma mFT_sol_proper M : (M \proper G) = solvable M. Proof. apply/idP/idP; first exact: mFT_sol. -by rewrite properT; apply: contraL; move/eqP->; exact: mFT_nonSolvable. +by rewrite properT; apply: contraL; move/eqP->; apply: mFT_nonSolvable. Qed. Lemma mFT_pgroup_proper p P : p.-group P -> P \proper G. @@ -162,7 +168,7 @@ exact: mFT_norm_proper. Qed. Lemma mFT_cent1_proper x : x != 1 -> 'C[x] \proper G. -Proof. by rewrite -cycle_eq1 -cent_cycle; exact: mFT_cent_proper. Qed. +Proof. by rewrite -cycle_eq1 -cent_cycle; apply: mFT_cent_proper. Qed. Lemma mFT_quo_sol M H : H :!=: 1 -> solvable (M / H). Proof. @@ -212,7 +218,7 @@ Proof. by move=> maxM sHM; apply: sub_proper_trans (mmax_proper maxM). Qed. Lemma mmax_norm X M : M \in 'M -> X :!=: 1 -> X \proper G -> M \subset 'N(X) -> 'N(X) = M. -Proof. by move=> maxM ntX prX; exact: mmax_max (mFT_norm_proper _ _). Qed. +Proof. by move=> maxM ntX prX; apply: mmax_max (mFT_norm_proper _ _). Qed. Lemma mmax_normal_subset A M : M \in 'M -> A <| M -> ~~ (A \subset [1]) -> 'N(A) = M. @@ -300,7 +306,7 @@ by move/eq_mmax=> ->. Qed. Lemma mmax_uniq_id : {subset 'M <= 'U}. -Proof. by move=> M maxM; apply/uniq_mmaxP; exists M; exact: mmax_sup_id. Qed. +Proof. by move=> M maxM; apply/uniq_mmaxP; exists M; apply: mmax_sup_id. Qed. Lemma def_uniq_mmaxJ M K x : 'M(K) = [set M] -> 'M(K :^ x) = [set M :^ x]%G. Proof. @@ -311,8 +317,8 @@ Qed. Lemma uniq_mmaxJ K x :((K :^ x)%G \in 'U) = (K \in 'U). Proof. apply/uniq_mmaxP/uniq_mmaxP=> [] [M uK_M]. - exists (M :^ x^-1)%G; rewrite -(conjsgK x K); exact: def_uniq_mmaxJ. -by exists (M :^ x)%G; exact: def_uniq_mmaxJ. + by exists (M :^ x^-1)%G; rewrite -(conjsgK x K); apply: def_uniq_mmaxJ. +by exists (M :^ x)%G; apply: def_uniq_mmaxJ. Qed. Lemma uniq_mmax_norm_sub (M U : {group gT}) : @@ -465,10 +471,9 @@ suffices: P \subset K. by case: eqP p_pr => // ->. suffices sP_pAC: P \subset 'O_pi^'(A <*> 'C(A)). rewrite (subset_trans sP_pAC) ?pcore_max ?pcore_pgroup //. - rewrite /normal (char_norm_trans (pcore_char _ _)) ?normsG ?joing_subr //. - rewrite andbT -quotient_sub1; last first. - rewrite (subset_trans (pcore_sub _ _)) // join_subG normG cents_norm //. - by rewrite centsC. + rewrite /normal gFnorm_trans ?normsG ?joing_subr // andbT. + rewrite -quotient_sub1; last first. + by rewrite gFsub_trans // join_subG !(normG, norms_cent). rewrite /= -(setIidPr (pcore_sub _ _)) quotientGI ?joing_subr //=. rewrite {1}cent_joinEr // quotientMidr coprime_TIg // coprime_morph //. by rewrite coprime_pi' ?cardG_gt0 //= -/pi [pnat _ _]pcore_pgroup. @@ -503,9 +508,8 @@ wlog defH: H prHG sAH ntHQ1 ntHQ2 / Q1 :&: Q2 != 1 -> H :=: 'N(Q1 :&: Q2). by rewrite subsetI normG (subsetIl, subsetIr). pose L := 'O_pi^'(H); have sLH: L \subset H := pcore_sub _ _. have [nLA coLA solL]: [/\ A \subset 'N(L), coprime #|L| #|A| & solvable L]. -- rewrite (char_norm_trans (pcore_char _ _)) ?normsG //. - rewrite coprime_sym coprime_pi' ?cardG_gt0 ?[pnat _ _]pcore_pgroup //. - by rewrite (solvableS sLH) ?mFT_sol. + rewrite gFnorm_trans ?normsG // coprime_sym coprime_pi' ?cardG_gt0 //. + by rewrite -pgroupE pcore_pgroup (solvableS sLH) ?mFT_sol. have Qsyl Q: Q \in |/|*(A; q) -> Q :&: H != 1 -> exists R : {group _}, [/\ q.-Sylow(L) R, A \subset 'N(R) & Q :&: H \subset R]. - case/mem_max_normed=> qQ nQA ntQH. @@ -525,8 +529,7 @@ have [Q3 maxQ3 sR2Q3] := max_normed_exists (pHall_pgroup sylR2) nR2A. have maxQ1h: (Q1 :^ h)%G \in |/|*(A; q) by rewrite actsKmax. case: (eqsVneq Q1 Q2) => [| neQ12]; first by exists 1; rewrite ?group1 ?conjsg1. have ntHQ3: Q3 :&: H != 1. - apply: contra ntHQ2; rewrite -!subG1; apply: subset_trans. - by rewrite subsetI subsetIr (subset_trans sQR2). + by apply: subG1_contra ntHQ2; rewrite subsetI subsetIr (subset_trans sQR2). have ntHQ1h: (Q1 :^ h) :&: H != 1. by move: ntHQ1; rewrite !trivg_card1 -(cardJg _ h) conjIg (conjGid Hh). suff [prI1 prI2]: Q1 :&: Q2 \proper Q1 :&: R1 /\ Q1 :&: Q2 \proper Q2 :&: R2. @@ -541,8 +544,8 @@ suff [prI1 prI2]: Q1 :&: Q2 \proper Q1 :&: R1 /\ Q1 :&: Q2 \proper Q2 :&: R2. case/(IHm H) => // k2 Kk2 defQ2; case/(IHm H) => // k3 Kk3 defQ3. by exists (h * k3 * k2); rewrite ?groupM ?conjsgM // -defQ3. case: (eqVneq (Q1 :&: Q2) 1) => [-> | ntQ12]. - rewrite !proper1G; split; [apply: contra ntHQ1 | apply: contra ntHQ2]; - by rewrite -!subG1; apply: subset_trans; rewrite subsetI subsetIl. + by rewrite !proper1G; split; [move: ntHQ1 | move: ntHQ2]; + apply: subG1_contra; rewrite subsetI subsetIl. rewrite -(setIidPr (subset_trans (pHall_sub sylR1) sLH)) setIA. rewrite -(setIidPr (subset_trans (pHall_sub sylR2) sLH)) setIA. rewrite (setIidPl sQR1) (setIidPl sQR2) {}defH //. @@ -591,18 +594,17 @@ case/existsP=> z; rewrite !inE => /and3P[ntzQ2 ntz Cz]. have prCz: 'C[z] \proper G by rewrite -cent_cycle mFT_cent_proper ?cycle_eq1. have sACz: A \subset 'C[z] by rewrite sub_cent1 (subsetP cAB) ?(subsetP sCB). have [|//|k Kk defQ2]:= normed_constrained_meet_trans sACz prCz maxQ1 maxQ2. - apply: contra ntCQ1; rewrite -!subG1; apply: subset_trans. - by rewrite setIS //= -cent_cycle centS ?cycle_subG. -exists k => //; exact: val_inj. + by apply: subG1_contra ntCQ1; rewrite setIS //= -cent_cycle centS ?cycle_subG. +by exists k => //; apply: val_inj. Qed. (* This is B & G, Theorem 7.3. *) Theorem normed_constrained_rank2_trans : q %| #|'C(A)| -> 'r('Z(A)) >= 2 -> [transitive K, on |/|*(A; q) | 'JG]. Proof. -move=> qC; case/rank_geP=> B; case/nElemP=> p; do 2![case/setIdP]. -rewrite subsetI; case/andP=> sBA cAB abelB mB2; have [_ cBB _] := and3P abelB. -have{abelB mB2} ncycB: ~~ cyclic B by rewrite (abelem_cyclic abelB) (eqP mB2). +move=> qC /rank_geP[B /nElemP[p /setIdP[/setIdP[/subsetIP[sBA cAB] abelB] oB]]]. +have [_ cBB _] := and3P abelB. +have{abelB oB} ncycB: ~~ cyclic B by rewrite (abelem_cyclic abelB) (eqP oB). have [R0 sylR0] := Sylow_exists q 'C(A); have [cAR0 qR0 _] := and3P sylR0. have nR0A: A \subset 'N(R0) by rewrite cents_norm // centsC. have{nR0A} [R maxR sR0R] := max_normed_exists qR0 nR0A. @@ -616,12 +618,12 @@ have ntQ: Q != 1%G. by apply: contra ntR => Q1; rewrite trivg_max_norm -(eqP Q1) // inE in maxR *. have ntRC: 'C_R(A) != 1. have sR0CR: R0 \subset 'C_R(A) by rewrite subsetI sR0R. - suffices: R0 :!=: 1 by rewrite -!proper1G; move/proper_sub_trans->. + suffices: R0 :!=: 1 by apply: subG1_contra. move: ntR; rewrite -!cardG_gt1 -(part_pnat_id qR) (card_Hall sylR0). - by rewrite !p_part_gt1 !mem_primes !cardG_gt0 qC; case/and3P=> ->. + by rewrite !p_part_gt1 !mem_primes !cardG_gt0 qC => /and3P[->]. have: [exists (z | 'C_Q[z] != 1), z \in B^#]. apply: contraR ntQ => trQ; have:= subset_trans sBA nQA. - rewrite -[_ == _]subG1; move/coprime_abelian_gen_cent1 <- => //; last first. + rewrite -[_ == _]subG1=> /coprime_abelian_gen_cent1 <- //; last first. by rewrite coprime_sym (coprimeSg sBA) ?coprime_pi' /pgroup ?(pi_pnat qQ). rewrite gen_subG; apply/bigcupsP=> z Cz; rewrite subG1. by apply: contraR trQ => ntCz; apply/existsP; exists z; rewrite ntCz. @@ -629,9 +631,9 @@ case/existsP=> z; rewrite 2!inE => /and3P[ntzQ ntz Bz]. have prCz: 'C[z] \proper G by rewrite -cent_cycle mFT_cent_proper ?cycle_eq1. have sACz: A \subset 'C[z] by rewrite sub_cent1 (subsetP cAB). have [|//|k Kk defQ2]:= normed_constrained_meet_trans sACz prCz maxR maxQ. - apply: contra ntRC; rewrite -!subG1; apply: subset_trans. - by rewrite setIS //= -cent_cycle centS // cycle_subG (subsetP sBA). -exists k => //; exact: val_inj. + apply: subG1_contra ntRC; rewrite setIS //=. + by rewrite -cent_cycle centS // cycle_subG (subsetP sBA). +by exists k => //; apply: val_inj. Qed. (* This is B & G, Theorem 7.4. *) @@ -644,12 +646,11 @@ Theorem normed_trans_superset P : /\ 'N(P) = 'C_K(P) * 'N_('N(P))(Q)}]. Proof. move=> snAP piP trnK; set KP := 'O_pi^'('C(P)). -have defK: forall B, A \subset B -> 'C_K(B) = 'O_pi^'('C(B)). - move=> B sAB; apply/eqP; rewrite eqEsubset {1}setIC pcoreS ?centS //. - rewrite subsetI pcore_sub (sub_Hall_pcore hallK) ?pcore_pgroup //. - by rewrite (subset_trans (pcore_sub _ _)) ?centS. +have defK B: A \subset B -> 'C_K(B) = 'O_pi^'('C(B)). + move=> sAB; apply/eqP; rewrite eqEsubset {1}setIC pcoreS ?centS // subsetI. + by rewrite gFsub (sub_Hall_pcore hallK) ?pcore_pgroup // gFsub_trans ?centS. suffices: [transitive KP, on |/|*(P; q) | 'JG] /\ |/|*(P; q) \subset |/|*(A; q). - have nsKPN: KP <| 'N(P) := char_normal_trans (pcore_char _ _) (cent_normal _). + have nsKPN: KP <| 'N(P) := gFnormal_trans _ (cent_normal _). case=> trKP smnPA; rewrite (defK _ (subnormal_sub snAP)); split=> // Q maxQ. have defNP: KP * 'N_('N(P))(Q) = 'N(P). rewrite -(astab1JG Q) -normC; last by rewrite subIset 1?normal_norm. @@ -667,7 +668,7 @@ wlog{snAP} [B maxnB snAB]: / {B : grT | maxnormal B P P & A <|<| B}. apply; exists B => //; apply: subnormal_trans snAD (normal_subnormal _). by apply: normalS sDB _ nDP; case/andP: (maxgroupp maxnB); case/andP. have [prBP nBP] := andP (maxgroupp maxnB); have sBP := proper_sub prBP. -have{lePm}: #|B| < m by exact: leq_trans (proper_card prBP) _. +have{lePm}: #|B| < m by apply: leq_trans (proper_card prBP) _. case/IHm=> {IHm}// [|trnB smnBA]; first by rewrite (pgroupS sBP). have{maxnB} abelPB: is_abelem (P / B). apply: charsimple_solvable (maxnormal_charsimple _ maxnB) _ => //. @@ -678,7 +679,7 @@ have{abelPB} [p p_pr pPB]: exists2 p, prime p & p.-group (P / B). have{prBP} pi_p: p \in pi. case/pgroup_pdiv: pPB => [|_ pPB _]. by rewrite -subG1 quotient_sub1 // proper_subn. - by apply: pgroupP p_pr pPB; exact: quotient_pgroup. + by apply: pgroupP p_pr pPB; apply: quotient_pgroup. pose S := |/|*(B; q); have p'S: #|S| %% p != 0. have pi'S: pi^'.-nat #|S| := pnat_dvd (atrans_dvd trnB) (pcore_pgroup _ _). by rewrite -prime_coprime // (pnat_coprime _ pi'S) ?pnatE. @@ -717,13 +718,13 @@ have smnP_S: |/|*(P; q) \subset S. have qNQ2: q.-group 'N_Q2(Q1) by rewrite (pgroupS _ qQ2) ?subsetIl. pose KN := 'O_pi^'('N(Q1)); have sNQ2_KN: 'N_Q2(Q1) \subset KN. rewrite hyp71 // inE normsI ?norms_norm ?(subset_trans sAB nQ2B) //=. - by rewrite /psubgroup subsetIr andbT; exact: pi_pnat qNQ2 _. + by rewrite /psubgroup subsetIr andbT; apply: pi_pnat qNQ2 _. rewrite -Sylow_subnorm (pHall_subl _ sNQ2_KN) ?subsetI ?sQ12 ?normG //= -/KN. suff: exists Q3 : grT, [/\ q.-Sylow(KN) Q3, P \subset 'N(Q3) & Q1 \subset Q3]. move: maxQ1; rewrite inE; case/maxgroupP=> _ maxQ1 [Q3 [sylQ3 nQ3P sQ13]]. by rewrite -(maxQ1 Q3) // (pHall_pgroup sylQ3). apply: coprime_Hall_subset; rewrite //= -/KN. - - by rewrite (char_norm_trans (pcore_char _ _)) ?norms_norm. + - by rewrite gFnorm_trans ?norms_norm. - by rewrite coprime_sym (pnat_coprime piP (pcore_pgroup _ _)). - by rewrite (solvableS (pcore_sub _ _)) ?mFT_sol. by rewrite pcore_max ?normalG // /pgroup (pi_pnat qQ1). @@ -732,9 +733,8 @@ apply/imsetP; exists Q0 => //; apply/setP=> Q2. apply/idP/imsetP=> [maxQ2 | [k Pk ->]]; last by rewrite (actsP actsKmnP). have [S_Q0 S_Q2]: Q0 \in S /\ Q2 \in S by rewrite !(subsetP smnP_S). pose KB := 'O_pi^'('C(B)); pose KBP := KB <*> P. -have pi'KB: pi^'.-group KB by exact: pcore_pgroup. -have nKB_P: P \subset 'N(KB). - by rewrite (char_norm_trans (pcore_char _ _)) ?norms_cent. +have pi'KB: pi^'.-group KB by apply: pcore_pgroup. +have nKB_P: P \subset 'N(KB) by rewrite gFnorm_trans ?norms_cent. have [k KBk defQ2]:= atransP2 trnB S_Q0 S_Q2. have [qQ2 nQ2P] := mem_max_normed maxQ2. have hallP: pi.-Hall('N_KBP(Q2)) P. @@ -799,14 +799,14 @@ have [p_pr pdvA [r oApr]] := pgroup_pdiv pA ntA. have{r oApr} def_pi: pi =i (p : nat_pred). by move=> p'; rewrite !inE oApr primes_exp // primes_prime ?inE. have def_pi' := eq_negn def_pi; have defK := eq_pcore _ def_pi'. -pose Z := 'Ohm_1('Z(P)); have sZ_ZP: Z \subset 'Z(P) by exact: Ohm_sub. +pose Z := 'Ohm_1('Z(P)); have sZ_ZP: Z \subset 'Z(P) by apply: Ohm_sub. have sZP_A: 'Z(P) \subset A by rewrite -defCA setIS ?centS. have sZA := subset_trans sZ_ZP sZP_A. -have nsA1: 'Ohm_1(A) <| P by exact: (char_normal_trans (Ohm_char _ _)). +have nsA1: 'Ohm_1(A) <| P by apply: gFnormal_trans. pose inZor1 B := B \subset Z \/ #|Z| = p /\ Z \subset B. have [B [E2_B nsBP sBZ]]: exists B, [/\ B \in 'E_p^2(A), B <| P & inZor1 B]. - have pZP: p.-group 'Z(P) by exact: pgroupS (center_sub _) pP. - have pZ: p.-group Z by exact: pgroupS sZ_ZP pZP. + have pZP: p.-group 'Z(P) by apply: pgroupS (center_sub _) pP. + have pZ: p.-group Z by apply: pgroupS sZ_ZP pZP. have abelZ: p.-abelem Z by rewrite Ohm1_abelem ?center_abelian. have nsZP: Z <| P := sub_center_normal sZ_ZP; have [sZP nZP] := andP nsZP. case: (eqVneq Z 1). @@ -890,7 +890,7 @@ wlog Zb: b X Y defX B'b p'Y nYA sYX / b \in Z. case: (primeP p_pr) => _ dv_p; move/dv_p=> {dv_p}. case/pred2P=> oP21; first by rewrite -(index1g sP12 oP21) normal_refl. by rewrite (p_maximal_normal pP2) ?p_index_maximal ?oP21. - have nsZP1_2: 'Z(P1) <| P2 by rewrite (char_normal_trans (center_char _)). + have nsZP1_2: 'Z(P1) <| P2 by rewrite gFnormal_trans. have sZKp: Z \subset 'O_{p^', p}(X). suff: 'Z(P1) \subset 'O_{p^', p}(X). apply: subset_trans; rewrite subsetI {1}defP1 (subset_trans sZB). @@ -904,8 +904,7 @@ wlog Zb: b X Y defX B'b p'Y nYA sYX / b \in Z. have <-: [~: Y, Z] * 'C_Y(Z) = Y. exact: coprime_cent_prod (solvableS sYX solX). set K := 'O_p^'(X); have [nKY nKZ]: Y \subset 'N(K) /\ Z \subset 'N(K). - rewrite !(char_norm_trans (pcore_char _ _)) ?(subset_trans sZA) ?normsG //. - by rewrite -defX cBA. + by rewrite !gFnorm_trans ?(subset_trans sZA) ?normsG // -defX cBA. rewrite mul_subG //. have coYZK: coprime #|Y / K| #|'O_p(X / K)|. by rewrite coprime_sym coprime_morphr ?(pnat_coprime (pcore_pgroup _ _)). @@ -923,14 +922,14 @@ wlog Zb: b X Y defX B'b p'Y nYA sYX / b \in Z. - rewrite !inE -cycle_eq1 -defZ trivg_card_le1 oZ -ltnNge prime_gt1 //=. by rewrite (subsetP sZB). by rewrite normsI // norms_cent // cents_norm // centsC (subset_trans sZA). -set K := 'O_p^'(X); have nsKX: K <| X by exact: pcore_normal. +set K := 'O_p^'(X); have nsKX: K <| X by apply: pcore_normal. case/setD1P: B'b => ntb Bb. have [sAX solX]: A \subset X /\ solvable X by rewrite -defX cBA ?solCB. have sPX: P \subset X. by rewrite -defX -cent_set1 centsC sub1set; case/setIP: (subsetP sZ_ZP b Zb). have [nKA nKY nKP]: [/\ A \subset 'N(K), Y \subset 'N(K) & P \subset 'N(K)]. by rewrite !(subset_trans _ (normal_norm nsKX)). -have sylPX: p.-Sylow(X) P by exact: pHall_subl (subsetT _) sylP. +have sylPX: p.-Sylow(X) P by apply: pHall_subl (subsetT _) sylP. have sAKb: A \subset 'O_{p^', p}(X). exact: (odd_p_abelian_constrained (mFT_odd _)) abA nsAP. have coYZK: coprime #|Y / K| #|'O_p(X / K)|. @@ -942,8 +941,8 @@ have cYAq: A / K \subset 'C_('O_p(X / K))(Y / K). by rewrite commg_subl /= gFnorm. have cYKq: Y / K \subset 'C('O_p(X / K)). apply: coprime_nil_faithful_cent_stab => /=. - - by rewrite (char_norm_trans (pcore_char _ _)) ?normsG ?quotientS. - - by rewrite coprime_morphr ?(pnat_coprime (pcore_pgroup _ _)). + - by rewrite gFnorm_trans ?normsG ?quotientS. + - by rewrite coprime_sym. - exact: pgroup_nil (pcore_pgroup _ _). apply: subset_trans (cYAq); rewrite -defCA -['C_P(A) / K](morphim_restrm nKP). rewrite injm_cent ?ker_restrm ?ker_coset ?morphim_restrm -?quotientE //. diff --git a/mathcomp/odd_order/BGsection8.v b/mathcomp/odd_order/BGsection8.v index 8e306fa..60d8a67 100644 --- a/mathcomp/odd_order/BGsection8.v +++ b/mathcomp/odd_order/BGsection8.v @@ -1,7 +1,12 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div fintype path. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div fintype path. +From mathcomp Require Import finset prime fingroup automorphism action gproduct gfunctor. +From mathcomp Require Import center commutator pgroup gseries nilpotent sylow abelian maximal. +From mathcomp Require Import BGsection1 BGsection5 BGsection6 BGsection7. (******************************************************************************) @@ -41,17 +46,16 @@ have nilF: nilpotent F := Fitting_nil _. have nilZ := nilpotentS (center_sub _) nilF. have piZ: \pi('Z(F)) = \pi(F) by rewrite pi_center_nilpotent. have def_pi: pi = \pi(F). - by apply/eq_piP=> q; apply/idP/idP; last rewrite -piZ; exact: piSg. + by apply/eq_piP=> q; apply/idP/idP; last rewrite -piZ; apply: piSg. have def_nZq: forall q, q \in pi -> 'N('Z(F)`q) = M. move=> q; rewrite def_pi -piZ -p_part_gt1. rewrite -(card_Hall (nilpotent_pcore_Hall _ nilZ)) cardG_gt1 /= -/F => ntZ. - apply: mmax_normal => //=; apply: char_normal_trans (Fitting_normal _). - exact: char_trans (pcore_char _ _) (center_char _). + by apply: mmax_normal => //=; rewrite !gFnormal_trans. have sCqM: forall q, q \in pi -> 'C(A`q) \subset M. - move=> q; move/def_nZq <-; rewrite cents_norm // centS //. + move=> q /def_nZq <-; rewrite cents_norm // centS //. rewrite (sub_Hall_pcore (nilpotent_pcore_Hall _ _)) ?pcore_pgroup //. - by apply: nilpotentS (Fitting_nil M); exact: subsetIl. - exact: subset_trans (pcore_sub _ _) _. + by apply: nilpotentS (Fitting_nil M); apply: subsetIl. + exact: gFsub_trans. have sA0A: A0 \subset A by rewrite subsetI sA0F. have pi_p: p \in pi. by apply: (piSg sA0A); rewrite -[p \in _]logn_gt0 (leq_trans _ dimA0_3). @@ -76,33 +80,31 @@ have{p'F} pi_alt q: exists2 r, r \in pi & r != q. by apply/pgroupP=> q q_pr qF; rewrite !inE pF // mem_primes q_pr cardG_gt0. have sNZqXq' q X: A \subset X -> X \proper G -> 'O_q^'('N_X('Z(F)`q)) \subset 'O_q^'(X). -- move=> sAX prX; have sZqX: 'Z(F)`q \subset X. - exact: subset_trans (pcore_sub _ _) (subset_trans sZA sAX). +- move=> sAX prX. + have sZqX: 'Z(F)`q \subset X by apply: gFsub_trans (subset_trans sZA sAX). have cZqNXZ: 'O_q^'('N_X('Z(F)`q)) \subset 'C('Z(F)`q). have coNq'Zq: coprime #|'O_q^'('N_X('Z(F)`q))| #|'Z(F)`q|. by rewrite coprime_sym coprime_pcoreC. rewrite (sameP commG1P trivgP) -(coprime_TIg coNq'Zq) subsetI commg_subl /=. - rewrite commg_subr /= andbC (subset_trans (pcore_sub _ _)) ?subsetIr //=. - by rewrite (char_norm_trans (pcore_char _ _)) ?normsG // subsetI sZqX normG. + rewrite commg_subr /= andbC gFsub_trans ?subsetIr //=. + by rewrite gFnorm_trans ?normsG // subsetI sZqX normG. have: 'O_q^'('C_X(('Z(F))`q)) \subset 'O_q^'(X). by rewrite p'core_cent_pgroup ?mFT_sol // /psubgroup sZqX pcore_pgroup. apply: subset_trans; apply: subset_trans (pcoreS _ (subcent_sub _ _)). - by rewrite !subsetI subxx cZqNXZ (subset_trans (pcore_sub _ _)) ?subsetIl. + by rewrite !subsetI subxx cZqNXZ gFsub_trans ?subsetIl. have sArXq' q r X: q \in pi -> q != r -> A \subset X -> X \proper G -> A`r \subset 'O_q^'(X). - move=> pi_q r'q sAX prX; apply: subset_trans (sNZqXq' q X sAX prX). apply: subset_trans (pcoreS _ (subsetIr _ _)). - rewrite -setIA (setIidPr (pcore_sub _ _)) subsetI. - rewrite (subset_trans (pcore_sub _ _)) //= def_nZq //. + rewrite -setIA (setIidPr (pcore_sub _ _)) subsetI gFsub_trans //= def_nZq //. apply: subset_trans (pcore_Fitting _ _); rewrite -/F. - rewrite (sub_Hall_pcore (nilpotent_pcore_Hall _ nilF)) //; last first. - exact: subset_trans (pcore_sub _ _) sAF. + rewrite (sub_Hall_pcore (nilpotent_pcore_Hall _ nilF)) ?gFsub_trans //. by apply: (pi_pnat (pcore_pgroup _ _)); rewrite !inE eq_sym. have cstrA: normed_constrained A. split=> [||X Y sAX prX]. - by apply/eqP=> A1; rewrite /pi /= A1 cards1 in pi_p. - exact: sub_proper_trans (subset_trans sAF (Fitting_sub _)) prM. - rewrite !inE -/pi -andbA; case/and3P=> sYX pi'Y nYA. + rewrite !inE -/pi -andbA => /and3P[sYX pi'Y nYA]. rewrite -bigcap_p'core subsetI sYX; apply/bigcapsP=> [[q /= _] pi_q]. have [r pi_r q'r] := pi_alt q. have{sArXq'} sArXq': A`r \subset 'O_q^'(X) by apply: sArXq'; rewrite 1?eq_sym. @@ -116,13 +118,12 @@ have cstrA: normed_constrained A. rewrite -(setIid Y) setIAC coprime_TIg // (coprimeSg cA_CYr) //. by rewrite (pnat_coprime piCA). have{CYr1} ->: Y :=: [~: Y, A`r]. - rewrite -(mulg1 [~: Y, _]) -CYr1 coprime_cent_prod //. - - by rewrite (subset_trans (pcore_sub _ _)). - - rewrite coprime_sym (coprimeSg (pcore_sub _ _)) //= -/A. + rewrite -(mulg1 [~: Y, _]) -CYr1 coprime_cent_prod ?gFsub_trans //. + rewrite coprime_sym (coprimeSg (pcore_sub _ _)) //= -/A. by rewrite coprime_pi' ?cardG_gt0. by rewrite mFT_sol // (sub_proper_trans sYX). - rewrite (subset_trans (commgS _ sArXq')) // commg_subr. - by rewrite (char_norm_trans (pcore_char _ _)) ?normsG. + rewrite (subset_trans (commgS _ sArXq')) //. + by rewrite commg_subr gFnorm_trans ?normsG. have{cstrA} nbyApi'1 q: q \in pi^' -> |/|*(A; q) = [set 1%G]. move=> pi'q; have trA: [transitive 'O_pi^'('C(A)), on |/|*(A; q) | 'JG]. apply: normed_constrained_rank3_trans; rewrite //= -/A. @@ -144,7 +145,7 @@ have{cstrA} nbyApi'1 q: q \in pi^' -> |/|*(A; q) = [set 1%G]. have nQM: M \subset 'N(Q). apply/normsP=> x Mx; apply: congr_group; apply/set1P. rewrite -defFmax (acts_act (norm_acts_max_norm _ _)) ?defFmax ?set11 //. - by apply: subsetP Mx; exact: gFnorm. + by apply: subsetP Mx; apply: gFnorm. have{nQM} nsQM: Q <| M. rewrite inE in maxM; case/maxgroupP: maxM => _ maxM. rewrite -(maxM 'N(Q)%G) ?normalG ?mFT_norm_proper //. @@ -161,8 +162,7 @@ have piD: \pi(D) = pi. move=> q; rewrite -p_part_gt1 -card_pcore_nil // cardG_gt1 /= -/D. apply: contraR => /nbyApi'1 defAmax. have nDqA: A \subset 'N(D`q). - rewrite (char_norm_trans (pcore_char _ _)) //. - by rewrite (subset_trans sAH) ?gFnorm. + by rewrite gFnorm_trans // (subset_trans sAH) ?gFnorm. have [Q]:= max_normed_exists (pcore_pgroup _ _) nDqA. by rewrite defAmax -subG1; move/set1P->. apply/eq_piP=> q; apply/idP/idP=> [|pi_q]; first exact: pi_sig. @@ -171,20 +171,19 @@ have piD: \pi(D) = pi. have <-: 'O_sigma^'(H) = 1. apply/eqP; rewrite -trivg_Fitting ?(solvableS (pcore_sub _ _)) //. rewrite Fitting_pcore -(setIidPr (pcore_sub _ _)) coprime_TIg //. - by rewrite coprime_pi' ?cardG_gt0 //; exact: pcore_pgroup. - rewrite -bigcap_p'core subsetI (subset_trans (pcore_sub _ _)) //=. - apply/bigcapsP=> [[r /= _] sig_r]; apply: sArXq' => //; first exact: pi_sig. - by apply: contra sig'q; move/eqP <-. + by rewrite coprime_pi' ?cardG_gt0 //; apply: pcore_pgroup. + rewrite -bigcap_p'core subsetI gFsub_trans //=. + apply/bigcapsP=> -[r /= _] sig_r; apply: sArXq' => //; first exact: pi_sig. + by apply: contraNneq sig'q => <-. have cAD q r: q != r -> D`q \subset 'C(A`r). move=> r'q; have [-> |] := eqVneq D`q 1; first by rewrite sub1G. rewrite -cardG_gt1 card_pcore_nil // p_part_gt1 piD => pi_q. have sArHq': A`r \subset 'O_q^'(H) by rewrite sArXq'. have coHqHq': coprime #|D`q| #|'O_q^'(H)| by rewrite coprime_pcoreC. rewrite (sameP commG1P trivgP) -(coprime_TIg coHqHq') commg_subI //. - rewrite subsetI subxx /= p_core_Fitting (subset_trans (pcore_sub _ _)) //. - exact: gFnorm. - rewrite subsetI sArHq' (subset_trans (subset_trans (pcore_sub _ _) sAH)) //. - by rewrite /= p_core_Fitting gFnorm. + by rewrite subsetI subxx /= p_core_Fitting gFsub_trans ?gFnorm. + rewrite subsetI sArHq' gFsub_trans ?(subset_trans sAH) //=. + by rewrite p_core_Fitting gFnorm. have sDM: D \subset M. rewrite [D]FittingEgen gen_subG; apply/bigcupsP=> [[q /= _] _]. rewrite -p_core_Fitting -/D; have [r pi_r r'q] := pi_alt q. @@ -194,7 +193,7 @@ have cApHp': A`p \subset 'C('O_p^'(H)). by rewrite coprime_sym coprime_pcoreC. have solHp': solvable 'O_p^'(H) by rewrite (solvableS (pcore_sub _ _)). have nHp'Ap: A`p \subset 'N('O_p^'(H)). - by rewrite (subset_trans (subset_trans (pcore_sub _ _) sAH)) ?gFnorm. + by rewrite gFsub_trans ?gFnorm_trans ?normsG. apply: subset_trans (coprime_cent_Fitting nHp'Ap coApHp' solHp'). rewrite subsetI subxx centsC /= FittingEgen gen_subG. apply/bigcupsP=> [[q /= _] _]; have [-> | /cAD] := eqVneq q p. @@ -211,14 +210,14 @@ have sHp'_NMDp': 'O_p^'(H) \subset 'O_p^'('N_M(D`p)). by rewrite /= p_core_Fitting pcore_normal. have{sHp'_NMDp'} sHp'Mp': 'O_p^'(H) \subset 'O_p^'(M). have pM_D: p.-subgroup(M) D`p. - by rewrite /psubgroup pcore_pgroup (subset_trans (pcore_sub _ _)). + by rewrite /psubgroup pcore_pgroup gFsub_trans. apply: subset_trans (p'core_cent_pgroup pM_D (mFT_sol prM)). apply: subset_trans (pcoreS _ (subcent_sub _ _)). rewrite !subsetI sHp'_NMDp' sHp'M andbT /= (sameP commG1P trivgP). have coHp'Dp: coprime #|'O_p^'(H)| #|D`p|. by rewrite coprime_sym coprime_pcoreC. rewrite -(coprime_TIg coHp'Dp) subsetI commg_subl commg_subr /=. - by rewrite p_core_Fitting !(subset_trans (pcore_sub _ _)) ?gFnorm. + by rewrite p_core_Fitting !gFsub_trans ?gFnorm. have sMp'H: 'O_p^'(M) \subset H. rewrite -(mmax_normal maxH (pcore_normal p H)) /= -p_core_Fitting //. rewrite -/D (subset_trans _ (cent_sub _)) // centsC. @@ -226,7 +225,7 @@ have sMp'H: 'O_p^'(M) \subset H. have coMp'Dp: coprime #|'O_p^'(M)| #|D`p|. by rewrite coprime_sym coprime_pcoreC. have nMp'Dp: D`p \subset 'N('O_p^'(M)). - by rewrite (subset_trans (subset_trans (pcore_sub _ _) sDM)) ?gFnorm. + by rewrite gFsub_trans ?(subset_trans sDM) ?gFnorm. apply: subset_trans (coprime_cent_Fitting nMp'Dp coMp'Dp solMp'). rewrite subsetI subxx centsC /= FittingEgen gen_subG. apply/bigcupsP=> [[q /= _] _]; have [<- | /cAD] := eqVneq p q. @@ -234,7 +233,7 @@ have sMp'H: 'O_p^'(M) \subset H. rewrite centsC; apply: subset_trans. rewrite -p_core_Fitting Fitting_pcore pcore_max ?pcore_pgroup //=. rewrite /normal subsetI -pcoreI pcore_sub subIset ?gFnorm //=. - rewrite pcoreI (subset_trans (pcore_sub _ _)) //= -/F centsC. + rewrite pcoreI gFsub_trans //= -/F centsC. case/dprodP: (nilpotent_pcoreC p nilF) => _ _ /= cFpp' _. rewrite centsC (subset_trans cFpp' (centS _)) //. have hallFp := nilpotent_pcore_Hall p nilF. @@ -244,8 +243,7 @@ have{sHp'Mp' sMp'H} eqHp'Mp': 'O_p^'(H) = 'O_p^'(M). apply: subset_trans (sNZqXq' p H sAH prH). apply: subset_trans (pcoreS _ (subsetIr _ _)). rewrite -setIA (setIidPr (pcore_sub _ _)) subsetI sMp'H /=. - rewrite (mmax_normal maxM (char_normal_trans (pcore_char _ _) _)) //. - by rewrite (char_normal_trans (center_char _)) ?Fitting_normal. + rewrite (mmax_normal maxM) ?gFnormal_trans //. by rewrite -cardG_gt1 card_pcore_nil // p_part_gt1 piZ -def_pi. have ntHp': 'O_p^'(H) != 1. have [q pi_q p'q] := pi_alt p; have: D`q \subset 'O_p^'(H). @@ -280,7 +278,7 @@ have sAF: A \subset F. have sZA: 'Z(F) \subset A. by rewrite -defCA setISS ?centS // defF pcore_sub_Hall. have sCAM: 'C(A) \subset M. - have nsZM: 'Z(F) <| M := char_normal_trans (center_char _) (Fitting_normal _). + have nsZM: 'Z(F) <| M by rewrite !gFnormal_trans. rewrite -(mmax_normal maxM nsZM); last first. rewrite /= -(setIidPr (center_sub _)) meet_center_nil ?Fitting_nil //. by rewrite -proper1G (proper_sub_trans _ sAF) ?proper1G. @@ -288,18 +286,17 @@ have sCAM: 'C(A) \subset M. have nsZL_M: 'Z('L(P)) <| M. by rewrite (Puig_center_normal (mFT_odd _) solM sylP). have sNPM: 'N(P) \subset M. - rewrite -(mmax_normal maxM nsZL_M). - by rewrite (char_norm_trans (center_Puig_char P)). - apply/eqP; move/(trivg_center_Puig_pgroup (pHall_pgroup sylP))=> P1. + rewrite -(mmax_normal maxM nsZL_M) ?gFnorm_trans //. + apply/eqP => /(trivg_center_Puig_pgroup (pHall_pgroup sylP))-P1. by rewrite -subG1 -P1 sAP in ntA. have sylPG: p.-Sylow(G) P := mmax_sigma_Sylow maxM sylP sNPM. split; rewrite // (uniq_mmax_subset1 maxM sAM). have{scn3_A} scn3_A: A \in 'SCN_3[p] by apply/bigcupP; exists P; rewrite // inE. pose K := 'O_p^'('C(A)); have sKF: K \subset F. - have sKM: K \subset M := subset_trans (pcore_sub _ _) sCAM. + have sKM: K \subset M := gFsub_trans _ sCAM. apply: subset_trans (cent_sub_Fitting solM). rewrite subsetI sKM coprime_nil_faithful_cent_stab ?Fitting_nil //. - - by rewrite (subset_trans (subset_trans (pcore_sub _ _) sCAM)) ?gFnorm. + - by rewrite gFsub_trans ?(subset_trans sCAM) ?gFnorm. - by rewrite /= -/F defF coprime_pcoreC. have sACK: A \subset 'C_F(K) by rewrite subsetI sAF centsC pcore_sub. by rewrite /= -/F -/K (subset_trans _ sACK) //= -defCA setISS ?centS. @@ -331,14 +328,13 @@ have{p'nbyA_1} p'nbyA_1 X: apply/bigcupsP=> [[q /= _] _]; have [-> | p'q] := eqVneq q p. rewrite -(setIidPl (pcore_sub _ _)) coprime_TIg //. by rewrite (pnat_coprime (pcore_pgroup _ _)). - have [|R] := max_normed_exists (pcore_pgroup q X) (char_norm_trans _ nXA). - exact: pcore_char. + have [R] := max_normed_exists (pcore_pgroup q X) (gFnorm_trans _ nXA). by rewrite p'nbyA_1 // => /set1P->. -apply/subsetPn=> [[H0 MA_H0 neH0M]]. -have:= erefl [arg max_(H > H0 | (H \in 'M(A)) && (H != M)) #|H :&: M|`_p]. -case: arg_maxP => [|H {H0 MA_H0 neH0M}]; first by rewrite MA_H0 -in_set1. -rewrite /= inE -andbA; case/and3P=> maxH sAH neHM maxHM _. -have prH: H \proper G by rewrite inE in maxH; exact: maxgroupp maxH. +apply/subsetPn=> -[H0 MA_H0 neH0M]. +pose H := [arg max_(H > H0 | (H \in 'M(A)) && (H != M)) #|H :&: M|`_p]. +case: arg_maxP @H => [|H {H0 MA_H0 neH0M}]; first by rewrite MA_H0 -in_set1. +rewrite /= inE -andbA => /and3P[maxH sAH neHM] maxHM. +have prH: H \proper G by rewrite inE in maxH; apply: maxgroupp maxH. have sAHM: A \subset H :&: M by rewrite subsetI sAH. have [R sylR_HM sAR]:= Sylow_superset sAHM (pgroupS sAP pP). have [/subsetIP[sRH sRM] pR _] := and3P sylR_HM. @@ -370,7 +366,7 @@ have ntZLR: 'Z('L(R)) != 1. have defH: 'N('Z('L(R))) = H := mmax_normal maxH nsZLR_H ntZLR. have{sylR_H} sylR: p.-Sylow(G) R. rewrite -Sylow_subnorm setTI (pHall_subl _ _ sylR_H) ?normG //=. - by rewrite -defH (char_norm_trans (center_Puig_char R)). + by rewrite -defH !gFnorm_trans. have nsZLR_M: 'Z('L(R)) <| M. have sylR_M := pHall_subl sRM (subsetT _) sylR. exact: Puig_center_normal (mFT_odd _) solM sylR_M _. @@ -389,7 +385,7 @@ have [pF | npF] := boolP (p.-group 'F(M)). rewrite (p_rank_Sylow sylP) (leq_trans dimF3) //. by rewrite p_rankS ?Fitting_sub. have [A] := rank3_SCN3 pP (mFT_odd _) dimP3. - by case/(SCN_Fitting_Uniqueness maxM pF)=> // _ sAF; exact: uniq_mmaxS. + by case/(SCN_Fitting_Uniqueness maxM pF)=> // _ sAF; apply: uniq_mmaxS. case/p_rank_geP: dimF3 => A /setIdP[EpA dimA3]. have [A0 maxA0 sAA0] := @maxgroup_exists _ [pred X in 'E_p('F(M))] _ EpA. have [_ abelA] := pElemP EpA; have pmaxA0: A0 \in 'E*_p('F(M)) by rewrite inE. diff --git a/mathcomp/odd_order/BGsection9.v b/mathcomp/odd_order/BGsection9.v index dba9344..2153024 100644 --- a/mathcomp/odd_order/BGsection9.v +++ b/mathcomp/odd_order/BGsection9.v @@ -1,9 +1,16 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq div fintype path. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div fintype path. +From mathcomp Require Import finset prime fingroup action automorphism quotient cyclic. +From mathcomp Require Import gproduct gfunctor pgroup center commutator gseries nilpotent. +From mathcomp Require Import sylow abelian maximal hall. +From mathcomp Require Import BGsection1 BGsection4 BGsection5 BGsection6. +From mathcomp Require Import BGsection7 BGsection8. (******************************************************************************) @@ -46,7 +53,7 @@ have sHp'M: 'O_p^'(H) \subset M. apply: subset_trans snbBp'_M; rewrite (bigcup_max 'O_p^'(H)%G) // inE -andbA. by rewrite subsetT pcore_pgroup (subset_trans sBH) ?gFnorm. have{snbBp'_M} defMp': <<\bigcup_(K in |/|_G(P; p^')) K>> = 'O_p^'(M). - have nMp'M: M \subset 'N('O_p^'(M)) by exact: gFnorm. + have nMp'M: M \subset 'N('O_p^'(M)) by apply: gFnorm. have nMp'P := subset_trans sPM nMp'M. apply/eqP; rewrite eqEsubset gen_subG sub_gen ?andbT; last first. by rewrite (bigcup_max 'O_p^'(M)%G) // inE -andbA subsetT pcore_pgroup. @@ -70,9 +77,8 @@ have{defMp'} sNPM: 'N(P) \subset M. have [Mp'1 | ntMp'] := eqVneq 'O_p^'(M) 1. have nsZLP: 'Z('L(P)) <| M. by apply: Puig_center_normal Mp'1 => //; apply: mFT_odd. - rewrite -(mmax_normal maxM nsZLP). - exact: char_norm_trans (center_Puig_char P) _. - apply: contraNneq ntR => /(trivg_center_Puig_pgroup pP) P1. + rewrite -(mmax_normal maxM nsZLP) ?gFnorm_trans //. + apply: contraNneq ntR => /(trivg_center_Puig_pgroup pP)-P1. by rewrite -subG1 -P1. rewrite -(mmax_normal maxM (pcore_normal _ _) ntMp') /= -defMp' norms_gen //. apply/subsetP=> x nPx; rewrite inE sub_conjg; apply/bigcupsP=> K. @@ -222,7 +228,7 @@ have pP := pHall_pgroup sylP; have pA := pgroupS sAP pP. have ntA: A :!=: 1 by rewrite -rank_gt0 -(subnKC Age3). have [p_pr _ [e oA]] := pgroup_pdiv pA ntA. have{e oA} def_piA: \pi(A) =i (p : nat_pred). - by rewrite /= oA pi_of_exp //; exact: pi_of_prime. + by rewrite /= oA pi_of_exp //; apply: pi_of_prime. have FmCAp_le2 M: M \in 'M('C(A)) -> 'r_p('F(M)) <= 2. case/setIdP=> maxM cCAM; rewrite leqNgt; apply: contra uA' => Fge3. exact: (any_rank3_Fitting_Uniqueness maxM Fge3). @@ -278,7 +284,7 @@ have sNP_mCA M: M \in 'M('C(A)) -> 'N(P) \subset M. have [y cRy [defQx]] := atransP2 trCRq' maxQ maxQx. rewrite -(mulgKV y x) groupMr. by rewrite (subsetP sNQ_M) // inE conjsgM defQx conjsgK. - apply: subsetP cRy; apply: (subset_trans (pcore_sub _ _)). + apply: subsetP cRy; apply: gFsub_trans. exact: subset_trans (centS _) sCAM. have sNA_M: 'N(A) \subset M. by rewrite sNR_M // subsetI sAP (subset_trans cAA). @@ -304,8 +310,7 @@ have uNP0_mCA M: M \in 'M('C(A)) -> 'M('N(P0)) = [set M]. have cDP0: P0 \subset 'C(D). have sA1A := Ohm_sub 1 A. have nDA1: 'Ohm_1(A) \subset 'N(D). - apply: subset_trans sA1A (subset_trans sAM (char_norm _)). - exact: char_trans (pcore_char _ _) (Fitting_char _). + by rewrite !gFnorm_trans // gFsub_trans // normsG. have abelA1: p.-abelem 'Ohm_1(A) by rewrite Ohm1_abelem. have dimA1ge3: logn p #|'Ohm_1(A)| >= 3. by rewrite -(rank_abelem abelA1) rank_Ohm1. @@ -348,7 +353,7 @@ have uNP0_mCA M: M \in 'M('C(A)) -> 'M('N(P0)) = [set M]. rewrite -(def_uniq_mmax uE maxM (subset_trans sEF (Fitting_sub _))). by rewrite inE maxL. have cDL_P0: P0 \subset 'C(D :&: L). - have nsDM: D <| M:= char_normal_trans (pcore_char _ _) (Fitting_normal M). + have nsDM: D <| M by rewrite !gFnormal_trans. have{nsDM} [sDM nDM] := andP nsDM. have sDL: D :&: L \subset L :&: M by rewrite setIC setIS. have nsDL: D :&: L <| L :&: M by rewrite /normal sDL setIC normsIG. @@ -388,7 +393,7 @@ have uNP0_mCA M: M \in 'M('C(A)) -> 'M('N(P0)) = [set M]. rewrite (mmax_normal maxM) ?mmax_sup_id //. have sNP_M := sNP_mCA M mCA_M; have sPM := subset_trans (normG P) sNP_M. rewrite /normal comm_subG //= -/P0. - have nFP: P \subset 'N(F) by rewrite (subset_trans _ (gFnorm _ _)). + have nFP: P \subset 'N(F) by apply: subset_trans (gFnorm _ _). have <-: F <*> P * 'N_M(P) = M. apply: Frattini_arg (pHall_subl (joing_subr _ _) (subsetT _) sylP). rewrite -(quotientGK (Fitting_normal M)) /= norm_joinEr //= -/F. diff --git a/mathcomp/odd_order/Make b/mathcomp/odd_order/Make index d7ed459..a7ba4be 100644 --- a/mathcomp/odd_order/Make +++ b/mathcomp/odd_order/Make @@ -1,4 +1,3 @@ -all.v BGappendixAB.v BGappendixC.v BGsection10.v diff --git a/mathcomp/odd_order/PFsection1.v b/mathcomp/odd_order/PFsection1.v index dc5ce95..d5e8ea3 100644 --- a/mathcomp/odd_order/PFsection1.v +++ b/mathcomp/odd_order/PFsection1.v @@ -1,9 +1,16 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime ssralg finset fingroup morphism. -Require Import perm automorphism quotient action zmodp center commutator. +From mathcomp +Require Import perm automorphism quotient action zmodp finalg center commutator. +From mathcomp Require Import poly cyclic pgroup nilpotent matrix mxalgebra mxrepresentation. +From mathcomp Require Import vector falgebra fieldext ssrnum algC rat algnum galois. +From mathcomp Require Import classfun character inertia integral_char vcharacter. Require ssrint. @@ -28,110 +35,54 @@ Variable gT : finGroupType. Lemma odd_eq_conj_irr1 (G : {group gT}) t : odd #|G| -> (('chi[G]_t)^*%CF == 'chi_t) = ('chi_t == 1). Proof. -move=> OG; apply/eqP/eqP=> [Ht | ->]; last exact: cfConjC_cfun1. -pose a := (@Zp1 1). -have Aito: - is_action <[a]> (fun (t : Iirr G) v => if v == a then conjC_Iirr t else t). - split=> [[[|[]]] //= _ t1 t2 Hj |j [[|[]]] // HH1 [[|[]]] // HH2 ] //=. - by apply: (inv_inj (@conjC_IirrK _ _)). - by rewrite conjC_IirrK. -pose ito := Action Aito. -have Acto: - is_action <[a]> (fun (c : {set gT}) v => if v == a then c^-1%g else c). - split=> [[[|[]]] //= _ t1 t2 Hj |j [[|[]]] // HH1 [[|[]]] // HH2 ] //=. - by rewrite -[t1]invgK Hj invgK. - by rewrite invgK. -pose cto := Action Acto. -have F1: [acts <[a]>, on (classes G) | cto]. - apply/subsetP=> j Hj. - rewrite !inE Hj; apply/subsetP=> u. - case/imsetP=> g GiG ->. - by rewrite inE /=; case: (_ == _) => //; - rewrite -?classVg mem_classes // ?groupV. -have F2 u x y: x \in G -> y \in cto (x ^: G) a -> 'chi_u x = 'chi_(ito u a) y. - rewrite mem_invg -{2}[y]invgK => Gx {y}/imsetP[y Gy ->]. +rewrite -coprimen2 => oddG; pose A := <[1 : 'Z_2]>. +have Z2P (a : 'Z_2): a = 0 \/ a = 1 by apply/pred2P; case: a => -[|[]]. +pose Ito (t : Iirr G) := [fun a : 'Z_2 => iter a (@conjC_Iirr _ G) t]. +pose Cto (C : {set gT}) := [fun a : 'Z_2 => iter a invg C]. +have IactP: is_action A Ito. + split=> [|i /Z2P[]->] /Z2P[]-> //=; last by rewrite conjC_IirrK. + exact/inv_inj/conjC_IirrK. +have CactP: is_action A Cto. + by split=> [|C /Z2P[]->] /Z2P[]-> //=; [apply: invg_inj | rewrite invgK]. +pose Iact := Action IactP; pose Cact := Action CactP. +have n_cG_A: [acts A, on classes G | Cact]. + rewrite cycle_subG !inE cycle_id; apply/subsetP=> _ /imsetP[x Gx ->]. + by rewrite !inE /= -classVg mem_classes ?groupV. +transitivity (t \in [set 0]); last by rewrite inE irr_eq1. +suffices{t} /eqP->: [set 0] == 'Fix_Iact[1]. + by rewrite !inE sub1set inE -(inj_eq irr_inj) conjC_IirrE. +rewrite eqEcard !(sub1set, inE) conjC_Iirr_eq0 eqxx /=. +rewrite (card_afix_irr_classes (cycle_id _) n_cG_A) => [|i x xy Gx]; last first. + rewrite inE => {xy}/imsetP[y Gy /(canRL invgK)->]. by rewrite -conjVg cfunJ {y Gy}//= conjC_IirrE cfunE -irr_inv invgK. -have F3: forall c, c \in classes G -> c^-1%g = c -> c = 1%g. - move=> c; case/imsetP => g GiG ->; rewrite -classVg => Hg. - move: (class_refl G g^-1); rewrite Hg; case/imsetP=> x XiG Hx. - have F4: (x ^+ 2)%g \in 'C_G[g]. - apply/subcent1P; split; rewrite ?groupM //. - apply: (mulgI (x * x * g)^-1)%g. - rewrite mulVg !invMg Hx conjgE !mulgA mulgK. - rewrite -[(_ * g * x)%g]mulgA -[(_ * (g * _))%g]mulgA -conjgE. - by rewrite -Hx mulgK mulVg. - have F5 : x \in 'C_G[g]. - suff->: (x = (x ^+ 2) ^+ (#|G| %/2).+1)%g by apply: groupX. - rewrite -expgM -[(_%/_).+1]addn1 mulnDr muln1 -{3}addn1 addnA. - move: (modn2 #|G|); rewrite {1}OG /= => HH; rewrite -{3}HH. - rewrite [(2 * _)%N]mulnC -divn_eq expgD expg1. - by move: (order_dvdG XiG); rewrite order_dvdn; move/eqP->; rewrite mul1g. - move: Hx; rewrite conjgE; case/subcent1P: F5=> _ ->. - rewrite mulgA mulVg mul1g => HH. - have F6: (g ^+ 2 == 1)%g by rewrite expgS -{1}HH expg1 mulVg. - suff: #[g] == 1%N by rewrite order_eq1; move/eqP->; apply: class1G. - move: F6 (order_gt0 g) (order_dvdG GiG); rewrite -order_dvdn. - move/(dvdn_leq (isT : (0 < 2)%N)); case: #[_]=> // [[|[]]] //. - by rewrite dvdn2 OG. -apply/eqP; case: (boolP (t == 0))=> // Hd. - by move/eqP: Hd->; rewrite irr0. -have:= card_afix_irr_classes (cycle_id a) F1 F2. -have->: #|'Fix_(classes G | cto)[a]| = 1%N. - apply: (@eq_card1 _ 1%g)=> c; apply/idP/idP; rewrite !inE. - case/andP=> GiG HH; apply/eqP; apply: F3=> //; apply/eqP. - by move/subsetP: HH; move/(_ a); rewrite !inE eqxx; apply. - move/eqP->; rewrite classes1. - apply/subsetP=> b; rewrite !inE; move/eqP=> -> /=. - by rewrite invg1. -rewrite (cardD1 (0 : Iirr _)). -have->: 0 \in 'Fix_ito[a]. - apply/afixP=> b; rewrite !inE; move/eqP->; rewrite /=. - apply: irr_inj; apply/cfunP=> g. - by rewrite conjC_IirrE cfConjCE irr0 cfun1E conjC_nat. -rewrite (cardD1 t) //. -suff->: t \in [predD1 'Fix_ito[a] & 0] by []. -rewrite inE /= Hd. -apply/afixP=> b; rewrite !inE; move/eqP->; rewrite /=. -apply: irr_inj; apply/cfunP=> g. -by rewrite conjC_IirrE Ht. +have ->: #|[set 0 : Iirr G]| = #|[1 {set gT}]| by rewrite !cards1. +apply/subset_leq_card/subsetP=> _ /setIdP[/imsetP[x Gx ->] /afix1P-DxGV]. +have /imsetP[y Gy DxV]: x^-1%g \in x ^: G by rewrite -DxGV memV_invg class_refl. +have{Gy} cxy: y \in 'C[x]. + suffices cxy2: (y ^+ 2)%g \in 'C[x] by rewrite -(expgK oddG Gy) groupX. + by rewrite cent1C cent1E conjgC conjgM -DxV conjVg -DxV invgK. +rewrite inE classG_eq1 -in_set1 -(expgK oddG Gx) groupX // inE. +by rewrite -eq_invg_mul DxV conjgE -(cent1P cxy) mulKg. Qed. Variables G H : {group gT}. (* This is Peterfalvi (1.2). *) -Lemma not_in_ker_char0 t g : g \in G -> +Lemma irr_reg_off_ker_0 t g : g \in G -> H <| G -> ~~ (H \subset cfker 'chi[G]_t) -> 'C_H[g] = 1%g -> 'chi_t g = 0. Proof. -move=> GiG HnG nHsC CH1. -have: (#|'C_G[g]| <= #|'C_(G/H)[coset H g]|)%N. - suff->: #|'C_G[g]| = #|'C_G[g] / H|%G. - by apply: (subset_leq_card (quotient_subcent1 H G g)). - apply: card_isog. - apply: isog_trans (second_isog _); last first. - apply: subset_trans (normal_norm HnG). - by apply: subcent1_sub. - suff->: H :&: 'C_G[g] = 1%g by exact: quotient1_isog. - rewrite -CH1. - apply/setP=> h; rewrite inE. - apply/andP/subcent1P; case=> H1 H2; split=> //. - by move/subcent1P: H2; case. - apply/subcent1P; split=> //. - by apply: (subsetP (normal_sub HnG)). -have F1: coset H g \in (G / H)%g by exact: mem_quotient. -rewrite -leC_nat. -have:= second_orthogonality_relation g GiG. -rewrite mulrb class_refl => <-. -have:= second_orthogonality_relation (coset H g) F1. -rewrite mulrb class_refl => <-; rewrite -!(eq_bigr _ (fun _ _ => normCK _)). -rewrite sum_norm_irr_quo // (bigID (fun i => H \subset cfker 'chi_i)) //=. -set S := \sum_(i | ~~ _) _; set S' := \sum_(i | _) _ => HH. -have /eqP F2: S == 0. - rewrite eqr_le -(ler_add2l S') addr0 HH /=. - by apply: sumr_ge0 => j _; rewrite mulr_ge0 ?normr_ge0. -apply/eqP; have: `|'chi_t g| ^+ 2 == 0. - apply/eqP; apply: (psumr_eq0P _ F2) nHsC => j _. - by rewrite mulr_ge0 ?normr_ge0. -by rewrite mulf_eq0 orbb normr_eq0. +pose kerH i := H \subset cfker 'chi[G]_i => Gg nsHG kerH't regHg; apply/eqP. +pose sum_norm2 x := \sum_i `|'chi_i x| ^+ 2. +have norm2_ge0 a: 0 <= `|a| ^+ 2 :> algC by rewrite exprn_ge0 ?normr_ge0. +have{regHg}: sum_norm2 gT G g <= sum_norm2 _ (G / H)%G (coset H g). + rewrite ). + rewrite !second_orthogonality_relation ?mem_quotient // !class_refl ler_nat. + suffices /card_isog->: 'C_G[g] \isog 'C_G[g] / H. + exact/subset_leq_card/quotient_subcent1. + by apply/quotient_isog; rewrite ?subIset 1?normal_norm // setICA regHg setIg1. +rewrite /sum_norm2 (bigID kerH) ?sum_norm_irr_quo //= -ler_subr_addl subrr. +rewrite ler_eqVlt psumr_eq0 ?ler_gtF ?sumr_ge0 // orbF => /allP/(_ t)/implyP. +by rewrite mem_index_enum kerH't expf_eq0 normr_eq0. Qed. (* This is Peterfalvi (1.3)(a). *) @@ -247,13 +198,13 @@ Proof. case: m Chi => [|[|m]] // Chi _ irrChi Chifree Chi1 ChiCF [iso_tau Ztau]. rewrite -(tnth_nth 0 _ 0); set chi := tnth Chi. have chiE i: chi i = Chi`_i by rewrite -tnth_nth. -have inChi i: chi i \in Chi by exact: mem_tnth. -have{irrChi} irrChi i: chi i \in irr H by exact: irrChi. +have inChi i: chi i \in Chi by apply: mem_tnth. +have{irrChi} irrChi i: chi i \in irr H by apply: irrChi. have eq_chi i j: (chi i == chi j) = (i == j). by rewrite /chi !(tnth_nth 0) nth_uniq ?size_tuple ?free_uniq. have dot_chi i j: '[chi i, chi j] = (i == j)%:R. rewrite -eq_chi; have [/irrP[{i}i ->] /irrP[{j}j ->]] := (irrChi i,irrChi j). - by rewrite cfdot_irr inj_eq //; exact: irr_inj. + by rewrite cfdot_irr inj_eq //; apply: irr_inj. pose F i j := chi i - chi j. have DF i j : F i j = F i 0 - F j 0 by rewrite /F opprB addrA subrK. have ZF i j: F i j \in 'Z[Chi, L]. @@ -371,7 +322,7 @@ Lemma cfclass_Ind_irrP i j : Proof. move=> nsHG; have [sHG _] := andP nsHG. case: ifP (cfclass_Ind_cases j i nsHG) => [|_ Oji]; first by left. -right=> eq_chijG; have /negP[]: 'Ind[G] 'chi_i != 0 by exact: Ind_irr_neq0. +right=> eq_chijG; have /negP[]: 'Ind[G] 'chi_i != 0 by apply: Ind_irr_neq0. by rewrite -cfnorm_eq0 {1}eq_chijG Oji. Qed. @@ -611,7 +562,7 @@ have CIr: i \in irr_constt ('Ind[G] 'chi_i1). by rewrite inE /= -Frobenius_reciprocity /= cfdotC conjC_eq0. have BsKi : B \subset cfker 'chi_i1. suff BsKri: B \subset cfker ('Res[C] 'chi_i). - by apply: (subset_trans BsKri); exact: (cfker_constt _ Hi1). + by apply: (subset_trans BsKri); apply: (cfker_constt _ Hi1). apply/subsetP=> g GiG. have F: g \in C by rewrite (subsetP (subset_trans BsD _)). rewrite cfkerEchar // inE F !cfResE //. diff --git a/mathcomp/odd_order/PFsection10.v b/mathcomp/odd_order/PFsection10.v index d380e47..3e717c6 100644 --- a/mathcomp/odd_order/PFsection10.v +++ b/mathcomp/odd_order/PFsection10.v @@ -1,13 +1,24 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime ssralg poly finset center. +From mathcomp Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +From mathcomp Require Import gfunctor gproduct cyclic commutator gseries nilpotent pgroup. +From mathcomp Require Import sylow hall abelian maximal frobenius. +From mathcomp Require Import matrix mxalgebra mxrepresentation mxabelem vector. +From mathcomp Require Import BGsection1 BGsection3 BGsection7 BGsection15 BGsection16. +From mathcomp Require Import ssrnum algC classfun character integral_char inertia vcharacter. +From mathcomp Require Import PFsection1 PFsection2 PFsection3 PFsection4. +From mathcomp Require Import PFsection5 PFsection6 PFsection7 PFsection8 PFsection9. (******************************************************************************) @@ -163,7 +174,7 @@ Lemma FTtypeP_ref_irr : Proof. have [_ /has_nonprincipal_irr[s nz_s] _ _ _] := Frobenius_context frobMbar. exists ('Ind 'chi_s %% M'')%CF; split. -- exact/cfMod_irr/(irr_induced_Frobenius_ker (FrobeniusWker frobMbar)). +- by rewrite cfMod_irr ?irr_induced_Frobenius_ker ?(FrobeniusWker frobMbar). - by rewrite -cfIndMod ?normal_sub // -mod_IirrE // mem_calS mod_Iirr_eq0. rewrite -cfIndMod ?cfInd1 ?normal_sub // -(index_sdprod defM) cfMod1. by rewrite lin_char1 ?mulr1 //; apply/char_abelianP/sub_der1_abelian. @@ -706,8 +717,9 @@ have lb_rho: 1 - w1%:R / #|M'|%:R <= '[rho chi]. rewrite -leC_nat in lb_M'bar; apply: ler_trans lb_M'bar _. rewrite ler_subr_addl -mulrS prednK ?cardG_gt0 // leC_nat. by rewrite dvdn_leq ?dvdn_quotient. -have{lb_rho ub_rho}: 1 - #|G1|%:R/ #|G|%:R - w1%:R^-1 < w1%:R / #|M'|%:R :> algC. - rewrite -addrA -opprD ltr_subl_addr -ltr_subl_addl. +have{lb_rho ub_rho}: + 1 - #|G1|%:R / #|G|%:R - w1%:R^-1 < w1%:R / #|M'|%:R :> algC. +- rewrite -addrA -opprD ltr_subl_addr -ltr_subl_addl. apply: ler_lt_trans (ler_trans lb_rho ub_rho) _; rewrite addrC ltr_add2l. rewrite ltr_pdivr_mulr ?gt0CG // mulrC -(sdprod_card defM) natrM. by rewrite mulfK ?neq0CG // defA ltC_nat (cardsD1 1%g M') group1. @@ -882,20 +894,22 @@ suffices [tau1]: coherent calS M^# tau by case/FTtype345_noncoherence_main. have [[_ U_M_1] MtypeV] := compl_of_typeV maxM MtypeP Mtype5. have [_ [_ _ _ defH] _ [_ _ _ sW2H' _] _] := MtypeP. have{U_M_1 defH} defMF: M`_\F = H by rewrite /= -defH U_M_1 sdprodg1. -have nilH: nilpotent `H by rewrite -defMF Fcore_nil. -have scohS := subset_subcoherent scohS0 sSS0. -have [|//|[[_]]] := (non_coherent_chief nsM'M (nilpotent_sol nilH) scohS) 1%G. - split; rewrite ?mFT_odd ?normal1 ?sub1G ?quotient_nil //. - by rewrite joingG1 (FrobeniusWker frobMbar). -rewrite /= joingG1 -(index_sdprod defM) /= -/w1 -[H^`(1)%g]/`H' => ubHbar [p[]]. -rewrite -(isog_abelian (quotient1_isog H)) -(isog_pgroup p (quotient1_isog H)). -rewrite subn1 => pH not_cHH /negP not_w1_dv_p'1. +have nilH := Fcore_nil M; rewrite defMF -/w1 in MtypeV nilH. +without loss [p [pH not_cHH ubHbar not_w1_dv_p1]]: / exists p : nat, + [/\ p.-group H, ~~ abelian H, #|H : H'| <= 4 * w1 ^ 2 + 1 & ~ w1 %| p.-1]%N. +- have [isoH1 solH] := (quotient1_isog H, nilpotent_sol nilH). + have /non_coherent_chief-IHcoh := subset_subcoherent scohS0 sSS0. + apply: IHcoh (fun coh _ => coh) _ => // [|[[_ ubH] [p [pH ab'H] /negP-dv'p]]]. + split; rewrite ?mFT_odd ?normal1 ?sub1G ?quotient_nil //. + by rewrite joingG1 (FrobeniusWker frobMbar). + apply; exists p; rewrite (isog_abelian isoH1) (isog_pgroup p isoH1) -subn1. + by rewrite /= joingG1 -(index_sdprod defM) in ubH dv'p. have ntH: H :!=: 1%g by apply: contraNneq not_cHH => ->; apply: abelian1. have [sH'H nH'H] := andP nsM''M'; have sW2H := subset_trans sW2H' sH'H. have def_w2: w2 = p by apply/eqP; have:= pgroupS sW2H pH; rewrite pgroupE pnatE. -have [p_pr _ [e oH]] := pgroup_pdiv pH ntH. -rewrite -/w1 /= defMF oH pi_of_exp {e oH}// /pi_of primes_prime // in MtypeV. -have [tiHG | [_ /predU1P[->[]|]]// | [_ /predU1P[->|//] [oH w1p1 _]]] := MtypeV. +have piHp q: q \in \pi(H) -> q = p. + by rewrite /= -(part_pnat_id pH) pi_of_part // => /andP[_ /eqnP]. +have [tiHG | [_ /piHp-> []//] | [_ /piHp-> [oH w1_dv_p1 _]]] := MtypeV. suffices [tau1 [Itau1 Dtau1]]: coherent (seqIndD H M H 1) M^# 'Ind[G]. exists tau1; split=> // phi Sphi; rewrite {}Dtau1 //. rewrite zcharD1_seqInd // -subG1 -setD_eq0 -defA in Sphi tiHG ntH. @@ -903,8 +917,8 @@ have [tiHG | [_ /predU1P[->[]|]]// | [_ /predU1P[->|//] [oH w1p1 _]]] := MtypeV. apply: (@Sibley_coherence _ [set:_] M H W1); first by rewrite mFT_odd. right; exists W2 => //; exists 'A0(M), W, defW. by rewrite -defA -{2}(group_inj defMs). -rewrite pcore_pgroup_id // in oH. -have esH: extraspecial H. +have [p_pr _ _] := pgroup_pdiv pH ntH; rewrite (pcore_pgroup_id pH) in oH. +have{not_cHH} esH: extraspecial H. by apply: (p3group_extraspecial pH); rewrite // oH pfactorK. have oH': #|H'| = p. by rewrite -(card_center_extraspecial pH esH); have [[_ <-]] := esH. diff --git a/mathcomp/odd_order/PFsection11.v b/mathcomp/odd_order/PFsection11.v index 3584dbe..e981181 100644 --- a/mathcomp/odd_order/PFsection11.v +++ b/mathcomp/odd_order/PFsection11.v @@ -1,13 +1,24 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime ssralg poly finset center. +From mathcomp Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +From mathcomp Require Import gfunctor gproduct cyclic commutator gseries nilpotent pgroup. +From mathcomp Require Import sylow hall abelian maximal frobenius. +From mathcomp Require Import matrix mxalgebra mxrepresentation mxabelem vector. +From mathcomp Require Import BGsection1 BGsection3 BGsection7 BGsection15 BGsection16. +From mathcomp Require Import ssrnum ssrint algC classfun character inertia vcharacter. +From mathcomp Require Import PFsection1 PFsection2 PFsection3 PFsection4 PFsection5. +From mathcomp Require Import PFsection6 PFsection7 PFsection8 PFsection9 PFsection10. (******************************************************************************) @@ -180,10 +191,10 @@ Let nH0M: M \subset 'N(H0). Proof. by have /andP[/maxgroupp/andP[]] := chiefH0. Qed. Let nsH0H : H0 <| H. -Proof. by rewrite /normal sH0H (subset_trans (Fcore_sub _)). Qed. +Proof. by rewrite /normal sH0H gFsub_trans. Qed. Let nsH0C_M : H0C <| M. -Proof. by rewrite !normalY ?gFnormal /normal ?(subset_trans sH0H) ?gFsub. Qed. +Proof. by rewrite normalY // /normal ?(subset_trans sH0H) ?gFsub. Qed. Let defH0C : H0 \x C = H0C. Proof. @@ -206,7 +217,7 @@ have nH0_C := subset_trans sCM nH0M. have sH0C_HC: H0C \subset HC by apply: genS (setSU _ _). have defF: HC :=: 'F(M) by have [/dprodWY] := typeP_context MtypeP. have{defF} nilHC: nilpotent (HC / 1) by rewrite defF quotient_nil ?Fitting_nil. -have /bounded_seqIndD_coherent-bounded_coh1 := scoh1. +have /bounded_seqIndD_coherence-bounded_coh1 := scoh1. apply: bounded_coh1 nilHC cohH0C _; rewrite ?sub1G ?normal1 //. have[_ _ /= oHbar] := Ptype_Fcore_factor_facts maxM MtypeP notMtype5. rewrite -(index_sdprod defM) -divgS // -(dprod_card defHC) -(dprod_card defH0C). @@ -220,17 +231,15 @@ Lemma bounded_proper_coherent H1 : (#|HU : H1| <= 2 * q * #|U : C| + 1)%N. Proof. move=> nsH1_M psH1_M' cohH1; have [nsHHU _ _ _ _] := sdprod_context defHU. -rewrite -leC_nat natrD -ler_subl_addr. -have ->: (2 * q * #|U : C|)%:R = 2%:R * #|M : HC|%:R * sqrtC #|HC : HC|%:R. - rewrite indexgg sqrtC1 mulr1 -mulnA natrM; congr (_ * _%:R). - apply/eqP; rewrite // -(eqn_pmul2l (cardG_gt0 HC)) Lagrange ?normal_sub //. - rewrite mulnCA -(dprod_card defHC) -mulnA (Lagrange (subsetIl _ _)). - by rewrite -(sdprod_card defM) -(sdprod_card defHU) mulnC. -have ns_M: [/\ H1 <| M, H0C <| M, HC <| M & HC <| M] by []. -case: (coherent_seqIndD_bound _ _ scoh1 ns_M) FTtype34_noncoherence => //. -suffices /center_idP->: abelian (HC / H0C) by rewrite genS ?setSU. -suffices /isog_abelian->: HC / H0C \isog H / H0 by apply: abelem_abelian p _ _. -by rewrite /= (joingC H0) isog_sym quotient_sdprodr_isog ?(dprodWsdC defHC). +suffices: #|HU : H1|%:R - 1 <= 2%:R * #|M : HC|%:R * sqrtC #|HC : HC|%:R. + rewrite indexgg sqrtC1 mulr1 -leC_nat natrD -ler_subl_addr -mulnA natrM. + congr (_ <= _ * _%:R); apply/eqP; rewrite -(eqn_pmul2l (cardG_gt0 HC)). + rewrite Lagrange ?normal_sub // mulnCA -(dprod_card defHC) -mulnA mulnC. + by rewrite Lagrange ?subsetIl // (sdprod_card defHU) (sdprod_card defM). +apply/negP/(coherent_seqIndD_bound _ _ scoh1 _ _ _ FTtype34_noncoherence) => //. +suffices /center_idP->: abelian (HC / H0C) by rewrite genS ?setSU. +suffices /isog_abelian<-: Hbar \isog HC / H0C by apply: abelem_abelian abelHbar. +by rewrite /= [`H0C]joingC quotient_sdprodr_isog ?(dprodWsdC defHC). Qed. (* This is Peterfalvi (11.5). *) @@ -295,8 +304,8 @@ have{isomHU} defC: C :=: U'. by rewrite (dprodW defHC) -defM'' -quotient_der // -mulHU mul_subG ?normG. have{coH_UW1} defH0: H0 :=: H'. pose Hhat := (H / H')%g; pose Uhat := (U / H')%g; pose HUhat := (HU / H')%g. - have nH'H: H \subset 'N(H') := gFnorm _ _. - have nH'U: U \subset 'N(H') := char_norm_trans (der_char _ _) nHU. + have nH'H: H \subset 'N(H') by apply: gFnorm. + have nH'U: U \subset 'N(H') by apply: gFnorm_trans. apply/eqP; rewrite eqEsubset andbC. rewrite der1_min ?(abelem_abelian abelHbar) ?normal_norm //=. rewrite -quotient_sub1 /= -/H'; last exact: subset_trans sH0H nH'H. @@ -342,7 +351,7 @@ suffices defRHpU: R \x ('O_p(H) <*> U) = HU. rewrite -(sdprodWY defHU) -[in RHS](dprodWY defRHp) -joingA. have [_ _ cRHp tiRHp] := dprodP defRHp. rewrite dprodEY //= -/R; first by rewrite join_subG cRHp centsC. -rewrite joingC (norm_joinEl (char_norm_trans (pcore_char p H) nHU)). +rewrite joingC norm_joinEl 1?gFnorm_trans //. by rewrite -(setIidPl sRH) -setIA -group_modr ?gFsub // tiHU mul1g. Qed. @@ -441,7 +450,7 @@ have{pHhat} gal'M: ~~ typeP_Galois MtypeP. rewrite minUHbar //= -/Hbar -?fHhat 1?morphim_injm_eq1 ?morphimS // -subG1. rewrite quotient_sub1 ?(normal_norm nsH0hatZ) // not_esHhat -['Z(_)]cosetpreK. rewrite im_f ?sub_cosetpre_quo // quotient_norms ?norm_quotient_pre //. - by rewrite (char_norm_trans (center_char _)) ?quotient_norms. + by rewrite gFnorm_trans ?quotient_norms. have [H1 []] := typeP_Galois_Pn maxM notMtype5 gal'M. rewrite def_p => oH1 nH1Ubar _ /bigdprodWY-defHbar _. have /cyclicP[xbar defH1]: cyclic H1 by rewrite prime_cyclic ?oH1. @@ -1081,7 +1090,7 @@ have bridgeS1: {in S1, forall zeta, eq_proj_eta (tau (bridge0 zeta)) eta0row}. rewrite -Dtau1 ?zcharD1_seqInd //. rewrite cfdot_suml big_map big1_seq // => _ /(cycTIirrP defW)[i [j ->]]. apply/eqP; rewrite cfdotC fmorph_eq0 cfdotZr raddfB cfdotBl. - by rewrite !o_tau1_eta ?cfAut_seqInd ?irr_aut // subrr mulr0. + by rewrite !o_tau1_eta ?cfAut_seqInd ?cfAut_irr // subrr mulr0. have a2_ge0 i j: 0 <= a_ i j ^+ 2 by rewrite -realEsqr Creal_Cint. have a11_0: a11 = 0. have: ('[X] < (2 * q.-1)%:R). diff --git a/mathcomp/odd_order/PFsection12.v b/mathcomp/odd_order/PFsection12.v index 2b3a3e1..fa68ea5 100644 --- a/mathcomp/odd_order/PFsection12.v +++ b/mathcomp/odd_order/PFsection12.v @@ -1,17 +1,32 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime ssralg finset center. +From mathcomp Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +From mathcomp Require Import gfunctor gproduct cyclic commutator gseries nilpotent pgroup. +From mathcomp Require Import sylow hall abelian maximal frobenius. +From mathcomp Require Import matrix mxalgebra mxpoly mxrepresentation mxabelem vector. +From mathcomp Require Import falgebra fieldext finfield. +From mathcomp Require Import BGsection1 BGsection2 BGsection3 BGsection4 BGsection7. +From mathcomp Require Import BGsection14 BGsection15 BGsection16. +From mathcomp Require Import ssrnum ssrint algC cyclotomic algnum. +From mathcomp Require Import classfun character inertia vcharacter. +From mathcomp Require Import PFsection1 PFsection2 PFsection3 PFsection4 PFsection5. +From mathcomp Require Import PFsection6 PFsection7 PFsection8 PFsection9 PFsection10. +From mathcomp Require Import PFsection11. Set Implicit Arguments. @@ -49,10 +64,8 @@ Let calI := [seq 'chi_i | i in calX]. (* This does not actually use the Ltype1 assumption. *) Lemma FTtype1_ref_irr : exists2 phi, phi \in calS & phi 1%g = #|L : H|%:R. Proof. -have [solH ntH] := (nilpotent_sol (Fcore_nil L), mmax_Fcore_neq1 maxL). -have [s Ls nzs] := solvable_has_lin_char ntH solH. -exists ('Ind 'chi_s); last by rewrite cfInd1 ?gFsub // lin_char1 ?mulr1. -by rewrite mem_seqInd ?gFnormal ?normal1 // !inE sub1G subGcfker -irr_eq1 nzs. +have solH: solvable H := nilpotent_sol (Fcore_nil L). +by apply: exists_linInd; rewrite ?normal1 // proper1G mmax_Fcore_neq1. Qed. Let mem_calI i : i \in calX -> 'chi_i \in calI. @@ -122,11 +135,11 @@ have CHy1: 'C_H[y] = 1%g. rewrite FTsupp1_type1 Ltype1 //=; exists z; first by rewrite !inE ntz. by rewrite 3!inE nty Ly cent1C. have: j \in calX by apply/FTtype1_irrP; exists chi. -by rewrite !inE => /andP[/not_in_ker_char0->]. +by rewrite !inE => /andP[/irr_reg_off_ker_0->]. Qed. (* This is Peterfalvi (12.2)(a), second part. *) -Lemma FPtype1_irr_isometry : +Lemma FTtype1_irr_isometry : {in 'Z[calI, L^#], isometry tau, to 'Z[irr G, G^#]}. Proof. apply: (sub_iso_to _ _ (Dade_Zisometry _)) => // phi. @@ -136,21 +149,21 @@ apply: zchar_trans_on phi S_phi => _ /imageP[i /FTtype1_irrP[j calSj Sj_i] ->]. by rewrite zchar_split irr_vchar; have [_ _ ->] := FTtype1_seqInd_facts calSj. Qed. -Lemma FPtype1_irr_subcoherent : +Lemma FTtype1_irr_subcoherent : {R : 'CF(L) -> seq 'CF(G) | subcoherent calI tau R}. Proof. -apply: irr_subcoherent; last exact: FPtype1_irr_isometry. +apply: irr_subcoherent; last exact: FTtype1_irr_isometry. have UcalI: uniq calI by apply/dinjectiveP; apply: in2W irr_inj. split=> // _ /imageP[i Ii ->]; rewrite !inE in Ii; first exact: mem_irr. by apply/imageP; exists (conjC_Iirr i); rewrite ?inE conjC_IirrE ?cfker_aut. apply/hasPn=> psi; case/imageP => i; rewrite !inE => /andP[kerH'i _] ->. rewrite /cfReal odd_eq_conj_irr1 ?mFT_odd // irr_eq1 -subGcfker. -by apply: contra kerH'i; apply: subset_trans; apply: gFsub. +by apply: contra kerH'i; apply: gFsub_trans. Qed. -Local Notation R1gen := FPtype1_irr_subcoherent. +Local Notation R1gen := FTtype1_irr_subcoherent. (* This is Peterfalvi (12.2)(b). *) -Lemma FPtype1_subcoherent (R1 := sval R1gen) : +Lemma FTtype1_subcoherent (R1 := sval R1gen) : {R : 'CF(L) -> seq 'CF(G) | [/\ subcoherent calS tau R, {in Iirr_kerD L H 1%G, forall i (phi := 'chi_i), @@ -161,7 +174,7 @@ Lemma FPtype1_subcoherent (R1 := sval R1gen) : Proof. have nrS: ~~ has cfReal calS by apply: seqInd_notReal; rewrite ?mFT_odd. have U_S: uniq calS by apply: seqInd_uniq. -have ccS: conjC_closed calS by apply: cfAut_seqInd. +have ccS: cfConjC_closed calS by apply: cfAut_seqInd. have conjCS: cfConjC_subset calS (seqIndD H L H 1) by split. case: R1gen @R1 => /= R1 subc1. have [[chi_char nrI ccI] tau_iso oI h1 hortho] := subc1. @@ -180,7 +193,8 @@ exists R; split => //= => [| i Ii]; last first. split=> //; apply/eqP; rewrite -eqC_nat -cfnorm_orthonormal // -{}tau_im. have ?: 'chi_i - ('chi_i)^*%CF \in 'Z[calI, L^#]. have hchi : 'chi_i \in 'Z[calI, L] by rewrite mem_zchar_on // cfun_onG. - rewrite sub_aut_zchar ?cfAut_zchar // => _ /mapP[j _ ->]; exact: irr_vchar. + rewrite sub_aut_zchar ?cfAut_zchar // => _ /mapP[j _ ->]. + exact: irr_vchar. have [-> // _] := tau_iso; rewrite cfnormBd ?cfnorm_conjC ?cfnorm_irr //. by have [_ ->] := pairwise_orthogonalP oI; rewrite ?ccI // eq_sym (hasPn nrI). have calS_portho : pairwise_orthogonal calS by apply: seqInd_orthogonal. @@ -226,8 +240,8 @@ Qed. End Twelve2. -Local Notation R1gen := FPtype1_irr_subcoherent. -Local Notation Rgen := FPtype1_subcoherent. +Local Notation R1gen := FTtype1_irr_subcoherent. +Local Notation Rgen := FTtype1_subcoherent. (* This is Peterfalvi (12.3) *) Lemma FTtype1_seqInd_ortho L1 L2 (maxL1 : L1 \in 'M) (maxL2 : L2 \in 'M) @@ -506,8 +520,7 @@ split=> //; have [U [MtypeF MtypeI]] := FTtypeP 1 maxL Ltype1. have [[ntH ntU defL] _ _] := MtypeF; have nsHL: H <| L := gFnormal _ L. have nilH: nilpotent H := Fcore_nil L; have solH := nilpotent_sol nilH. have frobHU: [Frobenius L = H ><| U] := set_Frobenius_compl defL frobL. -pose R := sval (Rgen maxL Ltype1). -have scohS: subcoherent calS tau R by case: (svalP (Rgen maxL Ltype1)). +have [R [scohS _ _]] := Rgen maxL Ltype1; rewrite -/calS -/tau in scohS. have [tiH | [cHH _] | [expUdvH1 _]] := MtypeI. - have /Sibley_coherence := And3 (mFT_odd L) nilH tiH. case/(_ U)=> [|tau1 [IZtau1 Dtau1]]; first by left. @@ -516,17 +529,13 @@ have [tiH | [cHH _] | [expUdvH1 _]] := MtypeI. - apply/(uniform_degree_coherence scohS)/(@all_pred1_constant _ #|L : H|%:R). apply/allP=> _ /mapP[_ /seqIndP[s _ ->] ->] /=. by rewrite cfInd1 ?gFsub // lin_char1 ?mulr1 //; apply/char_abelianP. -have isoHbar := quotient1_isog H. -have /(_ 1%G)[|//|[_ [p [pH _] /negP[]]]] := non_coherent_chief nsHL solH scohS. - split; rewrite ?mFT_odd ?normal1 ?sub1G -?(isog_nil isoHbar) //= joingG1. +apply: (non_coherent_chief _ _ scohS) id _ => // [|[_ [p [pH _] /negP[]]]]. + split; rewrite ?mFT_odd ?normal1 ?sub1G ?quotient_nil //= joingG1. apply/existsP; exists (U / H')%G. - rewrite Frobenius_proper_quotient ?(sol_der1_proper solH) //. - exact: char_normal_trans (der_char 1 H) nsHL. -rewrite -(isog_pgroup p isoHbar) in pH. -have [pr_p p_dv_H _] := pgroup_pdiv pH ntH. -rewrite subn1 -(index_sdprod defL). -have [-> *] := typeF_context MtypeF; last by split; rewrite ?(sdprodWY defL). -by rewrite expUdvH1 // mem_primes pr_p cardG_gt0. + by rewrite Frobenius_proper_quotient ?(sol_der1_proper solH) ?gFnormal_trans. +rewrite subn1 -(index_sdprod defL) -(isog_pgroup p (quotient1_isog H)) in pH *. +have /expUdvH1: p \in \pi(H) by rewrite -p_part_gt1 part_pnat_id ?cardG_gt1. +by have [-> //] := typeF_context MtypeF; split; rewrite ?(sdprodWY defL). Qed. End Twelve_4_to_6. @@ -585,7 +594,7 @@ have abelP1: p.-abelem P1 := Ohm1_abelem pP0 abP0. have [pP1 abP1 _] := and3P abelP1. have sP10: P1 \subset P0 := Ohm_sub 1 P0; have sP1M := subset_trans sP10 sP0M. have nKP1: P1 \subset 'N(K) by rewrite (subset_trans sP1M) ?gFnorm. -have nK'P1: P1 \subset 'N(K') := char_norm_trans (der_char 1 K) nKP1. +have nK'P1: P1 \subset 'N(K') by apply: gFnorm_trans. have{coKP0} coKP1: coprime #|K| #|P1| := coprimegS sP10 coKP0. have solK: solvable K := nilpotent_sol (Fcore_nil M). have isoP1: P1 \isog P1 / K'. @@ -627,6 +636,7 @@ Hypotheses (P0_1s_x : x \in 'Ohm_1(P0)^#) (not_sCxK' : ~~ ('C_K[x] \subset K')). Hypotheses (sNxM : 'N(<[x]>) \subset M) (not_sCxL : ~~ ('C[x] \subset L)). Let H := L`_\F%G. +Local Notation "` 'H'" := (gval L)`_\F (at level 0, format "` 'H'"). Let nsHL : H <| L. Proof. exact: gFnormal. Qed. (* This is Peterfalvi (12.10). *) @@ -664,7 +674,7 @@ have [Ltype1 | notLtype1] := boolP (FTtype L == 1)%N; last first. by have [|y _ /cyclicS->] := Hall_psubJ hallU piK'p _ pP0; rewrite ?cyclicJ. have sP0H: P0 \subset H by rewrite /= -FTcore_type1. have [U [LtypeF /= LtypeI]] := FTtypeP 1 maxL Ltype1. -have [[_ _ defL] _ _] := LtypeF; have [_ sUL _ nHU _] := sdprod_context defL. +have [[_ _ defL] _ _] := LtypeF; have [/=_ sUL _ nHU _] := sdprod_context defL. have not_tiH: ~ normedTI H^# G L. have H1x: x \in H^# by rewrite !inE ntx (subsetP sP0H). by case/cent1_normedTI/(_ x H1x)/idPn; rewrite setTI. @@ -696,22 +706,19 @@ suffices: q %| p ^ 2 - 1 ^ 2. rewrite subn_sqr addn1 subn1 Euclid_dvdM //. by case/orP; [exists false | exists true]. pose P := 'O_p(H); pose P1 := 'Ohm_1(P). -have chP1H: P1 \char H := char_trans (Ohm_char 1 _) (pcore_char p H). have sylP: p.-Sylow(H) P := nilpotent_pcore_Hall p (Fcore_nil L). -have [sPH pP _] := and3P sylP. +have [sPH pP _] := and3P sylP; have sP1H: P1 \subset H by rewrite 2?gFsub_trans. have abelP1: p.-abelem P1 by rewrite Ohm1_abelem ?(abelianS sPH). -have [pP1 _] := andP abelP1. have prankP1: 'r_p(P1) = 2. - apply/eqP; rewrite p_rank_Ohm1 eqn_leq -{1}rank_pgroup // -{1}rankH rankS //=. - by rewrite -prankP0 (p_rank_Sylow sylP) p_rankS. -have ntP1: P1 != 1%g by rewrite -rank_gt0 (rank_pgroup pP1) prankP1. + apply/anti_leq; rewrite p_rank_Ohm1 (p_rank_Sylow sylP). + by rewrite -[in (_ <= 2)%N]rankH p_rank_le_rank -prankP0 p_rankS. +have ntP1: P1 != 1%g by rewrite -rank_gt0 ltnW // -prankP1 p_rank_le_rank. have [_ _ [U0 [sU0U expU0 frobHU0]]] := LtypeF. -have nP1U0: U0 \subset 'N(P1). - by rewrite (char_norm_trans chP1H) ?(subset_trans sU0U). -rewrite subn1 -prankP1 p_rank_abelem // -card_pgroup //. -have frobP1U0 := Frobenius_subl ntP1 (char_sub chP1H) nP1U0 frobHU0. +have nP1U0: U0 \subset 'N(P1) by rewrite (subset_trans sU0U) 2?gFnorm_trans. +rewrite subn1 -prankP1 p_rank_abelem -?card_pgroup //= -/P1 ?abelem_pgroup //. +have frobP1U0 := Frobenius_subl ntP1 sP1H nP1U0 frobHU0. apply: dvdn_trans (Frobenius_dvd_ker1 frobP1U0). -by have:= piUq; rewrite -expU0 pi_of_exponent mem_primes => /and3P[]. +by do [rewrite -expU0 pi_of_exponent mem_primes => /and3P[] //] in piUq. Qed. Let Ltype1 : FTtype L == 1%N. Proof. exact: FT_Frobenius_type1 frobL. Qed. @@ -745,8 +752,7 @@ have nP0A: A \subset 'N(P0). rewrite [_ :&: _](sub_pHall sylP0_HM) ?setSI ?pcore_sub //. by rewrite (pgroupS (subsetIl _ _)) ?pcore_pgroup. by rewrite subsetI sP0Hp. - have chHpL: 'O_p(H) \char L := char_trans (pcore_char p H) (Fcore_char L). - by rewrite normsI ?(char_norm_trans chHpL) ?normsG // cycle_subG. + by rewrite normsI ?gFnorm_trans ?normsG // cycle_subG. apply: wlog_neg => piH'q. have coHQ: coprime #|H| #|A| by rewrite -orderE coprime_pi' // oz pnatE. have frobP0A: [Frobenius P0 <*> A = P0 ><| A]. @@ -823,22 +829,20 @@ have sZP0: 'Z(P) \subset P0. by rewrite -cent_set1 setIS ?centS // sub1set (subsetP sP0P). by rewrite -defP0 setIS // (subset_trans _ sNxM) // cents_norm ?cent_cycle. have ntT: T :!=: 1%g. - rewrite Ohm1_eq1 center_nil_eq1 ?(pgroup_nil pP) // (subG1_contra sP0P) //. - by apply/trivgPn; exists x. + rewrite Ohm1_eq1 center_nil_eq1 ?(pgroup_nil pP) //. + by apply/trivgPn; exists x; rewrite ?(subsetP sP0P). have [_ sEL _ nHE tiHE] := sdprod_context defL. -have charTP: T \char P := char_trans (Ohm_char 1 _) (center_char P). have{ntT} [V minV sVT]: {V : {group gT} | minnormal V E & V \subset T}. - apply: mingroup_exists; rewrite ntT (char_norm_trans charTP) //. - exact: char_norm_trans (pcore_char p H) nHE. + by apply: mingroup_exists; rewrite ntT 3?gFnorm_trans. have abelT: p.-abelem T by rewrite Ohm1_abelem ?center_abelian ?(pgroupS sZP0). -have sTP := subset_trans (Ohm_sub 1 _) sZP0. +have sTP0: T \subset P0 by apply: gFsub_trans. have rankT: ('r_p(T) <= 2)%N by rewrite -prankP0 p_rankS. have [abelV /andP[ntV nVE]] := (abelemS sVT abelT, mingroupp minV). have pV := abelem_pgroup abelV; have [pr_p _ [n oV]] := pgroup_pdiv pV ntV. have frobHE: [Frobenius L = H ><| E] by rewrite /E; case: (sigW _). have: ('r_p(V) <= 2)%N by rewrite (leq_trans (p_rankS p sVT)). rewrite (p_rank_abelem abelV) // oV pfactorK // ltnS leq_eqVlt ltnS leqn0 orbC. -have sVH := subset_trans sVT (subset_trans (char_sub charTP) sPH). +have sVH: V \subset H by rewrite (subset_trans sVT) 3?gFsub_trans. have regVE: 'C_E(V) = 1%g. exact: cent_semiregular (Frobenius_reg_compl frobHE) sVH ntV. case/pred2P=> dimV; rewrite {n}dimV in oV. @@ -1168,8 +1172,7 @@ have lb_psiM: '[rhoM psi] >= #|K :\: K'|%:R / #|M|%:R * e.-1%:R ^+ 2. have psi_xg: (psi (x * g)%g == e%:R %[mod 1 - eps])%A. have [-> // _] := rhoL_psi; rewrite -[x]mulg1 -chi1. rewrite (vchar_ker_mod_prim prim_eps) ?group1 ?(seqInd_vcharW Schi) //. - rewrite (subsetP _ _ P01x) // (subset_trans (Ohm_sub 1 _)) //. - by rewrite (subset_trans sP0H) ?gFsub. + by rewrite (subsetP _ _ P01x) // gFsub_trans ?(subset_trans sP0H) ?gFsub. have{psi_xg} /dvdCP[a Za /(canRL (subrK _))->]: (p %| psi g - e%:R)%C. rewrite (int_eqAmod_prime_prim prim_eps) ?rpredB ?rpred_nat // eqAmod0. apply: eqAmod_trans psi_xg; rewrite eqAmod_sym. @@ -1190,7 +1193,7 @@ have lb_psiM: '[rhoM psi] >= #|K :\: K'|%:R / #|M|%:R * e.-1%:R ^+ 2. have irrS: {subset calS <= irr L} by have [] := FT_Frobenius_coherence maxL. have lb_psiL: '[rhoL psi] >= 1 - e%:R / #|H|%:R. have irr_chi := irrS _ Schi. - have Sgt1: (1 < size calS)%N by apply: seqInd_nontrivial (mFT_odd _) _ _ Schi. + have Sgt1: (1 < size calS)%N by apply: seqInd_nontrivial (mFT_odd L) _ Schi. have De: #|L : H| = e by rewrite -(index_sdprod defL). have [|_] := Dade_Ind1_sub_lin cohS_H Sgt1 irr_chi Schi; rewrite De //=. by rewrite -De odd_Frobenius_index_ler ?mFT_odd // => -[_ _ []//]. @@ -1247,7 +1250,7 @@ have{lb_psiM lb_psiL ub_rhoML ubM} ubK: (#|K / K'|%g < 4)%N. have [U [MtypeF _]] := FTtypeP 1 maxM Mtype1. have{U MtypeF} [_ _ [U0 [sU0ML expU0 frobU0]]] := compl_of_typeF defM MtypeF. have [/sdprodP[_ _ nKU0 tiKU0] ntK _ _ _] := Frobenius_context frobU0. -have nK'U0: U0 \subset 'N(K') := char_norm_trans (der_char 1 K) nKU0. +have nK'U0: U0 \subset 'N(K') by apply: gFnorm_trans. have frobU0K': [Frobenius K <*> U0 / K' = (K / K') ><| (U0 / K')]%g. have solK: solvable K by rewrite ?nilpotent_sol ?Fcore_nil. rewrite Frobenius_proper_quotient ?(sol_der1_proper solK) // /(_ <| _). @@ -1257,10 +1260,9 @@ have isoU0: U0 \isog U0 / K'. have piU0p: p \in \pi(U0 / K')%g. rewrite /= -(card_isog isoU0) -pi_of_exponent expU0 pi_of_exponent. rewrite mem_primes pr_p cardG_gt0 /= -ox order_dvdG // (subsetP _ _ P01x) //. - rewrite (subset_trans (Ohm_sub 1 _)) // subsetI sP0M. - by rewrite (subset_trans sP0H) ?gFsub. + by rewrite gFsub_trans // subsetI sP0M (subset_trans sP0H) ?gFsub. have /(Cauchy pr_p)[z U0z oz]: p %| #|U0 / K'|%g. - by rewrite mem_primes in piU0p; case/and3P: piU0p. + by do [rewrite mem_primes => /and3P[]//] in piU0p. have frobKz: [Frobenius (K / K') <*> <[z]> = (K / K') ><| <[z]>]%g. rewrite (Frobenius_subr _ _ frobU0K') ?cycle_subG //. by rewrite cycle_eq1 -order_gt1 oz ltnW. @@ -1289,21 +1291,20 @@ set K := M`_\F => maxM Mtype1; have [U [MtypeF _]] := FTtypeP 1 maxM Mtype1. have hallU: \pi(K)^'.-Hall(M) U. by rewrite -(compl_pHall U (Fcore_Hall M)) sdprod_compl; have [[]] := MtypeF. apply: FrobeniusWker (U) _ _; have{MtypeF} [_ -> _] := typeF_context MtypeF. -apply/forall_inP=> P0 /SylowP[p _ sylP0]. -rewrite (odd_pgroup_rank1_cyclic (pHall_pgroup sylP0)) ?mFT_odd // leqNgt. -apply/negP=> prankP0. -have piUp: p \in \pi(U) by rewrite -p_rank_gt0 -(p_rank_Sylow sylP0) ltnW. +apply/forall_inP=> P /SylowP[p _ sylP]. +rewrite (odd_pgroup_rank1_cyclic (pHall_pgroup sylP)) ?mFT_odd // leqNgt. +apply/negP=> prankP. +have piUp: p \in \pi(U) by rewrite -p_rank_gt0 -(p_rank_Sylow sylP) ltnW. have{piUp} K'p: p \in \pi(K)^' := pnatPpi (pHall_pgroup hallU) piUp. -have{U hallU sylP0} sylP0: p.-Sylow(M) P0 := subHall_Sylow hallU K'p sylP0. -have{P0 sylP0 prankP0} prankM: (1 < 'r_p(M))%N by rewrite -(p_rank_Sylow sylP0). -case/negP: K'p => /=. -elim: {p}_.+1 {-2}p M @K (ltnSn p) maxM Mtype1 prankM => // p IHp q M K. -rewrite ltnS leq_eqVlt => /predU1P[->{q} | /(IHp q M)//] maxM Mtype1 prankM. -apply/idPn; rewrite -p'groupEpi /= -/K => p'K. -have [P0 sylP0] := Sylow_exists p M. -have [] := non_Frobenius_FTtype1_witness maxM Mtype1 prankM p'K sylP0. -move=> abP0 prankP0 [L [maxL sP0_Ls [x P01s_x []]]]. -exact: (FTtype1_nonFrobenius_contradiction IHp) P01s_x. +have{U hallU sylP} sylP: p.-Sylow(M) P := subHall_Sylow hallU K'p sylP. +have{P sylP prankP} prankM: (1 < 'r_p(M))%N by rewrite -(p_rank_Sylow sylP). +case/negP: K'p => /=; move: {2}p.+1 (ltnSn p) => n ltpn. +elim: n => // n IHn in p M @K ltpn maxM Mtype1 prankM *. +move: ltpn; rewrite ltnS leq_eqVlt => /predU1P[Dp | /IHn-> //]. +apply/idPn=> p'K; rewrite -p'groupEpi /= -/K -{n}Dp in p'K IHn. +have [P sylP] := Sylow_exists p M. +case/non_Frobenius_FTtype1_witness: (sylP) => // cPP prankP [L [maxL sPLs]]. +by case=> x P1s_x []; apply: (FTtype1_nonFrobenius_contradiction IHn) P1s_x. Qed. (* This is Peterfalvi, Theorem (12.17). *) diff --git a/mathcomp/odd_order/PFsection13.v b/mathcomp/odd_order/PFsection13.v index 58e0142..6547dfd 100644 --- a/mathcomp/odd_order/PFsection13.v +++ b/mathcomp/odd_order/PFsection13.v @@ -1,16 +1,30 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime binomial ssralg poly finset. +From mathcomp Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +From mathcomp Require Import gfunctor gproduct center cyclic commutator gseries nilpotent. +From mathcomp Require Import pgroup sylow hall abelian maximal frobenius. +From mathcomp Require Import matrix mxalgebra mxrepresentation mxabelem vector. +From mathcomp Require Import BGsection1 BGsection3 BGsection7. +From mathcomp Require Import BGsection14 BGsection15 BGsection16. +From mathcomp Require Import ssrnum rat algC cyclotomic algnum. +From mathcomp Require Import classfun character integral_char inertia vcharacter. +From mathcomp Require Import PFsection1 PFsection2 PFsection3 PFsection4. +From mathcomp Require Import PFsection5 PFsection6 PFsection7 PFsection8 PFsection9. +From mathcomp Require Import PFsection10 PFsection11 PFsection12. (******************************************************************************) @@ -168,7 +182,7 @@ Lemma Ptype_Fcore_kernel_trivial : H0 :=: 1%g. Proof. have [/type2facts[_ oP _]| /type34ker1[]//] := boolP (FTtype S == 2). have [/and3P[]] := Ptype_Fcore_kernel_exists maxS StypeP notStype5. -case/maxgroupp/andP=> /proper_sub sH0P nH0S /subset_trans/(_ nH0S)nH0P _ _. +case/maxgroupp/andP=> /proper_sub-sH0P nH0S /subset_trans/(_ nH0S)nH0P _ _. apply: card1_trivg; rewrite -(divg_indexS sH0P) -card_quotient //. have [_ _ ->] := Ptype_Fcore_factor_facts maxS StypeP notStype5. by rewrite pHbar_p -{}oP divnn ?cardG_gt0. @@ -469,7 +483,7 @@ have Hzeta0 := sS1H _ S1zeta0. have dH_1 zeta: zeta \in calH -> (zeta - zeta0) 1%g == 0. by move=> Tzeta; rewrite 2!cfunE !calHuq // subrr eqxx. have H1dzeta zeta: zeta \in calH -> zeta - zeta0 \in 'CF(S, H^#). - have HonH: {subset calH <= 'CF(S, H)} by exact: seqInd_on. + have HonH: {subset calH <= 'CF(S, H)} by apply: seqInd_on. by move=> Hzeta; rewrite cfun_onD1 rpredB ?HonH ?dH_1. pose calH1 := rem zeta1 (rem zeta0 (filter [mem calS1] calH)). pose calH2 := filter [predC calS1] calH. @@ -1533,10 +1547,9 @@ have ->: 'N(W2) = 'N_S(W2). rewrite -{1}(sdprodW defS) setIC -group_modr ?cents_norm 1?centsC //=. rewrite mulG_subG joing_subr /= -(sdprodW defPU) setIC. rewrite -group_modl ?cents_norm //= mulG_subG joing_subl /= andbT. -set K := 'N_U(W2). -have nPKW1: K <*> W1 \subset 'N(P). - rewrite (subset_trans _ (gFnorm _ _)) // -(sdprodWY defS) genS ?setSU //. - by rewrite subIset ?sUPU. +set K := 'N_U(W2); have sKPU: K \subset PU by rewrite subIset ?sUPU. +have{sKPU} nPKW1: K <*> W1 \subset 'N(P). + by rewrite gFnorm_trans ?normsG // -(sdprodWY defS) genS ?setSU. have nW2KW1: K <*> W1 \subset 'N(W2). by rewrite join_subG subsetIr cents_norm // centsC. have coPKW1: coprime #|P| #|K <*> W1|. @@ -2076,7 +2089,7 @@ have RealGammaL: cfReal GammaL. rewrite rmorphB /= cfAutInd rmorph1 addrC opprB addrA subrK. by rewrite (cfConjC_Dade_coherent cohL) ?mFT_odd // -raddfB Dtau1 // ZsubL. have:= Dade_Ind1_sub_lin cohL _ irr_phi Lphi; rewrite -/betaL -/tauL -/calL. -rewrite (seqInd_nontrivial _ _ _ irr_phi) ?odd_Frobenius_index_ler ?mFT_odd //. +rewrite (seqInd_nontrivial _ _ _ Lphi) ?odd_Frobenius_index_ler ?mFT_odd //. case=> // -[o_tauL_1 o_betaL_1 ZbetaL] ub_betaL _. have{o_tauL_1 o_betaL_1} o_GaL_1: '[GammaL, 1] = 0. by rewrite !cfdotBl cfnorm1 o_betaL_1 (orthoPr o_tauL_1) ?map_f ?subr0 ?subrr. diff --git a/mathcomp/odd_order/PFsection14.v b/mathcomp/odd_order/PFsection14.v index bd7ae60..0056312 100644 --- a/mathcomp/odd_order/PFsection14.v +++ b/mathcomp/odd_order/PFsection14.v @@ -1,16 +1,30 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime binomial ssralg poly finset. +From mathcomp Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +From mathcomp Require Import gfunctor gproduct center cyclic commutator gseries nilpotent. +From mathcomp Require Import pgroup sylow hall abelian maximal frobenius. +From mathcomp Require Import matrix mxalgebra mxrepresentation mxabelem vector. +From mathcomp Require Import BGsection1 BGsection3 BGsection7. +From mathcomp Require Import BGsection14 BGsection15 BGsection16 BGappendixC. +From mathcomp Require Import ssrnum rat algC cyclotomic algnum. +From mathcomp Require Import classfun character integral_char inertia vcharacter. +From mathcomp Require Import PFsection1 PFsection2 PFsection3 PFsection4. +From mathcomp Require Import PFsection5 PFsection6 PFsection7 PFsection8 PFsection9. +From mathcomp Require Import PFsection10 PFsection11 PFsection12 PFsection13. (******************************************************************************) @@ -331,7 +345,7 @@ have [[Itau1 Ztau1] Dtau1] := cohM. have o1M: orthonormal (map tau1M calM). apply: map_orthonormal Itau1 _. exact: sub_orthonormal (undup_uniq _) (irr_orthonormal M). -have Lgt1: (1 < size calL)%N by apply: seqInd_nontrivial (mFT_odd _ ) _ _ Lphi. +have Lgt1: (1 < size calL)%N by apply: seqInd_nontrivial (mFT_odd _ ) _ Lphi. have [[_ _]] := Dade_Ind1_sub_lin cohL Lgt1 irr_phi Lphi phi1. rewrite -/tauL -/betaL -/calL => ZbetaL [Gamma [_ _ [b _ Dbeta]]]. rewrite odd_Frobenius_index_ler ?mFT_odd // -/u => -[]// [_ ub_Ga] _ nz_a. @@ -343,7 +357,7 @@ suffices ->: '[X] = (a / v) ^+ 2 * (\sum_(xi <- calM) xi 1%g ^+ 2 / '[xi]). rewrite exprMn !mulrA divfK ?neq0CiG // mulrAC -mulrA. by rewrite ler_pemull ?sqr_Cint_ge1 // divr_ge0 ?ler0n. have [_ -> defX] := orthonormal_span o1M M_X. -have Mgt1: (1 < size calM)%N by apply: seqInd_nontrivial (mFT_odd _ ) _ _ Mpsi. +have Mgt1: (1 < size calM)%N by apply: seqInd_nontrivial (mFT_odd _ ) _ Mpsi. have [[oM1 _ _] _ _] := Dade_Ind1_sub_lin cohM Mgt1 irr_psi Mpsi psi1. rewrite exprMn -(Cint_normK Za) -[v]normr_nat -normfV -/v mulr_sumr. rewrite defX cfnorm_sum_orthonormal // big_map; apply: eq_big_seq => xi Mxi. @@ -596,7 +610,7 @@ have{x ntx R0x ntCPx} sZR_R0: 'Z(R) \subset R0. by rewrite subsetI sR0R sub_cent1 (subsetP abR0). by rewrite subIset ?sCxS ?orbT. pose R1 := 'Ohm_1('Z(R))%G; pose m := logn r #|R1|. -have sR10: R1 \subset R0 by rewrite (subset_trans (Ohm_sub 1 _)). +have sR10: R1 \subset R0 by apply: gFsub_trans. have oR1: #|R1| = (r ^ m)%N by rewrite -card_pgroup ?(pgroupS sR10). have{sZR_R0 rR0_2} m12: pred2 1%N 2 m. transitivity (0 < m < 1 + 2)%N; first by rewrite -mem_iota !inE. @@ -605,9 +619,8 @@ have{sZR_R0 rR0_2} m12: pred2 1%N 2 m. by rewrite (subG1_contra sR0R) // -rank_gt0 rR0_2. have [y Qy defLy] := defL; have [_ _ /joing_subP[_ nHW2y] _] := sdprodP defLy. have chR1H: R1 \char H. - apply: char_trans (char_trans (Ohm_char 1 _) (center_char R)) _. - by rewrite (nilpotent_Hall_pcore (Fcore_nil L) sylR) gFchar. -have nR1W2y: W2 :^ y \subset 'N(R1) := char_norm_trans chR1H nHW2y. + by rewrite !gFchar_trans // (nilpotent_Hall_pcore (Fcore_nil L) sylR) gFchar. +have nR1W2y: W2 :^ y \subset 'N(R1) by apply: char_norm_trans chR1H nHW2y. have regR1W2y: semiregular R1 (W2 :^ y). have /Frobenius_reg_ker regHW12y := set_Frobenius_compl defLy frobL. exact: semiregularS (char_sub chR1H) (joing_subr _ _) regHW12y. @@ -619,10 +632,10 @@ apply: leq_trans (_ : p.-1 <= r)%N. by rewrite -divn2 ltn_divLR // -{1}[p.-1]muln1 -(subnKC pgt2) ltn_pmul2l. have: p %| (r ^ m).-1. by have:= regular_norm_dvd_pred nR1W2y regR1W2y; rewrite cardJg oR1. -rewrite -[p.-1]subn1 leq_subLR predn_exp Euclid_dvdM // => /orP[]/dvdn_leq. - by rewrite -(subnKC (prime_gt1 pr_r)) => /implyP/leq_trans->; rewrite 2?ltnW. -move/implyP; case/pred2P: m12 => ->; rewrite !big_ord_recl big_ord0 ?addn0 //=. -by rewrite -(subnKC pgt2). +rewrite -[p.-1]subn1 leq_subLR predn_exp Euclid_dvdM // => /orP[]/dvdn_leq. + by rewrite -(subnKC (prime_gt1 pr_r)) => /(_ isT)/leq_trans->; rewrite 2?ltnW. +case/pred2P: m12 => ->; rewrite ?(big_ord_recl 1) big_ord1 => /(_ isT) //. +by move/leq_trans->. Qed. (* This is Peterfalvi (14.7). *) diff --git a/mathcomp/odd_order/PFsection2.v b/mathcomp/odd_order/PFsection2.v index 9eef9e8..9ec3104 100644 --- a/mathcomp/odd_order/PFsection2.v +++ b/mathcomp/odd_order/PFsection2.v @@ -1,9 +1,16 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime ssralg poly finset center. +From mathcomp Require Import fingroup morphism perm automorphism quotient action zmodp. +From mathcomp Require Import gfunctor gproduct cyclic pgroup frobenius ssrnum. +From mathcomp Require Import matrix mxalgebra mxrepresentation vector algC classfun character. +From mathcomp Require Import inertia vcharacter PFsection1. (******************************************************************************) @@ -55,7 +62,7 @@ Lemma partition_cent_rcoset (H : {group gT}) g (C := 'C_H[g]) (Cg := C :* g) : partition (Cg :^: H) (H :* g) /\ #|Cg :^: H| = #|H : C|. Proof. move=> nHg coHg; pose pi := \pi(#[g]). -have notCg0: Cg != set0 by apply/set0Pn; exists g; exact: rcoset_refl. +have notCg0: Cg != set0 by apply/set0Pn; exists g; apply: rcoset_refl. have id_pi: {in Cg, forall u, u.`_ pi = g}. move=> _ /rcosetP[u /setIP[Hu cgu] ->]; rewrite consttM; last exact/cent1P. rewrite (constt_p_elt (pgroup_pi _)) (constt1P _) ?mul1g //. @@ -238,7 +245,7 @@ Qed. Lemma Dade_support_norm : G \subset 'N(Atau). Proof. -by rewrite norms_bigcup //; apply/bigcapsP=> a _; exact: class_support_norm. +by rewrite norms_bigcup //; apply/bigcapsP=> a _; apply: class_support_norm. Qed. Lemma Dade_support_normal : Atau <| G. @@ -253,7 +260,7 @@ Fact Dade_subproof (alpha : 'CF(L)) : Proof. rewrite genGid; apply: intro_class_fun => [x y Gx Gy | x notGx]. congr (oapp _ _); apply: eq_pick => a; rewrite memJ_norm //. - apply: subsetP Gy; exact: class_support_norm. + by apply: subsetP Gy; apply: class_support_norm. case: pickP => // a /andP[Aa Ha_u]. by rewrite (subsetP Dade_support_sub) // in notGx; apply/bigcupP; exists a. Qed. @@ -355,9 +362,9 @@ have {1}<-: cover P_A = A. by exists (a ^: L); rewrite ?class_refl // -rLid; do 2!apply: mem_imset. have [tiP_A injFA]: trivIset P_A /\ {in T &, injective (class^~ L)}. apply: trivIimset => [_ _ /imsetP[a Aa ->] /imsetP[b Ab ->] |]; last first. - by apply/imsetP=> [[a _ /esym/eqP/set0Pn[]]]; exists a; exact: class_refl. + by apply/imsetP=> [[a _ /esym/eqP/set0Pn[]]]; exists a; apply: class_refl. rewrite !rLid; apply: contraR => /pred0Pn[c /andP[/=]]. - by do 2!move/class_transr <-. + by do 2!move/class_eqP <-. rewrite big_trivIset //= big_imset {P_A tiP_A injFA}//=. apply: canRL (mulKf (neq0CG G)) _; rewrite mulrA big_distrr /=. apply: eq_bigr => a /sTA=> {T sTA}Aa. @@ -446,7 +453,7 @@ Fact Dade_restrM B : {in 'M(B) &, {morph Dade_restrm B : x y / x * y}%g}. Proof. rewrite /Dade_restrm; case: ifP => calP_B; last exact: morphM. have defM := Dade_set_sdprod calP_B; have [nsHM _ _ _ _] := sdprod_context defM. -by apply: remgrM; first exact: sdprod_compl. +by apply: remgrM; first apply: sdprod_compl. Qed. Canonical Dade_restr_morphism B := Morphism (@Dade_restrM B). Definition Dade_cfun_restriction B := @@ -573,7 +580,7 @@ rewrite big1 ?add0r => [|x /andP[calAx not_nBaLx]]; last first. apply: supp_aBgP => //; apply: contra not_nBaLx. set b := fBg x => /andP[Ab Hb_g]; have [Gx MBx] := setIdP calAx. rewrite inE mem_remgr ?mulHNB //; apply/imsetP/Dade_support1_TI => //. - by apply/pred0Pn; exists g; exact/andP. + by apply/pred0Pn; exists g; apply/andP. rewrite (partition_big fBg (mem nBaL)) /= => [|x]; last by case/andP. apply: eq_bigr => b; case/setIP=> Nb aLb; rewrite mulr_natr -sumr_const. apply: eq_big => x; rewrite ![x \in _]inE -!andbA. @@ -581,7 +588,7 @@ apply: eq_big => x; rewrite ![x \in _]inE -!andbA. by rewrite mem_rcoset mem_divgr ?mulHNB. suffices ->: fBg x = b. by rewrite inE Nb (subsetP _ _ HBb_gx) // -mulHNB mulgS ?sub1set. - by rewrite /fBg; have [h Hh ->] := rcosetP HBb_gx; exact: remgrMid. + by rewrite /fBg; have [h Hh ->] := rcosetP HBb_gx; apply: remgrMid. move/and4P=> [_ Mgx _ /eqP def_fx]. rewrite rDadeE // Mgx -/(fBg x) def_fx; case/imsetP: aLb => y Ly ->. by rewrite cfunJ // (subsetP sAL). @@ -605,9 +612,9 @@ transitivity (- (\sum_(B in calP) n1 B * aa1 B)); last first. by move=> _ /imsetP[x Lx ->]; rewrite dBJ. have dB: B \in calP := dB1L B B1L_B. rewrite (eq_bigl (mem (B :^: L))) => [|B2 /=]; last first. - apply/andP/idP=> [[_ /eqP <-] | /(orbit_trans B1L_B) B1L_B2]. + apply/andP/idP=> [[_ /eqP <-] | /orbit_trans/(_ B1L_B)-B1L_B2]. by rewrite orbit_sym (mem_repr B2) ?orbit_refl. - by rewrite [B2 :^: L](orbit_transl B1L_B2) -defB dB1L. + by rewrite [B2 :^: L](orbit_eqP B1L_B2) -defB dB1L. rewrite (eq_bigr (fun _ => n1 B * aa1 B)) => [|_ /imsetP[x Lx ->]]. rewrite cfunE sumr_const -mulr_natr mulrAC card_orbit astab1Js divfK //. by rewrite pnatr_eq0 -lt0n indexg_gt0. @@ -660,7 +667,7 @@ rewrite (bigID [pred B : {set gT} | a \in B]) /= mulrDl addrA. apply: canRL (subrK _) _; rewrite -mulNr -sumrN; congr (_ + _ * _). symmetry. rewrite (reindex_onto (fun B => a |: B) (fun B => B :\ a)) /=; last first. - by move=> B; case/andP=> _; exact: setD1K. + by move=> B; case/andP=> _; apply: setD1K. symmetry; apply: eq_big => B. rewrite setU11 andbT -!andbA; apply/and3P/and3P; case. do 2![case/setIdP] => sBA ntB /setIP[La nBa] _ notBa. @@ -672,7 +679,7 @@ symmetry; apply: eq_big => B. rewrite eq_sym eqEcard subsetUl cards1 (cardsD1 a) setU11 ltnS leqn0 /=. rewrite cards_eq0 => notB0 /eqP defB. have notBa: a \notin B by rewrite -defB setD11. - split=> //; last by apply: contraNneq notBa => ->; exact: set11. + split=> //; last by apply: contraNneq notBa => ->; apply: set11. rewrite !inE sBA La -{1 3}defB notB0 subsetD1 sBa_aB. by rewrite mem_conjg /(a ^ _) invgK mulgA mulgK. do 2![case/andP] => /setIdP[dB Na] _ notBa. @@ -728,7 +735,7 @@ Qed. (* This summarizes Peterfalvi (2.6). *) Lemma Dade_Zisometry : {in 'Z[irr L, A], isometry Dade, to 'Z[irr G, G^#]}. Proof. -split; first by apply: sub_in2 Dade_isometry; exact: zchar_on. +split; first by apply: sub_in2 Dade_isometry; apply: zchar_on. by move=> phi Zphi; rewrite /= zchar_split Dade_vchar ?Dade_cfun. Qed. @@ -747,7 +754,7 @@ Proof. have [/andP[sAL nAL] notA_1 sLG conjAG [H defCa coHL]] := ddA. have nsA1L: A1 <| L by rewrite /normal (subset_trans sA1A). split; rewrite ?(contra (@ssA1A _)) //; first exact: sub_in2 conjAG. -by exists H; [exact: sub_in1 defCa | exact: sub_in2 coHL]. +by exists H; [apply: sub_in1 defCa | apply: sub_in2 coHL]. Qed. Local Notation ddA1 := restr_Dade_hyp. diff --git a/mathcomp/odd_order/PFsection3.v b/mathcomp/odd_order/PFsection3.v index 063aacc..c5633a2 100644 --- a/mathcomp/odd_order/PFsection3.v +++ b/mathcomp/odd_order/PFsection3.v @@ -1,11 +1,20 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime ssralg matrix poly finset. +From mathcomp Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +From mathcomp Require Import gfunctor center gproduct cyclic pgroup abelian frobenius. +From mathcomp Require Import mxalgebra mxrepresentation vector falgebra fieldext galois. +From mathcomp Require Import ssrnum rat algC algnum classfun character. +From mathcomp Require Import integral_char inertia vcharacter. +From mathcomp Require Import PFsection1 PFsection2. (******************************************************************************) @@ -634,7 +643,7 @@ Proof. by case: set_cl => th1 [Uth /Uth]. Qed. Lemma unsat_wlog_cases th1 th2 : (unsat th1 -> unsat th2) -> unsat th1 -> (true /\ unsat th1) /\ unsat th2. -Proof. by move=> Uth2 Uth1; split; last exact: Uth2. Qed. +Proof. by move=> Uth2 Uth1; split; last apply: Uth2. Qed. (* Extend the orthonormal basis *) diff --git a/mathcomp/odd_order/PFsection4.v b/mathcomp/odd_order/PFsection4.v index 816ac05..887e6dd 100644 --- a/mathcomp/odd_order/PFsection4.v +++ b/mathcomp/odd_order/PFsection4.v @@ -1,9 +1,16 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime ssralg poly finset fingroup. +From mathcomp Require Import morphism perm automorphism quotient action gfunctor gproduct. +From mathcomp Require Import center commutator zmodp cyclic pgroup nilpotent hall frobenius. +From mathcomp Require Import matrix mxalgebra mxrepresentation vector ssrnum algC classfun. +From mathcomp Require Import character inertia vcharacter PFsection1 PFsection2 PFsection3. (******************************************************************************) @@ -649,7 +656,7 @@ rewrite (card_afix_irr_classes Lz actsL_KK) => [|k x y Kx /=]; last first. apply: leq_trans (subset_leq_card _) (leq_imset_card (class^~ K) _). apply/subsetP=> _ /setIP[/imsetP[x Kx ->] /afix1P/normP nxKz]. suffices{Kx} /pred0Pn[t /setIP[xKt czt]]: #|'C_(x ^: K)[z]| != 0%N. - rewrite -(class_transr xKt); apply: mem_imset; have [y Ky Dt] := imsetP xKt. + rewrite -(class_eqP xKt); apply: mem_imset; have [y Ky Dt] := imsetP xKt. by rewrite -(@prKW1 z) ?(czt, inE) ?ntz // Dt groupJ. have{coKp}: ~~ (p %| #|K|) by rewrite -prime_coprime // coprime_sym. apply: contraNneq => /(congr1 (modn^~ p))/eqP; rewrite mod0n. @@ -776,7 +783,7 @@ Lemma prDade_irr_on k : Proof. move=> kerH'i; apply/cfun_onP=> g; rewrite !inE => /norP[ntg A'g]. have [Kg | /cfun0-> //] := boolP (g \in K). -apply: not_in_ker_char0 (normalS _ _ nsHL) kerH'i _ => //. +apply: irr_reg_off_ker_0 (normalS _ _ nsHL) kerH'i _ => //. apply/trivgP/subsetP=> h /setIP[Hh cgh]; apply: contraR A'g => nth. apply/(subsetP sIH_A)/bigcupP; exists h; first exact/setDP. by rewrite 3!inE ntg Kg cent1C. @@ -894,7 +901,7 @@ Qed. (* in part (a). *) Theorem uniform_prTIred_coherent k (calT := uniform_prTIred_seq ptiWL k) : k != 0 -> - (*a*) [/\ pairwise_orthogonal calT, ~~ has cfReal calT, conjC_closed calT, + (*a*) [/\ pairwise_orthogonal calT, ~~ has cfReal calT, cfConjC_closed calT, 'Z[calT, L^#] =i 'Z[calT, A] & exists2 psi, psi != 0 & psi \in 'Z[calT, A]] (*b*) /\ (exists2 tau1 : {linear 'CF(L) -> 'CF(G)}, @@ -903,7 +910,7 @@ Theorem uniform_prTIred_coherent k (calT := uniform_prTIred_seq ptiWL k) : /\ {in 'Z[calT, L^#], tau1 =1 tau}). Proof. have uniqT: uniq calT by apply/dinjectiveP; apply: in2W; apply: prTIred_inj. -have sTmu: {subset calT <= codom mu_} by exact: image_codom. +have sTmu: {subset calT <= codom mu_} by apply: image_codom. have oo_mu: pairwise_orthogonal (codom mu_). apply/pairwise_orthogonalP; split=> [|_ _ /codomP[j1 ->] /codomP[j2 ->]]. apply/andP; split; last by apply/injectiveP; apply: prTIred_inj. @@ -911,7 +918,7 @@ have oo_mu: pairwise_orthogonal (codom mu_). by rewrite cfdot_prTIred; case: (j1 =P j2) => // -> /eqP. have real'T: ~~ has cfReal calT. by apply/hasPn=> _ /imageP[j /andP[nzj _] ->]; apply: prTIred_not_real. -have ccT: conjC_closed calT. +have ccT: cfConjC_closed calT. move=> _ /imageP[j Tj ->]; rewrite -prTIred_aut image_f // inE aut_Iirr_eq0. by rewrite prTIred_aut cfunE conj_Cnat ?Cnat_char1 ?prTIred_char. have TonA: 'Z[calT, L^#] =i 'Z[calT, A]. @@ -938,7 +945,7 @@ have iso_f1: {in codom mu_, isometry f1, to 'Z[irr G]}. rewrite !f1mu cfdotZl cfdotZr rmorph_sign signrMK !cfdot_suml. apply: eq_bigr => i1 _; rewrite !cfdot_sumr; apply: eq_bigr => i2 _. by rewrite cfdot_cycTIiso cfdot_prTIirr. -have [tau1 Dtau1 Itau1] := Zisometry_of_iso oo_mu iso_f1. +have [tau1 Dtau1 Itau1] := Zisometry_of_iso (orthogonal_free oo_mu) iso_f1. exists tau1 => [j|]; first by rewrite Dtau1 ?codom_f ?f1mu. split=> [|psi]; first by apply: sub_iso_to Itau1 => //; apply: zchar_subset. rewrite zcharD1E => /andP[/zchar_expansion[//|z _ Dpsi] /eqP psi1_0]. diff --git a/mathcomp/odd_order/PFsection5.v b/mathcomp/odd_order/PFsection5.v index 0c3b1eb..2031785 100644 --- a/mathcomp/odd_order/PFsection5.v +++ b/mathcomp/odd_order/PFsection5.v @@ -1,10 +1,18 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime ssralg poly finset center. +From mathcomp Require Import fingroup morphism perm automorphism quotient action zmodp. +From mathcomp Require Import gfunctor gproduct cyclic pgroup frobenius. +From mathcomp Require Import matrix mxalgebra mxrepresentation vector ssrint. +From mathcomp Require Import ssrnum algC classfun character inertia vcharacter. +From mathcomp Require Import PFsection1 PFsection2 PFsection3 PFsection4. (******************************************************************************) @@ -26,7 +34,7 @@ Require Import PFsection1 PFsection2 PFsection3 PFsection4. (* isometry extending a subcoherent tau on 'Z[S] with *) (* size S = 2, then so is dual_iso nu. *) (* We provide a set of definitions that cover the various \cal S notations *) -(* introduces in Peterfalvi sections 5, 6, 7, and 9 to 14. *) +(* introduced in Peterfalvi sections 5, 6, 7, and 9 to 14. *) (* Iirr_ker K A == the set of all i : Iirr K such that the kernel of *) (* 'chi_i contains A. *) (* Iirr_kerD K B A == the set of all i : Iirr K such that the kernel of *) @@ -122,10 +130,10 @@ Lemma seqInd_uniq : uniq S. Proof. exact: undup_uniq. Qed. Lemma seqIndP phi : reflect (exists2 i, i \in calX & phi = 'Ind[L] 'chi_i) (phi \in S). -Proof. by rewrite mem_undup; exact: imageP. Qed. +Proof. by rewrite mem_undup; apply: imageP. Qed. Lemma seqInd_on : {subset S <= 'CF(L, K)}. -Proof. by move=> _ /seqIndP[i _ ->]; exact: cfInd_normal. Qed. +Proof. by move=> _ /seqIndP[i _ ->]; apply: cfInd_normal. Qed. Lemma seqInd_char : {subset S <= character}. Proof. by move=> _ /seqIndP[i _ ->]; rewrite cfInd_char ?irr_char. Qed. @@ -137,7 +145,7 @@ Lemma Cint_seqInd1 phi : phi \in S -> phi 1%g \in Cint. Proof. by rewrite CintE; move/Cnat_seqInd1->. Qed. Lemma seqInd_neq0 psi : psi \in S -> psi != 0. -Proof. by move=> /seqIndP[i _ ->]; exact: Ind_irr_neq0. Qed. +Proof. by move=> /seqIndP[i _ ->]; apply: Ind_irr_neq0. Qed. Lemma seqInd1_neq0 psi : psi \in S -> psi 1%g != 0. Proof. by move=> Spsi; rewrite char1_eq0 ?seqInd_char ?seqInd_neq0. Qed. @@ -205,23 +213,23 @@ Qed. Lemma sub_seqInd_on phi psi : phi \in S -> psi \in S -> psi 1%g *: phi - phi 1%g *: psi \in 'CF(L, K^#). -Proof. by move=> Sphi Spsi; exact: zchar_on (sub_seqInd_zchar Sphi Spsi). Qed. +Proof. by move=> Sphi Spsi; apply: zchar_on (sub_seqInd_zchar Sphi Spsi). Qed. Lemma size_irr_subseq_seqInd S1 : subseq S1 S -> {subset S1 <= irr L} -> (#|L : K| * size S1 = #|[set i | 'Ind 'chi[K]_i \in S1]|)%N. Proof. -move=> sS1S irrS1; rewrite (card_imset_Ind_irr nsKL) => [|i|i y]; first 1 last. +move=> sS1S irrS1; have uniqS1: uniq S1 := subseq_uniq sS1S seqInd_uniq. +rewrite (card_imset_Ind_irr nsKL) => [|i|i y]; first 1 last. - by rewrite inE => /irrS1. -- rewrite !inE => S1iG Ly; congr (_ \in S1): S1iG. - by apply: cfclass_Ind => //; apply/cfclassP; exists y; rewrite ?conjg_IirrE. -congr (_ * _)%N; rewrite -(size_map (@cfIirr _ _)) -(card_uniqP _); last first. - rewrite map_inj_in_uniq ?(subseq_uniq sS1S) ?seqInd_uniq //. - by apply: (@can_in_inj _ _ _ _ (tnth (irr L))) => phi /irrS1/cfIirrE. -apply: eq_card => s; apply/mapP/imsetP=> [[phi S1phi ->] | [i]]. - have /seqIndP[i _ Dphi] := mem_subseq sS1S S1phi. +- by rewrite !inE => *; rewrite conjg_IirrE -(cfConjgInd _ _ nsKL) ?cfConjg_id. +congr (_ * _)%N; transitivity #|map cfIirr S1|. + rewrite (card_uniqP _) ?size_map ?map_inj_in_uniq //. + exact: sub_in2 irrS1 _ (can_in_inj (@cfIirrE _ _)). +apply: eq_card => s; apply/idP/imsetP=> [/mapP[phi S1phi] | [i S1iG]] {s}->. + have /seqIndP[i _ Dphi]: phi \in S := mem_subseq sS1S S1phi. by exists i; rewrite ?inE -Dphi. -by rewrite inE => S1iG ->; exists ('Ind 'chi_i). +by apply: map_f; rewrite inE in S1iG. Qed. Section Beta. @@ -333,10 +341,19 @@ move=> phi Sphi /=; rewrite sub_aut_zchar ?seqInd_zchar ?cfAut_seqInd //. exact: seqInd_vcharW. Qed. +Lemma seqIndD_nonempty : H <| K -> M <| K -> M \proper H -> {phi | phi \in S}. +Proof. +move=> nsHK nsMK /andP[sMH ltMH]; pose X := Iirr_kerD H M. +suffices: \sum_(i in X) 'chi_i 1%g ^+ 2 > 0. + have [->|[i Xi]] := set_0Vmem X; first by rewrite big_set0 ltrr. + by exists ('Ind 'chi_i); apply/seqIndP; exists i. +by rewrite sum_Iirr_kerD_square ?mulr_gt0 ?gt0CiG ?subr_gt0 // ltr1n indexg_gt1. +Qed. + Hypothesis sHK : H \subset K. Lemma seqInd_sub : {subset S <= seqIndD K 1}. -Proof. by apply: seqIndS; exact: Iirr_kerDS (sub1G M) sHK. Qed. +Proof. by apply: seqIndS; apply: Iirr_kerDS (sub1G M) sHK. Qed. Lemma seqInd_ortho_Ind1 : {in S, forall phi, '[phi, 'Ind[L, K] 1] = 0}. Proof. @@ -372,7 +389,7 @@ have/cfunP/(_ 1%g) := scaled_cfResInd_sum_cfclass i nsKL. rewrite !cfunE sum_cfunE -def_phi cfResE // mulrAC => ->; congr (_ * _). rewrite reindex_cfclass //=; apply/esym/eq_big => j; last by rewrite !cfunE. rewrite (sameP (cfclass_Ind_irrP _ _ nsKL) eqP) -def_phi -mem_seqInd //. -by apply/andP/eqP=> [[/(nth_index 0){2}<- /eqP->] | -> //]; exact: nth_index. +by apply/andP/eqP=> [[Sj /eqP/(congr1 (nth 0 S))] | ->]; rewrite ?nth_index. Qed. Section Odd. @@ -381,7 +398,7 @@ Hypothesis oddL : odd #|L|. Lemma seqInd_conjC_ortho : {in S, forall phi, '[phi, phi^*] = 0}. Proof. -by move=> _ /seqInd_sub/seqIndC1P[i nzi ->]; exact: odd_induced_orthogonal. +by move=> _ /seqInd_sub/seqIndC1P[i nzi ->]; apply: odd_induced_orthogonal. Qed. Lemma seqInd_conjC_neq : {in S, forall phi, phi^* != phi}%CF. @@ -393,6 +410,13 @@ Qed. Lemma seqInd_notReal : ~~ has cfReal S. Proof. exact/hasPn/seqInd_conjC_neq. Qed. +Lemma seqInd_nontrivial chi : chi \in S -> (1 < size S)%N. +Proof. +move=> Schi; pose S2 := chi^*%CF :: chi. +have: {subset S2 <= S} by apply/allP/and3P; rewrite /= cfAut_seqInd. +by apply: uniq_leq_size; rewrite /= inE seqInd_conjC_neq. +Qed. + Variable chi : 'CF(L). Hypotheses (irr_chi : chi \in irr L) (Schi : chi \in S). @@ -408,14 +432,6 @@ rewrite !inE -(inj_eq irr_inj) conjC_IirrE -Dchi seqInd_conjC_neq //. by rewrite cfAut_seqInd Schi. Qed. -Lemma seqInd_nontrivial : (size S > 1)%N. -Proof. -apply: (@leq_trans (size [seq 'chi_i | i in [pred i | 'chi_i \in S]])). - by rewrite size_map -cardE -cardsE seqInd_nontrivial_irr. -apply: uniq_leq_size => [| _ /imageP[i Schi_i ->] //]. -exact/dinjectiveP/(in2W irr_inj). -Qed. - End Odd. End SeqIndD. @@ -445,18 +461,19 @@ Variables L G : {group gT}. (* really used in (7.8), where it is equivalent to the simpler condition *) (* (size S > 1). For us the empty S is coherent; this avoids duplicate *) (* work in some inductive proofs, e.g., subcoherent_norm - Lemma (5.4) - *) -(* belom. *) +(* below. *) (* - The preconditions for coherence (A < L, S < Z[irr L], and tau Z-linear *) (* on some E < Z[irr L]) are not part of the definition of "coherent". *) (* These will be captured as separate requirements; in particular in the *) (* Odd Order proof tau will always be C-linear on all of 'CF(L). *) (* - By contrast, our "coherent" only supplies an additive (Z-linear) *) -(* isometry, where the source text ambiguously specifies "linear" one. *) +(* isometry, where the source text ambiguously specifies a "linear" one. *) (* When S consists of virtual characters this implies the existence of *) (* a C-linear one: the linear extension of the restriction of the *) -(* isometry to a basis of the Z-module Z[S]; the latter being given by *) -(* the Smith normal form (see intdiv.v). The weaker requirement lets us *) -(* use the dual_iso construction when size S = 2. *) +(* isometry to a basis of the Z-module Z[S]. The latter can be found from *) +(* the Smith normal form (see intdiv.v) of the coordinate matrix of S. *) +(* The weaker Z-linearity lets us use the dual_iso construction when *) +(* size S = 2. *) (* Finally, note that although we have retained the A parameter, in the *) (* sequel we shall always take A = L^#, as in the text it is always the case *) (* that Z[S, A] = Z[S, L^#]. *) @@ -469,7 +486,7 @@ Definition coherent S A tau := exists tau1, coherent_with S A tau tau1. (* The Z-linearity constraint on tau will be expressed by an additive or *) (* linear structure on tau. *) Definition subcoherent S tau R := - [/\ (*a*) [/\ {subset S <= character}, ~~ has cfReal S & conjC_closed S], + [/\ (*a*) [/\ {subset S <= character}, ~~ has cfReal S & cfConjC_closed S], (*b*) {in 'Z[S, L^#], isometry tau, to 'Z[@irr gT G, G^#]}, (*c*) pairwise_orthogonal S, (*d*) {in S, forall xi : 'CF(L : {set gT}), @@ -493,7 +510,7 @@ Lemma subgen_coherent S1 S2 A tau: Proof. move/zchar_trans=> sS21 [tau1 [[Itau1 Ztau1] def_tau]]. exists tau1; split; last exact: sub_in1 def_tau. -by split; [exact: sub_in2 Itau1 | exact: sub_in1 Ztau1]. +by split; [apply: sub_in2 Itau1 | apply: sub_in1 Ztau1]. Qed. Lemma subset_coherent S1 S2 A tau: @@ -507,7 +524,7 @@ Lemma subset_coherent_with S1 S2 A tau (tau1 : {additive 'CF(L) -> 'CF(G)}) : coherent_with S1 A tau tau1. Proof. move=> /zchar_subset sS12 [Itau1 Dtau1]. -by split=> [|xi /sS12/Dtau1//]; exact: sub_iso_to Itau1. +by split=> [|xi /sS12/Dtau1//]; apply: sub_iso_to Itau1. Qed. Lemma perm_eq_coherent S1 S2 A tau: @@ -555,12 +572,66 @@ have Schi2: {subset chi2 <= 'Z[S]} by apply/allP; rewrite /= !sSZ ?ccS. have Schi_diff: chi - chi^*%CF \in 'Z[S, L^#]. by rewrite sub_aut_zchar // zchar_onG sSZ ?ccS. split=> // [_ /mapP[xi /Schi2/Znu ? -> //]||]. - apply: map_orthonormal; first by apply: sub_in2 Inu; exact: zchar_trans_on. + apply: map_orthonormal; first by apply: sub_in2 Inu; apply: zchar_trans_on. rewrite orthonormalE (conjC_pair_orthogonal ccS) //=. by rewrite cfnorm_conjC !cfnorm_irr !eqxx. by rewrite -raddfB -cfunD1E Dnu // irr_vchar_on ?Ztau. Qed. +(* There is a simple, direct way of establishing that S is coherent when S *) +(* has a pivot character eta1 whose degree divides the degree of all other *) +(* eta_i in S, as then (eta_i - a_i *: eta1)_i>1 will be a basis of Z[S, L^#] *) +(* for some integers a_i. In that case we just need to find a virtual *) +(* character zeta1 of G with the same norm as eta1, and the same dot product *) +(* on the image of the eta_i - a_i *: eta1 under tau, for then the linear *) +(* extension of tau that assigns zeta1 to eta1 is an isometry. *) +(* This is used casually by Peterfalvi, e.g., in (5.7), but a rigorous *) +(* proof does require some work, which is best factored as a Lemma. *) +Lemma pivot_coherence S (tau : {additive 'CF(L) -> 'CF(G)}) R eta1 zeta1 : + subcoherent S tau R -> eta1 \in S -> zeta1 \in 'Z[irr G] -> + {in [predD1 S & eta1], forall eta : 'CF(L), + exists2 a, a \in Cnat /\ eta 1%g = a * eta1 1%g + & '[tau (eta - a *: eta1), zeta1] = - a * '[eta1]} -> + '[zeta1] = '[eta1] -> + coherent S L^# tau. +Proof. +case=> -[N_S _ _] [Itau Ztau] oSS _ _ Seta1 Zzeta1 isoS Izeta1. +have freeS := orthogonal_free oSS; have uniqS := free_uniq freeS. +have{oSS} [/andP[S'0 _] oSS] := pairwise_orthogonalP oSS. +pose d := eta1 1%g; pose a (eta : 'CF(L)) := truncC (eta 1%g / d). +have{S'0} nzd: d != 0 by rewrite char1_eq0 ?N_S ?(memPn S'0). +pose S1 := eta1 :: [seq eta - eta1 *+ a eta | eta <- rem eta1 S]. +have sS_ZS1: {subset S <= 'Z[S1]}; last apply: (subgen_coherent sS_ZS1). + have Zeta1: eta1 \in 'Z[S1] by rewrite mem_zchar ?mem_head. + apply/allP; rewrite (eq_all_r (perm_eq_mem (perm_to_rem Seta1))) /= Zeta1. + apply/allP=> eta Seta; rewrite -(rpredBr eta (rpredMn (a eta) Zeta1)). + exact/mem_zchar/mem_behead/map_f. +have{sS_ZS1} freeS1: free S1. + have Sgt0: (0 < size S)%N by case: (S) Seta1. + rewrite /free eqn_leq dim_span /= size_map size_rem ?prednK // -(eqnP freeS). + by apply/dimvS/span_subvP => eta /sS_ZS1/zchar_span. +pose iso_eta1 zeta := zeta \in 'Z[S, L^#] /\ '[tau zeta, zeta1] = '[zeta, eta1]. +have{isoS} isoS: {in behead S1, forall zeta, iso_eta1 zeta}. + rewrite /iso_eta1 => _ /mapP[eta Seta ->]; rewrite mem_rem_uniq // in Seta. + have{Seta} [/isoS[q [Nq Dq] Itau_eta1] [eta1'eta Seta]] := (Seta, andP Seta). + rewrite zcharD1E rpredB ?rpredMn ?mem_zchar //= -scaler_nat /a Dq mulfK //. + by rewrite truncCK // !cfunE Dq subrr cfdotBl cfdotZl -mulNr oSS ?add0r. +have isoS1: {in S1, isometry [eta tau with eta1 |-> zeta1], to 'Z[irr G]}. + split=> [xi eta | eta]; rewrite !in_cons /=; last first. + by case: eqP => [-> | _ /isoS[/Ztau/zcharW]]. + do 2!case: eqP => [-> _|_ /isoS[? ?]] //; last exact: Itau. + by apply/(can_inj conjCK); rewrite -!cfdotC. +have [nu Dnu IZnu] := Zisometry_of_iso freeS1 isoS1. +exists nu; split=> // phi; rewrite zcharD1E => /andP[]. +case/(zchar_expansion (free_uniq freeS1)) => b Zb {phi}-> phi1_0. +have{phi1_0} b_eta1_0: b eta1 = 0. + have:= phi1_0; rewrite sum_cfunE big_cons big_seq big1 ?addr0 => [|zeta]. + by rewrite !cfunE (mulIr_eq0 _ (mulIf nzd)) => /eqP. + by case/isoS; rewrite cfunE zcharD1E => /andP[_ /eqP->] _; rewrite mulr0. +rewrite !raddf_sum; apply/eq_big_seq=> xi S1xi; rewrite !raddfZ_Cint //=. +by rewrite Dnu //=; case: eqP => // ->; rewrite b_eta1_0 !scale0r. +Qed. + End SubsetCoherent. (* This is Peterfalvi (5.3)(a). *) @@ -569,12 +640,11 @@ Lemma irr_subcoherent (L G : {group gT}) S tau : {in 'Z[S, L^#], isometry tau, to 'Z[irr G, G^#]} -> {R | subcoherent S tau R}. Proof. -case=> U_S irrS ccS nrS [isoL Ztau]. -have N_S: {subset S <= character} by move=> _ /irrS/irrP[i ->]; exact: irr_char. -have vcS: {subset S <= 'Z[irr L]} by move=> chi /N_S/char_vchar. -have o1SS: orthonormal S by exact: sub_orthonormal (irr_orthonormal L). -have [[_ dotSS] oSS] := (orthonormalP o1SS, orthonormal_orthogonal o1SS). -have freeS := orthogonal_free oSS. +case=> uniqS irrS ccS nrS [isoL Ztau]. +have N_S: {subset S <= character} by move=> _ /irrS/irrP[i ->]; apply: irr_char. +have Z_S: {subset S <= 'Z[irr L]} by move=> chi /N_S/char_vchar. +have o1S: orthonormal S by apply: sub_orthonormal (irr_orthonormal L). +have [[_ dotSS] oS] := (orthonormalP o1S, orthonormal_orthogonal o1S). pose beta chi := tau (chi - chi^*)%CF; pose eqBP := _ =P beta _. have Zbeta: {in S, forall chi, chi - (chi^*)%CF \in 'Z[S, L^#]}. move=> chi Schi; rewrite /= zcharD1E rpredB ?mem_zchar ?ccS //= !cfunE. @@ -588,17 +658,17 @@ have R chi: {R : 2.-tuple 'CF(G) | (chi \in S) ==> sum_beta chi R && Zortho R}. rewrite isoL // cfnormBd ?dotSS ?ccS ?eqxx // eq_sym -/(cfReal _). by rewrite (negPf (hasPn nrS _ _)). case/zchar_small_norm; rewrite ?(zcharW (Ztau _ _)) // => R [oR ZR sumR]. - by exists R; apply/and3P; split; [exact/eqP | exact/allP | ]. + by exists R; apply/and3P; split; [apply/eqP | apply/allP | ]. exists (fun xi => val (val (R xi))); split=> // [chi Schi | chi phi Schi Sphi]. by case: (R chi) => Rc /=; rewrite Schi => /and3P[/eqBP-> /allP]. -case/andP => /and3P[/= /eqP opx /eqP opx' _] _. +case/andP => /and3P[/=/eqP-opx /eqP-opx' _] _. have{opx opx'} obpx: '[beta phi, beta chi] = 0. rewrite isoL ?Zbeta // cfdotBl !cfdotBr -{3}[chi]cfConjCK. by rewrite !cfdot_conjC opx opx' rmorph0 !subr0. case: (R phi) => [[[|a [|b []]] //= _]]. rewrite Sphi => /and3P[/eqBP sum_ab Zab o_ab]. case: (R chi) => [[[|c [|d []]] //= _]]. -rewrite Schi => /and3P[/eqBP sum_cd Zcd o_cd]. +rewrite Schi => /and3P[/eqBP-sum_cd Zcd o_cd]. suffices: orthonormal [:: a; - b; c; d]. rewrite (orthonormal_cat [:: a; _]) => /and3P[_ _]. by rewrite /orthogonal /= !cfdotNl !oppr_eq0. @@ -644,9 +714,9 @@ have Itau: {in 'Z[S, L^#], isometry tau, to 'Z[irr G, G^#]}. have orthoS: pairwise_orthogonal S. exact: sub_pairwise_orthogonal sSS0 uS (seqInd_orthogonal nsKL _). pose S1 := filter (mem (irr L)) S. -have sS1S: {subset S1 <= S} by apply: mem_subseq; exact: filter_subseq. +have sS1S: {subset S1 <= S} by apply/mem_subseq/filter_subseq. have sZS1S: {subset 'Z[S1, L^#] <= 'Z[S, L^#]}. - by apply: zchar_subset sS1S; exact: orthogonal_free. + by apply: zchar_subset sS1S; apply: orthogonal_free. have [||R1 cohR1] := irr_subcoherent _ _ (sub_iso_to sZS1S sub_refl Itau). - split=> [|phi|phi]; rewrite ?mem_filter ?filter_uniq //; try case/andP=> //. by case/irrP=> i {2}-> /=/ccS->; rewrite cfConjC_irr. @@ -709,7 +779,7 @@ exists R; split=> [|phi w S1phi irr_w|j]; first 1 last. apply/orthoPr=> aa R1aa; rewrite (orthogonalP (oS1sigma phi _)) ?map_f //. by rewrite mem_filter andbC. - by rewrite /R; case: pickP => /= [k /eqP/(prTIred_inj ddA)-> | /(_ j)/eqP]. -have Zw i j: w_ i j \in 'Z[irr W] by exact: irr_vchar. +have Zw i j: w_ i j \in 'Z[irr W] by apply: irr_vchar. have{oS1sigma} oS1dsw psi j: psi \in S1 -> orthogonal (R1 psi) (dsw _ j). move/oS1sigma/orthogonalP=> opsiW. apply/orthogonalP=> aa _ R1aa /codomP[i ->]. @@ -718,19 +788,19 @@ have odsw j1 j2: j1 != j2 -> orthogonal (dsw _ j1) (dsw _ j2). move/negPf=> j2'1; apply/orthogonalP=> _ _ /codomP[i1 ->] /codomP[i2 ->]. by rewrite cfdotZl cfdotZr (cfdot_cycTIiso ddA) j2'1 andbF !mulr0. split=> // [|phi Sphi|phi xi Sphi Sxi]. -- by split=> // phi /sSS0; exact: seqInd_char. +- by split=> // phi /sSS0; apply: seqInd_char. - rewrite /R; case: pickP => [j /eqP Dphi /= | /inS1/(_ Sphi)/R1ok//]. have nz_j: j != 0 by rewrite Smu_nz -?Dphi. have [Isig Zsig]: {in 'Z[irr W], isometry sigma, to 'Z[irr G]}. exact: cycTI_Zisometry. split=> [aa | |]. - - rewrite mem_cat -map_comp => /orP. - by case=> /codomP[i ->]; rewrite ?rpredN rpredZsign Zsig. + - rewrite mem_cat -map_comp. + by case/orP=> /codomP[i ->]; rewrite ?rpredN rpredZsign Zsig. - rewrite orthonormal_cat orthogonal_oppr odsw ?andbT; last first. rewrite -(inj_eq (prTIred_inj ddA)) (prTIred_aut ddA) -/mu -Dphi. by rewrite eq_sym (hasPn nrS). suffices oNdsw k: orthonormal (dsw j k). - by rewrite map_orthonormal ?oNdsw //; apply: in2W; exact: opp_isometry. + by rewrite map_orthonormal ?oNdsw //; apply: in2W; apply: opp_isometry. apply/orthonormalP; split=> [|_ _ /codomP[i1 ->] /codomP[i2 ->]]. rewrite map_inj_uniq ?enum_uniq // => i1 i2 /(can_inj (signrZK _))/eqP. by rewrite (cycTIiso_eqE ddA) eqxx andbT => /eqP. @@ -755,9 +825,8 @@ have [Dj1 | j2'1] := eqVneq j1 j2. by rewrite {2}Dxi Dj1 -Dphi o_xi_phi in nz_xi. have [Dj1 | j2c'1] := eqVneq j1 (conjC_Iirr j2). by rewrite {2}Dxi Dj1 /mu (prTIred_aut ddA) -/mu -Dphi o_xi_phi' in nz_xi. -rewrite orthogonal_catl andbC orthogonal_oppl. -rewrite !orthogonal_catr !orthogonal_oppr !odsw ?(inj_eq (aut_Iirr_inj _)) //. -by rewrite (inv_eq (@conjC_IirrK _ _)). +rewrite orthogonal_catl orthogonal_oppl !orthogonal_catr !orthogonal_oppr. +by rewrite !odsw ?(inv_eq (@conjC_IirrK _ _)) ?conjC_IirrK. Qed. Section SubCoherentProperties. @@ -780,7 +849,7 @@ Proof. case=> uS1 sS1 ccS1; have [[N_S nrS _] Itau oS defR oR] := cohS. split; last 1 [exact: sub_in1 defR | exact: sub_in2 oR]. - split=> // [xi /sS1/N_S// | ]. - by apply/hasPn; exact: sub_in1 (hasPn nrS). + by apply/hasPn; apply: sub_in1 (hasPn nrS). - by apply: sub_iso_to Itau => //; apply: zchar_subset. exact: sub_pairwise_orthogonal oS. Qed. @@ -822,7 +891,7 @@ Lemma subcoherent_norm chi psi (tau1 : {additive 'CF(L) -> 'CF(G)}) X Y : [/\ '[X] = '[chi], '[Y] = '[psi] & exists2 E, subseq E (R chi) & X = \sum_(xi <- E) xi]]. Proof. -case=> Schi Zpsi /and3P[/andP[/eqP ochi_psi _] /andP[/eqP ochic_psi _] _] S0. +case=> Schi Zpsi /and3P[/andP[/eqP-ochi_psi _] /andP[/eqP-ochic_psi _] _] S0. move=> [Itau1 Ztau1] tau1dchi [defXY oXY oYR]. have [[ZS nrS ccS] [tS Zt] oS /(_ _ Schi)[ZR o1R tau_dchi] _] := cohS. have [/=/andP[S'0 uS] oSS] := pairwise_orthogonalP oS. @@ -834,34 +903,32 @@ have dotS00R xi: xi \in R chi -> '[tau1 S0`_0, xi] = a xi. move=> Rxi; rewrite dXY1 cfdotBl (orthoPl oY1R) // subr0. by rewrite defX1 cfproj_sum_orthonormal. have nchi: '[chi] = \sum_(xi <- R chi) a xi. - transitivity '[S0`_0, S0`_1]. - rewrite [rhs in _ = rhs]cfdotC cfdotBl !cfdotBr ochi_psi ochic_psi. - by rewrite (oSS _ _ Schic) // !subr0 -cfdotC. - rewrite -Itau1 ?mem_zchar ?mem_nth // tau1dchi tau_dchi cfdot_sumr. - exact: eq_big_seq. -have nX: '[X1] <= '[X] ?= iff (X == X1). - rewrite -subr_eq0 -{1 2}[X](subrK X1) cfnormDd. - rewrite -lerif_subLR subrr -cfnorm_eq0 eq_sym. - by apply: lerif_eq; apply: cfnorm_ge0. - rewrite defX1 cfdot_sumr big1_seq // => xi Rxi; rewrite cfdotZr cfdotBl. - rewrite cfproj_sum_orthonormal // -[X](subrK Y) cfdotDl -defXY dotS00R //. - by rewrite (orthoPl oYR) // addr0 subrr mulr0. + transitivity '[tau1 S0`_0, tau1 S0`_1]; last first. + by rewrite tau1dchi tau_dchi cfdot_sumr; apply: eq_big_seq dotS00R. + rewrite [RHS]cfdotC Itau1 ?mem_zchar ?mem_nth // cfdotBl !cfdotBr. + by rewrite ochi_psi ochic_psi (oSS chi^*%CF) // !subr0 -cfdotC. +have normX: '[X1] <= '[X] ?= iff (X == X1). + rewrite -[in '[X]](subrK X1 X) -subr_eq0 cfnormDd. + by rewrite -lerif_subLR subrr -cfnorm_eq0 eq_sym; apply/lerif_eq/cfnorm_ge0. + rewrite defX1 cfdot_sumr big1_seq // => xi Rxi. + rewrite cfdotZr cfdotBl cfproj_sum_orthonormal // -{2}dotS00R // defXY. + by rewrite cfdotBl (orthoPl oYR) // subr0 subrr mulr0. pose is01a xi := a xi == (a xi != 0)%:R. have leXa xi: a xi <= `|a xi| ^+ 2 ?= iff is01a xi. - apply/lerifP; rewrite /is01a; have /CintP[b ->] := Za xi. - rewrite -intr_norm -rmorphX ltr_int intr_eq0 pmulrn !eqr_int. - by case: b => [[|[|n]]|] //=; rewrite ltr_eexpr. -have{nchi nX} part_a: '[chi] <= '[X] ?= iff all is01a (R chi) && (X == X1). - apply: lerif_trans nX; rewrite nchi defX1 cfnorm_sum_orthonormal //. + rewrite Cint_normK //; split; first by rewrite Cint_ler_sqr. + rewrite eq_sym -subr_eq0 -[lhs in _ - lhs]mulr1 -mulrBr mulf_eq0 subr_eq0. + by rewrite /is01a; case a_xi_0: (a xi == 0). +have{nchi normX} part_a: '[chi] <= '[X] ?= iff all is01a (R chi) && (X == X1). + apply: lerif_trans normX; rewrite nchi defX1 cfnorm_sum_orthonormal //. by rewrite -big_all !(big_tnth _ _ (R chi)) big_andE; apply: lerif_sum. split=> [|/lerif_eq part_b]; first by case: part_a. have [_ /esym] := lerif_add part_a part_b; rewrite -!cfnormBd // -defXY. rewrite Itau1 ?mem_zchar ?mem_head // eqxx => /andP[a_eq /eqP->]. split=> //; first by apply/esym/eqP; rewrite part_a. have{a_eq} [/allP a01 /eqP->] := andP a_eq; rewrite defX1. -exists (filter [preim a of predC1 0] (R chi)); first exact: filter_subseq. +exists [seq xi <- R chi | a xi != 0]; first exact: filter_subseq. rewrite big_filter [rhs in _ = rhs]big_mkcond /=. -by apply: eq_big_seq => xi /a01/eqP{1}->; rewrite scaler_nat -mulrb. +by apply: eq_big_seq => xi Rxi; rewrite -mulrb -scaler_nat -(eqP (a01 _ _)). Qed. (* This is Peterfalvi (5.5). *) @@ -871,7 +938,7 @@ Lemma coherent_sum_subseq chi (tau1 : {additive 'CF(L) -> 'CF(G)}) : tau1 (chi - chi^*%CF) = tau (chi - chi^*%CF) -> exists2 E, subseq E (R chi) & tau1 chi = \sum_(a <- E) a. Proof. -set S1 := (chi :: _) => Schi [iso_t1 Zt1] t1cc'. +set S1 := chi :: _ => Schi [iso_t1 Zt1] t1cc'. have freeS1: free S1. have [[_ nrS ccS] _ oS _ _] := cohS. by rewrite orthogonal_free ?(conjC_pair_orthogonal ccS). @@ -882,7 +949,7 @@ have Zt1c: tau1 (chi - 0) \in 'Z[irr G]. by rewrite subr0 Zt1 ?mem_zchar ?mem_head. have [X R_X [Y defXY]] := subcoherent_split Schi Zt1c. case/subcoherent_norm: (defXY); last 2 [by []]. -- by rewrite /orthogonal /= !cfdot0r eqxx Schi cfun0_zchar. +- by rewrite Schi rpred0 /orthogonal /= !cfdot0r eqxx. - by split; [apply: sub_in2 iso_t1 | apply: sub_in1 Zt1]. move=> _ [|_ /eqP]; rewrite cfdot0l ?cfnorm_ge0 // cfnorm_eq0 => /eqP Y0. case=> E sER defX; exists E => //; rewrite -defX -[X]subr0 -Y0 -[chi]subr0. @@ -895,11 +962,11 @@ Corollary mem_coherent_sum_subseq S1 chi (tau1 : {additive 'CF(L) -> 'CF(G)}) : exists2 E, subseq E (R chi) & tau1 chi = \sum_(a <- E) a. Proof. move=> uccS1 [Itau1 Dtau1] S1chi; have [uS1 sS1S ccS1] := uccS1. -have S1chi_s: chi^*%CF \in S1 by exact: ccS1. +have S1chi_s: chi^*%CF \in S1 by apply: ccS1. apply: coherent_sum_subseq; first exact: sS1S. by apply: sub_iso_to Itau1 => //; apply: zchar_subset; apply/allP/and3P. -apply: Dtau1; rewrite sub_aut_zchar ?zchar_onG ?mem_zchar // => phi /sS1S. -by have [[charS _ _] _ _ _ _] := cohS => /charS/char_vchar. +apply: Dtau1; rewrite sub_aut_zchar ?zchar_onG ?mem_zchar // => phi /sS1S-Sphi. +by apply: char_vchar; have [[->]] := cohS. Qed. (* A frequently used consequence of (5.5). *) @@ -946,58 +1013,50 @@ have [[uS1 sS1S _] [uS2 sS2S _]] := (uccS1, uccS2). have [[[Itau1 Ztau1] Dtau1] [[Itau2 Ztau2] Dtau2]] := (cohS1, cohS2). have [[N_S1 _ _] _ oS11 _ _] := subset_subcoherent uccS1. have [_ _ oS22 _ _] := subset_subcoherent uccS2. -have{N_S1} nz_chi1: chi 1%g != 0; last move/mem_zchar in S1chi. +have nz_chi1: chi 1%g != 0; last move/mem_zchar in S1chi. by rewrite char1_eq0 ?N_S1 //; have [/memPn->] := andP oS11. have oS12: orthogonal S1 S2. apply/orthogonalP=> xi1 xi2 Sxi1 Sxi2; apply: orthoPr xi1 Sxi1. by rewrite subset_ortho_subcoherent ?sS2S //; apply: S1'2. -pose S3 := S1 ++ S2; pose Y := map tau1 S1 ++ map tau2 S2. -have oS33: pairwise_orthogonal S3 by rewrite pairwise_orthogonal_cat oS11 oS22. -have oYY: pairwise_orthogonal Y. +set S3 := S1 ++ S2; pose Y := map tau1 S1 ++ map tau2 S2. +have oS3: pairwise_orthogonal S3 by rewrite pairwise_orthogonal_cat oS11 oS22. +have oY: pairwise_orthogonal Y. by rewrite pairwise_orthogonal_cat !map_pairwise_orthogonal ?coherent_ortho. have Z_Y: {subset Y <= 'Z[irr G]}. - move=> xi_tau; rewrite mem_cat => /orP[] /mapP[xi Sxi ->] {xi_tau}. - by rewrite Ztau1 ?mem_zchar. - by rewrite Ztau2 ?mem_zchar. -have nY: map cfnorm Y = map cfnorm (S1 ++ S2). + move=> psi; rewrite mem_cat. + by case/orP=> /mapP[xi /mem_zchar] => [/Ztau1 | /Ztau2]-Zpsi ->. +have normY: map cfnorm Y = map cfnorm (S1 ++ S2). rewrite !map_cat -!map_comp; congr (_ ++ _). by apply/eq_in_map => xi S1xi; rewrite /= Itau1 ?mem_zchar. by apply/eq_in_map => xi S2xi; rewrite /= Itau2 ?mem_zchar. -have [tau3 /eqP defY ZItau3] := Zisometry_of_cfnorm oS33 oYY nY Z_Y. -exists tau3; split=> {ZItau3}// xi; rewrite zcharD1E /= => /andP[S3xi]. +have [tau3 defY ZItau3] := Zisometry_of_cfnorm oS3 oY normY Z_Y. have{defY} [defY1 defY2]: {in S1, tau3 =1 tau1} /\ {in S2, tau3 =1 tau2}. - have:= defY; rewrite map_cat eqseq_cat ?size_map // => /andP[]. - by split; apply/eq_in_map/eqP. -have{S3xi} [xi1 [xi2 [Sxi1 Sxi2 ->] {xi}]]: - exists xi1, exists xi2, [/\ xi1 \in 'Z[S1], xi2 \in 'Z[S2] & xi = xi1 + xi2]. -- have uS3 := free_uniq (orthogonal_free oS33). - have [z Zz ->] := zchar_expansion uS3 S3xi; rewrite big_cat. - pose Y_ S4 := \sum_(mu <- S4) z mu *: mu. - suffices ZS_Y S4: Y_ S4 \in 'Z[S4] by exists (Y_ S1), (Y_ S2). - by rewrite /Y_ big_seq rpred_sum // => psi /mem_zchar/rpredZ_Cint->. -rewrite cfunE addrC addr_eq0 linearD => /eqP xi2_1. -transitivity (tau1 xi1 + tau2 xi2). - have [z1 Zz1 ->] := zchar_nth_expansion Sxi1. - have [z2 Zz2 ->] := zchar_nth_expansion Sxi2. - rewrite !raddf_sum; congr(_ + _); apply: eq_bigr => i _; - by rewrite !raddfZ_Cint -?(defY1, defY2) ?mem_nth. -have Z_S1_1 zeta: zeta \in 'Z[S1] -> zeta 1%g \in Cint. - move=> Szeta; rewrite Cint_vchar1 // (zchar_sub_irr _ Szeta) {zeta Szeta}//. - by move=> zeta /sS1S Szeta; apply: char_vchar; have [[->]] := cohS. -have [Zchi1 Zxi1] := (Z_S1_1 _ S1chi, Z_S1_1 _ Sxi1). -apply: (scalerI nz_chi1); rewrite scalerDr -!raddfZ_Cint // scalerDr. -rewrite -[_ *: _](subrK (xi1 1%g *: chi)) raddfD -[_ + _]addrA. -rewrite -[rhs in _ = tau rhs]addrA linearD Dtau1; last first. - by rewrite zcharD1E rpredB ?rpredZ_Cint ?Z_S1_1 //= !cfunE mulrC subrr. -congr (_ + _); rewrite -[_ *: xi2](addKr (xi1 1%g *: phi)) (raddfD tau2). -rewrite [_ + _]addrA [rhs in tau rhs]addrA linearD; congr (_ + _); last first. - rewrite Dtau2 // zcharD1E rpredD ?rpredZ_Cint ?Z_S1_1 //=. - by rewrite !cfunE mulrC xi2_1 chi1_phi mulrN subrr. -rewrite raddfN (raddfZ_Cint tau1) // (raddfZ_Cint tau2) // -!scalerBr linearZ. -by congr (_ *: _). + have/eqP := defY; rewrite map_cat eqseq_cat ?size_map //. + by case/andP; split; apply/eq_in_map/eqP. +exists tau3; split=> {ZItau3}// eta; rewrite zcharD1E. +case/andP=> /(zchar_expansion (free_uniq (orthogonal_free oS3)))[b Zb {eta}->]. +pose bS Si := \sum_(xi <- Si) b xi *: xi. +have ZbS Si: bS Si \in 'Z[Si]. + by rewrite /bS big_seq rpred_sum // => eta /mem_zchar/rpredZ_Cint->. +rewrite big_cat /= -!/(bS _) cfunE addrC addr_eq0 linearD => /eqP-bS2_1. +transitivity (tau1 (bS S1) + tau2 (bS S2)). + by rewrite !raddf_sum; congr (_ + _); apply/eq_big_seq=> xi Si_xi; + rewrite !raddfZ_Cint // -(defY1, defY2). +have Z_S1_1 psi: psi \in 'Z[S1] -> psi 1%g \in Cint. + by move/zchar_sub_irr=> Zpsi; apply/Cint_vchar1/Zpsi => ? /N_S1/char_vchar. +apply/(scalerI nz_chi1)/(addIr (- bS S1 1%g *: tau (chi - phi))). +rewrite [in LHS]tau_chi_phi !scalerDr -!raddfZ_Cint ?rpredN ?Z_S1_1 //=. +rewrite addrACA -!raddfD -raddfB !scalerDr !scaleNr scalerN !opprK. +rewrite Dtau2 ?Dtau1 ?zcharD1E ?cfunE; first by rewrite -raddfD addrACA. + by rewrite mulrC subrr rpredB ?rpredZ_Cint ?Z_S1_1 /=. +by rewrite mulrC bS2_1 -chi1_phi mulNr addNr rpredD ?rpredZ_Cint ?Z_S1_1 /=. Qed. -(* This is essentially Peterfalvi (5.6.3), which gets reused in (9.11.8). *) +(* This is essentially Peterfalvi (5.6.3), which gets reused in (9.11.8). *) +(* While the assumptions are similar to those of the pivot_coherence lemma, *) +(* the two results are mostly independent: here S1 need not have a pivot, and *) +(* extend_coherent_with does not apply to the base case (size S = 2) of *) +(* pivot_coherence, which is almost as hard to prove as the general case. *) Lemma extend_coherent_with S1 (tau1 : {additive 'CF(L) -> 'CF(G)}) chi phi a X : cfConjC_subset S1 S -> coherent_with S1 L^# tau tau1 -> [/\ phi \in S1, chi \in S & chi \notin S1] -> @@ -1061,15 +1120,16 @@ apply: (bridge_coherent sS20 cohS2 sS10 cohS1) => //. by rewrite mem_head (zchar_on Zbeta) rpredZ_Cint ?mem_zchar. Qed. -(* This is Peterfalvi (5.6). *) +(* This is Peterfalvi (5.6): Feit's result that a coherent set can always be *) +(* extended by a character whose degree is below a certain threshold. *) Lemma extend_coherent S1 xi1 chi : - cfConjC_subset S1 S -> [/\ xi1 \in S1, chi \in S & chi \notin S1] -> + cfConjC_subset S1 S -> xi1 \in S1 -> chi \in S -> chi \notin S1 -> [/\ (*a*) coherent S1 L^# tau, (*b*) (xi1 1%g %| chi 1%g)%C & (*c*) 2%:R * chi 1%g * xi1 1%g < \sum_(xi <- S1) xi 1%g ^+ 2 / '[xi]] -> coherent (chi :: chi^*%CF :: S1) L^# tau. Proof. -move=> ccsS1S [S1xi1 Schi notS1chi] [[tau1 cohS1] xi1_dv_chi1 ub_chi1]. +move=> ccsS1S S1xi1 Schi notS1chi [[tau1 cohS1] xi1_dv_chi1 ub_chi1]. have [[uS1 sS1S ccS1] [[Itau1 Ztau1] Dtau1]] := (ccsS1S, cohS1). have{xi1_dv_chi1} [a Za chi1] := dvdCP _ _ xi1_dv_chi1. have [[N_S nrS ccS] ZItau oS R_P oR] := cohS; have [Itau Ztau] := ZItau. @@ -1077,7 +1137,7 @@ have [Sxi1 [ZRchi o1Rchi sumRchi]] := (sS1S _ S1xi1, R_P _ Schi). have ocS1 xi: xi \in S1 -> '[chi, xi] = 0. by apply: orthoPl; rewrite orthogonal_sym subset_ortho_subcoherent. have /andP[/memPn/=nzS _] := oS; have [Nchi nz_chi] := (N_S _ Schi, nzS _ Schi). -have oS1: pairwise_orthogonal S1 by exact: sub_pairwise_orthogonal oS. +have oS1: pairwise_orthogonal S1 by apply: sub_pairwise_orthogonal oS. have [freeS freeS1] := (orthogonal_free oS, orthogonal_free oS1). have nz_nS1 xi: xi \in S1 -> '[xi] != 0 by rewrite cfnorm_eq0 => /sS1S/nzS. have nz_xi11: xi1 1%g != 0 by rewrite char1_eq0 ?N_S ?nzS. @@ -1093,7 +1153,9 @@ have Zachi: chi - a *: xi1 \in 'Z[S, L^#]. have Ztau_achi := zcharW (Ztau _ Zachi). have [X R_X [Y defXY]] := subcoherent_split Schi Ztau_achi. have [eqXY oXY oYRchi] := defXY; pose X1 := map tau1 (in_tuple S1). -have oX1: pairwise_orthogonal X1 by exact: map_pairwise_orthogonal. +suffices defY: Y = a *: tau1 xi1. + by move: eqXY; rewrite defY; apply: extend_coherent_with; rewrite -?defY. +have oX1: pairwise_orthogonal X1 by apply: map_pairwise_orthogonal. have N_S1_1 xi: xi \in S1 -> xi 1%g \in Cnat by move/sS1S/N_S/Cnat_char1. have oRchiX1 psi: psi \in 'Z[R chi] -> orthogonal psi X1. move/zchar_span=> Rpsi; apply/orthoPl=> chi2 /memv_span. @@ -1124,66 +1186,39 @@ have [lam Zlam [Z oZS1 defY]]: rewrite -eqXY Dtau1 // Itau // ?(zchar_subset sS1S) //. rewrite cfdotBl !cfdotBr !cfdotZr !ocS1 // !mulr0 subrr add0r !cfdotZl. by rewrite opprB addrAC subrK subrr. -have [|| leXchi _] := subcoherent_norm _ _ (erefl _) defXY. +have [||leXchi _] := subcoherent_norm _ _ (erefl _) defXY. - rewrite Schi scale_zchar ?char_vchar ?N_S /orthogonal //= !cfdotZr ocS1 //. by rewrite -[xi1]cfConjCK cfdot_conjC ocS1 ?ccS1 // conjC0 mulr0 eqxx. - apply: sub_iso_to ZItau; [apply: zchar_trans_on; apply/allP | exact: zcharW]. rewrite /= Zachi sub_aut_zchar ?zchar_onG ?mem_zchar ?ccS //. by move=> xi /N_S/char_vchar. -have{defY leXchi lam Z Zlam oZS1 ub_chi1} defY: Y = a *: tau1 xi1. - have nXY: '[X] + '[Y] = '[chi] + '[a *: xi1]. - by rewrite -!cfnormBd // ?cfdotZr ?ocS1 ?mulr0 // -eqXY Itau. - have{leXchi nXY}: '[Y] <= a ^+ 2 * '[xi1]. - by rewrite -(ler_add2l '[X]) nXY cfnormZ Cint_normK // ler_add2r. - rewrite defY cfnormDd; last first. - rewrite cfdotC (span_orthogonal oZS1) ?rmorph0 ?memv_span1 //. - rewrite big_seq memvB ?memvZ ?memv_suml ?memv_span ?map_f //. - by move=> theta S1theta; rewrite memvZ ?memv_span. - rewrite -subr_ge0 cfnormB cfnormZ Cint_normK // Itau1 ?Z_S1 //. - rewrite -2!addrA (opprD (_ * _)) addNKr cfnormZ Cint_normK // oppr_ge0. - rewrite cfnorm_sum_orthogonal //; set sum_a := \sum_(xi <- _) _. - rewrite -cfdotC cfdotC cfdotZl cfdotZr cfproj_sum_orthogonal ?map_f // a_xi1. - rewrite Itau1 ?Z_S1 // 3!rmorphM !(aut_Cint _ Za) fmorphV aut_Cint //. - rewrite -cfdotC -mulr2n 2!mulrA divfK ?nz_nS1 // -mulrnAr addrA => ub_lam. - have [lam0 | nz_lam] := eqVneq lam 0. - suffices /eqP->: Z == 0 by rewrite lam0 scale0r subr0 addr0. - rewrite -cfnorm_eq0 eqr_le cfnorm_ge0 andbT. - by rewrite lam0 -mulrA !mul0r subrr add0r in ub_lam. - set d := \sum_(xi <- _) _ in ub_chi1; pose b := 2%:R * chi 1%g * xi1 1%g / d. - have pos_S1_1 := Cnat_ge0 (Cnat_char1 (N_S _ (sS1S _ _))). - have xi11_gt0: 0 < xi1 1%g by rewrite ltr_def nz_xi11 pos_S1_1. - have d_gt0: 0 < d. - have a_xi_ge0 xi: xi \in S1 -> 0 <= xi 1%g ^+ 2 / '[xi]. - by move/pos_S1_1 => xi_1_pos; rewrite 2?mulr_ge0 // invr_ge0 cfnorm_ge0. - rewrite [d]big_seq; case defS1: {1 2}S1 S1xi1 => // [xi S1'] _. - have{defS1} S1xi: xi \in S1 by rewrite defS1 mem_head. - rewrite big_cons S1xi ltr_spaddl ?sumr_ge0 // ltr_def a_xi_ge0 //=. - by rewrite !mulf_neq0 ?invr_eq0 ?char1_eq0 -?cfnorm_eq0 ?nz_nS1 ?N_S ?sS1S. - have nz_d: d != 0 by rewrite eqr_le ltr_geF. - have b_gt0: 0 < b. - rewrite !pmulr_rgt0 ?ltr0n ?invr_gt0 // lt0r. - by rewrite Cnat_ge0 ?Cnat_char1 ?char1_eq0 ?N_S // nzS. - have{ub_chi1} b_lt1: b < 1 by rewrite ltr_pdivr_mulr ?mul1r. - have{ub_lam} ub_lam: lam ^+ 2 <= b * lam. - rewrite -(ler_pmul2r d_gt0) (mulrAC b) divfK //. - rewrite -[d](divfK (mulf_neq0 nz_xi11 nz_xi11)) chi1 mulr_natl -mulrnAl. - rewrite !mulrA 2!(mulrAC _ _ lam) 2?ler_pmul2r // -mulrA -expr2. - have ->: d / xi1 1%g ^+ 2 = sum_a. - rewrite big_distrl /sum_a big_map !big_seq; apply: eq_bigr => xi S1xi /=. - rewrite a_E // Itau1 ?Z_S1 //= (normr_idP _); last first. - by rewrite !(cfnorm_ge0, mulr_ge0, invr_ge0) ?pos_S1_1. - rewrite mulrAC 2!exprMn -!exprVn [p in p * '[xi]]mulrA. - by rewrite divfK ?nz_nS1. - rewrite -subr_ge0 -opprB oppr_ge0 (ler_trans _ ub_lam) //. - by rewrite (mulrC lam) -{1}[_ - _]addr0 ler_add2l cfnorm_ge0. - have lam_gt0: 0 < lam. - rewrite ltr_def nz_lam -(ler_pmul2l b_gt0) mulr0. - by apply: ler_trans ub_lam; rewrite -Cint_normK // mulr_ge0 ?normr_ge0. - rewrite ler_pmul2r // ltr_geF // in ub_lam. - rewrite (ltr_le_trans b_lt1) //; have:= lam_gt0. - have /CnatP[n ->]: lam \in Cnat by rewrite CnatEint Zlam ltrW. - by rewrite ltr0n ler1n. -by move: eqXY; rewrite defY; apply: extend_coherent_with => //; rewrite -defY. +have normXY: '[X] + '[Y] = '[chi] + '[a *: xi1]. + by rewrite -!cfnormBd // ?cfdotZr ?ocS1 ?mulr0 // -eqXY Itau. +have{leXchi normXY}: '[Y] <= a ^+ 2 * '[xi1]. + by rewrite -(ler_add2l '[X]) normXY cfnormZ Cint_normK // ler_add2r. +rewrite {}defY cfnormDd; last first. + rewrite cfdotC (span_orthogonal oZS1) ?rmorph0 ?memv_span1 //. + rewrite big_seq memvB ?memvZ ?memv_suml ?memv_span ?map_f //. + by move=> theta S1theta; rewrite memvZ ?memv_span. +rewrite -cfnormN opprB cfnormB !cfnormZ !Cint_normK // addrAC ler_subl_addl. +rewrite cfdotZl cfdotZr cfnorm_sum_orthogonal ?cfproj_sum_orthogonal ?map_f //. +rewrite a_xi1 Itau1 ?Z_S1 // addrAC ler_add2r !(divfK, mulrA) ?nz_nS1 //. +rewrite !conj_Cint ?rpredM // => /ler_gtF-lb_2_lam_a. +suffices lam0: lam = 0; last apply: contraFeq lb_2_lam_a => nz_lam. + suffices ->: Z = 0 by rewrite lam0 scale0r subrK. + by apply: contraFeq lb_2_lam_a; rewrite -cfnorm_gt0 lam0 expr0n !mul0r !add0r. +rewrite ltr_paddr ?cfnorm_ge0 // -mulr2n -mulr_natl mulrCA. +have xi11_gt0: xi1 1%g > 0 by rewrite char1_gt0 ?N_S ?sS1S -?cfnorm_eq0 ?nz_nS1. +have a_gt0: a > 0 by rewrite -(ltr_pmul2r xi11_gt0) mul0r -chi1 char1_gt0. +apply: ler_lt_trans (_ : lam ^+ 2 * (2%:R * a) < _). + by rewrite ler_pmul2r ?mulr_gt0 ?ltr0n ?Cint_ler_sqr. +rewrite ltr_pmul2l ?(ltr_le_trans ltr01) ?sqr_Cint_ge1 {lam Zlam nz_lam}//. +rewrite -(ltr_pmul2r xi11_gt0) -mulrA -chi1 -(ltr_pmul2r xi11_gt0). +congr (_ < _): ub_chi1; rewrite -mulrA -expr2 mulr_suml big_map. +apply/eq_big_seq=> xi S1xi; rewrite a_E // Itau1 ?mem_zchar //. +rewrite ger0_norm ?divr_ge0 ?cfnorm_ge0 ?char1_ge0 ?N_S ?sS1S //. +rewrite [_ / _ / _]mulrAC [RHS]mulrAC -exprMn divfK //. +by rewrite [RHS]mulrAC divfK ?nz_nS1 // mulrA. Qed. (* This is Peterfalvi (5.7). *) @@ -1191,153 +1226,119 @@ Qed. (* isometry, which would necessarily map irreducibles to signed irreducibles. *) (* It would then only remain to show that the signs are chosen consistently, *) (* by considering the degrees of the differences. *) +(* This result is complementary to (5.6): it follow from it when S has 4 or *) +(* fewer characters, or reducible characters. On the contrary, (5.7) can be *) +(* used to provide an initial set of characters with a threshold high enough *) +(* to enable (repeated) application of (5.6), as in seqIndD_irr_coherence. *) Lemma uniform_degree_coherence : constant [seq chi 1%g | chi : 'CF(L) <- S] -> coherent S L^# tau. Proof. -case defS: {1}S => /= [|chi S1] szS; first by rewrite defS; exact: nil_coherent. -have{szS} unifS xi: xi \in S -> xi 1%g = chi 1%g. +case defS: {1}S => /= [|chi1 S1] szS; first by rewrite defS; apply nil_coherent. +have{szS} unifS xi: xi \in S -> xi 1%g = chi1 1%g. by rewrite defS => /predU1P[-> // | S'xi]; apply/eqP/(allP szS)/map_f. -have Schi: chi \in S by rewrite defS mem_head. +have{S1 defS} Schi1: chi1 \in S by rewrite defS mem_head. have [[N_S nrS ccS] IZtau oS R_P oR] := cohS; have [Itau Ztau] := IZtau. -have freeS := orthogonal_free oS. have Zd: {in S &, forall xi1 xi2, xi1 - xi2 \in 'Z[S, L^#]}. move=> xi1 xi2 Sxi1 Sxi2 /=. by rewrite zcharD1E rpredB ?mem_zchar //= !cfunE !unifS ?subrr. -have [neq_chic Schic] := (hasPn nrS _ Schi, ccS _ Schi). -have [/andP[/memPn notS0 _] ooS] := pairwise_orthogonalP oS. -pose S' xi := [predD1 S & xi]; pose S'c xi := predD1 (S' xi) xi^*%CF. -have{oR} oR xi1 xi2: xi1 \in S -> xi2 \in S'c xi1 -> orthogonal (R xi1) (R xi2). - move=> Sxi1 /and3P[/= neq_xi21c neq_xi21 Sxi2]. - by rewrite orthogonal_sym oR // /orthogonal /= !ooS ?eqxx // ccS. +pose chi2 := chi1^*%CF; have Schi2: chi2 \in S by rewrite ccS. +have ch1'2: chi2 != chi1 by apply/(hasPn nrS). +have [_ oSS] := pairwise_orthogonalP oS. +pose S1 xi := [predD1 S & xi]; pose S2 xi := [predD1 (S1 xi) & xi^*%CF]. +have{oR} oR xi1 xi2: xi1 \in S -> xi2 \in S2 xi1 -> orthogonal (R xi1) (R xi2). + move=> Sxi1 /and3P[/= xi1J'2 xi1'2 Sxi2]. + by rewrite orthogonal_sym oR // /orthogonal /= !oSS ?eqxx // ccS. have oSc xi: xi \in S -> '[xi, xi^*] = 0. - by move=> Sxi; rewrite ooS ?ccS // -[_ == _]negbK eq_sym (hasPn nrS). -pose D xi := tau (chi - xi). -have Z_D xi: xi \in S -> D xi \in 'Z[irr G] by move/(Zd _ _ Schi)/Ztau/zcharW. -have /CnatP[N defN]: '[chi] \in Cnat by rewrite Cnat_cfdot_char ?N_S. -have dotD: {in S' chi &, forall xi1 xi2, '[D xi1, D xi2] = N%:R + '[xi1, xi2]}. -- move=> xi1 xi2 /andP[ne_xi1chi Sxi1] /andP[ne_xi2chi Sxi2]. + by move=> Sxi; rewrite oSS ?ccS // eq_sym (hasPn nrS). +pose D xi := tau (chi1 - xi). +have Z_D xi: xi \in S -> D xi \in 'Z[irr G] by move/(Zd _ _ Schi1)/Ztau/zcharW. +have /CnatP[N defN]: '[chi1] \in Cnat by rewrite Cnat_cfdot_char ?N_S. +have dotD: {in S1 chi1 &, forall xi1 xi2, '[D xi1, D xi2] = N%:R + '[xi1, xi2]}. + move=> xi1 xi2 /andP[ch1'xi1 Sxi1] /andP[chi1'xi2 Sxi2]. rewrite Itau ?Zd // cfdotBl !cfdotBr defN. - by rewrite 2?ooS // 1?eq_sym // opprB !subr0. -have /R_P[ZRchi o1Rchi defRchi] := Schi; have frRchi := orthonormal_free o1Rchi. -have szRchi: size (R chi) = (N + N)%N. + by rewrite 2?oSS // 1?eq_sym // opprB !subr0. +have /R_P[ZRchi oRchi defRchi] := Schi1. +have szRchi: size (R chi1) = (N + N)%N. apply: (can_inj natCK); rewrite -cfnorm_orthonormal // -defRchi. by rewrite dotD ?inE ?ccS ?(hasPn nrS) // cfnorm_conjC defN -natrD. -pose sub_Rchi X := exists2 E, subseq E (R chi) & X = \sum_(a <- E) a. -pose Xspec X := [/\ X \in 'Z[R chi], '[X]_G = N%:R & sub_Rchi X]. -pose Xi_spec X xi := X - D xi \in 'Z[R xi] /\ '[X, D xi] = N%:R. -have haveX xi: xi \in S'c chi -> exists2 X, Xspec X & Xi_spec X xi. - move=> S'xi; have /and3P[/= ne_xi_chi' ne_xi_chi Sxi] := S'xi. +pose subRchi1 X := exists2 E, subseq E (R chi1) & X = \sum_(a <- E) a. +pose Xspec X := [/\ X \in 'Z[R chi1], '[X] = N%:R & subRchi1 X]. +pose Xi_spec (X : 'CF(G)) xi := X - D xi \in 'Z[R xi] /\ '[X, D xi] = N%:R. +have haveX xi: xi \in S2 chi1 -> exists2 X, Xspec X & Xi_spec X xi. + move=> S2xi; have /and3P[/= chi2'xi ch1'xi Sxi] := S2xi. have [neq_xi' Sxi'] := (hasPn nrS xi Sxi, ccS xi Sxi). - have [X RchiX [Y1 defXY1]] := subcoherent_split Schi (Z_D _ Sxi). - have [eqXY1 oXY1 oY1chi] := defXY1; have sRchiX := zchar_span RchiX. + have [X RchiX [Y1 defXY1]] := subcoherent_split Schi1 (Z_D _ Sxi). + have [[eqXY1 oXY1 oY1chi] sRchiX] := (defXY1, zchar_span RchiX). have Z_Y1: Y1 \in 'Z[irr G]. - rewrite -[Y1](subrK X) -opprB -eqXY1 addrC rpredB ?Z_D //. - exact: (zchar_trans ZRchi). - have [X1 RxiX1 [Y defX1Y]] := subcoherent_split Sxi Z_Y1; pose Y2 := X + Y. - have [eqX1Y oX1Y oYxi] := defX1Y; pose D2 := tau (xi - chi). - have oY2Rxi: orthogonal Y2 (R xi). - apply/orthogonalP=> _ phi /predU1P[-> | //] Rxi_phi. - rewrite cfdotDl (orthoPl oYxi) // addr0. - by rewrite (span_orthogonal (oR _ _ _ S'xi)) // (memv_span Rxi_phi). + by rewrite -rpredN -(rpredDl _ (zchar_trans ZRchi RchiX)) -eqXY1 Z_D. + have [X1 RxiX1 [Y defX1Y]] := subcoherent_split Sxi Z_Y1. + have [[eqX1Y oX1Y oYxi] sRxiX1] := (defX1Y, zchar_span RxiX1). + pose Y2 : 'CF(G) := X + Y; pose D2 : 'CF(G) := tau (xi - chi1). + have oY2Rxi: orthogonal Y2 (R xi). + apply/orthoPl=> phi Rxi_phi; rewrite cfdotDl (orthoPl oYxi) // addr0. + by rewrite (span_orthogonal (oR chi1 xi _ _)) // memv_span. have{oY2Rxi} defX1Y2: [/\ D2 = X1 - Y2, '[X1, Y2] = 0 & orthogonal Y2 (R xi)]. rewrite -opprB -addrA -opprB -eqX1Y -eqXY1 -linearN opprB cfdotC. by rewrite (span_orthogonal oY2Rxi) ?conjC0 ?memv_span1 ?(zchar_span RxiX1). have [||minX eqX1] := subcoherent_norm _ _ (erefl _) defXY1. - - by rewrite char_vchar ?N_S /orthogonal //= !ooS ?eqxx // eq_sym. + - by rewrite char_vchar ?N_S /orthogonal //= !oSS ?eqxx // eq_sym. - apply: sub_iso_to IZtau; last exact: zcharW. by apply: zchar_trans_on; apply/allP; rewrite /= !Zd. have [||minX1 _]:= subcoherent_norm _ _ (erefl _) defX1Y2. - - rewrite char_vchar ?N_S /orthogonal //= !ooS ?eqxx //. - by rewrite (inv_eq (@cfConjCK _ _)). + - rewrite char_vchar ?N_S /orthogonal //= !oSS ?eqxx ?inv_eq //. + exact: cfConjCK. - apply: sub_iso_to IZtau; last exact: zcharW. by apply: zchar_trans_on; apply/allP; rewrite /= !Zd. - have span_head := memv_span (mem_head _ _); have sRxiX1 := zchar_span RxiX1. - have Y0: Y = 0. - apply/eqP; rewrite -cfnorm_eq0 eqr_le cfnorm_ge0 andbT. - rewrite -(ler_add2l ('[X] + '[X1])) -!addrA. - rewrite -2?cfnormBd -?eqX1Y -?eqXY1 ?addr0; first last. - - by rewrite cfdotC (span_orthogonal oYxi) ?rmorph0 ?span_head. - - by rewrite cfdotC (span_orthogonal oY1chi) ?rmorph0 ?span_head. - by rewrite dotD ?inE ?ne_xi_chi // -defN ler_add. - rewrite eqX1Y Y0 subr0 defN in eqX1. - have [nX _ defX] := eqX1 minX1; exists X => //; red. - rewrite eqXY1 eqX1Y Y0 subr0 opprD opprK addNKr cfdotBr nX. - by rewrite (span_orthogonal (oR _ _ _ S'xi)) ?subr0 ?(zchar_span RxiX1). -pose X_spec X := forall xi, X - D xi \in 'Z[irr G] /\ '[X, D xi] = N%:R. -have [X [RchiX nX defX] X_S'c]: exists2 X, Xspec X & {in S'c chi, X_spec X}. - have [S_chi | /allPn[xi1 Sxi1]] := altP (@allP _ (pred2 chi chi^*%CF) S). - pose E := take N (R chi); pose Ec := drop N (R chi). - have eqRchi: E ++ Ec = R chi by rewrite cat_take_drop. - have:= o1Rchi; rewrite -eqRchi orthonormal_cat => /and3P[onE onEc oEEc]. - exists (\sum_(a <- E) a) => [|xi /and3P[? ? /S_chi/norP[] //]]. - split; last by exists E; rewrite // -[E]cats0 -eqRchi cat_subseq ?sub0seq. - rewrite big_seq rpred_sum // => a Ea. - by rewrite mem_zchar // -eqRchi mem_cat Ea. - by rewrite cfnorm_orthonormal //= size_takel ?szRchi ?leq_addl. - case/norP=> ne_xi1chi ne_xi1chi'; have S'xi1: xi1 \in S'c chi by exact/and3P. - have [X [RchiX nX defX] [Rxi1X1 XD_N]] := haveX _ S'xi1. - exists X => // xi S'xi; have [ne_xi_chi' ne_xi_chi /= Sxi] := and3P S'xi. - have /R_P[ZRxi _ _] := Sxi; have /R_P[ZRxi1 _ defRxi1] := Sxi1. - have [-> | ne_xi_xi1] := eqVneq xi xi1; first by rewrite (zchar_trans ZRxi1). + rewrite eqX1Y cfnormBd // defN in eqX1. + have{eqX1} [|nX n_xi defX] := eqX1; first by rewrite ler_paddr ?cfnorm_ge0. + exists X => //; split; last by rewrite eqXY1 cfdotBr oXY1 subr0. + suffices Y0: Y = 0 by rewrite eqXY1 eqX1Y Y0 subr0 opprB addrC subrK. + apply/eqP; rewrite -cfnorm_eq0 lerif_le ?cfnorm_ge0 //. + by rewrite -(ler_add2l '[X1]) addr0 n_xi. +pose XDspec X := {in S2 chi1, forall xi, '[X, D xi] = N%:R}. +have [X [RchiX nX defX] XD_N]: exists2 X, Xspec X & XDspec X. + have [sSchi | /allPn[xi1 Sxi1]] := altP (@allP _ (pred2 chi1 chi2) S). + pose E := take N (R chi1). + exists (\sum_(a <- E) a) => [|xi]; last by case/and3P=> ? ? /sSchi/norP[]. + have defE: E ++ drop N (R chi1) = R chi1 by rewrite cat_take_drop. + have sER: subseq E (R chi1) by rewrite -defE prefix_subseq. + split; last by [exists E]; move/mem_subseq in sER. + by rewrite big_seq rpred_sum // => a Ea; rewrite mem_zchar ?sER. + rewrite cfnorm_orthonormal ?size_takel ?szRchi ?leq_addl //. + by have:= oRchi; rewrite -defE orthonormal_cat => /andP[]. + case/norP=> chi1'xi1 chi2'xi1'; have S2xi1: xi1 \in S2 chi1 by apply/and3P. + pose xi2 := xi1^*%CF; have /haveX[X [RchiX nX defX] [Rxi1X1 XD_N]] := S2xi1. + exists X => // xi S2xi; have [chi1'xi chi2'xi /= Sxi] := and3P S2xi. + have /R_P[_ _ defRxi1] := Sxi1; have [-> // | xi1'xi] := eqVneq xi xi1. have [sRchiX sRxi1X1] := (zchar_span RchiX, zchar_span Rxi1X1). - have [-> | ne_xi_xi1'] := eqVneq xi xi1^*%CF. - rewrite /D -[chi](subrK xi1) -addrA linearD cfdotDr XD_N opprD addrA. - rewrite defRxi1 big_seq (span_orthogonal (oR _ _ _ S'xi1)) ?addr0 //. - by rewrite rpredB ?rpred_sum // (zchar_trans ZRxi1). - by rewrite memv_suml // => a /memv_span. - have [X' [RchiX' nX' _] [RxiX' X'D_N]] := haveX _ S'xi. - have [ZXi sRxiX'] := (zchar_trans ZRxi RxiX', zchar_span RxiX'). + have [-> | xi2'xi] := eqVneq xi xi2. + rewrite /D -[chi1](subrK xi1) -addrA linearD cfdotDr XD_N defRxi1 big_seq. + rewrite (span_orthogonal (oR chi1 xi1 _ _)) ?addr0 ?rpred_sum //. + exact/memv_span. + have /haveX[X' [RchiX' nX' _] [Rxi3X' X'D_N]] := S2xi. + have [sRchiX' sRxi1X'] := (zchar_span RchiX', zchar_span Rxi3X'). suffices: '[X - X'] == 0 by rewrite cfnorm_eq0 subr_eq0 => /eqP->. - rewrite cfnormB subr_eq0 nX nX' aut_Cint -?mulr2n; last first. - by rewrite Cint_cfdot_vchar ?(zchar_trans ZRchi). - apply/eqP; congr (_ *+ _); transitivity '[D xi1, D xi]. - by rewrite dotD ?inE ?ne_xi_chi ?ne_xi1chi ?ooS ?addr0 // eq_sym. - rewrite -[D xi](subrK X') -opprB addrC -[D _](subrK X) -opprB addrC. - rewrite cfdotBr cfdotBl -addrA addrC -addrA addrCA cfdotBl opprB. + have ZXX': '[X, X'] \in Cint by rewrite Cint_cfdot_vchar ?(zchar_trans ZRchi). + rewrite cfnormB subr_eq0 nX nX' aut_Cint {ZXX'}//; apply/eqP/esym. + congr (_ *+ 2); rewrite -(addNKr (X - D xi1) X) cfdotDl cfdotC. + rewrite (span_orthogonal (oR chi1 xi1 _ _)) // conjC0. + rewrite -(subrK (D xi) X') cfdotDr cfdotDl cfdotNl opprB subrK. rewrite (span_orthogonal (oR xi1 xi _ _)) //; last exact/and3P. - rewrite (span_orthogonal (oR chi xi _ _)) // subrr add0r. - rewrite cfdotC (span_orthogonal (oR chi xi1 _ _)) ?rmorph0 ?oppr0 ?add0r //. - exact: (zchar_span RchiX'). -have ZX: X \in 'Z[irr G] := zchar_trans ZRchi RchiX. -have{defX X_S'c} X_S': {in S' chi, X_spec X}. - move=> xi. - have [-> _| ne_xi_chi' S'xi] := eqVneq xi chi^*%CF; last exact/X_S'c/andP. - rewrite /D defRchi {1}big_seq rpredB ?rpred_sum //. - have{defX} [E sER defX] := defX; pose Ec := filter [predC E] (R chi). - have eqRchi: perm_eq (R chi) (E ++ Ec). - by rewrite -(perm_filterC (mem E)) -(subseq_uniqP _ _) ?free_uniq. - have:= o1Rchi; rewrite (eq_orthonormal eqRchi) orthonormal_cat. - case/and3P=> onE _ oEEc. - rewrite (eq_big_perm _ eqRchi) big_cat /= -defX cfdotDr nX defX !big_seq. - by rewrite (span_orthogonal oEEc) ?addr0 // memv_suml // => ? /memv_span. -pose X_ xi := X - D xi. -have X_chi: X_ chi = X by rewrite /X_ /D subrr linear0 subr0. -have{X_S'} ZI_X: {in S, isometry X_, to 'Z[irr G]}. - have dotXD_N xi: xi \in S' chi -> '[X, D xi] = N%:R by case/X_S'. - have S_S': {subset S <= [predU1 chi & [predD1 S' chi & chi]]}. - by move=> xi; rewrite !inE; case: eqP. - split=> [xi1 xi2 Sxi1 Sxi2 | xi]; last first. - by case/S_S'/predU1P=> [-> | /andP[_ /X_S'[]//]]; rewrite X_chi. - have /predU1P[-> | /andP[chi'xi1 S'xi1]] := S_S' _ Sxi1. - have /predU1P[->|/andP[chi'xi2 S'xi2]] := S_S' _ Sxi2; rewrite X_chi ?nX //. - by rewrite cfdotBr nX dotXD_N // subrr ooS // eq_sym. - have /predU1P[->|/andP[chi'xi2 S'xi2]] := S_S' _ Sxi2. - by rewrite X_chi cfdotBl nX cfdotC dotXD_N // rmorph_nat subrr ooS. - rewrite cfdotBl !cfdotBr nX (cfdotC _ X) !dotXD_N // conjC_nat. - by rewrite opprB subrr add0r dotD // addrC addKr. -have [tau1 Dtau1 Itau1] := Zisometry_of_iso oS ZI_X. -exists tau1; split=> // xi; rewrite zcharD1E. -case/andP=> /zchar_expansion[|z Zz ->{xi}]; first exact: free_uniq. -rewrite defS big_cons /= !cfunE addr_eq0 => eq_z. -have{eq_z} ->: z chi = - \sum_(xi <- S1) z xi. - have nz_chi1: chi 1%g != 0 by rewrite char1_eq0 ?N_S // notS0. - apply: (mulIf nz_chi1); rewrite (eqP eq_z) sum_cfunE mulNr mulr_suml. - congr (- _); apply: eq_big_seq => xi S1xi. - by rewrite cfunE unifS // defS !inE S1xi orbT. -rewrite scaleNr scaler_suml addrC -opprB -sumrB !linearN !linear_sum /=. -apply: eq_big_seq => xi S1xi; rewrite -scalerBr !linearZ /= -/(D _). -congr (_ *: - _); rewrite linearB !Dtau1 // ?defS 1?mem_behead //. -by rewrite X_chi opprD addNKr opprK. + rewrite (span_orthogonal (oR chi1 xi _ _)) // oppr0 !add0r. + by rewrite dotD ?oSS ?addr0 1?eq_sym //; apply/andP. +have{RchiX} ZX: X \in 'Z[irr G] := zchar_trans ZRchi RchiX. +apply: (pivot_coherence cohS Schi1 ZX); rewrite defN //. +move=> xi /andP[chi1'xi Sxi]; exists 1; first by rewrite rpred1 mul1r unifS. +rewrite scale1r mulN1r -conjC_nat -opprB raddfN cfdotNl cfdotC; congr (- _^*). +have [-> /= | chi2'xi] := eqVneq xi chi2; last exact/XD_N/and3P. +have{defX}[E ssER defX] := defX; pose Ec := filter [predC E] (R chi1). +have eqRchi: perm_eq (R chi1) (E ++ Ec). + rewrite -(perm_filterC (mem E)) -(subseq_uniqP _ _) //. + exact/free_uniq/orthonormal_free. +have /and3P[oE _ oEEc]: [&& orthonormal E, orthonormal Ec & orthogonal E Ec]. + by rewrite (eq_orthonormal eqRchi) orthonormal_cat in oRchi. +rewrite defRchi (eq_big_perm _ eqRchi) big_cat -defX cfdotDr nX defX !big_seq. +by rewrite (span_orthogonal oEEc) ?addr0 ?rpred_sum //; apply: memv_span. Qed. End SubCoherentProperties. @@ -1512,7 +1513,7 @@ Local Notation "alpha ^u" := (cfAut u alpha). (* contain non-irreducible characters; for groups of odd order, the second *) (* assumption holds uniformly for all calS of the form seqIndD. *) (* We have stated the coherence assumption directly over L^#; this lets us *) -(* drop the Z{S, A] = Z{S, L^#] assumption, and is more consistent with the *) +(* drop the Z[S, A] = Z[S, L^#] assumption, and is more consistent with the *) (* rest of the proof. *) Lemma cfAut_Dade_coherent calS tau1 chi : coherent_with calS L^# tau tau1 -> diff --git a/mathcomp/odd_order/PFsection6.v b/mathcomp/odd_order/PFsection6.v index d74185a..7654743 100644 --- a/mathcomp/odd_order/PFsection6.v +++ b/mathcomp/odd_order/PFsection6.v @@ -1,11 +1,20 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime ssralg poly finset center. +From mathcomp Require Import fingroup morphism perm automorphism quotient action zmodp. +From mathcomp Require Import gfunctor gproduct cyclic pgroup commutator gseries nilpotent. +From mathcomp Require Import sylow abelian maximal hall frobenius. +From mathcomp Require Import matrix mxalgebra mxrepresentation vector ssrnum algC algnum. +From mathcomp Require Import classfun character inertia vcharacter integral_char. +From mathcomp Require Import PFsection1 PFsection2 PFsection3 PFsection4 PFsection5. (******************************************************************************) @@ -13,7 +22,7 @@ Require Import PFsection1 PFsection2 PFsection3 PFsection4 PFsection5. (* Some Coherence Theorems *) (* Defined here: *) (* odd_Frobenius_quotient K L M <-> *) -(* L has odd order, M <| L, K with K / M is nilpotent, and L / H1 is a *) +(* L has odd order, M <| L, K with K / M nilpotent, and L / H1 is a *) (* Frobenius group with kernel K / H1, where H1 / M = (K / M)^(1). *) (* This is the statement of Peterfalvi, Hypothesis (6.4), except for *) (* the K <| L and subcoherence assumptions, to be required separately. *) @@ -30,7 +39,7 @@ Local Open Scope ring_scope. Section Six. Variables (gT : finGroupType) (G : {group gT}). -Implicit Types H K L M : {group gT}. +Implicit Types H K L P M W Z : {group gT}. (* Grouping lemmas that assume Hypothesis (6.1). *) Section GeneralCoherence. @@ -51,7 +60,7 @@ Let sKL : K \subset L. Proof. exact: normal_sub. Qed. Let nKL : L \subset 'N(K). Proof. exact: normal_norm. Qed. Let orthS: pairwise_orthogonal calS. Proof. by case: scohS. Qed. Let sSS M : {subset S M <= calS}. Proof. exact: seqInd_sub. Qed. -Let ccS M : conjC_closed (S M). Proof. exact: cfAut_seqInd. Qed. +Let ccS M : cfConjC_closed (S M). Proof. exact: cfAut_seqInd. Qed. Let uniqS M : uniq (S M). Proof. exact: seqInd_uniq. Qed. Let nrS : ~~ has cfReal calS. Proof. by case: scohS => [[]]. Qed. @@ -60,13 +69,11 @@ Lemma exists_linInd M : Proof. move=> ltMK nsMK; have [sMK nMK] := andP nsMK. have ntKM: (K / M)%g != 1%g by rewrite -subG1 quotient_sub1 // proper_subn. -have [r /andP[_ r1] ntr] := solvable_has_lin_char ntKM (quotient_sol M solK). -exists ('Ind[L, K] ('chi_r %% M)%CF); last first. - by rewrite cfInd1 // cfModE // morph1 (eqP r1) mulr1. -apply/seqIndP; exists (mod_Iirr r); last by rewrite mod_IirrE. -rewrite !inE subGcfker mod_IirrE ?cfker_mod //= andbT. -apply: contraNneq ntr => /(canRL (mod_IirrK nsMK))->. -by rewrite quo_IirrE // irr0 ?cfker_cfun1 ?cfQuo_cfun1. +have [r lin_r ntr] := solvable_has_lin_char ntKM (quotient_sol M solK). +pose i := mod_Iirr r; exists ('Ind[L] 'chi_i); last first. + by rewrite cfInd1 ?mod_IirrE // cfMod1 lin_char1 ?mulr1. +apply/seqIndP; exists i; rewrite // !inE subGcfker mod_IirrE ?cfker_mod //=. +by rewrite mod_Iirr_eq0 // -irr_eq1 ntr. Qed. (* This is Peterfalvi (6.2). *) @@ -74,58 +81,42 @@ Lemma coherent_seqIndD_bound (A B C D : {group gT}) : [/\ A <| L, B <| L, C <| L & D <| L] -> (*a*) [/\ A \proper K, B \subset D, D \subset C, C \subset K & D / B \subset 'Z(C / B)]%g -> - (*b1*) coherent (S A) L^# tau -> - (*b2*) coherent (S B) L^# tau - \/ #|K : A|%:R - 1 <= 2%:R * #|L : C|%:R * sqrtC #|C : D|%:R. + (*b*) coherent (S A) L^# tau -> \unless coherent (S B) L^# tau, + #|K : A|%:R - 1 <= 2%:R * #|L : C|%:R * sqrtC #|C : D|%:R. Proof. move=> [nsAL nsBL nsCL nsDL] [ltAK sBD sDC sCK sDbZC] cohA. -have [|not_ineq] := boolP (_ <= _); [by right | left]. have sBC := subset_trans sBD sDC; have sBK := subset_trans sBC sCK. have [sAK nsBK] := (proper_sub ltAK, normalS sBK sKL nsBL). have{sBC} [nsAK nsBC] := (normalS sAK sKL nsAL, normalS sBC sCK nsBK). -pose wf S1 := [/\ uniq S1, {subset S1 <= calS} & conjC_closed S1]. -pose S1 := [::] ++ S A; set S2 := [::] in S1; rewrite -[S A]/S1 in cohA. -have wfS1: wf S1 by split; [apply: uniqS | apply: sSS | apply: ccS]. -move: {2}_.+1 (ltnSn (size calS - size S1)) => n. -elim: n => // n IHn in (S2) S1 wfS1 cohA *; rewrite ltnS => leSnS1. -have [uniqS1 sS1S ccS1] := wfS1. -have [sAB1 | /allPn[psi /= SBpsi notS1psi]] := altP (@allP _ (mem S1) (S B)). - by apply: subset_coherent cohA. -have [neq_psi_c SBpsic] := (hasPn nrS _ (sSS SBpsi), ccS SBpsi). -have wfS1': wf [:: psi, psi^* & S1]%CF. - split=> [|xi|xi]; rewrite /= !inE 1?andbC. - - rewrite negb_or eq_sym neq_psi_c notS1psi uniqS1 (contra (ccS1 _)) //. - by rewrite cfConjCK. - - by case/predU1P=> [|/predU1P[|/sS1S]] -> //; rewrite (@sSS B). - do 2![case/predU1P=> [-> |]; first by rewrite ?cfConjCK eqxx ?orbT // eq_sym]. - by move/ccS1=> ->; rewrite !orbT. -apply: (IHn [:: psi, psi^* & S2]%CF) => //; last first. - rewrite -subSn ?uniq_leq_size //; try by case: wfS1'. - by rewrite /= subSS (leq_trans _ leSnS1) // leq_sub2l ?leqW. +rewrite real_lerNgt ?rpredB ?ger0_real ?mulr_ge0 ?sqrtC_ge0 ?ler0n //. +apply/unless_contra; rewrite negbK -(Lagrange_index sKL sCK) natrM => lb_KA. +pose S2 : seq 'CF(L) := [::]; pose S1 := S2 ++ S A; rewrite -[S A]/S1 in cohA. +have ccsS1S: cfConjC_subset S1 calS by apply: seqInd_conjC_subset1. +move: {2}_.+1 (leq_addr (size S1) (size calS).+1) => n. +elim: n => [|n IHn] in (S2) S1 ccsS1S cohA * => lb_n. + by rewrite ltnNge uniq_leq_size // in lb_n; case: ccsS1S. +without loss /allPn[psi /= SBpsi S1'psi]: / ~~ all (mem S1) (S B). + by case: allP => [sAB1 _ | _]; [apply: subset_coherent cohA | apply]. +have [[_ sS1S _] Spsi] := (ccsS1S, sSS SBpsi). +apply (IHn [:: psi, psi^* & S2]%CF); rewrite ?addnS 1?leqW {n lb_n IHn}//= -/S1. + exact: extend_cfConjC_subset. have [phi SAphi phi1] := exists_linInd ltAK nsAK. -have: [/\ phi \in S1, psi \in calS & psi \notin S1]. - by rewrite mem_cat SAphi orbT (@sSS B). -have /seqIndP[i /setDP[kBi _] def_psi] := SBpsi; rewrite inE in kBi. -move/(extend_coherent scohS); apply; rewrite // {phi SAphi}phi1; split=> //. - by rewrite def_psi cfInd1 // dvdC_mulr // CintE Cnat_irr1. -have Spos xi: xi \in calS -> 0 <= xi 1%g by move/Cnat_seqInd1/Cnat_ge0. -rewrite big_cat sum_seqIndD_square //= -subr_gt0 -addrA ltr_paddl //=. - rewrite big_seq sumr_ge0 // => xi S2xi. - by rewrite !mulr_ge0 ?invr_ge0 ?cfnorm_ge0 ?Spos ?sS1S // mem_cat S2xi. -rewrite mulrC -mulrBl pmulr_rgt0 ?gt0CiG // subr_gt0. -rewrite real_ltrNge ?rpredB ?rpredM ?rpred_nat ?rpred1 //; last first. - by rewrite realE Spos ?(sSS SBpsi). -apply: contra not_ineq => /ler_trans-> //. -rewrite -mulrA ler_pmul2l ?ltr0n // def_psi cfInd1 //. -rewrite -(Lagrange_index sKL sCK) natrM -mulrA ler_pmul2l ?gt0CiG //. -exact: irr1_bound_quo sDbZC. +have{SAphi} S1phi: phi \in S1 by rewrite mem_cat SAphi orbT. +apply: (extend_coherent scohS) ccsS1S S1phi Spsi S1'psi _. +have{SBpsi} /seqIndP[i /setDP[kBi _] {psi}->] := SBpsi; rewrite inE in kBi. +rewrite {phi}phi1 cfInd1 // dvdC_mulr //; last by rewrite CintE Cnat_irr1. +split; rewrite // big_cat sum_seqIndD_square // big_seq ltr_paddl //=. + apply/sumr_ge0=> xi S2xi; rewrite divr_ge0 ?cfnorm_ge0 ?exprn_ge0 //. + by rewrite Cnat_ge0 // (Cnat_seqInd1 (sS1S _ _)) // mem_cat S2xi. +rewrite mulrC ltr_pmul2l ?gt0CiG //; apply: ler_lt_trans lb_KA. +by rewrite -!mulrA !ler_wpmul2l ?ler0n // (irr1_bound_quo nsBC). Qed. (* This is Peterfalvi, Theorem (6.3). *) -Theorem bounded_seqIndD_coherent M H H1 : +Theorem bounded_seqIndD_coherence M H H1 : [/\ M <| L, H <| L & H1 <| L] -> [/\ M \subset H1, H1 \subset H & H \subset K] -> - (*a*) nilpotent (H / M)%g -> + (*a*) nilpotent (H / M) -> (*b*) coherent (S H1) L^# tau -> (*c*) (#|H : H1| > 4 * #|L : K| ^ 2 + 1)%N -> coherent (S M) L^# tau. @@ -144,41 +135,36 @@ suffices{m IHm leAm} cohB: coherent (S B) L^# tau. have /andP[sHL nHL] := nsHL. have sAbZH: (A / B \subset 'Z(H / B))%g. have nBA := subset_trans sAL nBL; have nsBA : B <| A by apply/andP. - have minBb: minnormal (A / B)%g (L / B)%g. - apply/mingroupP; split=> [|Db /andP[ntDb nDLb] sDAb]. + set Zbar := 'Z(H / B); set Abar := (A / B)%g; pose Lbar := (L / B)%g. + have nZHbar: Lbar \subset 'N(Zbar) by rewrite gFnorm_trans ?quotient_norms. + have /mingroupP[/andP[ntAbar nALbar] minBbar]: minnormal Abar Lbar. + apply/mingroupP; split=> [|Dbar /andP[ntDbar nDLbar] sDAbar]. by rewrite -subG1 quotient_sub1 // not_sAB quotient_norms. - have: Db <| (L / B)%g by rewrite /normal (subset_trans sDAb) ?quotientS. - case/(inv_quotientN nsBL)=> D defDb sBD /andP[sDL nDL]. - apply: contraNeq ntDb => neqDAb; rewrite defDb quotientS1 //. - case/maxgroupP: maxB => /= _ /(_ D) {1}<- //. - rewrite -(quotient_proper (normalS sBD sDL nsBL)) // -defDb. - by rewrite properEneq sDAb neqDAb. - apply/setIidPl; case/mingroupP: minBb => /andP[ntAb nALb]. - apply; rewrite ?subsetIl //. - have nZHb := char_norm_trans (center_char (H / B)) (quotient_norms _ nHL). - rewrite andbC normsI //= meet_center_nil //=; last first. - by rewrite quotient_normal // (normalS sAH sHL). + have: Dbar <| Lbar by rewrite /normal (subset_trans sDAbar) ?quotientS. + case/(inv_quotientN nsBL)=> D defDbar sBD /andP[sDL nDL]. + apply: contraNeq ntDbar => neqDAbar; rewrite defDbar quotientS1 //. + have [_ /(_ D) {1}<- //] := maxgroupP maxB. + rewrite -(quotient_proper (normalS sBD sDL nsBL)) // -defDbar. + by rewrite properEneq sDAbar neqDAbar. + apply/setIidPl/minBbar; rewrite ?subsetIl {minBbar}//= andbC -/Abar -/Zbar. + rewrite normsI ?meet_center_nil ?quotient_normal ?(normalS sAH sHL) //=. suffices /homgP[f /= <-]: (H / B)%g \homg (H / M)%g by rewrite morphim_nil. by apply: homg_quotientS; rewrite ?(subset_trans sHL) ?normal_norm. -have [defA | ltAH] := eqVproper sAH. - by rewrite addn1 defA indexgg in lbHA. -have [sAK ltAK] := (subset_trans sAH sHK, proper_sub_trans ltAH sHK). -case: (@coherent_seqIndD_bound A B H A) => // /idPn[]. -apply: contraL lbHA; rewrite -ltnNge -ltC_nat -(Lagrange_index sHK sAH) natrM. -set x := #|H : A|%:R => ub_x. -have nz_x2: sqrtC x != 0 by rewrite sqrtC_eq0 neq0CiG. -have x2_gt0: 0 < sqrtC x by rewrite ltr_def nz_x2 sqrtC_ge0 ler0n. -have{ub_x}: sqrtC x - (sqrtC x)^-1 <= (2 * #|L : K|)%:R. - rewrite -(ler_pmul2r (gt0CiG K H)) -natrM -mulnA Lagrange_index //. - rewrite natrM -(ler_pmul2r x2_gt0) mulrC mulrBl mulrBr. - rewrite !mulrA -expr2 sqrtCK divff // (ler_trans _ ub_x) // mulrC. - by rewrite ler_add2l ler_opp2 mul1r ler1n. -rewrite -(@ler_pexpn2r _ 2) -?topredE /= ?ler0n //; last first. - rewrite subr_ge0 -(ler_pmul2r x2_gt0) -expr2 mulVf // sqrtCK. - by rewrite ler1n. -rewrite -natrX expnMn -(ler_add2r 2%:R) -addnS natrD. -apply: ltr_le_trans; rewrite sqrrB // exprVn sqrtCK divff //. -by rewrite addrAC subrK addrC -subr_gt0 addrK invr_gt0 gt0CiG. +have ltAH: A \proper H. + by rewrite properEneq sAH (contraTneq _ lbHA) // => ->; rewrite indexgg addn1. +set x := sqrtC #|H : A|%:R. +have [nz_x x_gt0]: x != 0 /\ 0 < x by rewrite gtr_eqF sqrtC_gt0 gt0CiG. +without loss{cohA} ubKA: / #|K : A|%:R - 1 <= 2%:R * #|L : H|%:R * x. + have [sAK ltAK] := (subset_trans sAH sHK, proper_sub_trans ltAH sHK). + exact: coherent_seqIndD_bound id. +suffices{lbHA}: (x - x^-1) ^+ 2 <= (2 * #|L : K|)%:R ^+ 2. + rewrite ltr_geF // sqrrB divff // sqrtCK ltr_spaddr ?exprn_gt0 ?invr_gt0 //. + by rewrite ler_subr_addr -natrX -natrD ler_nat expnMn addnS lbHA. +rewrite ler_pexpn2r ?unfold_in /= ?ler0n //; last first. + by rewrite subr_ge0 -div1r ler_pdivr_mulr // -expr2 sqrtCK ler1n. +rewrite -(ler_pmul2l x_gt0) -(ler_pmul2l (gt0CiG K H)) 2!mulrBr -expr2 sqrtCK. +rewrite !mulrA mulfK // mulrAC natrM mulrCA -2!natrM [in _ * x]mulnC. +by rewrite !Lagrange_index // (ler_trans _ ubKA) // ler_add2l ler_opp2 ler1n. Qed. (* This is the statement of Peterfalvi, Hypothesis (6.4). *) @@ -189,101 +175,92 @@ Definition odd_Frobenius_quotient M (H1 := K^`(1) <*> M) := (* This is Peterfalvi (6.5). *) Lemma non_coherent_chief M (H1 := (K^`(1) <*> M)%G) : - odd_Frobenius_quotient M -> - coherent (S M) L^# tau -\/ [/\ (*a*) chief_factor L H1 K /\ (#|K : H1| <= 4 * #|L : K| ^ 2 + 1)%N - & (*b*) exists2 p : nat, p.-group (K / M)%g /\ ~~ abelian (K / M)%g + odd_Frobenius_quotient M -> \unless coherent (S M) L^# tau, + [/\ (*a*) chief_factor L H1 K /\ (#|K : H1| <= 4 * #|L : K| ^ 2 + 1)%N + & (*b*) exists2 p : nat, p.-group (K / M)%g /\ ~~ abelian (K / M) & (*c*) ~~ (#|L : K| %| p - 1)]. Proof. -case=> oddL [nsML sMK nilKb]; rewrite /= -(erefl (gval H1)) => frobLb. +case=> oddL [nsML sMK nilKM]; rewrite /= -(erefl (gval H1)) => frobLb. set e := #|L : K|; have odd_e: odd e := dvdn_odd (dvdn_indexg L K) oddL. -have{odd_e} mod1e_lb m: (odd m -> m > 1 -> m == 1 %[mod e] -> 2 * e + 1 <= m)%N. - move=> odd_m m_gt1; rewrite eqn_mod_dvd ?(ltnW m_gt1) //. - rewrite -[m]odd_double_half odd_m subn1 /= -mul2n addn1 ltnS leq_pmul2l //. - rewrite Gauss_dvdr; last by rewrite coprime_sym prime_coprime // dvdn2 odd_e. - by apply: dvdn_leq; rewrite -(subnKC m_gt1). -have nsH1L: H1 <| L by rewrite normalY // (char_normal_trans (der_char 1 K)). -have sH1K: H1 \subset K by rewrite join_subG der_sub. +have{odd_e} mod1e_lb m: odd m -> m == 1 %[mod e] -> (m > 1 -> 2 * e + 1 <= m)%N. + move=> odd_m e_dv_m1 m_gt1; rewrite eqn_mod_dvd 1?ltnW // subn1 in e_dv_m1. + by rewrite mul2n addn1 dvdn_double_ltn. +have nsH1L: H1 <| L by rewrite normalY // gFnormal_trans. +have nsH1K: H1 <| K by rewrite (normalS _ sKL nsH1L) // join_subG der_sub. +have [sH1K nH1K] := andP nsH1K; have sMH1: M \subset H1 by apply: joing_subr. have cohH1: coherent (S H1) L^# tau. apply: uniform_degree_coherence (subset_subcoherent scohS _) _ => //. - apply/(@all_pred1_constant _ #|L : K|%:R)/allP=> _ /mapP[chi Schi ->] /=. + apply/(@all_pred1_constant _ e%:R)/allP=> _ /mapP[chi Schi ->] /=. have [i /setIdP[_]] := seqIndP Schi; rewrite inE join_subG -lin_irr_der1. - by do 2![case/andP]=> _ /eqP chi1 _ ->; rewrite cfInd1 // chi1 mulr1. -have sMH1: M \subset H1 by apply: joing_subr. -have [ubK | lbK] := leqP; last by left; apply: bounded_seqIndD_coherent lbK. + by case/andP=> lin_chi _ ->; rewrite cfInd1 ?lin_char1 ?mulr1. +apply/unlessP; have [/val_inj-> | ltMH1] := eqVproper sMH1; first by left. +have [lbK|ubK] := ltnP; [by left; apply: bounded_seqIndD_coherence lbK | right]. have{ubK} ubK: (#|K : H1| < (2 * e + 1) ^ 2)%N. - rewrite sqrnD expnMn (leq_ltn_trans ubK) // -subn_gt0 addKn. + apply: leq_ltn_trans ubK _; rewrite -subn_gt0 sqrnD expnMn addKn. by rewrite !muln_gt0 indexg_gt0. -have [-> | neqMH1] := eqVneq M H1; [by left | right]. -have{neqMH1} ltMH1: M \proper H1 by rewrite properEneq neqMH1. have{frobLb} [[E1b frobLb] [sH1L nH1L]] := (existsP frobLb, andP nsH1L). have [defLb ntKb _ _ /andP[sE1L _]] := Frobenius_context frobLb. -have nH1K: K \subset 'N(H1) := subset_trans sKL nH1L. -have chiefH1: chief_factor L H1 K. +have iH1_mod1e H2: + H1 \subset H2 -> H2 \subset K -> L \subset 'N(H2) -> #|H2 : H1| == 1 %[mod e]. +- move=> sH12 sH2K nPL; have sH2L := subset_trans sH2K sKL. + rewrite eqn_mod_dvd // subn1 -card_quotient ?(subset_trans sH2L) //. + have [-> | ntH2b] := eqVneq (H2 / H1)%g 1%g; first by rewrite cards1. + have ->: e = #|E1b|. + by rewrite (index_sdprod defLb) index_quotient_eq ?(setIidPr sH1L). + have /Frobenius_subl/Frobenius_dvd_ker1-> := frobLb; rewrite ?quotientS //. + by rewrite (subset_trans sE1L) ?quotient_norms. +have{iH1_mod1e} chiefH1: chief_factor L H1 K. have ltH1K: H1 \proper K by rewrite /proper sH1K -quotient_sub1 ?subG1. rewrite /chief_factor nsKL andbT; apply/maxgroupP; rewrite ltH1K. split=> // H2 /andP[ltH2K nH2L] sH12; have sH2K := proper_sub ltH2K. have /eqVproper[// | ltH21] := sH12; case/idPn: ubK; rewrite -leqNgt. - have dv_e H3: H1 \subset H3 -> H3 \subset K -> L \subset 'N(H3) -> - #|H3 : H1| == 1 %[mod e]. - - move=> sH13 sH3K nH3L; rewrite eqn_mod_dvd // subn1. - rewrite /e -(index_quotient_eq _ sKL nH1L) ?subIset ?sH1K ?orbT //. - rewrite -[#|_ : _|]divgS ?quotientS // -(sdprod_card defLb) mulKn //. - rewrite -card_quotient ?(subset_trans (subset_trans sH3K sKL)) //. - rewrite regular_norm_dvd_pred ?(subset_trans sE1L) ?quotient_norms //. - apply: semiregular_sym; apply: sub_in1 (Frobenius_reg_compl frobLb). - by apply/subsetP; rewrite setSD ?quotientS. - have dv_H21 := dv_e H2 sH12 sH2K nH2L. - have dv_KH2: #|K : H2| == 1 %[mod e]. - have:= dv_e K sH1K (subxx K) nKL; rewrite -(Lagrange_index sH2K sH12). - by rewrite -modnMmr (eqP dv_H21) modnMmr muln1. + have iKH1: (#|K : H2| * #|H2 : H1|)%N = #|K : H1| by apply: Lagrange_index. + have iH21_mod1e: #|H2 : H1| == 1 %[mod e] by apply/iH1_mod1e. + have iKH1_mod1e: #|K : H1| = 1 %[mod e] by apply/eqP/iH1_mod1e. + have iKH2_mod1e: #|K : H2| == 1 %[mod e]. + by rewrite -iKH1_mod1e -iKH1 -modnMmr (eqP iH21_mod1e) modnMmr muln1. have odd_iK := dvdn_odd (dvdn_indexg _ _) (oddSg (subset_trans _ sKL) oddL). - have iK_gt1 H3 H4: H4 \proper H3 -> (#|H3 : H4| > 1)%N. - by rewrite indexg_gt1 => /andP[]. - by rewrite -(Lagrange_index sH2K sH12) leq_mul ?mod1e_lb ?odd_iK ?iK_gt1. -split=> //; have nMK := subset_trans sKL (normal_norm nsML). + by rewrite -iKH1 leq_mul ?mod1e_lb ?odd_iK ?indexg_gt1 ?proper_subn. +have nMK: K \subset 'N(M) := subset_trans sKL (normal_norm nsML). +have nMK1: K^`(1)%g \subset 'N(M) by apply: gFsub_trans. have not_abKb: ~~ abelian (K / M). - apply: contra (proper_subn ltMH1) => /derG1P/trivgP. - rewrite /= join_subG subxx andbT -quotient_der ?quotient_sub1 //. - exact: subset_trans (der_sub 1 K) nMK. -have /is_abelemP[p p_pr /and3P[pKb _ _]]: is_abelem (K / H1)%g. + apply: contra (proper_subn ltMH1) => /derG1P/trivgP/=. + by rewrite join_subG subxx andbT -quotient_der ?quotient_sub1. +have /is_abelemP[p p_pr /and3P[pKb _ _]]: is_abelem (K / H1). have: solvable (K / H1)%g by apply: quotient_sol solK. by case/(minnormal_solvable (chief_factor_minnormal chiefH1)). -have [_ p_dv_Kb _] := pgroup_pdiv pKb ntKb. -have iso3M := third_isog sMH1 (normalS sMK sKL nsML) (normalS sH1K sKL nsH1L). -have pKM: p.-group (K / M)%g. - have /dprodP[_ defKM cKMpp' tiKMpp'] := nilpotent_pcoreC p nilKb. - rewrite -defKM (eqP (forall_inP (nilpotent_sol nilKb) 'O_p^'(_)%G _)). - by rewrite mulg1 pcore_pgroup. - have /isomP[inj_quo im_quo] := quotient_isom (cents_norm cKMpp') tiKMpp'. - rewrite subsetI pcore_sub /= -(injmSK inj_quo) // (morphim_der _ 1) //. - rewrite {inj_quo}im_quo /= -[Q in Q^`(1)%g]quotientMidl defKM. - rewrite -quotient_der ?gFnorm ?quotientS //. - rewrite -quotient_sub1 ?(subset_trans (pcore_sub _ _) (der_norm _ _)) //. - rewrite -[(_ / _)%g]setIid coprime_TIg //. - apply: pnat_coprime (quotient_pgroup _ (pcore_pgroup _ _)). - apply: pgroupS (quotientS _ (pcore_sub _ _)) _. - rewrite /= -quotient_der // -(quotientYidr (subset_trans (der_sub 1 K) nMK)). - by rewrite (isog_pgroup _ iso3M) ?(normalS sMK sKL nsML). -exists p => //; apply: contra not_abKb => e_dv_p1. +have [[_ p_dv_Kb _] nsMK] := (pgroup_pdiv pKb ntKb, normalS sMK sKL nsML). +have isoKb: K / M / (H1 / M) \isog K / H1 := third_isog sMH1 nsMK nsH1K. +have{nilKM} pKM: p.-group (K / M)%g. + pose Q := 'O_p^'(K / M); have defKM: _ \x Q = _ := nilpotent_pcoreC p nilKM. + have nH1Q: Q \subset 'N(H1 / M) by rewrite gFsub_trans ?quotient_norms. + have hallQb := quotient_pHall nH1Q (nilpotent_pcore_Hall p^' nilKM). + have{nH1Q hallQb pKb} sQH1: (Q \subset H1 / M)%g. + rewrite -quotient_sub1 // subG1 trivg_card1 /= (card_Hall hallQb). + by rewrite partG_eq1 pgroupNK (isog_pgroup p isoKb). + suffices Q_1: Q = 1%g by rewrite -defKM Q_1 dprodg1 pcore_pgroup. + apply: contraTeq sQH1 => ntQ; rewrite quotientYidr ?quotient_der //. + rewrite (sameP setIidPl eqP) -(dprod_modr (der_dprod 1 defKM)) ?gFsub //= -/Q. + rewrite setIC coprime_TIg ?(coprimeSg (der_sub 1 _)) ?coprime_pcoreC //. + by rewrite dprod1g proper_neq ?(sol_der1_proper (nilpotent_sol nilKM)) ?gFsub. +split=> //; exists p => //; apply: contra not_abKb => e_dv_p1. rewrite cyclic_abelian // Phi_quotient_cyclic //. have /homgP[f <-]: (K / M / 'Phi(K / M) \homg K / H1)%g. - apply: homg_trans (isog_hom iso3M). - rewrite homg_quotientS ?gFnorm ?quotient_norms //=. - rewrite quotientYidr ?(subset_trans (der_sub 1 K)) // quotient_der //. - by rewrite (Phi_joing pKM) joing_subl. + apply: homg_trans (isog_hom isoKb). + rewrite homg_quotientS ?gFnorm ?quotient_norms //= quotientYidr //. + by rewrite quotient_der // (Phi_joing pKM) joing_subl. rewrite {f}morphim_cyclic // abelian_rank1_cyclic; last first. by rewrite sub_der1_abelian ?joing_subl. rewrite (rank_pgroup pKb) (leq_trans (p_rank_le_logn _ _)) //. rewrite -ltnS -(ltn_exp2l _ _ (prime_gt1 p_pr)) -p_part part_pnat_id //. rewrite card_quotient // (leq_trans ubK) // leq_exp2r //. -have odd_p: odd p := dvdn_odd p_dv_Kb (quotient_odd _ (oddSg sKL oddL)). -by rewrite mod1e_lb // ?eqn_mod_dvd ?prime_gt0 ?prime_gt1. +have odd_p: odd p by rewrite (dvdn_odd p_dv_Kb) ?quotient_odd ?(oddSg sKL). +by rewrite mod1e_lb ?eqn_mod_dvd ?prime_gt0 ?prime_gt1. Qed. (* This is Peterfalvi (6.6). *) -Lemma seqIndD_irr_coherence (Z : {group gT}) (calX := seqIndD K L Z 1) : - odd_Frobenius_quotient 1%G -> +Lemma seqIndD_irr_coherence Z (calX := seqIndD K L Z 1) : + odd_Frobenius_quotient 1 -> [/\ Z <| L, Z :!=: 1 & Z \subset 'Z(K)]%g -> {subset calX <= irr L} -> calX =i [pred chi in irr L | ~~ (Z \subset cfker chi)] @@ -306,16 +283,14 @@ split=> [chi|]. by rewrite cfker_constt // cfInd_char ?irr_char //. case/irrX/irrP: Xchi chi_r (Xchi) => r' ->. by rewrite cfdot_irr pnatr_eq0 -lt0n; case: eqP => // ->. -have [|[]] := non_coherent_chief Frob_quo1. - by apply: subset_coherent; apply: seqInd_sub. -have [oddL _] := Frob_quo1; rewrite /= joingG1 => frobLb _ [p []]. -set e := #|L : K|; have e_gt0: (e > 0)%N by apply: indexg_gt0. -have isoK1 := isog_symr (quotient1_isog K). -rewrite (isog_abelian isoK1) {isoK1}(isog_pgroup _ isoK1). -have [-> | ntK pK _ not_e_dv_p1] := eqsVneq K [1]; first by rewrite abelian1. -have{ntK} [p_pr p_dv_K _] := pgroup_pdiv pK ntK. +apply: non_coherent_chief (subset_coherent (seqInd_sub sZK)) _ => //= -[_ [p]]. +have [oddL _] := Frob_quo1; rewrite joingG1 -/calX => frobLb []. +rewrite -(isog_pgroup p (quotient1_isog K)) => pK ab'K. +set e := #|L : K| => not_e_dv_p1; have e_gt0: (e > 0)%N by apply: indexg_gt0. +have ntK: K != 1%G by apply: contraNneq ab'K => ->; rewrite quotient1 abelian1. +have{ab'K ntK} [p_pr p_dv_K _] := pgroup_pdiv pK ntK. set Y := calX; pose d (xi : 'CF(L)) := logn p (truncC (xi 1%g) %/ e). -have: conjC_closed Y by apply: cfAut_seqInd. +have: cfConjC_closed Y by apply: cfAut_seqInd. have: perm_eq (Y ++ [::]) calX by rewrite cats0. have: {in Y & [::], forall xi1 xi2, d xi1 <= d xi2}%N by []. elim: {Y}_.+1 {-2}Y [::] (ltnSn (size Y)) => // m IHm Y X' leYm leYX' defX ccY. @@ -345,7 +320,7 @@ have{homoY} /hasP[xi1 Yxi1 lt_xi1_chi]: has (fun xi => d xi < d chi)%N Y. rewrite eqn_pmul2l // eqn_exp2l ?prime_gt1 //. by rewrite eqn_leq leYchi //= leqNgt (hasPn geYchi). pose Y' := rem chi^*%CF (rem chi Y); pose X'' := [:: chi, chi^*%CF & X']. -have ccY': conjC_closed Y'. +have ccY': cfConjC_closed Y'. move=> xi; rewrite !(inE, mem_rem_uniq) ?rem_uniq //. by rewrite !(inv_eq (@cfConjCK _ _)) cfConjCK => /and3P[-> -> /ccY->]. have Xchi := sYX _ Ychi; have defY: perm_eq [:: chi, chi^*%CF & Y'] Y. @@ -354,22 +329,21 @@ have Xchi := sYX _ Ychi; have defY: perm_eq [:: chi, chi^*%CF & Y'] Y. apply: perm_eq_coherent (defY) _. have d_chic: d chi^*%CF = d chi. by rewrite /d cfunE conj_Cnat // (Cnat_seqInd1 Xchi). -have /andP[uniqY' Y'x1]: uniq Y' && (xi1 \in Y'). - rewrite !(inE, mem_rem_uniq) ?rem_uniq // Yxi1 andbT -negb_or. +have /and3P[uniqY' Y'xi1 notY'chi]: [&& uniq Y', xi1 \in Y' & chi \notin Y']. + rewrite !(inE, mem_rem_uniq) ?rem_uniq // Yxi1 eqxx andbF !andbT -negb_or. by apply: contraL lt_xi1_chi => /pred2P[] ->; rewrite ?d_chic ltnn. -have xi1P: [/\ xi1 \in Y', chi \in calS & chi \notin Y']. - by rewrite Y'x1 sYS ?(inE, mem_rem_uniq) ?rem_uniq // eqxx andbF. have sY'Y: {subset Y' <= Y} by move=> xi /mem_rem/mem_rem. -apply: (extend_coherent scohS) xi1P _; first by split=> // xi /sY'Y/sYS. +have sccY'S: cfConjC_subset Y' calS by split=> // xi /sY'Y/sYS. +apply: (extend_coherent scohS _ Y'xi1); rewrite ?sYS {sccY'S notY'chi}//. have{defX} defX: perm_eq (Y' ++ X'') calX. by rewrite (perm_catCA Y' [::_; _]) catA -(perm_eqrP defX) perm_cat2r. have{d_chic} le_chi_X'': {in X'', forall xi, d chi <= d xi}%N. by move=> xi /or3P[/eqP-> | /eqP-> | /leYX'->] //; rewrite d_chic. rewrite !Ndg ?sYX // dvdC_nat dvdn_pmul2l // dvdn_exp2l 1?ltnW //; split=> //. - apply: IHm defX ccY' => [|xi xi' /sY'Y/leYchi le_xi_chi /le_chi_X'']. + apply: IHm defX ccY' => [|xi xi' /sY'Y/leYchi-le_xi_chi /le_chi_X'']. by rewrite -ltnS // (leq_trans _ leYm) // -(perm_eq_size defY) ltnW. exact: leq_trans. -have pos_p n: (0 < p ^ n)%N by rewrite expn_gt0 prime_gt0. +have p_gt0 n: (0 < p ^ n)%N by rewrite expn_gt0 prime_gt0. rewrite -!natrM; apply: (@ltr_le_trans _ (e ^ 2 * (p ^ d chi) ^ 2)%:R). rewrite ltr_nat -expnMn -mulnn mulnAC !mulnA 2?ltn_pmul2r //. rewrite -mulnA mulnCA ltn_pmul2l // -(subnK lt_xi1_chi) addnS expnS. @@ -386,10 +360,10 @@ have def_sum_xi1 S: {subset S <= calX} -> sum_xi1 S = (e ^ 2 * sum_p2d S)%:R. rewrite expnM -expnMn natrX -Ndg //. by have /irrP[i ->] := irrX _ Xxi; rewrite cfnorm_irr divr1. rewrite -/(sum_xi1 _) def_sum_xi1 ?leC_nat 1?dvdn_leq => [|||_ /sY'Y/sYX] //. - by rewrite muln_gt0 expn_gt0 e_gt0 [_ Y'](bigD1_seq xi1) //= addn_gt0 pos_p. + by rewrite muln_gt0 expn_gt0 e_gt0 [_ Y'](bigD1_seq xi1) //= addn_gt0 p_gt0. have coep: coprime e p. have:= Frobenius_ker_coprime frobLb; rewrite coprime_sym. - have /andP[_ nK'L] := char_normal_trans (der_char 1 K) nsKL. + have /andP[_ nK'L]: K^`(1) <| L by apply: gFnormal_trans. rewrite index_quotient_eq ?subIset ?der_sub ?orbT {nK'L}// -/e. have ntKb: (K / K^`(1))%g != 1%g by case/Frobenius_kerP: frobLb. have [_ _ [k ->]] := pgroup_pdiv (quotient_pgroup _ pK) ntKb. @@ -412,9 +386,9 @@ Qed. End GeneralCoherence. (* This is Peterfalvi (6.7). *) -(* In (6.8) we only know initially the P is Sylow in L; perhaps the lemma *) -(* should be stated with this equivalent (but weaker) assumption. *) -Lemma constant_irr_mod_TI_Sylow (Z L P : {group gT}) p i : +(* In (6.8) we only know initially the P group is Sylow in L; perhaps this *) +(* lemma should be stated with this equivalent (but weaker) assumption. *) +Lemma constant_irr_mod_TI_Sylow Z L P p i : p.-Sylow(G) P -> odd #|L| -> normedTI P^# G L -> [/\ Z <| L, Z :!=: 1%g & Z \subset 'Z(P)] -> {in Z^# &, forall x y, #|'C_L[x]| = #|'C_L[y]| } -> @@ -461,7 +435,7 @@ have{actsGC} PdvKa i j s: suffices inZ k y: y \in C k -> ~~ dC k Z^# -> y ^ x = y -> y \in Z. apply/exists_inP; exists (u * v)%g => //=. by rewrite groupM // (inZ i u, inZ j v). - rewrite /dC /C; have /imsetP[_ _ ->{k} /class_transr <-] := enum_valP k. + rewrite /dC /C; have /imsetP[_ _ ->{k} /class_eqP <-] := enum_valP k. case/exists_inP=> _ /imsetP[g Gg ->] /setD1P[nt_yg Zyg] yx. have xy: (x ^ y = x)%g by rewrite /conjg (conjgCV x) -{2}yx conjgK mulKg. rewrite -(memJ_conjg _ g) (normsP nZL) //. @@ -479,8 +453,8 @@ have [[Pz Lz] Gz] := (subsetP sZP z Zz, subsetP sZL z Zz, subsetP sZG z Zz). pose inC y Gy := enum_rank_in (@mem_classes _ y G Gy) (y ^: G). have CE y Gy: C (inC y Gy) = y ^: G by rewrite /C enum_rankK_in ?mem_classes. pose i0 := inC _ (group1 G); pose i1 := inC z Gz; pose i2 := inC _ (groupVr Gz). -suffices Ea2 l (phi := 'chi[G]_l): - kerZ l -> (phi z *+ a2 i1 i1 == phi 1%g + phi z *+ a2 i1 i2 %[mod #|P|])%A. +suffices Ea2 l (phi := 'chi[G]_l) (kerZphi : kerZ l): + (phi z *+ a2 i1 i1 == phi 1%g + phi z *+ a2 i1 i2 %[mod #|P|])%A. - move=> l phi kerZphi. have Zphi1: phi 1%g \in Cint by rewrite irr1_degree rpred_nat. have chi0 x: x \in Z -> 'chi[G]_0 x = 1. @@ -502,10 +476,10 @@ suffices Ea2 l (phi := 'chi[G]_l): rewrite rmorph1 mulr1; congr (_ * (_ + _) \in Crat). rewrite -sumr_const; apply: eq_bigr => x Z1x; have [_ Zx] := setD1P Z1x. by rewrite cfun1E cfResE ?Zx // rmorph1 mulr1; apply: kerZphi. -move=> kerZphi; pose alpha := 'omega_l['K_i1]; pose phi1 := phi 1%g. +pose alpha := 'omega_l['K_i1]; pose phi1 := phi 1%g. have tiZG: {in Z^#, forall y, 'C_G[y] \subset L}. move=> y /setD1P[nty /(subsetP sZP)Py]. - apply/subsetP=> u /setIP[Gu /cent1P cuy]. + apply/subsetP=> u /setIP[Gu /cent1P-cuy]. by rewrite -(memJ_P1 y) // /conjg -?cuy ?mulKg !inE nty. have Dalpha s: ~~ dC s Z^# -> alpha = 'omega_l['K_s]. case/exists_inP=> x /= /gring_mode_class_sum_eq-> Z1x. @@ -540,7 +514,7 @@ have Dalpha2 i j: ~~ dC i Z^# -> ~~ dC j Z^# -> rewrite eqAmod0_rat ?rpred_nat // dvdC_nat PdvKa //. rewrite -(setD1K (group1 Z)) [dC _ _]disjoint_sym disjoints_subset. rewrite subUset sub1set inE -disjoints_subset disjoint_sym. - rewrite (contra _ ntCs) // [C s]defCs => /class_transr. + rewrite (contra _ ntCs) // [C s]defCs => /class_eqP. by rewrite -(inj_eq enum_val_inj) defCs -/(C _) CE => ->. have zG'z1: (z^-1 \notin z ^: G)%g. have genL2 y: y \in L -> <[y]> = <[y ^+ 2]>. @@ -554,7 +528,7 @@ have zG'z1: (z^-1 \notin z ^: G)%g. by rewrite -(memJ_P1 z) -?zy ?in_setD ?groupV ?inE ?ntz. have a110: a i1 i1 i0 = 0%N. apply: contraNeq zG'z1 => /existsP[[u v] /setIdP[/andP[/=]]]. - rewrite rCi10 -!/(C _) !CE -eq_invg_mul => /imsetP[x Gx ->] /class_transr <-. + rewrite rCi10 -!/(C _) !CE -eq_invg_mul => /imsetP[x Gx ->] /class_eqP <-. by move/eqP <-; rewrite -conjVg classGidl ?class_refl. have a120: a i1 i2 i0 = #|C i1|. rewrite -(card_imset _ (@can_inj _ _ (fun y => (y, y^-1)%g) (@fst _ _) _)) //. @@ -593,7 +567,7 @@ Qed. (* This is Peterfalvi, Theorem (6.8). *) (* We omit the semi-direct structure of L in assumption (a), since it is *) (* implied by our statement of assumption (c). *) -Theorem Sibley_coherence (L H W1 : {group gT}) : +Theorem Sibley_coherence L H W1 : (*a*) [/\ odd #|L|, nilpotent H & normedTI H^# G L] -> (*b*) let calS := seqIndD H L H 1 in let tau := 'Ind[G, L] in (*c*) [\/ (*c1*) [Frobenius L = H ><| W1] @@ -602,17 +576,17 @@ Theorem Sibley_coherence (L H W1 : {group gT}) : prime_Dade_hypothesis G L H H H^# A0 defW] -> coherent calS L^# tau. Proof. -set A := H^# => [][oddL nilH tiA] S tau structL. -set case_c1 := [Frobenius L = H ><| W1] in structL. -have sLG: L \subset G by have [_ _ /eqP <-] := and3P tiA; apply: subsetIl. -have [defL ntH ntW1]: [/\ H ><| W1 = L, H :!=: 1 & W1 :!=: 1]%g. - have [/Frobenius_context[]// | [W2 _ [A0 [W [defW []]]]]] := structL. - by move=> _ [[-> -> _ _] [ntW2 /subG1_contra->]]. +set case_c1 := [Frobenius L = H ><| W1]; pose case_c2 := ~~ case_c1. +set A := H^#; set H' := H^`(1)%G => -[oddL nilH tiA] S tau hyp_c. +have sLG: L \subset G by have [] := normedTI_memJ_P tiA. +have ntH: H :!=: 1%g by have [] := normedTI_P tiA; rewrite setD_eq0 subG1. +have [defL ntW1]: H ><| W1 = L /\ W1 :!=: 1%g. + by have [/Frobenius_context[] | [? _ [? [? [? [_ [[]]]]]]]] := hyp_c. have [nsHL _ /mulG_sub[sHL sW1L] _ _] := sdprod_context defL. have [uccS nrS]: cfConjC_subset S S /\ ~~ has cfReal S. by do 2?split; rewrite ?seqInd_uniq ?seqInd_notReal //; apply: cfAut_seqInd. have defZS: 'Z[S, L^#] =i 'Z[S, A] by apply: zcharD1_seqInd. -have c1_irr: case_c1 -> {subset S <= irr L}. +have c1_irrS: case_c1 -> {subset S <= irr L}. move/FrobeniusWker=> frobL _ /seqIndC1P[i nz_i ->]. exact: irr_induced_Frobenius_ker. move defW2: 'C_H(W1)%G => W2; move defW: (W1 <*> W2)%G => W. @@ -620,608 +594,485 @@ have{defW} defW: W1 \x W2 = W. rewrite -defW dprodEY // -defW2 ?subsetIr // setICA setIA. by have [_ _ _ ->] := sdprodP defL; rewrite setI1g. pose V := cyclicTIset defW; pose A0 := A :|: class_support V L. -pose c2hyp := prime_Dade_hypothesis G L H H A A0 defW. +pose ddA0hyp := prime_Dade_hypothesis G L H H A A0 defW. have c1W2: case_c1 -> W2 = 1%G by move/Frobenius_trivg_cent/group_inj <-. -have{structL} c2W2: ~~ case_c1 -> [/\ prime #|W2|, W2 \subset H^`(1)%G & c2hyp]. - case: structL => [-> // | [W20 [prW20 sW20H'] W20hyp] _]. - have{W20hyp} [A00 [W0 [defW0 W20hyp]]] := W20hyp. - suffices /group_inj defW20: W2 :=: W20. - have eqW0: W0 = W by apply: group_inj; rewrite -defW0 -defW20. - rewrite -defW20 eqW0 in prW20 sW20H' defW0 W20hyp; split=> //. - rewrite /c2hyp (eq_irrelevance defW defW0) /A0. - by have [_ _ <-] := prDade_def W20hyp. - have [[_ _ _ cycW1] [_ _ _ prW120] _] := prDade_prTI W20hyp. - have [x defW1] := cyclicP cycW1; rewrite -defW2 /= defW1 cent_cycle prW120 //. - by rewrite !inE defW1 cycle_id -cycle_eq1 -defW1 ntW1. -have{c2W2} [prW2 sW2H' c2W2] := all_and3 c2W2. -have{sW2H'} sW2H': W2 \subset H^`(1)%G. - by have [/c1W2-> | /sW2H'//] := boolP case_c1; apply: sub1G. -pose sigma := cyclicTIiso (c2W2 _). -have [R scohS oRW]: exists2 R, subcoherent S tau R & forall c2 : ~~ case_c1, +have{hyp_c} hyp_c2: case_c2 -> [/\ prime #|W2|, W2 \subset H' & ddA0hyp]. + case: hyp_c => [/idPn// | [W2_ [prW2_ sW2_H'] [A0_ [W_ [defW_ ddA0_]]]] _]. + have idW2_: W2_ = W2. + have [[_ _ _ /cyclicP[x defW1]] [_ _ _ prW12] _] := prDade_prTI ddA0_. + have W1x: x \in W1^# by rewrite !inE -cycle_eq1 -defW1 ntW1 defW1 cycle_id. + by apply/group_inj; rewrite -defW2 /= defW1 cent_cycle prW12. + have idW_: W_ = W by apply/group_inj; rewrite -defW_ idW2_. + rewrite {}/ddA0hyp {}/A0 {}/V; rewrite -idW2_ -idW_ in defW *. + by rewrite (eq_irrelevance defW defW_); have [_ _ <-] := prDade_def ddA0_. +have{hyp_c2} [c2_prW2 c2_sW2H' c2_ddA0] := all_and3 hyp_c2. +have c2_ptiL c2 := prDade_prTI (c2_ddA0 c2). +have{c2_sW2H'} sW2H': W2 \subset H'. + by have [/c1W2-> | /c2_sW2H'//] := boolP case_c1; apply: sub1G. +pose sigma c2 := cyclicTIiso (c2_ddA0 c2). +have [R scohS oRW]: exists2 R, subcoherent S tau R & forall c2 : case_c2, {in [predI S & irr L] & irr W, forall phi w, orthogonal (R phi) (sigma c2 w)}. -- have sAG: A \subset G^# by rewrite setSD // (subset_trans (normal_sub nsHL)). +- have sAG: A \subset G^# by rewrite setSD // (subset_trans sHL). have Itau: {in 'Z[S, L^#], isometry tau, to 'Z[irr G, G^#]}. - split=> [xi1 xi2 | xi]. - rewrite !defZS => /zchar_on Axi1 /zchar_on Axi2. + split=> [xi1 xi2|xi]; first rewrite !defZS => /zchar_on-Axi1 /zchar_on-Axi2. exact: normedTI_isometry Axi1 Axi2. - rewrite !zcharD1E => /andP[Zxi /eqP xi1_0]. - rewrite cfInd1 // xi1_0 mulr0 eqxx cfInd_vchar //. - by apply: zchar_trans_on Zxi; apply: seqInd_vcharW. - have [/= Hc1 | Hc2] := boolP (idfun case_c1). - suffices [R]: {R | subcoherent S tau R} by exists R => // /negP[]. - by apply: irr_subcoherent => //; first by case: uccS (c1_irr Hc1). - have ddA0 := c2W2 Hc2. + rewrite !zcharD1E cfInd1 // => /andP[Zxi /eqP->]; rewrite mulr0. + by rewrite cfInd_vchar ?(zchar_trans_on _ Zxi) //=; apply: seqInd_vcharW. + have [/= c1 | /c2_ddA0-ddA0] := boolP (idfun case_c1). + suffices [R scohS]: {R | subcoherent S tau R} by exists R => // /negP[]. + by apply: irr_subcoherent; first have [[]] := (uccS, c1_irrS c1). + have Dtau: {in 'CF(L, A), tau =1 Dade ddA0}. + have nAL: L \subset 'N(A) by have [_ /subsetIP[]] := normedTI_P tiA. + have sAA0: A \subset A0 by apply: subsetUl. + by move=> phi Aphi /=; rewrite -(restr_DadeE _ sAA0) // [RHS]Dade_Ind. have [R [subcohR oRW _]] := prDade_subcoherent ddA0 uccS nrS. - exists R => [|not_c1 phi w irrSphi irr_w]; last first. + exists R => [|c2 phi w irrSphi irr_w]; last first. by rewrite /sigma -(cycTIiso_irrel ddA0) oRW. - set tau0 := Dade _ in subcohR. - have Dtau: {in 'CF(L, A), tau =1 tau0}. - have nAL: L \subset 'N(A) by rewrite normD1 normal_norm. - move=> phi Aphi; rewrite /tau0 -(restr_DadeE ddA0 (subsetUl _ _) nAL) //. - by rewrite /restr_Dade Dade_Ind. - have [Sok _ oSS Rok oRR] := subcohR; split=> // phi Sphi. - have [ZR oNR <-] := Rok _ Sphi; split=> //. - by rewrite Dtau ?irr_vchar_on ?sub_conjC_vchar ?(seqInd_vchar nsHL Sphi). -have [nsH'H nsH'L] := (der_normal 1 H, char_normal_trans (der_char 1 H) nsHL). -have [nH'L solH] := (normal_norm nsH'L, nilpotent_sol nilH). -have ltH'H: H^`(1)%g \proper H by rewrite ?(nil_comm_properl nilH) ?subsetIidl. + have [Sok _ oSS Rok oRR] := subcohR; split=> {Sok oSS oRR}// phi Sphi. + have [ZR oNR <-] := Rok _ Sphi; split=> {ZR oNR}//. + exact/Dtau/(zchar_on (seqInd_sub_aut_zchar _ _ Sphi)). +have solH := nilpotent_sol nilH; have nsH'H: H' <| H := der_normal 1 H. +have ltH'H: H' \proper H by rewrite (sol_der1_proper solH). +have nsH'L: H' <| L by apply: gFnormal_trans. +have [sH'H [sH'L nH'L]] := (normal_sub nsH'H, andP nsH'L). have coHW1: coprime #|H| #|W1|. - have [/Frobenius_coprime// | /c2W2[_ [[_ _]]]] := boolP case_c1. - by rewrite (coprime_sdprod_Hall_r defL). -have oW1: #|W1| = #|L : H| by rewrite -divgS // -(sdprod_card defL) mulKn. -have frobL1: [Frobenius L / H^`(1) = (H / H^`(1)) ><| (W1 / H^`(1))]%g. + suffices: Hall L W1 by rewrite (coprime_sdprod_Hall_r defL). + by have [/Frobenius_compl_Hall | /c2_ddA0/prDade_prTI[[]]] := boolP case_c1. +have oW1: #|W1| = #|L : H| by rewrite (index_sdprod defL). +have frobL1: [Frobenius L / H' = (H / H') ><| (W1 / H')]. apply: (Frobenius_coprime_quotient defL nsH'L) => //; split=> // x W1x. - have [/Frobenius_reg_ker-> //|] := boolP case_c1; first exact: sub1G. - by case/c2W2=> _ [_ [_ _ _ ->]]. + have [frobL | /c2_ptiL[_ [_ _ _ -> //]]] := boolP case_c1. + by rewrite (Frobenius_reg_ker frobL) ?sub1G. have odd_frobL1: odd_Frobenius_quotient H L 1. - have ? := FrobeniusWker frobL1. - by split=> //=; rewrite ?joingG1 // normal1 sub1G quotient_nil. + split=> //=; last by rewrite joingG1 (FrobeniusWker frobL1). + by rewrite normal1 sub1G quotient_nil. without loss [/p_groupP[p p_pr pH] not_cHH]: / p_group H /\ ~~ abelian H. - have [//| [_] [p []]] := non_coherent_chief nsHL solH scohS odd_frobL1. - rewrite (isog_abelian (quotient1_isog H)) -(isog_pgroup p (quotient1_isog H)). - by move=> /pgroup_p-> -> _; apply. + apply: (non_coherent_chief _ _ scohS odd_frobL1) => // -[_ [p [pH ab'H] _]]. + have isoH := quotient1_isog H; rewrite -(isog_pgroup p isoH) in pH. + by apply; rewrite (isog_abelian isoH) (pgroup_p pH). have sylH: p.-Sylow(G) H. (* required for (6.7) *) - have sylH: p.-Sylow(L) H. - apply/and3P; split=> //; rewrite -oW1 p'natE // -prime_coprime //. - by case/pgroup_pdiv: pH coHW1 => // _ _ [m ->]; rewrite coprime_pexpl. - have [P sylP sHP] := Sylow_superset (subset_trans sHL sLG) pH. - have [sPG pP _] := and3P sylP; have nilP := pgroup_nil pP. - rewrite -(nilpotent_sub_norm nilP sHP) // (sub_normal_Hall sylH) //. - exact: pgroupS (subsetIl P _) pP. - by have [_ _ /eqP <-] := and3P tiA; rewrite normD1 setSI. -pose caseA := 'Z(H) :&: W2 == [1]. -have caseB_P: ~~ caseA -> [/\ ~~ case_c1, W2 :!=: [1] & W2 \subset 'Z(H)]. - rewrite /caseA; have [-> |] := eqsVneq W2 [1]; first by rewrite setIg1 eqxx. - have [/c1W2->/eqP[]// | /prW2 pW2 ->] := boolP case_c1. - by rewrite setIC => /prime_meetG->. -pose Z := if caseA then ('Z(H) :&: H^`(1))%G else W2. -have /subsetIP[sZZ sZH']: Z \subset 'Z(H) :&: H^`(1)%g. - by rewrite /Z; case: ifPn => // /caseB_P[/c2W2[]] *; apply/subsetIP. -have caseB_cZL: ~~ caseA -> Z \subset 'Z(L). - move=> inB; have [_ _ /subsetIP[sW2H cW2H]] := caseB_P inB. - have [_ mulHW1 _ _] := sdprodP defL. - rewrite /Z (negPf inB) subsetI (subset_trans sW2H) //. - by rewrite -mulHW1 centM subsetI cW2H -defW2 subsetIr. -have nsZL: Z <| L. - have [inA | /caseB_cZL/sub_center_normal//] := boolP caseA. - by rewrite /Z inA (char_normal_trans _ nsHL) // charI ?gFchar. + rewrite -Sylow_subnorm -normD1; have [_ _ /eqP->] := and3P tiA. + by apply/and3P; rewrite -oW1 -pgroupE (coprime_p'group _ pH) // coprime_sym. +pose caseA := 'Z(H) :&: W2 \subset [1]%g; pose caseB := ~~ caseA. +have caseB_P: caseB -> [/\ case_c2, W2 :!=: 1%g & W2 \subset 'Z(H)]. + rewrite /caseB /caseA; have [->|] := eqP; first by rewrite subsetIr. + rewrite /case_c2; have [/c1W2->// | /c2_prW2-prW2 _] := boolP case_c1. + by rewrite setIC subG1 => /prime_meetG->. +pose Z := (if caseA then 'Z(H) :&: H' else W2)%G. +have /subsetIP[sZZH sZH']: Z \subset 'Z(H) :&: H'. + by rewrite /Z; case: ifPn => // /caseB_P[_ _ sZZH]; apply/subsetIP. +have caseB_sZZL: caseB -> Z \subset 'Z(L). + move=> in_caseB; have [_ _ /subsetIP[sW2H cW2H]] := caseB_P in_caseB. + rewrite /Z ifN // subsetI (subset_trans sW2H sHL). + by rewrite -(sdprodW defL) centM subsetI cW2H -defW2 subsetIr. +have nsZL: Z <| L; last have [sZL nZL] := andP nsZL. + have [in_caseA | /caseB_sZZL/sub_center_normal//] := boolP caseA. + by rewrite /Z in_caseA normalI ?gFnormal_trans. have ntZ: Z :!=: 1%g. - rewrite /Z; case: ifPn => [_ | /caseB_P[]//]. + rewrite /Z; case: ifPn => [_ | /caseB_P[] //]. by rewrite /= setIC meet_center_nil // (sameP eqP derG1P). -have nsZH := sub_center_normal sZZ; have [sZH nZH] := andP nsZH. +have nsZH: Z <| H := sub_center_normal sZZH; have [sZH nZH] := andP nsZH. have regZL: {in Z^# &, forall x y, #|'C_L[x]| = #|'C_L[y]| }. - have [inA | /caseB_cZL cZL] := boolP caseA; last first. + have [in_caseA | /caseB_sZZL/subsetIP[_ cZL]] := boolP caseA; last first. suffices defC x: x \in Z^# -> 'C_L[x] = L by move=> x y /defC-> /defC->. - by case/setD1P=> _ /(subsetP cZL)/setIP[_]; rewrite -sub_cent1 => /setIidPl. + by case/setD1P=> _ /(subsetP cZL); rewrite -sub_cent1 => /setIidPl. suffices defC x: x \in Z^# -> 'C_L[x] = H by move=> x y /defC-> /defC->. - case/setD1P=> ntx Zx; have /setIP[Hx cHx] := subsetP sZZ x Zx. + case/setD1P=> ntx Zx; have /setIP[Hx cHx] := subsetP sZZH x Zx. have [_ <- _ _] := sdprodP defL; rewrite -group_modl ?sub_cent1 //=. suffices ->: 'C_W1[x] = 1%g by rewrite mulg1. - have [/Frobenius_reg_compl-> // | in_c2] := boolP case_c1; first exact/setD1P. - have [_ [_ [_ _ _ regW1] _] _ _] := c2W2 in_c2. - apply: contraNeq ntx => /trivgPn[y /setIP[W1y cxy] nty]. - rewrite -in_set1 -set1gE -((_ =P [1]) inA) -(regW1 y) 2!inE ?nty //. - by rewrite inE cent1C cHx Hx. -have Zconst_modH := - constant_irr_mod_TI_Sylow sylH oddL tiA (And3 nsZL ntZ sZZ) regZL. -pose X := seqIndD H L Z 1; pose Y := seqIndD H L H H^`(1). + have [/Frobenius_reg_compl-> // | c2] := boolP case_c1; first exact/setD1P. + have [_ [_ _ _ regW1] _] := c2_ptiL c2. + apply: contraTeq in_caseA => /trivgPn[y /setIP[W1y cxy] nty]; apply/subsetPn. + by exists x; rewrite inE // -(regW1 y) 2!inE ?nty // Hx cHx cent1C. +have{regZL} irrZmodH := + constant_irr_mod_TI_Sylow sylH oddL tiA (And3 nsZL ntZ sZZH) regZL. +pose X := seqIndD H L Z 1; pose Y := seqIndD H L H H'. have ccsXS: cfConjC_subset X S by apply: seqInd_conjC_subset1. have ccsYS: cfConjC_subset Y S by apply: seqInd_conjC_subset1. have [[uX sXS ccX] [uY sYS ccY]] := (ccsXS, ccsYS). have X'Y: {subset Y <= [predC X]}. move=> _ /seqIndP[i /setIdP[_ kH'i] ->]; rewrite inE in kH'i. - by rewrite !inE mem_seqInd ?normal1 // !inE sub1G (subset_trans sZH'). + by rewrite !inE mem_seqInd ?normal1 // !inE (subset_trans sZH'). +have oXY: orthogonal X Y. + apply/orthogonalP=> xi eta Xxi Yeta; apply: orthoPr xi Xxi. + exact: (subset_ortho_subcoherent scohS sXS (sYS _ Yeta) (X'Y _ Yeta)). have irrY: {subset Y <= irr L}. move=> _ /seqIndP[i /setIdP[not_kHi kH'i] ->]; rewrite !inE in not_kHi kH'i. - have kH'iInd: H^`(1)%g \subset cfker ('Ind[L] 'chi_i). - by rewrite sub_cfker_Ind_irr ?normal_norm. - rewrite -(cfQuoK nsH'L kH'iInd) -cfIndQuo // -quo_IirrE //. - set i1 := quo_Iirr _ i; have /irrP[k ->]: 'Ind 'chi_i1 \in irr (L / H^`(1)). - apply: irr_induced_Frobenius_ker; first exact: FrobeniusWker frobL1. - apply: contraNneq not_kHi; rewrite -(quo_IirrK nsH'H kH'i) -/i1 => ->. - by rewrite mod_IirrE // irr0 cfMod_cfun1 ?cfker_cfun1. - by rewrite -mod_IirrE ?mem_irr. + rewrite -(cfQuo_irr nsH'L) ?sub_cfker_Ind_irr -?cfIndQuo -?quo_IirrE //. + apply: (irr_induced_Frobenius_ker (FrobeniusWker frobL1)). + by rewrite quo_Iirr_eq0 -?subGcfker. +have oY: orthonormal Y by apply: sub_orthonormal (irr_orthonormal L). have uniY: {in Y, forall phi : 'CF(L), phi 1%g = #|W1|%:R}. move=> _ /seqIndP[i /setIdP[_ kH'i] ->]; rewrite inE -lin_irr_der1 in kH'i. - rewrite cfInd1 // -divgS // -(sdprod_card defL) mulKn //. - by case/andP: kH'i => _ /eqP->; rewrite mulr1. + by rewrite cfInd1 // -divgS // -(sdprod_card defL) mulKn // lin_char1 ?mulr1. have scohY: subcoherent Y tau R by apply: (subset_subcoherent scohS). have [tau1 cohY]: coherent Y L^# tau. apply/(uniform_degree_coherence scohY)/(@all_pred1_constant _ #|W1|%:R). by apply/allP=> _ /mapP[phi Yphi ->]; rewrite /= uniY. have [[Itau1 Ztau1] Dtau1] := cohY. -have [eta1 Yeta1]: exists eta1, eta1 \in Y. - pose IY := Iirr_kerD H H H^`(1)%G. - have [IY0 | [j IYj]] := set_0Vmem IY; last first. - by exists ('Ind 'chi_j); apply/seqIndP; exists j. - have /idPn[]: \sum_(j in IY) ('chi_j 1%g) ^+ 2 == 0 by rewrite IY0 big_set0. - rewrite sum_Iirr_kerD_square ?der_sub // indexgg mul1r subr_eq0. - by rewrite pnatr_eq1 indexg_eq1 proper_subn. -have caseA_coh12: caseA -> coherent (X ++ Y) L^# tau. - move=> haveA. +have oYtau: orthonormal (map tau1 Y) by apply: map_orthonormal. +have [[_ oYY] [_ oYYt]] := (orthonormalP oY, orthonormalP oYtau). +have [eta1 Yeta1]: {eta1 | eta1 \in Y} by apply: seqIndD_nonempty. +pose m : algC := (size Y)%:R; pose m_ub2 a := (a - 1) ^+ 2 + (m - 1) * a ^+ 2. +have m_ub2_lt2 a: a \in Cint -> m_ub2 a < 2%:R -> a = 0 \/ a = 1 /\ size Y = 2. + move=> Za ub_a; have [|nza] := eqVneq a 0; [by left | right]. + have ntY: (1 < size Y)%N by apply: seqInd_nontrivial Yeta1. + have m1_ge1: 1 <= m - 1 by rewrite ler_subr_addr (ler_nat _ 2). + have a1: a = 1. + apply: contraFeq (ltr_geF ub_a); rewrite -subr_eq0 /m_ub2 => nz_a1. + by rewrite ler_add ?(mulr_ege1 m1_ge1) // sqr_Cint_ge1 ?rpredB. + rewrite /m_ub2 a1 subrr expr0n add0r expr1n mulr1 in ub_a. + rewrite ltr_subl_addr -mulrSr ltr_nat ltnS in ub_a. + by split; last apply/anti_leq/andP. +have{odd_frobL1} caseA_cohXY: caseA -> coherent (X ++ Y) L^# tau. + move=> in_caseA. have scohX: subcoherent X tau R by apply: subset_subcoherent ccsXS. have irrX: {subset X <= irr L}. - have [/c1_irr irrS | in_c2] := boolP case_c1. - move=> phi Xphi; apply: irrS; apply: seqIndS phi Xphi. - by rewrite Iirr_kerDS // (subset_trans sZH') ?der_sub. - move/(_ in_c2) in prW2; have [_ ptiL _ _] := c2W2 in_c2. - have [[_ _ _ cycW1] [ntW2 sW2H cycW2 prW1H] oddW] := ptiL. - have nZL := normal_norm nsZL; have nZW1 := subset_trans sW1L nZL. - have isoW2: (W2 / Z)%g \isog W2. + have [c1 | c2] := boolP case_c1; first by move=> phi /sXS/c1_irrS->. + have ptiL := c2_ptiL c2; have [_ [ntW2 sW2H _ _] _] := ptiL. + have{sW2H} isoW2: W2 / Z \isog W2. apply/isog_symr/quotient_isog; first exact: subset_trans sW2H nZH. - by rewrite -(setIidPr sZZ) setIAC ((_ =P [1]) haveA) setI1g. - have [|defWb ptiLZ] := primeTIhyp_quotient ptiL _ sZH nsZL. - by rewrite (isog_eq1 isoW2). - pose Ichi := primeTI_Ires ptiL; pose IchiZ := primeTI_Ires ptiLZ. - have eq_Ichi: codom (fun j1 => mod_Iirr (IchiZ j1)) =i codom Ichi. - apply/subset_cardP. - rewrite !card_codom; last first; try exact: prTIres_inj. - apply: inj_comp (prTIres_inj ptiLZ). - exact: can_inj (mod_IirrK (sub_center_normal sZZ)). - by rewrite !card_ord !NirrE (nclasses_isog isoW2). - apply/subsetP=> _ /codomP[/= j1 ->]. - have [[j2 /irr_inj->] | ] := prTIres_irr_cases ptiL (mod_Iirr (IchiZ j1)). - exact: codom_f. - case=> /idPn[]; rewrite mod_IirrE // cfIndMod // cfInd_prTIres. - apply: contra (prTIred_not_irr ptiLZ j1) => /irrP[ell Dell]. - by rewrite -[_ j1]cfModK // Dell -quo_IirrE ?mem_irr // -Dell cfker_mod. - move=> _ /seqIndP[k /setDP[_ kZ'k] ->]. - have [[j /irr_inj Dk] | [] //] := prTIres_irr_cases ptiL k. - case/negP: kZ'k; have: k \in codom Ichi by rewrite Dk codom_f. - by rewrite -eq_Ichi => /codomP[j1 ->]; rewrite !inE mod_IirrE ?cfker_mod. - have [//|] := seqIndD_irr_coherence nsHL solH scohS odd_frobL1 _ irrX. - rewrite -/X => defX [tau2 cohX]; have [[Itau2 Ztau2] Dtau2] := cohX. - have [xi1 Xxi1 Nd]: - exists2 xi1, xi1 \in X & forall xi, xi \in X -> (xi1 1%g %| xi 1%g)%C. - - pose IX := Iirr_kerD H Z 1%G; have [i0 IXi0]: exists i0, i0 \in IX. - apply/set0Pn; apply: contraNneq ntZ => IX_0. - have: \sum_(i in IX) ('chi_i 1%g) ^+ 2 == 0 by rewrite IX_0 big_set0. - rewrite sum_Iirr_kerD_square ?normal1 ?sub1G // indexg1 mulf_eq0. - by rewrite (negPf (neq0CiG H Z)) subr_eq0 trivg_card1 -eqC_nat. - have:= erefl [arg min_(i < i0 in IX) truncC ('chi_i 1%g)]. - have [//|{i0 IXi0} i1 IXi1 min_i1 _] := arg_minP. - exists ('Ind 'chi_i1); first by apply/seqIndP; exists i1. - move=> _ /seqIndP[i /min_i1 le_i1_i] ->; rewrite !cfInd1 //. - have pHP := p_natP (pnat_dvd _ pH). - move: (dvd_irr1_cardG i1) (dvd_irr1_cardG i) le_i1_i. - rewrite !irr1_degree -!natrM !dvdC_nat => /pHP[m1 ->] /pHP[m ->]. - rewrite !natCK leq_exp2l ?prime_gt1 // => /subnKC <-. - by rewrite expnD mulnA dvdn_mulr. + exact/trivgP/(subset_trans _ in_caseA)/setSI. + have{ntW2} ntW2bar: (W2 / Z != 1)%g by rewrite (isog_eq1 isoW2). + have{ntW2bar} [defWbar ptiLZ] := primeTIhyp_quotient ptiL ntW2bar sZH nsZL. + pose IchiZ := [set mod_Iirr (primeTI_Ires ptiLZ j) | j : Iirr (W2 / Z)]. + suffices /eqP-eq_Ichi: IchiZ == [set primeTI_Ires ptiL j | j : Iirr W2]. + move=> _ /seqIndP[k /setDP[_ kZ'k] ->]. + have [[j /irr_inj-Dk] | [] //] := prTIres_irr_cases ptiL k. + have{j Dk} /imsetP[j _ Dk]: k \in IchiZ by rewrite eq_Ichi Dk mem_imset. + by rewrite !inE Dk mod_IirrE ?cfker_mod in kZ'k. + rewrite eqEcard !card_imset; last exact: prTIres_inj; first last. + exact: inj_comp (morph_Iirr_inj _) (prTIres_inj _). + apply/andP; split; last by rewrite !card_ord !NirrE (nclasses_isog isoW2). + apply/subsetP=> k /imsetP[j _ Dk]. + have [[j1 /irr_inj->]|] := prTIres_irr_cases ptiL k; first exact: mem_imset. + case=> /idPn[]; rewrite {k}Dk mod_IirrE ?cfIndMod ?cfMod_irr //. + by rewrite cfInd_prTIres prTIred_not_irr. + have [//|defX [tau2 cohX]]: X =i _ /\ coherent X L^# tau := + seqIndD_irr_coherence nsHL solH scohS odd_frobL1 _ irrX. + have [[Itau2 Ztau2] Dtau2] := cohX. + pose dvd_degrees_X (d : algC) := {in X, forall xi : 'CF(L), d %| xi 1%g}%C. + have [xi1 Xxi1 dvd_xi1_1]: exists2 xi1, xi1 \in X & dvd_degrees_X (xi1 1%g). + have /all_sig[e De] i: {e | 'chi[H]_i 1%g = (p ^ e)%:R}. + have:= dvd_irr1_cardG i; rewrite irr1_degree dvdC_nat => dv_chi1_H. + by have /p_natP[e ->] := pnat_dvd dv_chi1_H pH; exists e. + have [_ /seqIndP[i0 IXi0 _]]: {phi | phi \in X}. + by apply: seqIndD_nonempty; rewrite ?normal1 ?proper1G. + pose xi1 := 'Ind[L] 'chi_[arg min_(i < i0 in Iirr_kerD H Z 1%G) e i]. + case: arg_minP => {i0 IXi0}//= i1 IXi1 min_i1 in xi1. + exists xi1 => [|_ /seqIndP[i IXi ->]]; first by apply/seqIndP; exists i1. + rewrite !cfInd1 // !De -!natrM dvdC_nat dvdn_pmul2l //. + by rewrite dvdn_Pexp2l ?min_i1 ?prime_gt1. + have nz_xi1_1: xi1 1%g != 0 by apply: seqInd1_neq0 Xxi1. pose d (xi : 'CF(L)) : algC := (truncC (xi 1%g / xi1 1%g))%:R. - have{Nd} def_d xi: xi \in X -> xi 1%g = d xi * xi1 1%g. - rewrite /d => Xxi; move: Xxi (Nd _ Xxi) => /irrX/irrP[i ->]. - have /irrX/irrP[i1 ->] := Xxi1. - rewrite !irr1_degree dvdC_nat => /dvdnP[q ->]. - by rewrite natrM -irr1_degree mulfK ?irr1_neq0 // natCK. - have d_xi1: d xi1 = 1. - by apply: (mulIf (seqInd1_neq0 nsHL Xxi1)); rewrite mul1r -def_d. - have oXY: orthogonal X Y. - apply/orthogonalP=> xi eta Xxi Yeta; apply: orthoPr xi Xxi. - exact: (subset_ortho_subcoherent scohS sXS (sYS _ Yeta) (X'Y _ Yeta)). - have [_ [Itau Ztau] /orthogonal_free freeS _ _] := scohS. + have{dvd_xi1_1} def_d xi: xi \in X -> xi 1%g = d xi * xi1 1%g. + rewrite /d => Xxi; have Xge0 := Cnat_ge0 (Cnat_seqInd1 (_ : _ \in X)). + by have /dvdCP_nat[||q ->] := dvd_xi1_1 xi Xxi; rewrite ?Xge0 ?mulfK ?natCK. + have d_xi1: d xi1 = 1 by rewrite /d divff ?truncC1. + have [_ [Itau /(_ _ _)/zcharW-Ztau] _ _ _] := scohS. have o_tauXY: orthogonal (map tau2 X) (map tau1 Y). exact: (coherent_ortho scohS). - have [a Na Dxi11]: exists2 a, a \in Cnat & xi1 1%g = a * #|W1|%:R. - have [i1 _ ->] := seqIndP Xxi1. - exists ('chi_i1 1%g); first exact: Cnat_irr1. - by rewrite cfInd1 // -divgS // -(sdprod_card defL) ?mulKn // mulrC. - pose psi1 := xi1 - a *: eta1; have Za: a \in Cint by rewrite CintE Na. + have [a Na xi1_1]: exists2 a, a \in Cnat & xi1 1%g = a * #|W1|%:R. + have [i _ ->] := seqIndP Xxi1; rewrite cfInd1 // -oW1 mulrC. + by exists ('chi_i 1%g); first apply: Cnat_irr1. + pose psi1 := xi1 - a *: eta1. have Zpsi1: psi1 \in 'Z[S, L^#]. - rewrite zcharD1E !cfunE (uniY _ Yeta1) -Dxi11 subrr eqxx. - by rewrite rpredB ?scale_zchar ?mem_zchar ?(sXS _ Xxi1) // sYS. + rewrite zcharD1E !cfunE (uniY _ Yeta1) -xi1_1 subrr eqxx andbT. + by rewrite rpredB ?rpredZ_Cnat ?mem_zchar ?(sXS _ Xxi1) // sYS. have [Y1 dY1 [X1 [dX1 _ oX1tauY]]] := orthogonal_split (map tau1 Y)(tau psi1). - have oY: orthonormal Y by apply: sub_orthonormal (irr_orthonormal L). - have oYtau: orthonormal (map tau1 Y) by apply: map_orthonormal. - have{dX1 Y1 dY1} [b Zb Dpsi1]: {b | b \in Cint & + have{dX1 Y1 dY1 oYtau} [b Zb tau_psi1]: {b | b \in Cint & tau psi1 = X1 - a *: tau1 eta1 + b *: (\sum_(eta <- Y) tau1 eta)}. - exists ('[tau psi1, tau1 eta1] + a). - rewrite rpredD // Cint_cfdot_vchar ?Ztau1 ?seqInd_zcharW //. - exact: zcharW (Ztau _ Zpsi1). - rewrite {1}dX1 addrC -addrA; congr (_ + _). - have [_ -> ->] := orthonormal_span oYtau dY1. - rewrite -[Y1](addrK X1) -dX1 big_map !(big_rem eta1 Yeta1) /=. - rewrite cfdotBl (orthoPl oX1tauY) ?map_f // subr0. - rewrite scalerDr addrA; congr (_ + _). - by rewrite addrC -scaleNr -scalerDl addrK. - rewrite raddf_sum; apply: eq_big_seq => eta. - rewrite mem_rem_uniq ?seqInd_uniq // => /andP[eta1'eta /= Yeta]. - congr (_ *: _); rewrite cfdotBl (orthoPl oX1tauY) ?map_f // subr0 addrC. - apply: canRL (subrK _) _; rewrite -2!raddfB /=. - have Zeta: (eta - eta1) \in 'Z[Y, L^#]. + by rewrite rpredD ?Cint_cfdot_vchar ?Cint_Cnat ?Ztau ?Ztau1 ?mem_zchar. + rewrite [LHS]dX1 addrC -addrA; congr (_ + _). + have{dY1} [_ -> ->] := orthonormal_span oYtau dY1. + transitivity (\sum_(xi <- map tau1 Y) '[tau psi1, xi] *: xi). + by apply/eq_big_seq=> xi ?; rewrite dX1 cfdotDl (orthoPl oX1tauY) ?addr0. + rewrite big_map scaler_sumr !(big_rem eta1 Yeta1) /= addrCA addrA scalerDl. + rewrite addrK; congr (_ + _); apply: eq_big_seq => eta. + rewrite mem_rem_uniq // => /andP[eta1'eta /= Yeta]; congr (_ *: _). + apply/(canRL (addNKr _)); rewrite addrC -2!raddfB /=. + have Zeta: eta - eta1 \in 'Z[Y, L^#]. by rewrite zcharD1E rpredB ?seqInd_zcharW //= !cfunE !uniY ?subrr. - rewrite Dtau1 // Itau // ?(zchar_subset sYS) //. - rewrite cfdotBl cfdotZl !cfdotBr 2?(orthogonalP oXY) // subr0 add0r. - have [_ oYY] := orthonormalP oY; rewrite !oYY // eqxx. - by rewrite eq_sym (negPf eta1'eta) add0r mulrN mulr1 opprK. - pose psi := 'Res[L] (tau1 eta1). - have [X2 dX2 [xi' [dxi' _ oxi'X]]] := orthogonal_split X psi. + rewrite Dtau1 // Itau // ?(zchar_subset sYS) // cfdotBl cfdotZl. + rewrite (span_orthogonal oXY) ?rpredB ?memv_span // add0r cfdotBr. + by rewrite !oYY // !mulrb eqxx ifN_eqC // sub0r mulrN1 opprK. have oX: orthonormal X by apply: sub_orthonormal (irr_orthonormal L). - have Zpsi: psi \in 'Z[irr L] by rewrite cfRes_vchar ?Ztau1 ?seqInd_zcharW. - pose sumXd := \sum_(xi <- X) d xi *: xi. + have [_ oXX] := orthonormalP oX. have Zxi1Xd xi: xi \in X -> xi - d xi *: xi1 \in 'Z[X, L^#]. move=> Xxi; rewrite zcharD1E !cfunE -def_d // subrr eqxx. - by rewrite rpredB ?scale_zchar ?seqInd_zcharW ?rpred_nat. - have{dxi' X2 dX2} [c Zc Dpsi]: {c | c \in Cint & psi = c *: sumXd + xi'}. - exists '[psi, xi1]; first by rewrite Cint_cfdot_vchar ?(seqInd_vcharW Xxi1). - rewrite {1}dxi'; congr (_ + _); have [_ -> ->] := orthonormal_span oX dX2. - rewrite -[X2](addrK xi') -dxi' raddf_sum; apply: eq_big_seq => /= xi Xxi. - rewrite cfdotBl (orthoPl oxi'X) // subr0 scalerA; congr (_ *: _). - apply/eqP; rewrite -subr_eq0 mulrC -[d xi]conj_Cnat ?Cnat_nat //. - rewrite -cfdotZr -cfdotBr cfdot_Res_l -Dtau2 ?Zxi1Xd //. - rewrite cfdotC raddfB raddfZ_Cnat ?Cnat_nat // cfdotBl cfdotZl. - by rewrite !(orthogonalP o_tauXY) ?map_f // mulr0 subr0 conjC0. - have Exi' z: z \in Z -> xi' z = xi' 1%g. - move=> Zz; rewrite [xi']cfun_sum_cfdot !sum_cfunE; apply: eq_bigr => ell _. - have [Xell |] := boolP ('chi_ell \in X). - by rewrite !cfunE (orthoPl oxi'X) ?mul0r. - by rewrite !cfunE defX inE /= mem_irr negbK => /subsetP/(_ z Zz)/cfker1->. + by rewrite rpredB ?rpredZnat ?mem_zchar. + pose psi := 'Res[L] (tau1 eta1); move Dc: '[psi, xi1] => c. + have Zpsi: psi \in 'Z[irr L] by rewrite cfRes_vchar ?Ztau1 ?seqInd_zcharW. + pose sumXd : 'CF(L) := \sum_(xi <- X) d xi *: xi. + have{Dc} [xi2 Dpsi oxi2X]: {xi2 | psi = c *: sumXd + xi2 & orthogonal xi2 X}. + exists (psi - c *: sumXd); first by rewrite addrC subrK. + apply/orthoPl=> xi Xxi; rewrite cfdotBl cfdotZl cfproj_sum_orthonormal //. + rewrite mulrC -[d xi]conjCK -Dc -cfdotZr -cfdotBr cfdot_Res_l -conjC0. + rewrite -/tau rmorph_nat -Dtau2 ?Zxi1Xd // raddfB raddfZnat -/(d xi) cfdotC. + by rewrite (span_orthogonal o_tauXY) ?rpredB ?rpredZ ?memv_span ?map_f. + have Exi2 z: z \in Z -> xi2 z = xi2 1%g. + rewrite [xi2]cfun_sum_constt => Zz; apply/cfker1; apply: subsetP z Zz. + apply: subset_trans (cfker_sum _ _ _); rewrite subsetI sZL. + apply/bigcapsP=> i; rewrite inE => xi2_i; rewrite cfker_scale_nz //. + by apply: contraR xi2_i => X_i; rewrite (orthoPl oxi2X) // defX inE mem_irr. have Eba: '[psi, psi1] = b - a. - rewrite cfdot_Res_l -/tau Dpsi1 -addrA !cfdotDr cfdotNr cfdotZr. - rewrite cfdotC (orthoPl oX1tauY) ?map_f // conjC0 add0r addrC. - rewrite cfdotC raddf_sum cfproj_sum_orthonormal // !aut_Cint //. - by have [_ ->] := orthonormalP oYtau; rewrite ?map_f // eqxx mulr1. - have nz_xi11: xi1 1%g != 0 by have /irrX/irrP[i ->] := Xxi1; apply: irr1_neq0. - have {Eba} Ebc: (a %| b - c)%C. + rewrite cfdotC cfdot_Res_r -/tau tau_psi1 cfdotDl cfdotBl cfdotZl. + rewrite (orthoPl oX1tauY) 1?oYYt ?map_f // eqxx sub0r addrC mulr1 rmorphB. + by rewrite scaler_sumr cfproj_sum_orthonormal // aut_Cint // aut_Cnat. + have{Eba oxi2X} Ebc: (a %| b - c)%C. rewrite -[b](subrK a) -Eba cfdotBr {1}Dpsi cfdotDl cfdotZl. - rewrite cfproj_sum_orthonormal // (orthoPl oxi'X) // addr0 d_xi1 mulr1. - rewrite addrC -addrA addKr addrC rpredB ?dvdC_refl //= cfdotZr aut_Cint //. + rewrite cfproj_sum_orthonormal // (orthoPl oxi2X) // addr0 d_xi1 mulr1. + rewrite addrC -addrA addKr addrC rpredB ?dvdC_refl //= cfdotZr aut_Cnat //. by rewrite dvdC_mulr // Cint_cfdot_vchar ?(seqInd_vcharW Yeta1). - have DsumXd: sumXd = (xi1 1%g)^-1 *: (cfReg L - (cfReg (L / Z)%g %% Z)%CF). - apply: canRL (scalerK nz_xi11) _; rewrite !cfReg_sum 2!linear_sum /=. - pose F (xi : 'CF(L)) := xi 1%g *: xi; transitivity (\sum_(xi <- X) F xi). - by apply: eq_big_seq => xi Xxi; rewrite scalerA mulrC -def_d. - rewrite (bigID (mem (Iirr_ker L Z))) /=; apply: canRL (addrK _) _. - rewrite addrC; congr (_ + _). - rewrite (eq_bigl _ _ (in_set _)) (reindex _ (mod_Iirr_bij nsZL)) /=. + have DsumXd: sumXd = (xi1 1%g)^-1 *: (cfReg L - cfReg (L / Z) %% Z)%CF. + apply/(canRL (scalerK nz_xi1_1))/(canRL (addrK _)); rewrite !cfReg_sum. + pose kerZ := [pred i : Iirr L | Z \subset cfker 'chi_i]. + rewrite 2!linear_sum (bigID kerZ) (reindex _ (mod_Iirr_bij nsZL)) /= addrC. + congr (_ + _). apply: eq_big => [i | i _]; first by rewrite mod_IirrE ?cfker_mod. by rewrite linearZ mod_IirrE // cfMod1. - transitivity (\sum_(xi <- [seq 'chi_i | i in [predC Iirr_ker L Z]]) F xi). - apply: eq_big_perm; apply: uniq_perm_eq => // [|xi]. - by rewrite (map_inj_uniq irr_inj) ?enum_uniq. - rewrite defX inE /=; apply/andP/imageP=> [[/irrP[i ->] kerZ'i] | [i]]. - by exists i; rewrite ?inE. - by rewrite !inE => ? ->; rewrite mem_irr. - by rewrite big_map big_filter; apply: eq_bigl => i; rewrite !inE. - have eta1tauZ z: z \in Z^# -> tau1 eta1 z - tau1 eta1 1%g = - c * #|H|%:R / a. - case/setD1P=> ntz Zz; transitivity (psi z - psi 1%g). - by rewrite !cfResE ?(subsetP (normal_sub nsZL)). - rewrite Dpsi DsumXd !cfunE Exi' ?cfuniE ?normal1 // set11 inE (negPf ntz). - rewrite mulr0 mulr1 sub0r Dxi11 cfker1 ?cfker_reg_quo //. - set cc := c * _ + _; rewrite 2!mulrDr -[rhs in _ - rhs]addrA -/cc. - rewrite addrC opprD {cc}subrK -(sdprod_card defL) mulnC natrM. - by rewrite invfM !mulrA divfK ?neq0CG // mulrAC -2!mulNr. - have{eta1tauZ} dvHpsi: (#|H| %| - c * #|H|%:R / a)%C. - have /dirrP[e [i Deta1]]: tau1 eta1 \in dirr G. - rewrite dirrE Ztau1 ?Itau1 ?seqInd_zcharW //=. - by have [_ ->] := orthonormalP oY; rewrite ?eqxx. - have [z ntz Zz] := trivgPn _ ntZ; have Z1z: z \in Z^# by apply/setD1P. - have /(Zconst_modH i)[|_] := Z1z. - move=> z1 z2 Zz1 Zz2; rewrite -(canLR (signrZK e) Deta1) !cfunE. - by apply/eqP; do 2!rewrite eq_sym (canRL (subrK _) (eta1tauZ _ _)) //. - by rewrite -(canLR (signrZK e) Deta1) !cfunE -mulrBr eta1tauZ // rpredMsign. - have nz_a: a != 0 by rewrite Dxi11 mulf_eq0 negb_or neq0CG andbT in nz_xi11. - have{dvHpsi} dv_ac: (a %| c)%C. - move: dvHpsi; rewrite !mulNr mulrAC rpredN => /dvdCP[q Zq]. - by move/(mulIf (neq0CG H))/(canRL (divfK nz_a))->; apply: dvdC_mull. - have{Ebc dv_ac} /dvdCP[q Zq Db]: (a %| b)%C by rewrite -[b](subrK c) rpredD. - pose m : algC := (size Y)%:R. - have Da1: 1 + a ^+ 2 = '[X1] + a ^+ 2 * ((q - 1) ^+ 2 + (m - 1) * q ^+ 2). - transitivity '[psi1]. - rewrite cfnormBd; last by rewrite cfdotZr (orthogonalP oXY) ?mulr0. - rewrite cfnormZ Cint_normK //. - have [[_ -> //] [_ -> //]]:= (orthonormalP oX, orthonormalP oY). - by rewrite !eqxx mulr1. - rewrite -Itau // Dpsi1 -addrA cfnormDd; last first. - rewrite addrC cfdotBr !cfdotZr cfdot_sumr (orthoPl oX1tauY) ?map_f //. - rewrite big_seq big1 ?mulr0 ?subrr // => eta Yeta. - by rewrite (orthoPl oX1tauY) ?map_f //. - congr (_ + _); rewrite cfnormD cfnormN !cfnormZ. - have [_ ->] := orthonormalP oYtau; rewrite ?map_f // eqxx mulr1. - rewrite cfnorm_map_orthonormal // -/m !Cint_normK // cfdotNl cfdotZl. - rewrite linear_sum cfdotC cfproj_sum_orthonormal // rmorphN rmorphM. - rewrite conjCK !aut_Cint // -mulr2n mulNrn -[_ - _]addrAC. - rewrite mulrDr -{1}[m](addNKr 1) mulrDr mulr1 addrA -sqrrB. - congr (_ + _); last by rewrite mulrCA -exprMn (mulrC a) addrC -Db mulrC. - by rewrite -exprMn -sqrrN opprB mulrBr mulr1 (mulrC a) -Db. - have{Da1} maxq: ~~ (2%:R <= (q - 1) ^+ 2 + (m - 1) * q ^+ 2). - have a2_gt1: a ^+ 2 > 1. - have /seqIndP[i1 /setDP[_ not_kerH'i1] Dxi1] := Xxi1. - apply: contraR not_kerH'i1; rewrite inE expr_gt1 ?Cnat_ge0 //. - have [n Da] := CnatP a Na; rewrite Da ltr1n -leqNgt leq_eqVlt. - rewrite ltnNge lt0n -!eqC_nat -{n}Da nz_a orbF => /eqP a_eq1. - rewrite (subset_trans sZH') // -lin_irr_der1 qualifE irr_char. - rewrite -(inj_eq (mulfI (neq0CiG L H))) -cfInd1 // -Dxi1 Dxi11 a_eq1. - by rewrite mul1r mulr1 -divgS //= -(sdprod_card defL) mulKn. - rewrite -(ler_pmul2l (ltr_trans ltr01 a2_gt1)) ltr_geF // mulr_natr. - apply: ler_lt_trans (_ : 1 + a ^+ 2 < _); last by rewrite ltr_add2r. - by rewrite Da1 -subr_ge0 addrK cfnorm_ge0. - clear psi Dpsi Zpsi Zb c sumXd DsumXd Zc xi' Exi' oxi'X. - wlog{Dpsi1 Itau1 Ztau1 Dtau1 oYtau b q maxq Db Zq} Dpsi1: + transitivity (\sum_(xi <- X) xi 1%g *: xi). + by apply: eq_big_seq => xi Xxi; rewrite scalerA mulrC -def_d. + rewrite (eq_big_perm [seq 'chi_i | i in [predC kerZ]]). + by rewrite big_map big_filter. + apply: uniq_perm_eq => // [|xi]. + by rewrite (map_inj_uniq irr_inj) ?enum_uniq. + rewrite defX; apply/andP/imageP=> [[/irrP[i ->]] | [i]]; first by exists i. + by move=> kerZ'i ->; rewrite mem_irr. + have nz_a: a != 0 by have:= nz_xi1_1; rewrite xi1_1 mulf_eq0 => /norP[]. + have{psi Dpsi Zpsi xi2 Exi2 sumXd DsumXd} tau_eta1_Z z: + z \in Z^# -> tau1 eta1 z - tau1 eta1 1%g = - c * #|H|%:R / a. + - case/setD1P=> /negPf-ntz Zz; have Lz := subsetP sZL z Zz. + transitivity (psi z - psi 1%g); first by rewrite !cfResE. + rewrite Dpsi DsumXd !(cfRegE, cfunE) eqxx -opprB 2!mulrDr -[_ + xi2 _]addrA. + rewrite Exi2 ?cfModE ?morph1 ?coset_id // ntz add0r addrK -mulNr mulrAC. + by rewrite xi1_1 invfM -(sdprod_card defL) mulnC natrM !mulrA divfK ?neq0CG. + have{tau_eta1_Z} dvH_cHa: (#|H| %| c * #|H|%:R / a)%C. + have /dirrP[e [i /(canLR (signrZK e))Deta1]]: tau1 eta1 \in dirr G. + by rewrite dirrE Ztau1 ?seqInd_zcharW //= oYYt ?map_f ?eqxx. + have /set0Pn[z Zz]: Z^# != set0 by rewrite setD_eq0 subG1. + have [z1 z2 Zz1 Zz2|_] := irrZmodH i _ z Zz. + rewrite -Deta1 !cfunE; congr (_ * _); apply/(addIr (- tau1 eta1 1%g)). + by rewrite !tau_eta1_Z. + by rewrite -Deta1 !cfunE -mulrBr rpredMsign ?tau_eta1_Z ?mulNr ?rpredN. + have{dvH_cHa} dv_ac: (a %| c)%C. + by rewrite -(@dvdC_mul2r _ a) ?divfK // mulrC dvdC_mul2l ?neq0CG in dvH_cHa. + have{c Ebc dv_ac} /dvdCP[q Zq Db]: (a %| b)%C by rewrite rpredBr in Ebc. + have norm_psi1: '[psi1] = 1 + a ^+ 2. + rewrite cfnormBd; last by rewrite cfdotZr (orthogonalP oXY) ?mulr0. + by rewrite cfnormZ norm_Cnat // oXX // oYY // !eqxx mulr1. + have{Zb oYYt} norm_tau_psi1: '[tau psi1] = '[X1] + a ^+ 2 * m_ub2 q. + rewrite tau_psi1 -addrA cfnormDd /m_ub2; last first. + rewrite addrC big_seq (span_orthogonal oX1tauY) ?memv_span1 //. + by rewrite rpredB ?rpredZ ?rpred_sum // => *; rewrite memv_span ?map_f. + congr (_ + _); transitivity (b ^+ 2 * m + a ^+ 2 - a * b *+ 2); last first. + rewrite [RHS]mulrC [in RHS]addrC mulrBl sqrrB1 !addrA mulrDl !mul1r subrK. + by rewrite mulrBl [m * _]mulrC mulrnAl mulrAC Db exprMn (mulrCA a) addrAC. + rewrite addrC cfnormB !cfnormZ Cint_normK ?norm_Cnat // cfdotZr. + rewrite cfnorm_map_orthonormal // -/m linear_sum cfproj_sum_orthonormal //. + by rewrite oYYt ?map_f // eqxx mulr1 rmorphM conjCK aut_Cnat ?aut_Cint. + have{norm_tau_psi1} mq2_lt2: m_ub2 q < 2%:R. + suffices a2_gt1: a ^+ 2 > 1. + have /ltr_pmul2l <-: a ^+ 2 > 0 by apply: ltr_trans a2_gt1. + rewrite -(ltr_add2l '[X1]) -norm_tau_psi1 ltr_paddl ?cfnorm_ge0 //. + by rewrite Itau // mulr_natr norm_psi1 ltr_add2r. + suffices a_neq1: a != 1. + rewrite expr_gt1 ?Cnat_ge0 // ltr_neqAle eq_sym a_neq1. + by rewrite -(norm_Cnat Na) norm_Cint_ge1 ?Cint_Cnat. + have /seqIndP[i1 /setDP[_ not_kerH'i1] Dxi1] := Xxi1. + apply: contraNneq not_kerH'i1 => a_eq1; rewrite inE (subset_trans sZH') //. + rewrite -lin_irr_der1 qualifE irr_char /= -(inj_eq (mulfI (neq0CiG L H))). + by rewrite -cfInd1 // -Dxi1 xi1_1 a_eq1 mul1r mulr1 oW1. + without loss{tau_psi1 Itau1 Ztau1 Dtau1 b q Db mq2_lt2 Zq} tau_psi1: tau1 cohY o_tauXY oX1tauY / tau psi1 = X1 - a *: tau1 eta1. - - move=> IH; have [q0 | nz_q] := eqVneq q 0. - by apply: (IH tau1) => //; rewrite Dpsi1 Db q0 mul0r scale0r addr0. - have m1_ge1: 1 <= m - 1. - rewrite -(@ler_add2r _ 1) subrK (ler_nat _ 2). - exact: seqInd_nontrivial (irrY _ Yeta1) (Yeta1). - have q1: q = 1. - apply: contraNeq maxq; rewrite -subr_eq0 => nz_q1. - rewrite ler_add // ?sqr_Cint_ge1 ?rpredB //. - rewrite (ler_trans m1_ge1) // -{1}[m - 1]mulr1. - by rewrite ler_pmul2l ?sqr_Cint_ge1 // (ltr_le_trans ltr01). - have szY2: (size Y <= 2)%N. - move: maxq; rewrite q1 subrr exprS mul0r add0r mulrA !mulr1. - by rewrite -(ler_add2r 1) subrK -mulrSr ler_nat -leqNgt. + - move=> IH; have [q0 | [q1 /eq_leq-szY2]] := m_ub2_lt2 q Zq mq2_lt2. + by apply: (IH tau1) => //; rewrite tau_psi1 Db q0 mul0r scale0r addr0. have defY: perm_eq Y (eta1 :: eta1^*)%CF. have uYeta: uniq (eta1 :: eta1^*)%CF. - by rewrite /= andbT inE eq_sym; have [[_ /hasPn/=->]] := scohY. + by rewrite /= inE eq_sym (hasPn nrS) ?sYS. rewrite perm_eq_sym uniq_perm_eq //. have [|//]:= leq_size_perm uYeta _ szY2. by apply/allP; rewrite /= Yeta1 ccY. - have memYtau1c: {subset map (tau1 \o cfAut conjC) Y <= map tau1 Y}. + have memYtau1c: {subset [seq tau1 eta^* | eta <- Y]%CF <= map tau1 Y}. by move=> _ /mapP[eta Yeta ->]; rewrite /= map_f ?ccY. - apply: (IH _ (dual_coherence scohY cohY szY2)). + apply: IH (dual_coherence scohY cohY szY2) _ _ _. - rewrite (map_comp -%R) orthogonal_oppr. by apply/orthogonalP=> phi psi ? /memYtau1c; apply: (orthogonalP o_tauXY). - rewrite (map_comp -%R) orthogonal_oppr. by apply/orthoPl=> psi /memYtau1c; apply: (orthoPl oX1tauY). - rewrite Dpsi1 (eq_big_perm _ defY) Db q1 /= mul1r big_cons big_seq1. + rewrite tau_psi1 (eq_big_perm _ defY) Db q1 /= mul1r big_cons big_seq1. by rewrite scalerDr addrA subrK -scalerN opprK. - have [[[Itau1 Ztau1] Dtau1] [_ oXX]] := (cohY, orthonormalP oX). + have [[Itau1 Ztau1] Dtau1] := cohY. have n1X1: '[X1] = 1. - apply: (addIr '[a *: tau1 eta1]); rewrite -cfnormBd; last first. - by rewrite cfdotZr (orthoPl oX1tauY) ?mulr0 ?map_f. - rewrite -Dpsi1 Itau // cfnormBd; last first. - by rewrite cfdotZr (orthogonalP oXY) ?mulr0. - by rewrite !cfnormZ Itau1 ?seqInd_zcharW // oXX ?eqxx. + rewrite -(canLR (addrK _) norm_psi1) -Itau // tau_psi1. + rewrite cfnormBd; last by rewrite cfdotZr (orthoPl oX1tauY) ?map_f ?mulr0. + by rewrite cfnormZ norm_Cnat // Itau1 ?mem_zchar ?oYY // eqxx mulr1 addrK. without loss{Itau2 Ztau2 Dtau2} defX1: tau2 cohX o_tauXY / X1 = tau2 xi1. move=> IH; have ZX: {subset X <= 'Z[X]} by apply: seqInd_zcharW. have dirrXtau xi: xi \in X -> tau2 xi \in dirr G. by move=> Xxi; rewrite dirrE Ztau2 1?Itau2 ?ZX //= oXX ?eqxx. have dirrX1: X1 \in dirr G. - rewrite dirrE n1X1 eqxx -[X1](subrK (a *: tau1 eta1)) -Dpsi1. - rewrite rpredD ?scale_zchar ?(zcharW (Ztau _ _)) //. - by rewrite Ztau1 ?seqInd_zcharW. - have oX1_Xd xi: - xi \in X -> xi != xi1 -> '[d xi *: tau2 xi1 - tau2 xi, X1] = d xi. + rewrite dirrE n1X1 eqxx -(canLR (subrK _) tau_psi1). + by rewrite rpredD ?rpredZ_Cnat ?(zcharW (Ztau _ _)) ?Ztau1 ?seqInd_zcharW. + have{Zxi1Xd} oXdX1 xi: xi \in X -> xi != xi1 -> + '[d xi *: tau2 xi1 - tau2 xi, X1] = d xi. - move=> Xxi xi1'xi; have ZXxi := Zxi1Xd xi Xxi. - rewrite -[X1](subrK (a *: tau1 eta1)) -Dpsi1 cfdotDr cfdotZr addrC. - rewrite cfdotBl cfdotZl 2?(orthogonalP o_tauXY) ?map_f //. - rewrite !(mulr0, subr0, conjC0) add0r -{1}raddfZ_Cnat ?Cnat_nat //. - rewrite -opprB cfdotNl -raddfB Dtau2 //. - rewrite Itau //; last exact: zchar_subset ZXxi. - rewrite cfdotBr cfdotZr addrC !cfdotBl !cfdotZl. - rewrite 2?(orthogonalP oXY) // !(mulr0, oppr0, add0r, conjC0). - by rewrite !oXX // eqxx (negPf xi1'xi) add0r opprK mulr1. - have Xxi2: xi1^*%CF \in X by apply: ccX. - have xi1'2: xi1^*%CF != xi1 by have [[_ /hasPn->]] := scohX. - have xi2tau_irr: - tau2 xi1^*%CF \in dirr G by rewrite dirr_opp dirrXtau. - have d_xi2: d xi1^*%CF = 1. - by rewrite /d cfunE conj_Cnat // (Cnat_seqInd1 Xxi1). - have [||def_X1]:= cfdot_add_dirr_eq1 (dirrXtau _ Xxi1) xi2tau_irr dirrX1. - - by rewrite -[tau2 xi1]scale1r -d_xi2 oX1_Xd. - - exact: IH. - have sX_xi12: {subset X <= xi1 :: xi1^*%CF}. - apply/allP/allPn=> [[xi3 Xxi3 /norP[xi1'3 /norP[xi2'3 _]]]]. - suffices d3_0: d xi3 = 0. - by have:= seqInd1_neq0 nsHL Xxi3; rewrite def_d // d3_0 mul0r eqxx. - rewrite -oX1_Xd // def_X1 cfdotNr cfdotBl cfdotZl !Itau2 ?ZX //. - by rewrite !oXX // (negPf xi2'3) eq_sym (negPf xi1'2) mulr0 add0r opprK. - have{sX_xi12 defX} defX: perm_eq X (xi1 :: xi1^*%CF). - have uXxi: uniq (xi1 :: xi1^*)%CF by rewrite /= andbT inE eq_sym. - rewrite perm_eq_sym uniq_perm_eq // => xi. - by apply/idP/idP; [rewrite !inE => /pred2P[]-> | apply: sX_xi12]. - have szX2: (size X <= 2)%N by rewrite (perm_eq_size defX). - apply: (IH _ (dual_coherence scohX cohX szX2)) def_X1. - rewrite (map_comp -%R) orthogonal_oppl. - apply/orthogonalP=> _ eta /mapP[xi Xxi ->]. - by apply: (orthogonalP o_tauXY); rewrite map_f ?ccX. - move: Dpsi1; rewrite -raddfZ_Cnat // defX1. - apply: (bridge_coherent scohS ccsXS cohX ccsYS cohY X'Y). + rewrite -(canLR (subrK _) tau_psi1) cfdotDr addrC. + rewrite (span_orthogonal o_tauXY) ?rpredB ?rpredZ ?memv_span ?map_f //. + rewrite add0r -opprB cfdotNl -{1}raddfZ_Cnat ?Cnat_nat // -raddfB. + rewrite Dtau2 // Itau ?cfdotBr ?opprB //; last exact: zchar_subset ZXxi. + rewrite (span_orthogonal oXY) ?rpredB ?rpredZ ?memv_span // sub0r. + by rewrite cfdotBl cfdotZl opprB !oXX // eqxx mulr1 mulrb ifN ?subr0. + pose xi3 := xi1^*%CF; have Xxi3: xi3 \in X by apply: ccX. + have xi1'3: xi3 != xi1 by rewrite (hasPn nrS) ?sXS. + have [| defX1]: X1 = tau2 xi1 \/ X1 = - tau2 xi3; first 2 [exact : IH]. + have d_xi3: d xi3 = 1 by rewrite /d cfunE conj_Cnat ?(Cnat_seqInd1 Xxi1). + have:= oXdX1 xi3 Xxi3 xi1'3; rewrite d_xi3 scale1r. + by apply: cfdot_add_dirr_eq1; rewrite // ?rpredN dirrXtau. + have szX2: (size X <= 2)%N. + apply: uniq_leq_size (xi1 :: xi3) uX _ => // xi4 Xxi4; rewrite !inE. + apply: contraR (seqInd1_neq0 nsHL Xxi4) => /norP[xi1'4 xi3'4]. + rewrite def_d // -oXdX1 // defX1 cfdotNr cfdotBl cfdotZl opprB. + by rewrite !Itau2 ?ZX ?oXX // !mulrb ifN ?ifN_eqC // mulr0 subr0 mul0r. + apply: (IH _ (dual_coherence scohX cohX szX2)) defX1. + apply/orthogonalP=> _ psi2 /mapP[xi Xxi -> /=] Ytau_psi2. + by rewrite cfdotNl (orthogonalP o_tauXY) ?oppr0 // map_f ?ccX. + rewrite -raddfZ_Cnat // defX1 in tau_psi1. + apply: (bridge_coherent scohS ccsXS cohX ccsYS cohY X'Y) tau_psi1. by rewrite (zchar_on Zpsi1) rpredZ_Cnat ?mem_zchar. -have{caseA_coh12} cohXY: coherent (X ++ Y) L^# tau. - have [/caseA_coh12// | caseB] := boolP caseA. - have defZ: Z = W2 by rewrite /Z (negPf caseB). - have{caseB} [case_c2 _ _] := caseB_P caseB. - move/(_ case_c2) in oRW; pose PtypeL := c2W2 case_c2. - have{prW2} pr_w2 := prW2 case_c2; set w2 := #|W2| in pr_w2. +have{caseA_cohXY Itau1 Ztau1 Dtau1 oYYt} cohXY: coherent (X ++ Y) L^# tau. + have [in_caseA | in_caseB] := boolP caseA; first exact: caseA_cohXY. + have defZ: Z = W2 by rewrite /Z ifN. + have /subsetIP[_ cZL] := caseB_sZZL in_caseB. + have{in_caseB} [c2 _ _] := caseB_P in_caseB; move/(_ c2) in oRW. + pose PtypeL := c2_ddA0 c2; pose w2 := #|W2|. + have{c2_prW2} pr_w2: prime w2 := c2_prW2 c2. have /cyclicP[z0 cycZ]: cyclic Z by rewrite defZ prime_cyclic. - have idYZ: {in Y & Z^#, forall (eta : 'CF(L)) x, tau1 eta x = tau1 eta z0}. - move=> eta x Yeta; rewrite !inE andbC cycZ => /andP[/cyclePmin[k]]. - rewrite orderE -cycZ defZ -/w2 => lt_k_w2 -> nt_z0k. - have k_gt0: (0 < k)%N by rewrite lt0n (contraNneq _ nt_z0k) // => ->. - have cokw2: coprime k w2 by rewrite coprime_sym prime_coprime // gtnNdvd. - have sW2G: W2 \subset G by rewrite -defW2 subIset // (subset_trans sHL). - have [u Du _]:= make_pi_cfAut G cokw2. - rewrite -Du ?Ztau1 ?seqInd_zcharW //; last by rewrite orderE -cycZ defZ. - have nAL: L \subset 'N(A) by rewrite normD1 normal_norm. + have oz0: #[z0] = w2 by rewrite /w2 -defZ cycZ. + have regYZ: {in Y & Z^#, forall (eta : 'CF(L)) x, tau1 eta x = tau1 eta z0}. + rewrite cycZ => eta x Yeta /setD1P[ntx /cyclePmin[k lt_k_z0 Dx]]. + have{ntx} k_gt0: (0 < k)%N by case: (k) Dx ntx => // -> /eqP[]. + have{lt_k_z0} [cokw2 zz0_dv_w2]: coprime k w2 /\ #[z0] %| w2. + by rewrite coprime_sym prime_coprime // -oz0 // gtnNdvd. + have [u Du _]:= make_pi_cfAut G cokw2; rewrite Dx -Du ?Ztau1 ?mem_zchar //. + have nAL: L \subset 'N(A) by have [_ /subsetIP[]] := normedTI_P tiA. pose ddA := restr_Dade_hyp PtypeL (subsetUl _ _) nAL. - have cohY_Dade: coherent_with Y L^# (Dade ddA) tau1. - split=> // phi Yphi; rewrite Dtau1 ?Dade_Ind //. - by rewrite (@zchar_on _ _ Y) -?zcharD1_seqInd. + have{Dtau1} Dtau1: {in 'Z[Y, L^#], tau1 =1 Dade ddA}. + by move=> phi Yphi/=; rewrite Dtau1 ?Dade_Ind ?(zcharD1_seqInd_on _ Yphi). + have cohY_Dade: coherent_with Y L^# (Dade ddA) tau1 by []. rewrite (cfAut_Dade_coherent cohY_Dade) ?irrY //; last first. split; last exact: cfAut_seqInd. exact: seqInd_nontrivial_irr (irrY _ Yeta) (Yeta). - rewrite -[cfAut u _](subrK eta) raddfD cfunE. - apply: canLR (subrK _) _; rewrite subrr. - have [_ ->] := cohY_Dade; last first. - by rewrite -opprB rpredN zcharD1_seqInd // seqInd_sub_aut_zchar. - rewrite Dade_id; last first. - by rewrite !inE -cycle_eq1 -cycle_subG -cycZ ntZ. + rewrite -[cfAut u _](subrK eta) -opprB addrC raddfB !cfunE -[RHS]subr0. + congr (_ - _); rewrite Dtau1 ?zcharD1_seqInd ?seqInd_sub_aut_zchar //. + rewrite Dade_id; last by rewrite !inE -cycle_eq1 -cycle_subG -cycZ ntZ. rewrite !cfunE cfker1 ?aut_Cnat ?subrr ?(Cnat_seqInd1 Yeta) //. - rewrite -cycle_subG -cycZ (subset_trans sZH') //. - have [j /setDP[kerH'j _] ->] := seqIndP Yeta. - by rewrite inE in kerH'j; rewrite sub_cfker_Ind_irr. - have [_ [Itau _] oSS _ _] := scohS. - have oY: orthonormal Y by apply: sub_orthonormal (irr_orthonormal L). - have oYtau: orthonormal (map tau1 Y) by apply: map_orthonormal. - have oXY: orthogonal X Y. - apply/orthogonalP=> xi eta Xxi Yeta; apply: orthoPr xi Xxi. - exact: (subset_ortho_subcoherent scohS sXS (sYS _ Yeta) (X'Y _ Yeta)). - have [Y1 Dpsi1 defY1]: exists2 Y1, - forall i : Iirr Z, i != 0 -> + have [j /setDP[kerH'j _] Deta] := seqIndP Yeta; rewrite inE in kerH'j. + by rewrite -cycle_subG -cycZ (subset_trans sZH') // Deta sub_cfker_Ind_irr. + have [_ [Itau /(_ _ _)/zcharW-Ztau] oSS _ _] := scohS. + pose gamma i : 'CF(L) := 'Ind[L] 'chi[Z]_i - #|H : Z|%:R *: eta1. + have [Y1 tau_gamma defY1]: exists2 Y1 : 'CF(G), forall i : Iirr Z, i != 0 -> exists2 X1 : 'CF(G), orthogonal X1 (map tau1 Y) - & tau ('Ind 'chi_i - #|H : Z|%:R *: eta1) = X1 - #|H : Z|%:R *: Y1 + & tau (gamma i) = X1 - #|H : Z|%:R *: Y1 & Y1 = tau1 eta1 \/ size Y = 2 /\ Y1 = dual_iso tau1 eta1. - - have [i0 nz_i0]: exists i0 : Iirr Z, i0 != 0. - by apply: (ex_intro _ (Sub 1%N _)) => //; rewrite NirrE classes_gt1. - pose psi1 := tau1 eta1; pose b := psi1 z0. + - pose psi1 := tau1 eta1; pose b := psi1 z0. pose a := (psi1 1%g - b) / #|Z|%:R. - have sZL := normal_sub nsZL; have sZG := subset_trans sZL sLG. - have Dpsi1: 'Res psi1 = a *: cfReg Z + b%:A. - apply/cfun_inP=> z Zz. - rewrite cfResE // !cfunE cfun1E Zz mulr1 cfuniE ?normal1 // inE. - have [-> | ntz] := altP eqP; first by rewrite mulr1 divfK ?neq0CG ?subrK. - by rewrite !mulr0 add0r idYZ // !inE ntz. + have sZG := subset_trans sZL sLG. + have Dpsi1: 'Res[Z] psi1 = a *: cfReg Z + b%:A. + apply/cfun_inP=> z Zz; rewrite cfResE // !(cfRegE, cfunE) cfun1E Zz mulr1. + have [-> | ntz] := altP eqP; first by rewrite divfK ?neq0CG ?subrK. + by rewrite mulr0 add0r regYZ // !inE ntz. have /dvdCP[x0 Zx0 Dx0]: (#|H : Z| %| a)%C. - have /dvdCP[x Zx Dx]: (#|H| %| b - psi1 1%g)%C. - have psi1Z z: z \in Z^# -> psi1 z = b. - case/setD1P=> ntz Zz; rewrite -(cfResE _ _ Zz) // Dpsi1 !cfunE cfun1E. - by rewrite cfuniE ?normal1 // Zz inE (negPf ntz) !mulr0 mulr1 add0r. - have /dirrP[e [i /(canLR (signrZK e)) Epsi1]]: psi1 \in dirr G. - have [_ oYt] := orthonormalP oYtau. - by rewrite dirrE oYt ?map_f // !eqxx Ztau1 ?seqInd_zcharW. - have Zz: z0 \in Z^# by rewrite !inE -cycle_eq1 -cycle_subG -cycZ ntZ /=. - have/(Zconst_modH i)[z1 Zz1 z2 Zz2 |_] := Zz. - by rewrite -Epsi1 !cfunE !psi1Z. - by rewrite -Epsi1 !cfunE -mulrBr rpredMsign psi1Z. - apply/dvdCP; exists (- x); first by rewrite rpredN. - rewrite /a -opprB Dx -(Lagrange sZH) mulnC [p in x * p]natrM -!mulNr. - by rewrite !mulrA !mulfK ?neq0CG. + suffices dvH_p_psi1: (#|H| %| b - psi1 1%g)%C. + rewrite -(@dvdC_mul2r _ #|Z|%:R) ?divfK ?neq0CG // -opprB rpredN /=. + by rewrite -natrM mulnC Lagrange. + have psi1Z z: z \in Z^# -> psi1 z = b by apply: regYZ. + have /dirrP[e [i /(canLR (signrZK e))-Epsi1]]: psi1 \in dirr G. + have [_ oYt] := orthonormalP oYtau. + by rewrite dirrE oYt ?map_f // !eqxx Ztau1 ?seqInd_zcharW. + have Zz: z0 \in Z^# by rewrite !inE -cycle_eq1 -cycle_subG -cycZ ntZ /=. + have [z1 z2 Zz1 Zz2 |_] := irrZmodH i _ _ Zz. + by rewrite -Epsi1 !cfunE !psi1Z. + by rewrite -Epsi1 !cfunE -mulrBr rpredMsign psi1Z. pose x1 := '[eta1, 'Res psi1]; pose x := x0 + 1 - x1. have Zx: x \in Cint. rewrite rpredB ?rpredD // Cint_cfdot_vchar // ?(seqInd_vcharW Yeta1) //. by rewrite cfRes_vchar // Ztau1 ?seqInd_zcharW. pose Y1 := - \sum_(eta <- Y) (x - (eta == eta1)%:R) *: tau1 eta. - pose alpha i := 'Ind[L, Z] 'chi_i - #|H : Z|%:R *: eta1. - have IZfacts i: i != 0 -> - [/\ 'chi_i 1%g = 1, 'Ind[L, Z] 'chi_i \in 'Z[X] & alpha i \in 'Z[S, L^#]]. - - move=> nzi; have /andP[_ /eqP lin_i]: 'chi_i \is a linear_char. + have IndZfacts i: i != 0 -> + [/\ 'chi_i 1%g = 1, 'Ind 'chi_i \in 'Z[X] & gamma i \in 'Z[S, L^#]]. + - move=> nzi; have /andP[_ /eqP-lin_i]: 'chi_i \is a linear_char. by rewrite lin_irr_der1 (derG1P _) ?sub1G // cycZ cycle_abelian. have Xchi: 'Ind 'chi_i \in 'Z[X]. rewrite -(cfIndInd _ sHL) // ['Ind[H] _]cfun_sum_constt linear_sum. - apply: rpred_sum => k k_i; rewrite linearZ; apply: scale_zchar. + apply: rpred_sum => k k_i; rewrite linearZ rpredZ_Cint ?mem_zchar //=. by rewrite Cint_cfdot_vchar_irr // cfInd_vchar ?irr_vchar. - rewrite seqInd_zcharW //; apply/seqIndP; exists k => //. - rewrite !inE sub1G andbT; apply: contra k_i => kerZk. - rewrite -Frobenius_reciprocity. - have ->: 'Res[Z] 'chi_k = ('chi_k 1%g)%:A. - apply: cfun_inP => z Zz; rewrite cfunE cfun1E Zz mulr1 cfResE //. - by rewrite cfker1 ?(subsetP kerZk). - by rewrite cfdotZr -irr0 cfdot_irr (negPf nzi) mulr0. + rewrite mem_seqInd ?normal1 // !inE sub1G andbT. + by rewrite -(sub_cfker_constt_Ind_irr k_i) // subGcfker. split=> //; rewrite zcharD1E !cfunE cfInd1 // uniY // lin_i mulr1. - rewrite -divgS // -(sdprod_card defL) -(Lagrange sZH) -mulnA mulKn //. - rewrite -natrM subrr rpredB //=; first by rewrite (zchar_subset sXS). - by rewrite scale_zchar ?rpred_nat // seqInd_zcharW ?sYS. - have Dalpha (i : Iirr Z) (nzi : i != 0) : + rewrite oW1 -natrM mulnC Lagrange_index // subrr eqxx andbT. + by rewrite rpredB ?rpredZnat ?(zchar_subset sXS Xchi) ?mem_zchar ?sYS. + have Dgamma (i : Iirr Z) (nzi : i != 0): exists2 X1 : 'CF(G), orthogonal X1 (map tau1 Y) - & tau (alpha i) = X1 - #|H : Z|%:R *: Y1. - - have [lin_i Xchi Zalpha] := IZfacts i nzi. - have Da: '[tau (alpha i), psi1] = a - #|H : Z|%:R * x1. - rewrite !(=^~ Frobenius_reciprocity, cfdotBl) cfResRes // cfdotZl. - congr (_ - _); rewrite cfdotC Dpsi1 cfdotDl cfdotZl cfReg_sum. - rewrite cfdot_suml (bigD1 i) //= big1 => [|j i'j]; last first. - by rewrite cfdotZl cfdot_irr (negPf i'j) mulr0. - rewrite !cfdotZl cfnorm_irr lin_i addr0 !mulr1. - rewrite -irr0 cfdot_irr eq_sym (negPf nzi) mulr0 addr0. + & tau (gamma i) = X1 - #|H : Z|%:R *: Y1. + - have [lin_i Xchi Zgamma] := IndZfacts i nzi. + have Da: '[tau (gamma i), psi1] = a - #|H : Z|%:R * x1. + rewrite !(=^~ cfdot_Res_r, cfdotBl) cfResRes // cfdotZl -/x1 Dpsi1. + congr (_ - _); rewrite cfdotDr cfReg_sum cfdotC cfdotZl cfdotZr. + rewrite -(big_tuple _ _ _ xpredT (fun xi : 'CF(Z) => xi 1%g *: xi)). + rewrite cfproj_sum_orthonormal ?irr_orthonormal ?mem_irr // lin_i mulr1. + rewrite -irr0 cfdot_irr (negPf nzi) mulr0 addr0. by rewrite aut_Cint // Dx0 rpredM ?rpred_nat. - have [Y2 dY2 [X1 [dX1 _ oX1Yt]]] := - orthogonal_split (map tau1 Y) (tau (alpha i)). - exists X1 => //; rewrite dX1 addrC scalerN opprK scaler_sumr. - congr (_ + _); have [_ -> ->] := orthonormal_span oYtau dY2. - rewrite big_map; apply: eq_big_seq => eta Yeta. - rewrite scalerA -[Y2](addrK X1) -dX1 cfdotBl (orthoPl oX1Yt) ?map_f //. - congr (_ *: _); rewrite subr0 !mulrBr mulrDr mulrC -Dx0. - rewrite (addrAC a) -Da -addrA -mulrBr addrC; apply: canRL (subrK _) _. - have Zeta: eta - eta1 \in 'Z[Y, L^#]. + exists (tau (gamma i) + #|H : Z|%:R *: Y1); last by rewrite addrK. + apply/orthoPl=> _ /mapP[eta Yeta ->]. + rewrite scalerN cfdotBl cfdotZl cfproj_sum_orthonormal // [x]addrAC. + rewrite -addrA mulrDr mulrBr mulrC -Dx0 -Da opprD addrA -!raddfB /=. + have Yeta_1: eta - eta1 \in 'Z[Y, L^#]. by rewrite zcharD1E rpredB ?seqInd_zcharW //= !cfunE !uniY ?subrr. - rewrite -cfdotBr -raddfB Dtau1 // Itau //; last first. - by rewrite (zchar_subset sYS) ?seqInd_free. - rewrite cfdotBl (span_orthogonal oXY) ?(zchar_span Xchi)//; last first. - by rewrite memvB ?memv_span. - have [_ oYY] := orthonormalP oY; rewrite cfdotZl cfdotBr !oYY //. - by rewrite eqxx sub0r -mulrN opprB eq_sym. - exists Y1 => //; have{Dalpha} [X1 oX1Y Dalpha] := Dalpha i0 nz_i0. - have [lin_i Xchi Zalpha] := IZfacts i0 nz_i0. - have norm_alpha: '[tau (alpha i0)] = (#|L : Z| + #|H : Z| ^ 2)%:R. + rewrite Dtau1 ?Itau // ?(zchar_subset sYS) // cfdotBl cfdotZl. + rewrite (span_orthogonal oXY) ?(zchar_span Xchi) ?(zchar_span Yeta_1) //. + by rewrite cfdotBr -mulrN opprB !oYY // eqxx eq_sym addrK. + have [i0 nz_i0] := has_nonprincipal_irr ntZ. + exists Y1 => //; have{Dgamma} [X1 oX1Y Dgamma] := Dgamma i0 nz_i0. + have [lin_i Xchi Zgamma] := IndZfacts i0 nz_i0. + have norm_gamma: '[tau (gamma i0)] = (#|L : Z| + #|H : Z| ^ 2)%:R. rewrite natrD Itau // cfnormBd; last first. rewrite (span_orthogonal oXY) ?(zchar_span Xchi) //. by rewrite memvZ ?memv_span. - rewrite cfnorm_Ind_irr //; congr (#|_ : _|%:R + _). - apply/setIidPl; apply: subset_trans (cent_sub_inertia _). - rewrite -(sdprodW defL) mulG_subG (centsS sZZ) centsC ?subsetIr //=. - by rewrite defZ -defW2 subsetIr. - have [_ oYY] := orthonormalP oY; rewrite cfnormZ oYY // eqxx mulr1. - by rewrite normCK rmorph_nat -natrM. - have{norm_alpha} ub_norm_alpha: '[tau (alpha i0)] < (#|H : Z| ^ 2).*2%:R. - rewrite norm_alpha -addnn ltr_nat ltn_add2r. - rewrite -divgS // -(sdprod_card defL) -(Lagrange sZH) -mulnA mulKn //. - rewrite ltn_pmul2l //. + rewrite cfnorm_Ind_irr //; congr (#|_ : Z|%:R + _); last first. + by rewrite cfnormZ oYY // eqxx mulr1 normCK rmorph_nat -natrM. + by apply/setIidPl; rewrite (subset_trans _ (cent_sub_inertia _)) 1?centsC. + have{norm_gamma} ub_norm_gamma: '[tau (gamma i0)] < (#|H : Z| ^ 2).*2%:R. + rewrite norm_gamma -addnn ltr_nat ltn_add2r. + rewrite -(Lagrange_index sHL) ?ltn_pmul2r // -[#|H : Z| ]prednK // ltnS. have frobL2: [Frobenius L / Z = (H / Z) ><| (W1 / Z)]%g. apply: (Frobenius_coprime_quotient defL nsZL) => //. split=> [|y W1y]; first exact: sub_proper_trans ltH'H. by rewrite defZ; have [/= ? [_ [_ _ _ ->]]] := PtypeL. - have nZW1 := subset_trans sW1L (normal_norm nsZL). - rewrite (card_isog (quotient_isog nZW1 _)); last first. - by rewrite coprime_TIg ?(coprimeSg sZH). - rewrite -(prednK (indexg_gt0 H Z)) ltnS -card_quotient //. + have nZW1 := subset_trans sW1L nZL. + have tiZW1: Z :&: W1 = 1%g by rewrite coprime_TIg ?(coprimeSg sZH). + rewrite -oW1 (card_isog (quotient_isog nZW1 tiZW1)) -card_quotient //. rewrite dvdn_leq ?(Frobenius_dvd_ker1 frobL2) // -subn1 subn_gt0. by rewrite cardG_gt1; case/Frobenius_context: frobL2. - pose m : algC := (size Y)%:R. - have{ub_norm_alpha} ub_xm: ~~ (2%:R <= (x - 1) ^+ 2 + (m - 1) * x ^+ 2). - have: ~~ (2%:R <= '[Y1]). - rewrite -2!(ler_pmul2l (gt0CiG H Z)) -!natrM mulnA muln2. - rewrite ltr_geF //; apply: ler_lt_trans ub_norm_alpha. - rewrite Dalpha cfnormBd. + have{ub_norm_gamma} ub_xm: m_ub2 x < 2%:R. + have: '[Y1] < 2%:R. + rewrite -2!(ltr_pmul2l (gt0CiG H Z)) -!natrM mulnA muln2. + apply: ler_lt_trans ub_norm_gamma; rewrite Dgamma cfnormBd. by rewrite cfnormZ normCK rmorph_nat mulrA -subr_ge0 addrK cfnorm_ge0. - rewrite scalerN -scaleNr cfdotZr cfdot_sumr big_seq. - rewrite big1 ?mulr0 // => eta Yeta. - by rewrite cfdotZr (orthoPl oX1Y) ?map_f ?mulr0. + rewrite (span_orthogonal oX1Y) ?memv_span1 ?rpredZ // rpredN big_seq. + by apply/rpred_sum => eta Yeta; rewrite rpredZ ?memv_span ?map_f. rewrite cfnormN cfnorm_sum_orthonormal // (big_rem eta1) //= eqxx. - rewrite big_seq (eq_bigr (fun _ => (x ^+ 2))) => [|eta]; last first. - rewrite mem_rem_uniq // => /andP[/negPf-> _]. + congr (_ + _ < _); first by rewrite Cint_normK 1?rpredB ?rpred1. + transitivity (\sum_(eta <- rem eta1 Y) x ^+ 2). + rewrite rem_filter // !big_filter; apply/eq_bigr => eta /negPf->. by rewrite subr0 Cint_normK. - rewrite Cint_normK 1?rpredB //= -big_seq; congr (~~ (_ <= _ + _)). - rewrite big_const_seq count_predT // -Monoid.iteropE. - rewrite /m (perm_eq_size (perm_to_rem Yeta1)) /=. - by rewrite mulrSr addrK mulr_natl. - have [x_eq0 | nz_x] := eqVneq x 0. + rewrite big_const_seq count_predT // -Monoid.iteropE -[LHS]mulr_natl. + by rewrite /m (perm_eq_size (perm_to_rem Yeta1)) /= mulrSr addrK. + have [x_eq0 | [x_eq1 szY2]] := m_ub2_lt2 x Zx ub_xm. left; rewrite /Y1 x_eq0 (big_rem eta1) //= eqxx sub0r scaleN1r. - rewrite big_seq big1 => [|eta]; last first. - by rewrite mem_rem_uniq // => /andP[/negPf-> _]; rewrite subrr scale0r. - by rewrite addr0 opprK. - have m1_ge1: 1 <= m - 1. - rewrite -(@ler_add2r _ 1) subrK (ler_nat _ 2). - exact: seqInd_nontrivial (irrY _ Yeta1) (Yeta1). - right; have x_eq1: x = 1. - apply: contraNeq ub_xm; rewrite -subr_eq0 => nz_x1; apply: ler_add. - by rewrite sqr_Cint_ge1 // rpredB. - rewrite (ler_trans m1_ge1) // -{1}[m - 1]mulr1 ler_pmul2l. - exact: sqr_Cint_ge1. - exact: ltr_le_trans ltr01 m1_ge1. - have szY2: size Y = 2. - apply: contraNeq ub_xm => Yneq2; rewrite x_eq1 /m subrr !exprS mul0r. - rewrite add0r !mul1r mulr1 -(ler_add2r 1) subrK -mulrSr ler_nat. - by rewrite ltn_neqAle eq_sym Yneq2 -leC_nat -/m -[m](subrK 1) ler_add2r. + rewrite big_seq big1 ?addr0 ?opprK => // eta. + by rewrite mem_rem_uniq // => /andP[/negPf-> _]; rewrite subrr scale0r. have eta1'2: eta1^*%CF != eta1 by apply: seqInd_conjC_neq Yeta1. have defY: perm_eq Y (eta1 :: eta1^*%CF). have uY2: uniq (eta1 :: eta1^*%CF) by rewrite /= inE eq_sym eta1'2. @@ -1229,74 +1080,67 @@ have{caseA_coh12} cohXY: coherent (X ++ Y) L^# tau. have sY2Y: {subset (eta1 :: eta1^*%CF) <= Y}. by apply/allP; rewrite /= cfAut_seqInd ?Yeta1. by have [|//]:= leq_size_perm uY2 sY2Y; rewrite szY2. - split=> //; congr (- _); rewrite (eq_big_perm _ defY) /= x_eq1. - rewrite big_cons big_seq1 eqxx (negPf eta1'2) subrr scale0r add0r. - by rewrite subr0 scale1r. - have [a Za Dxa]: exists2 a, forall xi, a xi \in Cint - & forall xi, xi \in X -> xi 1%g = a xi * #|W1|%:R + right; split=> //; congr (- _); rewrite (eq_big_perm _ defY) /= x_eq1. + rewrite big_cons big_seq1 eqxx (negPf eta1'2) subrr scale0r add0r subr0. + by rewrite scale1r. + have normY1: '[Y1] = 1. + have [-> | [_ ->]] := defY1; first by rewrite oYYt ?eqxx ?map_f. + by rewrite cfnormN oYYt ?eqxx ?map_f ?ccY. + have YtauY1: Y1 \in 'Z[map tau1 Y]. + have [-> | [_ ->]] := defY1; first by rewrite mem_zchar ?map_f. + by rewrite rpredN mem_zchar ?map_f ?ccY. + have spanYtau1 := zchar_span YtauY1. + have norm_eta1: '[eta1] = 1 by rewrite oYY ?eqxx. + have /all_sig2[a Za Dxa] xi: {a | a \in Cnat + & xi \in X -> xi 1%g = a * #|W1|%:R /\ (exists2 X1 : 'CF(G), orthogonal X1 (map tau1 Y) - & tau (xi - a xi *: eta1) = X1 - a xi *: Y1). - - pose aX (xi : 'CF(L)) : algC := (truncC (xi 1%g / #|W1|%:R))%:R. - exists aX => [xi | xi Xxi]; first exact: rpred_nat. - have [k kerZ'k def_xi] := seqIndP Xxi; rewrite !inE sub1G andbT in kerZ'k. - set a := aX xi; have Dxi1: xi 1%g = a * #|W1|%:R. - rewrite /a /aX def_xi cfInd1 // -divgS // -(sdprod_card defL) mulKn //. - by rewrite mulrC mulfK ?neq0CG // irr1_degree natCK. - split=> //; have Da: a = 'chi_k 1%g. - apply: (mulIf (neq0CG W1)); rewrite -Dxi1 def_xi cfInd1 //. - by rewrite mulrC -divgS // -(sdprod_card defL) mulKn. - have [i0 nzi0 Res_k]: exists2 i: Iirr Z, i != 0 & 'Res 'chi_k = a *: 'chi_i. - have [chi /andP[Nchi lin_chi] defRkZ] := cfcenter_Res 'chi_k. + & tau (xi - a *: eta1) = X1 - a *: Y1)}. + - case Xxi: (xi \in X); last by exists 0; rewrite ?rpred0. + have /sig2_eqW[k /setDP[_ kerZ'k] def_xi] := seqIndP Xxi. + rewrite inE in kerZ'k. + pose a := 'chi_k 1%g; have Na: a \in Cnat by apply: Cnat_irr1. + have Dxi1: xi 1%g = a * #|W1|%:R by rewrite mulrC oW1 def_xi cfInd1. + exists a => // _; split=> //. + have [i0 nzi0 Res_k]: exists2 i, i != 0 & 'Res[Z] 'chi_k = a *: 'chi_i. + have [chi lin_chi defRkZ] := cfcenter_Res 'chi_k. have sZ_Zk: Z \subset 'Z('chi_k)%CF. - by rewrite (subset_trans sZZ) // -cap_cfcenter_irr bigcap_inf. - have{Nchi lin_chi} /irrP[i defRk] : 'Res chi \in irr Z. - by rewrite lin_char_irr // qualifE cfRes_char // cfRes1. - have{chi defRk defRkZ} defRk: 'Res 'chi_k = a *: 'chi_i. - by rewrite -defRk -linearZ -/a Da -defRkZ /= cfResRes ?cfcenter_sub. - exists i => //; apply: contra kerZ'k. - rewrite -subGcfker => /subsetP sZker. - apply/subsetP=> t Zt; rewrite cfkerEirr inE. - by rewrite -!(cfResE _ sZH) // defRk !cfunE cfker1 ?sZker. + by rewrite (subset_trans sZZH) // -cap_cfcenter_irr bigcap_inf. + have{lin_chi} /irrP[i defRk]: 'Res chi \in irr Z. + by rewrite lin_char_irr ?cfRes_lin_char. + have{chi defRk defRkZ} defRk: 'Res[Z] 'chi_k = a *: 'chi_i. + by rewrite -defRk -linearZ -defRkZ /= cfResRes ?cfcenter_sub. + exists i => //; apply: contra kerZ'k => i_0; apply/constt0_Res_cfker=> //. + by rewrite inE defRk cfdotZl cfdot_irr i_0 mulr1 irr1_neq0. set phi := 'chi_i0 in Res_k; pose a_ i := '['Ind[H] phi, 'chi_i]. pose rp := irr_constt ('Ind[H] phi). have defIphi: 'Ind phi = \sum_(i in rp) a_ i *: 'chi_i. exact: cfun_sum_constt. have a_k: a_ k = a. - by rewrite /a_ -cfdot_Res_r Res_k cfdotZr cfnorm_irr mulr1 rmorph_nat. - have rp_k: k \in rp by rewrite inE ['[_, _]]a_k Da irr1_neq0. + by rewrite /a_ -cfdot_Res_r Res_k cfdotZr cfnorm_irr mulr1 conj_Cnat. + have rp_k: k \in rp by rewrite inE ['[_, _]]a_k irr1_neq0. have resZr i: i \in rp -> 'Res[Z] 'chi_i = a_ i *: phi. - move=> r_i; rewrite ['Res _]cfun_sum_cfdot. - rewrite (bigD1 i0) // big1 => /= [|j i0'j]. - rewrite cfdot_Res_l addr0 -/phi cfdotC conj_Cnat //. - by rewrite Cnat_cfdot_char_irr ?cfInd_char ?irr_char. - apply/eqP; rewrite scaler_eq0 cfdot_Res_l. - rewrite -(inj_eq (mulfI r_i)) mulr0 -/(a_ i) -cfdotZl. - have: '['Ind[H] phi, 'Ind[H] 'chi_j] = 0. - apply: not_cfclass_Ind_ortho => //. - have defIj: 'I_H['chi_j] = H. - apply/setIidPl; apply: subset_trans (cent_sub_inertia _). - by rewrite centsC (subset_trans sZZ) ?subsetIr. - rewrite -(congr1 (cfclass _) defIj) cfclass_inertia inE. - by rewrite eq_sym (inj_eq irr_inj). - rewrite defIphi cfdot_suml => /psumr_eq0P-> //; first by rewrite eqxx. - move=> i1 _; rewrite cfdotZl. - by rewrite mulr_ge0 ?Cnat_ge0 ?Cnat_cfdot_char ?cfInd_char ?irr_char. + rewrite constt_Ind_Res -/phi => /Clifford_Res_sum_cfclass-> //. + have Na_i: a_ i \in Cnat by rewrite Cnat_cfdot_char ?cfInd_char ?irr_char. + rewrite -/phi cfdot_Res_l cfdotC conj_Cnat {Na_i}//; congr (_ *: _). + have <-: 'I_H['Res[Z] 'chi_k] = H. + apply/eqP; rewrite eqEsubset subsetIl. + by apply: subset_trans (sub_inertia_Res _ _); rewrite ?sub_Inertia. + by rewrite Res_k inertia_scale_nz ?irr1_neq0 // cfclass_inertia big_seq1. have lin_phi: phi 1%g = 1. apply: (mulfI (irr1_neq0 k)); have /resZr/cfunP/(_ 1%g) := rp_k. - by rewrite cfRes1 // cfunE mulr1 a_k Da. + by rewrite cfRes1 // cfunE mulr1 a_k. have Da_ i: i \in rp -> 'chi_i 1%g = a_ i. move/resZr/cfunP/(_ 1%g); rewrite cfRes1 // cfunE => ->. by rewrite lin_phi mulr1. - pose chi i := 'Ind[L, H] 'chi_i; pose alpha i := chi i - a_ i *: eta1. + pose chi i := 'Ind[L] 'chi[H]_i; pose alpha i := chi i - a_ i *: eta1. have Aalpha i: i \in rp -> alpha i \in 'CF(L, A). - move=> r_i; rewrite cfun_onD1 !cfunE cfInd1 // (uniY _ Yeta1). - rewrite -divgS // -(sdprod_card defL) mulKn // Da_ // mulrC subrr eqxx. + move=> r_i; rewrite cfun_onD1 !cfunE cfInd1 // (uniY _ Yeta1) -oW1. + rewrite Da_ // mulrC subrr eqxx. by rewrite memvB ?cfInd_normal ?memvZ // (seqInd_on _ Yeta1). - have [sum_alpha sum_a2]: - 'Ind phi - #|H : Z|%:R *: eta1 = \sum_(i in rp) a_ i *: alpha i - /\ \sum_(i in rp) a_ i ^+ 2 = #|H : Z|%:R. - + set rhs2 := _%:R; set lhs1 := _ - _; set rhs1 := \sum_(i | _) _. - set lhs2 := \sum_(i | _) _. + have [sum_alpha sum_a2]: gamma i0 = \sum_(i in rp) a_ i *: alpha i + /\ \sum_(i in rp) a_ i ^+ 2 = #|H : Z|%:R. + + set lhs1 := LHS; set lhs2 := (lhs in _ /\ lhs = _). + set rhs1 := RHS; set rhs2 := (rhs in _ /\ _ = rhs). have eq_diff: lhs1 - rhs1 = (lhs2 - rhs2) *: eta1. rewrite scalerBl addrAC; congr (_ - _). rewrite -(cfIndInd _ sHL sZH) defIphi linear_sum -sumrB scaler_suml. @@ -1308,340 +1152,191 @@ have{caseA_coh12} cohXY: coherent (X ++ Y) L^# tau. rewrite sum_cfunE big1 ?subrr // => i rp_i. by rewrite cfunE (cfun_on0 (Aalpha i rp_i)) ?mulr0 // !inE eqxx. rewrite eq_diff cfunE mulf_eq0 subr_eq0 (negPf (seqInd1_neq0 _ Yeta1)) //. - rewrite orbF => /eqP sum_a2; split=> //; apply/eqP; rewrite -subr_eq0. + rewrite orbF => /eqP-sum_a2; split=> //; apply/eqP; rewrite -subr_eq0. by rewrite eq_diff sum_a2 subrr scale0r. have Xchi i: i \in rp -> chi i \in X. move=> rp_i; apply/seqIndP; exists i => //; rewrite !inE sub1G andbT. - apply: contra rp_i => kerZi; rewrite -cfdot_Res_r. - suffices ->: 'Res[Z] 'chi_i = ('chi_i 1%g)%:A. - by rewrite cfdotZr -irr0 cfdot_irr (negPf nzi0) mulr0. - apply/cfun_inP=> t Zt; rewrite cfResE // cfunE cfun1E Zt mulr1. - by rewrite cfker1 ?(subsetP kerZi). + apply: contra rp_i => kerZi; rewrite -cfdot_Res_r cfRes_sub_ker //. + by rewrite cfdotZr -irr0 cfdot_irr (negPf nzi0) mulr0. have oRY i: i \in rp -> orthogonal (R (chi i)) (map tau1 Y). move/Xchi=> Xchi_i; rewrite orthogonal_sym. by rewrite (coherent_ortho_supp scohS) // ?sXS // (contraL (X'Y _)). - have n1Y1: '[Y1] = 1. - have [_ oYYt] := orthonormalP oYtau. - have [-> | [_ ->]] := defY1; - by rewrite ?cfnormN oYYt ?eqxx ?map_f // cfAut_seqInd. - have YtauY1: Y1 \in 'Z[map tau1 Y]. - have [-> | [_ ->]] := defY1; - by rewrite ?rpredN mem_zchar ?map_f ?cfAut_seqInd. - have /fin_all_exists[XbZ defXbZ] i: exists XbZ, let: (X1, b, Z1) := XbZ in - [/\ tau (alpha i) = X1 - b *: Y1 + Z1, - i \in rp -> X1 \in 'Z[R (chi i)], i \in rp -> b \is Creal, - '[Z1, Y1] = 0 & i \in rp -> orthogonal Z1 (R (chi i))]. - - have [X1 dX1 [YZ1 [dXYZ _ oYZ1R]]] := + have Za_ i: a_ i \in Cint. + by rewrite Cint_cfdot_vchar_irr // cfInd_vchar ?irr_vchar. + have Zeta1: eta1 \in 'Z[irr L] := seqInd_vcharW Yeta1. + have Ztau_alpha i: tau (alpha i) \in 'Z[irr G]. + by rewrite !(cfInd_vchar, rpredB) ?irr_vchar ?rpredZ_Cint. + have /all_tag2[X1 R_X1 /all_tag2[b Rb /all_sig2[Z1 oZ1R]]] i: + {X1 : 'CF(G) & i \in rp -> X1 \in 'Z[R (chi i)] + & {b : algC & i \in rp -> b \is Creal + & {Z1 : 'CF(G) | i \in rp -> orthogonal Z1 (R (chi i)) + & tau (alpha i) = X1 - b *: Y1 + Z1 /\ '[Z1, Y1] = 0}}}. + + have [X1 dX1 [YZ1 [dXYZ _ oYZ1R]]] := orthogonal_split (R (chi i)) (tau (alpha i)). - have [Y1b dY1b [Z1 [dYZ1 _ oZY1]]] := orthogonal_split Y1 YZ1. - have{dY1b} [|b Db dY1b] := orthogonal_span _ dY1b. - by rewrite /pairwise_orthogonal /= inE eq_sym -cfnorm_eq0 n1Y1 oner_eq0. - exists (X1, - b Y1, Z1); split. - + by rewrite dXYZ dYZ1 dY1b scaleNr big_seq1 opprK addrA. - + have [_ _ _ Rok _] := scohS => /Xchi/sXS/Rok[ZR oRR _]. + exists X1. + have [_ _ _ Rok _] := scohS => /Xchi/sXS/Rok[ZR oRR _]. have [_ -> ->] := orthonormal_span oRR dX1. rewrite big_seq rpred_sum // => aa Raa. - rewrite scale_zchar ?mem_zchar //. - rewrite -[X1](addrK YZ1) -dXYZ cfdotBl (orthoPl oYZ1R) // subr0. - rewrite Cint_cfdot_vchar ?(ZR aa) //. - rewrite !(rpredB, cfInd_vchar) ?irr_vchar //. - rewrite scale_zchar ?(seqInd_vcharW Yeta1) // Cint_cfdot_vchar_irr //. - by rewrite cfInd_vchar ?irr_vchar. - + move=> rp_i; rewrite Db -[Y1b](addrK Z1) -dYZ1 cfdotBl (orthoP oZY1). - rewrite subr0 n1Y1 divr1 -[YZ1](addKr X1) -dXYZ cfdotDl cfdotNl. - rewrite (span_orthogonal (oRY i rp_i)) ?(zchar_span YtauY1) //. - rewrite oppr0 add0r Creal_Cint // rpredN Cint_cfdot_vchar //. - rewrite !(cfInd_vchar, rpredB) ?irr_vchar //. - rewrite -Da_ // scale_zchar ?Cint_Cnat ?Cnat_irr1 //. - exact: (seqInd_vcharW Yeta1). - apply: zchar_trans_on YtauY1 => _ /mapP[eta Yeta ->]. - by rewrite Ztau1 ?seqInd_zcharW. - + exact: (orthoP oZY1). - move/oRY=> oRiY; apply/orthoPl=> aa Raa. - rewrite -[Z1](addKr Y1b) -dYZ1 cfdotDl cfdotNl cfdotC (orthoPl oYZ1R) //. - rewrite dY1b addr0 big_seq1 cfdotZr. - by have [-> | [_ ->]] := defY1; - rewrite ?cfdotNr (orthogonalP oRiY) ?map_f ?cfAut_seqInd //; - rewrite ?(oppr0, mulr0, rmorph0). - pose X1 i := (XbZ i).1.1; pose b i := (XbZ i).1.2; pose Z1 i := (XbZ i).2. - have R_X1 i: i \in rp -> X1 i \in 'Z[R (chi i)]. - by rewrite /X1; case: (XbZ i) (defXbZ i) => [[? ?] ?] []. - have Rb i: i \in rp -> b i \is Creal. - by rewrite /b; case: (XbZ i) (defXbZ i) => [[? ?] ?] []. - have oZY1 i: '[Z1 i, Y1] = 0. - by rewrite /Z1; case: (XbZ i) (defXbZ i) => [[? ?] ?] []. - have oZ1R i: i \in rp -> orthogonal (Z1 i) (R (chi i)). - by rewrite /Z1; case: (XbZ i) (defXbZ i) => [[? ?] ?] []. - have{defXbZ} defXbZ i: tau (alpha i) = X1 i - b i *: Y1 + Z1 i. - by rewrite /X1 /b /Z1; case: (XbZ i) (defXbZ i) => [[? ?] ?] []. + rewrite rpredZ_Cint ?mem_zchar // -(canLR (addrK _) dXYZ) cfdotBl. + by rewrite (orthoPl oYZ1R) // subr0 Cint_cfdot_vchar ?(ZR aa). + pose b := - '[YZ1, Y1]; exists b => [rp_i|]. + rewrite Creal_Cint // rpredN -(canLR (addKr _) dXYZ) cfdotDl. + rewrite (span_orthogonal (oRY i rp_i)) ?rpredN ?(zchar_span YtauY1) //. + rewrite add0r Cint_cfdot_vchar // (zchar_trans_on _ YtauY1) //. + by move=> _ /mapP[eta Yeta ->]; rewrite Ztau1 ?mem_zchar. + exists (YZ1 + b *: Y1) => [/oRY-oRiY|]; last first. + by rewrite addrCA subrK addrC cfdotDl cfdotZl normY1 mulr1 addrN. + apply/orthoPl=> aa Raa; rewrite cfdotDl (orthoPl oYZ1R) // add0r. + by rewrite cfdotC (span_orthogonal oRiY) ?conjC0 ?rpredZ // memv_span. + case/all_and2=> defXbZ oZY1; have spanR_X1 := zchar_span (R_X1 _ _). have ub_alpha i: i \in rp -> [/\ '[chi i] <= '[X1 i] & '[a_ i *: eta1] <= '[b i *: Y1 - Z1 i] -> [/\ '[X1 i] = '[chi i], '[b i *: Y1 - Z1 i] = '[a_ i *: eta1] & exists2 E, subseq E (R (chi i)) & X1 i = \sum_(aa <- E) aa]]. - - move=> rp_i; apply: (subcoherent_norm scohS) (erefl _) _. - + rewrite sXS ?Xchi // scale_zchar ?(seqInd_vcharW Yeta1) //; last first. - by rewrite Cint_cfdot_vchar_irr // cfInd_vchar ?irr_vchar. - split=> //; apply/orthoPr=> xi2; rewrite !inE => Dxi2. - rewrite cfdotZr (orthogonalP oXY) ?mulr0 //. - by case/pred2P: Dxi2 => ->; rewrite ?cfAut_seqInd ?Xchi. - + have [_ IZtau _ _ _] := scohS. + + move=> rp_i; apply: (subcoherent_norm scohS) (erefl _) _. + * rewrite sXS ?Xchi ?rpredZ_Cint /orthogonal //; split=> //=. + by rewrite !cfdotZr !(orthogonalP oXY) ?mulr0 ?eqxx ?ccX // Xchi. + * have [[/(_ _ _)/char_vchar-Z_S _ _] IZtau _ _ _] := scohS. apply: sub_iso_to IZtau; [apply: zchar_trans_on | exact: zcharW]. apply/allP; rewrite /= zchar_split (cfun_onS (setSD _ sHL)) ?Aalpha //. - rewrite rpredB ?scale_zchar ?seqInd_zcharW ?(sYS eta1) ?sXS ?Xchi //. - by rewrite sub_aut_zchar ?zchar_onG ?seqInd_zcharW ?cfAut_seqInd; - rewrite ?sXS ?Xchi //; apply: seqInd_vcharW. - by rewrite -Da_ // irr1_degree rpred_nat. + rewrite rpredB ?rpredZ_Cint ?mem_zchar ?(sYS eta1) // ?sXS ?Xchi //=. + by rewrite sub_aut_zchar ?zchar_onG ?mem_zchar ?sXS ?ccX ?Xchi. suffices oYZ_R: orthogonal (b i *: Y1 - Z1 i) (R (chi i)). rewrite opprD opprK addrA -defXbZ cfdotC. - rewrite (span_orthogonal oYZ_R) ?memv_span1 ?conjC0 //. - exact: (zchar_span (R_X1 i rp_i)). - apply/orthoPl=> aa Raa. - rewrite cfdotBl cfdotZl (orthoPl (oZ1R i _)) //. - by rewrite subr0 cfdotC; have [-> | [_ ->]] := defY1; - rewrite ?cfdotNr (orthogonalP (oRY i _)) ?map_f ?cfAut_seqInd //; - by rewrite ?(mulr0, oppr0, rmorph0). + by rewrite (span_orthogonal oYZ_R) ?memv_span1 ?spanR_X1 ?conjC0. + apply/orthoPl=> aa Raa; rewrite cfdotBl (orthoPl (oZ1R i _)) // cfdotC. + by rewrite subr0 (span_orthogonal (oRY i _)) ?conjC0 ?rpredZ // memv_span. have leba i: i \in rp -> b i <= a_ i. move=> rp_i; have ai_gt0: a_ i > 0 by rewrite -Da_ ?irr1_gt0. - have /orP [b_le0|b_ge0] := Rb i rp_i; last first. - by rewrite (ler_trans b_ge0) ?ltrW. - rewrite -(@ler_pexpn2r _ 2) //; last exact: ltrW. + rewrite (ler_trans (real_ler_norm (Rb i _))) //. + rewrite -(@ler_pexpn2r _ 2) ?qualifE ?(ltrW ai_gt0) ?norm_ger0 //. apply: ler_trans (_ : '[b i *: Y1 - Z1 i] <= _). rewrite cfnormBd; last by rewrite cfdotZl cfdotC oZY1 ?conjC0 ?mulr0. - rewrite cfnormZ (normr_idP _) // n1Y1 mulr1 addrC -subr_ge0 addrK. - exact: cfnorm_ge0. + by rewrite cfnormZ normY1 mulr1 ler_addl cfnorm_ge0. rewrite -(ler_add2l '[X1 i]) -cfnormBd; last first. - rewrite cfdotBr cfdotZr (span_orthogonal (oRY i _)) //; last first. - - exact: (zchar_span YtauY1). - - exact: (zchar_span (R_X1 i rp_i)). - rewrite mulr0 sub0r cfdotC (span_orthogonal (oZ1R i _)) ?raddf0 //. - exact: memv_span1. - exact: (zchar_span (R_X1 i rp_i)). + rewrite cfdotBr cfdotZr (span_orthogonal (oRY i _)) ?spanR_X1 //. + rewrite mulr0 sub0r cfdotC. + by rewrite (span_orthogonal (oZ1R i _)) ?raddf0 ?memv_span1 ?spanR_X1. have Salpha: alpha i \in 'Z[S, L^#]. - rewrite zchar_split (cfun_onS (setSD _ sHL)) ?Aalpha //. - rewrite rpredB ?scale_zchar ?seqInd_zcharW - ?(sYS _ Yeta1) ?sXS ?Xchi //. - by rewrite -Da_ // irr1_degree rpred_nat. - rewrite opprD opprK addrA -defXbZ // Itau ?Salpha //. - rewrite cfnormBd; last first. - by rewrite cfdotZr (orthogonalP oXY) ?mulr0 ?Xchi. - rewrite cfnormZ (normr_idP _) ?(ltrW ai_gt0) //. - have [_ /(_ eta1)->//] := orthonormalP oY; rewrite eqxx mulr1 ler_add2r. - by have [] := ub_alpha i rp_i. + rewrite zcharD1_seqInd // zchar_split Aalpha // andbT. + by rewrite rpredB ?rpredZ_Cint ?mem_zchar ?(sYS eta1) ?sXS ?Xchi. + rewrite opprD opprK addrA -defXbZ ?Itau //. + rewrite cfnormBd; last by rewrite cfdotZr (orthogonalP oXY) ?mulr0 ?Xchi. + rewrite cfnormZ Cint_normK ?(oYY eta1) // eqxx mulr1 ler_add2r. + by have lbX1i: '[chi i] <= '[X1 i] by have [] := ub_alpha i rp_i. have{leba} eq_ab: {in rp, a_ =1 b}. move=> i rp_i; apply/eqP; rewrite -subr_eq0; apply/eqP. apply: (mulfI (irr1_neq0 i)); rewrite mulr0 Da_ // mulrBr. move: i rp_i; apply: psumr_eq0P => [i rp_i | ]. by rewrite subr_ge0 ler_pmul2l ?leba // -Da_ ?irr1_gt0. - have [X2 oX2Y /(congr1 (cfdotr Y1))] := Dpsi1 i0 nzi0. + have [X2 oX2Y /(congr1 (cfdotr Y1))] := tau_gamma i0 nzi0. rewrite sumrB sum_a2 sum_alpha /tau linear_sum /= cfdot_suml cfdotBl. rewrite (span_orthogonal oX2Y) ?memv_span1 ?(zchar_span YtauY1) // add0r. - rewrite cfdotZl n1Y1 mulr1 => /(canLR (@opprK _)) <-. + rewrite cfdotZl normY1 mulr1 => /(canLR (@opprK _)) <-. rewrite -opprD -big_split big1 ?oppr0 //= => i rp_i. rewrite linearZ cfdotZl /= -/tau defXbZ addrC cfdotDl oZY1 addr0. - rewrite cfdotBl cfdotZl n1Y1 mulr1. - rewrite (span_orthogonal (oRY i _)) ?(zchar_span YtauY1) //. - by rewrite add0r mulrN subrr. - exact: (zchar_span (R_X1 i rp_i)). - exists (X1 k). + rewrite cfdotBl cfdotZl normY1 mulr1 mulrBr addrC subrK. + by rewrite (span_orthogonal (oRY i _)) ?spanR_X1 ?mulr0. + exists (X1 k). apply/orthoPl=> psi /memv_span Ypsi. by rewrite (span_orthogonal (oRY k _)) // (zchar_span (R_X1 k rp_k)). apply/eqP; rewrite -/a def_xi -a_k defXbZ addrC -subr_eq0 eq_ab // addrK. - have n1eta1: '[eta1] = 1 by have [_ -> //] := orthonormalP oY; rewrite eqxx. rewrite -cfnorm_eq0 -(inj_eq (addrI '[b k *: Y1])). have [_ [|_]] := ub_alpha k rp_k. rewrite cfnormBd; last by rewrite cfdotZl cfdotC oZY1 conjC0 mulr0. - by rewrite addrC !cfnormZ eq_ab // n1Y1 n1eta1 -subr_ge0 addrK cfnorm_ge0. + by rewrite addrC !cfnormZ eq_ab // normY1 norm_eta1 ler_addr cfnorm_ge0. rewrite cfnormBd; last by rewrite cfdotZl cfdotC oZY1 conjC0 mulr0. - by move=> -> _; rewrite addr0 !cfnormZ eq_ab // n1Y1 n1eta1. - have oX: pairwise_orthogonal X by rewrite (sub_pairwise_orthogonal sXS). - have [_ oYY] := orthonormalP oY. - have [[N_S _ _] [_ /(_ _ _)/zcharW/=Ztau] _ _ _] := scohS. - have n1eta: '[eta1] = 1 by rewrite oYY ?eqxx. - have n1Y1: '[Y1] = 1. - have [_ oYYt] := orthonormalP oYtau. - have [-> | [_ ->]] := defY1; - by rewrite ?cfnormN oYYt ?eqxx ?map_f // cfAut_seqInd. - have YtauY1: Y1 \in <<map tau1 Y>>%VS. - by have [-> | [_ ->]] := defY1; - rewrite ?memvN memv_span ?map_f ?cfAut_seqInd. + by move=> -> _; rewrite addr0 !cfnormZ eq_ab // normY1 norm_eta1. + have scohXY: subcoherent (X ++ Y) tau R. + apply/(subset_subcoherent scohS). + split; first by rewrite cat_uniq uX uY andbT; apply/hasPn. + by move=> xi; rewrite mem_cat => /orP[/sXS | /sYS]. + by move=> xi; rewrite !mem_cat => /orP[/ccX-> | /ccY->]; rewrite ?orbT. + have XYeta1: eta1 \in X ++ Y by rewrite mem_cat Yeta1 orbT. have Z_Y1: Y1 \in 'Z[irr G]. by case: defY1 => [|[_]] ->; rewrite ?rpredN Ztau1 ?mem_zchar ?ccY. - have Zalpha xi: xi \in X -> xi - a xi *: eta1 \in 'Z[S, L^#]. - move=> Xxi; rewrite zcharD1E rpredB ?scale_zchar; - rewrite ?seqInd_zcharW ?(sXS xi) ?sYS //. - by rewrite !cfunE (uniY eta1) //= subr_eq0; have [<-] := Dxa xi Xxi. - have Zbeta eta: eta \in Y -> eta - eta1 \in 'Z[S, L^#]. - move=> Yeta; rewrite zcharD1E rpredB ?seqInd_zcharW ?sYS //=. - by rewrite !cfunE !uniY ?subrr. - have nza xi: xi \in X -> a xi != 0. - move=> Xxi; have [/eqP Dxi1 _] := Dxa _ Xxi; apply: contraTneq Dxi1 => ->. - by rewrite mul0r (seqInd1_neq0 _ Xxi). - have alphaY xi: xi \in X -> '[tau (xi - a xi *: eta1), Y1] = - a xi. - case/Dxa=> _ [X1 oX1Y ->]; rewrite cfdotBl cfdotZl n1Y1 mulr1. - by rewrite (span_orthogonal oX1Y) ?memv_span1 ?add0r. - have betaY eta: eta \in Y -> '[tau (eta - eta1), Y1] = (eta == eta1)%:R - 1. - move=> Yeta; rewrite -Dtau1; last first. - by rewrite zchar_split (zchar_on (Zbeta eta _)) ?rpredB ?seqInd_zcharW. - rewrite raddfB cfdotBl. - have [-> | [szY2 ->]] := defY1. - by rewrite !{1}Itau1 ?seqInd_zcharW // !oYY // !eqxx. - rewrite !cfdotNr opprK !{1}Itau1 ?oYY ?seqInd_zcharW ?cfAut_seqInd //. - have defY: (eta1 :: eta1^*)%CF =i Y. - apply: proj1 (leq_size_perm _ _ _); last by rewrite szY2. - by rewrite /= inE eq_sym (seqInd_conjC_neq _ _ _ Yeta1). - by apply/allP; rewrite /= Yeta1 cfAut_seqInd. - rewrite -defY !inE in Yeta; case/pred2P: Yeta => ->. - rewrite eqxx eq_sym (negPf (seqInd_conjC_neq _ _ _ Yeta1)) //. - by rewrite addrC !subrr. - by rewrite eqxx eq_sym (negPf (seqInd_conjC_neq _ _ _ Yeta1)) ?add0r ?addr0. - pose tau2_X xi := tau (xi - a xi *: eta1) + a xi *: Y1. - pose tau3_Y eta := tau (eta - eta1) + Y1. - have Itau2_X: {in X, isometry tau2_X, to 'Z[irr G]}. - split=> [xi1 xi2 Xxi1 Xxi2 | xi Xxi]; last first. - by rewrite rpredD ?rpredZ_Cint ?Za ?Ztau ?Zalpha. - rewrite /= cfdotDl !cfdotDr Itau ?Zalpha // cfdotBl !cfdotBr opprB !cfdotZr. - rewrite !aut_Cint ?Za // !cfdotZl (cfdotC Y1) !alphaY // n1Y1 n1eta rmorphN. - rewrite aut_Cint // (cfdotC eta1) !(orthogonalP oXY _ eta1) // conjC0. - by rewrite !mulr0 !subr0 !mulr1 !mulrN mulrC !addrA subrK addrK. - have{Itau2_X} [tau2 Dtau2 Itau2] := Zisometry_of_iso oX Itau2_X. - have{Itau2} cohX: coherent_with X L^# tau tau2. - split=> // xi; rewrite zcharD1E => /andP[/zchar_expansion[]// z Zz ->{xi}]. - pose sum_za := \sum_(xi <- X) z xi * a xi => /eqP sum_xi_0. - have{sum_xi_0} sum_za_0: sum_za = 0. - apply: (mulIf (neq0CG W1)); rewrite mul0r -{}sum_xi_0 sum_cfunE mulr_suml. - by apply: eq_big_seq => xi /Dxa[xi_1 _]; rewrite !cfunE xi_1 mulrA. - rewrite -[rhs in tau rhs](subrK (sum_za *: eta1)) {1}scaler_suml -sumrB. - rewrite -[tau _](addrK (sum_za *: Y1)) {1 3}sum_za_0 !scale0r addr0 subr0. - rewrite scaler_suml !raddf_sum [tau _]raddf_sum -big_split /= -/tau. - apply: eq_big_seq => xi Xxi; rewrite raddfZ_Cint //= Dtau2 //. - by rewrite -!scalerA -scalerBr [tau _]linearZ -scalerDr. - have Itau3_Y: {in Y, isometry tau3_Y, to 'Z[irr G]}. - split=> [eta3 eta4 Yeta3 Yeta4 | eta Yeta]; last first. - by rewrite rpredD // Ztau ?Zbeta. - rewrite /= cfdotDl !cfdotDr n1Y1 (cfdotC Y1) !betaY // Itau ?Zbeta //. - rewrite cfdotBl !cfdotBr !oYY // eqxx rmorphB rmorph1 rmorph_nat subrK. - rewrite (eq_sym eta1) opprB !addrA 3!(addrAC _ (- _)) addrK. - by rewrite(addrAC _ 1) subrK addrK. - have{oY} oY := orthonormal_orthogonal oY. - have{Itau3_Y} [tau3 Dtau3 Itau3] := Zisometry_of_iso oY Itau3_Y. - have{Itau3 cohY} cohY: coherent_with Y L^# tau tau3. - split=> // eta; rewrite zcharD1E => /andP[/zchar_expansion[]//z Zz ->{eta}]. - pose sum_z := \sum_(eta <- Y) z eta => /eqP sum_eta_0. - have{sum_eta_0} sum_z_0: sum_z = 0. - apply: (mulIf (neq0CG W1)); rewrite mul0r -sum_eta_0 sum_cfunE mulr_suml. - by apply: eq_big_seq => xi /uniY eta_1; rewrite !cfunE eta_1. - rewrite -[rhs in tau rhs](subrK (sum_z *: eta1)) {1}scaler_suml -sumrB. - rewrite -[tau _](addrK (sum_z *: Y1)) {1 3}sum_z_0 !scale0r addr0 subr0. - rewrite scaler_suml !raddf_sum [tau _]raddf_sum -big_split /= -/tau. - apply: eq_big_seq => eta Yeta; rewrite raddfZ_Cint //= Dtau3 //. - by rewrite -scalerBr [tau _]linearZ -scalerDr. - have [-> | ] := altP (@nilP _ X); first by exists tau3. - rewrite -lt0n -has_predT => /hasP[xi1 Xxi1 _]. - have: tau (xi1 - a xi1 *: eta1) = tau2 xi1 - tau3 (a xi1 *: eta1). - rewrite [tau3 _]linearZ Dtau2 //= Dtau3 // /tau3_Y subrr [tau 0]linear0. - by rewrite add0r addrK. - apply: (bridge_coherent scohS ccsXS cohX ccsYS cohY X'Y). - by rewrite (zchar_on (Zalpha _ Xxi1)) // rpredZ_Cint ?mem_zchar. -pose wf S1 := cfConjC_subset S1 S /\ {subset X ++ Y <= S1}. + apply: pivot_coherence scohXY XYeta1 Z_Y1 _ _; rewrite norm_eta1 //. + move=> xi /andP[eta1'xi]; rewrite /= mem_cat => /orP[Xxi | Yxi]. + have [Da1 [X1 oX1Y tauX1]] := Dxa _ Xxi. + exists (a xi); first by rewrite (uniY _ Yeta1). + rewrite -/tau {}tauX1 cfdotBl cfdotZl normY1 !mulr1. + by rewrite (span_orthogonal oX1Y) ?add0r ?memv_span1. + exists 1; first by rewrite rpred1 mul1r !uniY. + rewrite scale1r mulr1 -/tau -Dtau1 ?raddfB ?cfdotBl; last first. + by rewrite zcharD1E rpredB ?mem_zchar //= !cfunE !uniY ?subrr. + have [-> | [szY2 ->]] := defY1; rewrite ?cfdotNr !Itau1 ?mem_zchar ?ccY //. + by rewrite !oYY // eqxx (negPf eta1'xi) add0r. + pose Y2 := eta1 :: eta1^*%CF; suffices: xi \in Y2. + rewrite opprK !inE (negPf eta1'xi) /= => /eqP->. + by rewrite !oYY ?ccY // !mulrb eqxx ifN_eqC ?(hasPn nrS) ?sYS ?addr0. + have /leq_size_perm: {subset Y2 <= Y} by apply/allP; rewrite /= Yeta1 ccY. + by case=> [||->]; rewrite ?szY2 //= inE eq_sym (hasPn nrS) ?sYS. pose S1 := [::] ++ X ++ Y; set S2 := [::] in S1; rewrite -[X ++ Y]/S1 in cohXY. -have wfS1: wf S1. - do 2!split=> //; rewrite /S1 /= ?cat_uniq ?uX ?uY ?(introT hasPn) //. +have ccsS1S: cfConjC_subset S1 S. + rewrite /S1 /=; split; first by rewrite cat_uniq uX uY andbT; apply/hasPn. by apply/allP; rewrite all_cat !(introT allP). - by move=> phi; rewrite !mem_cat => /orP[/ccX|/ccY]->; rewrite ?orbT. -move: {2}_.+1 (ltnSn (size S - size S1)) => n. -elim: n => // n IHn in (S2) S1 wfS1 cohXY *; rewrite ltnS => leSnS1. -have [ccsS1S sXYS1] := wfS1; have [uS1 sS1S ccS1] := ccsS1S. -have [sSS1 | /allPn[psi /= Spsi notS1psi]] := altP (@allP _ (mem S1) S). - exact: subset_coherent cohXY. -have [_ _ ccS] := uccS. -have [neq_psi_c Spsic] := (hasPn nrS _ Spsi, ccS _ Spsi). -have wfS1': wf [:: psi, psi^* & S1]%CF. - split; last by move=> xi XYxi; rewrite !inE sXYS1 ?orbT. - split=> [|xi|xi]; rewrite /= !inE. - - by rewrite negb_or eq_sym neq_psi_c notS1psi (contra (ccS1 _)) ?cfConjCK. - - by do 2?[case/predU1P=> [-> //|]] => /sS1S. - rewrite (inv_eq (@cfConjCK _ _)) (can_eq (@cfConjCK _ _)) orbCA !orbA. - by case: pred2P => // _ /ccS1. -apply: (IHn [:: psi, psi^* & S2]%CF) => //; last first. - rewrite -subSn ?uniq_leq_size //; try by have [[]] := wfS1'. - by rewrite /= subSS (leq_trans _ leSnS1) // leq_sub2l ?leqW. -have ltZH': Z \proper H^`(1)%g. - rewrite properEneq sZH' andbT; apply: contraNneq notS1psi => eqZH'. - have [i /setDP[_ nt_i] ->] := seqIndP Spsi; rewrite sXYS1 // mem_cat. - rewrite !mem_seqInd ?normal1 //= -eqZH' !inE in nt_i *. - by rewrite sub1G nt_i andbT orNb. -have: [/\ eta1 \in S1, psi \in S & psi \notin S1]. - by rewrite Spsi sXYS1 // mem_cat Yeta1 orbT. + by move=> xi; rewrite !mem_cat => /orP[/ccX|/ccY]->; rewrite ?orbT. +move: {2}_.+1 (leq_addr (size S1) (size S).+1) => n. +elim: n => // [|n IHn] in (S2) S1 ccsS1S cohXY * => lb_n. + by rewrite ltnNge ?uniq_leq_size // in lb_n; have [] := ccsS1S. +have sXYS1: {subset X ++ Y <= S1} by apply/mem_subseq/suffix_subseq. +without loss /allPn[psi /= Spsi notS1psi]: / ~~ all (mem S1) S. + by case: allP => [/subset_coherent-cohS _ | _ cohS]; apply: cohS. +apply: (IHn [:: psi, psi^* & S2]%CF) => [|{lb_n}|]; last by rewrite !addnS leqW. + by have [_ _ ccS] := uccS; apply: extend_cfConjC_subset. have /seqIndC1P[i nzi Dpsi] := Spsi. -move/(extend_coherent scohS ccsS1S); apply; split=> //. - rewrite (uniY _ Yeta1) Dpsi cfInd1 // (index_sdprod defL) dvdC_mulr //. +have ltZH': Z \proper H'. + rewrite properEneq (contraNneq _ notS1psi) // => eqZH'; apply: sXYS1. + rewrite mem_cat Dpsi !mem_seqInd ?normal1 //. + by rewrite !inE sub1G andbT subGcfker nzi eqZH' orNb. +have Seta1: eta1 \in S1 by rewrite !mem_cat Yeta1 !orbT. +apply: (extend_coherent scohS ccsS1S Seta1) => {Seta1}//; split=> //. + rewrite (uniY _ Yeta1) Dpsi cfInd1 // oW1 dvdC_mulr //. by rewrite Cint_Cnat ?Cnat_irr1. -rewrite !big_cat //= (big_rem _ Yeta1) /= addrC -!addrA -big_cat //=. -rewrite sum_seqIndD_square ?normal1 ?sub1G // indexg1 addrC. -rewrite -subr_gt0 -!addrA ltr_spaddl //. - have /irrY/irrP[j ->] := Yeta1. - by rewrite cfnorm_irr divr1 exprn_gt0 ?irr1_gt0. -rewrite big_seq addr_ge0 ?sumr_ge0 // => [phi Sphi|]. - rewrite mulr_ge0 ?invr_ge0 ?cfnorm_ge0 ?exprn_ge0 // ?char1_pos //. - suffices /seqInd_char: phi \in S by apply: char1_ge0. - rewrite sS1S // !mem_cat; rewrite mem_cat in Sphi. - by case/orP: Sphi => [/mem_rem-> | ->]; rewrite ?orbT. -rewrite subr_ge0 -(Lagrange_index sHL sZH) -oW1 natrM mulrC -mulrA. -rewrite uniY ?ler_pmul2l ?gt0CG //. -rewrite -(prednK (cardG_gt0 Z)) [zz in zz - 1]mulrSr addrK -natrM. -rewrite Dpsi cfInd1 // -oW1. -rewrite -(@ler_pexpn2r _ 2) -?topredE /= ?mulr_ge0 ?ler0n //; last first. - by rewrite char1_ge0 ?irr_char. -rewrite !exprMn -!natrX mulrA -natrM. -apply: ler_trans (_ : (4 * #|W1| ^ 2)%:R * #|H : Z|%:R <= _). - rewrite ler_pmul2l; last by rewrite ltr0n muln_gt0 !expn_gt0 cardG_gt0. - rewrite (ler_trans (irr1_bound i)) // ler_nat dvdn_leq // indexgS //. - by rewrite (subset_trans sZZ) // -cap_cfcenter_irr bigcap_inf. -rewrite -natrM ler_nat expnMn mulnC -mulnA leq_pmul2l //. +rewrite !big_cat /= addrCA sum_seqIndD_square ?normal1 ?sub1G // ltr_spaddr //. + have /irrY/irrP[j Deta1] := Yeta1; have [_ sS1S _] := ccsS1S. + rewrite (big_rem eta1 Yeta1) addrCA -big_cat big_seq ltr_spaddl //=. + by rewrite Deta1 cfnorm_irr divr1 exprn_gt0 ?irr1_gt0. + apply/sumr_ge0=> phi YS2phi; rewrite divr_ge0 ?cfnorm_ge0 ?exprn_ge0 //. + rewrite char1_ge0 ?(seqInd_char (sS1S _ _)) //. + by move: YS2phi; rewrite !mem_cat => /orP[-> | /mem_rem->]; rewrite ?orbT. +rewrite indexg1 -(Lagrange_index sHL sZH) -oW1 natrM mulrC -mulrA. +rewrite uniY ?ler_wpmul2l ?ler0n -?(@natrB _ _ 1) // -natrM. +suffices ubW1: (#|W1|.*2 ^ 2 <= #|H : Z| * (#|Z| - 1) ^ 2)%N. + have chi1_ge0: 0 <= 'chi_i 1%g by rewrite char1_ge0 ?irr_char. + rewrite Dpsi cfInd1 // -oW1 -(@ler_pexpn2r _ 2) ?rpredM ?rpred_nat //. + rewrite -natrX expnMn mulnAC natrM mulrA -natrM exprMn -natrX mul2n. + rewrite ler_pmul ?ler0n ?exprn_ge0 ?(ler_trans (irr1_bound i)) ?ler_nat //. + rewrite dvdn_leq ?indexgS ?(subset_trans sZZH) //=. + by rewrite -cap_cfcenter_irr bigcap_inf. +have nZW1 := subset_trans sW1L nZL. +have tiZW1: Z :&: W1 = 1%g by rewrite coprime_TIg ?(coprimeSg sZH). have [in_caseA | in_caseB] := boolP caseA. - have regW1Z: semiregular Z W1. - have [in_c1 | in_c2] := boolP case_c1. - move=> x /(Frobenius_reg_ker in_c1) regHx; apply/trivgP. - by rewrite -regHx setSI. - have [/= _ [_ [_ _ _ prW1H] _] _ _] := c2W2 in_c2. - move=> x /prW1H prHx; apply/trivgP; rewrite -((_ =P [1]) in_caseA) -prHx. - by rewrite subsetI subIset ?sZZ // setSI. - rewrite -(mul1n (4 * _)%N) leq_mul // -(expnMn 2 _ 2) leq_exp2r //. - rewrite dvdn_leq //; first by rewrite -subn1 subn_gt0 cardG_gt1. - rewrite Gauss_dvd ?(@pnat_coprime 2) -?odd_2'nat ?(oddSg sW1L) //. - rewrite dvdn2 -{1}subn1 odd_sub // (oddSg (normal_sub nsZL)) //=. - by rewrite regular_norm_dvd_pred // (subset_trans sW1L) ?normal_norm. -rewrite -(muln1 (4 * _)%N) leq_mul //; last first. - by rewrite expn_gt0 -subn1 subn_gt0 orbF cardG_gt1. -rewrite -(expnMn 2 _ 2) -(Lagrange_index (der_sub 1 H) sZH') leq_mul //. - rewrite -(prednK (indexg_gt0 _ _)) leqW // dvdn_leq //. - by rewrite -subn1 subn_gt0 indexg_gt1 proper_subn. - rewrite Gauss_dvd ?(@pnat_coprime 2) -?odd_2'nat ?(oddSg sW1L) //. - rewrite dvdn2 -{1}subn1 odd_sub // -card_quotient ?der_norm //. - rewrite quotient_odd ?(oddSg sHL) //=. - rewrite (card_isog (quotient_isog (subset_trans sW1L nH'L) _)); last first. - by rewrite coprime_TIg ?(coprimeSg (der_sub 1 H)). - exact: Frobenius_dvd_ker1 frobL1. -rewrite -(prednK (indexg_gt0 _ _)) leqW // dvdn_leq //. - by rewrite -subn1 subn_gt0 indexg_gt1 proper_subn. -rewrite Gauss_dvd ?(@pnat_coprime 2) -?odd_2'nat ?(oddSg sW1L) //. -rewrite dvdn2 -{1}subn1 odd_sub //. -rewrite -card_quotient ?(subset_trans (der_sub 1 H)) //. -rewrite quotient_odd ?(oddSg (der_sub 1 H)) ?(oddSg sHL) //=. -have /andP[sZL nZL] := nsZL. -rewrite (card_isog (quotient_isog (subset_trans sW1L nZL) _)); last first. - by rewrite coprime_TIg ?(coprimeSg sZH). -suffices: [Frobenius (H^`(1) / Z) <*> (W1 / Z) = (H^`(1) / Z) ><| (W1 / Z)]%g. - exact: Frobenius_dvd_ker1. -suffices: [Frobenius (L / Z) = (H / Z) ><| (W1 / Z)]%g. - apply: Frobenius_subl (quotientS Z (der_sub 1 H)) _. - by rewrite quotient_neq1 // (normalS sZH' (der_sub 1 H)). + rewrite (leq_trans _ (leq_pmull _ _)) ?leq_exp2r // subn1 -ltnS prednK //. + suffices frobZW1: [Frobenius Z <*> W1 = Z ><| W1]. + by apply: ltn_odd_Frobenius_ker frobZW1 (oddSg _ oddL); apply/joing_subP. + have [|/c2_ptiL[_ _ prW1H _]] := boolP case_c1; first exact: Frobenius_subl. + apply/Frobenius_semiregularP; rewrite ?sdprodEY // => x W1x; apply/trivgP. + by rewrite /= -(setIidPl sZH) -setIA -(trivgP in_caseA) prW1H ?setSI. +rewrite (leq_trans _ (leq_pmulr _ _)) ?expn_gt0 ?orbF ?subn_gt0 ?cardG_gt1 //. +rewrite -(Lagrange_index sH'H sZH') leq_mul // ltnW //. + have tiH'W1: H' :&: W1 = 1%g by rewrite coprime_TIg ?(coprimeSg sH'H). + rewrite (card_isog (quotient_isog (subset_trans sW1L nH'L) tiH'W1)). + rewrite -card_quotient ?gFnorm // (ltn_odd_Frobenius_ker frobL1) //. + exact: quotient_odd. +suffices frobHW1Z: [Frobenius (H' / Z) <*> (W1 / Z) = (H' / Z) ><| (W1 / Z)]. + rewrite (card_isog (quotient_isog nZW1 tiZW1)). + rewrite -card_quotient ?(subset_trans sH'H) //. + apply: ltn_odd_Frobenius_ker frobHW1Z (oddSg _ (quotient_odd Z oddL)). + by rewrite join_subG !quotientS. +suffices: [Frobenius (L / Z) = (H / Z) ><| (W1 / Z)]. + apply: Frobenius_subl (quotientS Z sH'H) _. + by rewrite quotient_neq1 // (normalS sZH' sH'H). by rewrite quotient_norms ?(subset_trans sW1L). -apply: (Frobenius_coprime_quotient defL nsZL) => //; split=> [|x W1x]. - exact: sub_proper_trans sZH' ltH'H. -have /caseB_P[/c2W2[_ [_ [_ _ _ -> //] _] _ _] _ _] := in_caseB. -by rewrite /Z (negPf in_caseB). +apply: (Frobenius_coprime_quotient defL nsZL) => //. +split=> [|x W1x]; first exact: sub_proper_trans sZH' ltH'H. +by rewrite /Z ifN //; have /caseB_P[/c2_ptiL[_ _ ->]] := in_caseB. Qed. End Six. diff --git a/mathcomp/odd_order/PFsection7.v b/mathcomp/odd_order/PFsection7.v index eed77b7..dcf7f8f 100644 --- a/mathcomp/odd_order/PFsection7.v +++ b/mathcomp/odd_order/PFsection7.v @@ -1,10 +1,18 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime ssralg poly finset center. +From mathcomp Require Import fingroup morphism perm automorphism quotient action zmodp. +From mathcomp Require Import gfunctor gproduct cyclic pgroup commutator nilpotent frobenius. +From mathcomp Require Import matrix mxalgebra mxrepresentation BGsection3 vector. +From mathcomp Require Import ssrnum algC classfun character inertia vcharacter. +From mathcomp Require Import PFsection1 PFsection2 PFsection4 PFsection5 PFsection6. (******************************************************************************) @@ -149,7 +157,7 @@ Qed. (* This is the norm expansion embedded in Peterfalvi (7.3). *) Lemma cfnormE_invDade chi : '[chi^\rho] = #|L|%:R^-1 * (\sum_(a in A) `|chi^\rho a| ^+ 2). -Proof. by apply: cfnormE; exact: invDade_on. Qed. +Proof. by apply: cfnormE; apply: invDade_on. Qed. End InverseDade. @@ -365,7 +373,7 @@ have nz_phi1: phi 1%g != 0 by rewrite (seqInd1_neq0 nsHL Sphi). have NatS1e xi (Sxi : xi \in calS) := dvd_index_seqInd1 nsHL Sxi. have oS1: {in calS, forall psi, '[psi, 1] = 0} by apply: seqInd_ortho_1. have oS1H: {in calS, forall psi, '[psi, Ind1H] = 0} by apply: seqInd_ortho_Ind1. -have InuS: {in calS &, isometry nu} by apply: sub_in2 Inu; exact: seqInd_zcharW. +have InuS: {in calS &, isometry nu} by apply: sub_in2 Inu; apply: seqInd_zcharW. have ZnuS xi (Sxi : xi \in calS) := Znu xi (seqInd_zcharW Sxi). have S_Se xi (Sxi : xi \in calS) := seqInd_sub_lin_vchar nsHL Szeta zeta1 Sxi. have oSnu1: orthogonal calSnu 1%CF. @@ -675,7 +683,7 @@ have /fin_all_exists[r lin_r] i: exists r, 'chi_r \in S i /\ 'chi_r 1%g = e_ i. have [xi Sxi lin_xi] := exists_linInd (nsHL i) solHi lt1Hi (normal1 _). by have /irrP[r def_xi] := irrS i xi Sxi; exists r; rewrite -def_xi. have{lin_r} [Sr r1] := all_and2 lin_r. -have ntS i: (size (S i) > 1)%N by apply: seqInd_nontrivial (mem_irr _) (Sr i). +have ntS i: (size (S i) > 1)%N by apply: seqInd_nontrivial (Sr i). have /fin_all_exists[nu cohS] i: coherent (S i) (L i)^# 'Ind[G, L i]. have [[[frobLi tiAiL] sLiG] oddLi] := (frobL i, normedTI_A i, sLG i, oddL i). have [defLi ntHi ntEi _ _] := Frobenius_context frobLi. @@ -797,8 +805,8 @@ rewrite ler_wpmul2r ?cfnorm_ge0 // (@ler_trans _ 1) //. by rewrite -{2}(mulVf (nzh i)) ler_wpmul2l ?invr_ge0 ?ler0n ?min_i1. rewrite mulrC -normCK expr_ge1 ?normr_ge0 // norm_Cint_ge1 //. rewrite Cint_cfdot_vchar ?Znu ?seqInd_zcharW ?Sr //. -suffices /orP: c i i1 != 0 \/ c i1 i != 0 by rewrite ci1_0. -apply: Dade_sub_lin_nonorthogonal; rewrite ?mem_irr ?Sr ?r1 //; try exact: cohS. +suffices []: c i i1 != 0 \/ c i1 i != 0 by rewrite ?ci1_0. +apply/Dade_sub_lin_nonorthogonal; rewrite ?mem_irr ?Sr ?r1 //; try exact: cohS. exact: disjoint_Atau. Qed. diff --git a/mathcomp/odd_order/PFsection8.v b/mathcomp/odd_order/PFsection8.v index 72a0d00..4c6f14d 100644 --- a/mathcomp/odd_order/PFsection8.v +++ b/mathcomp/odd_order/PFsection8.v @@ -1,14 +1,25 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime ssralg poly finset center. +From mathcomp Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +From mathcomp Require Import gfunctor gproduct cyclic commutator nilpotent pgroup. +From mathcomp Require Import sylow hall abelian maximal frobenius. +From mathcomp Require Import matrix mxalgebra mxrepresentation vector. +From mathcomp Require Import BGsection1 BGsection3 BGsection7 BGsection10. +From mathcomp Require Import BGsection14 BGsection15 BGsection16. Require ssrnum. +From mathcomp Require Import algC classfun character inertia vcharacter. +From mathcomp Require Import PFsection1 PFsection2 PFsection3 PFsection4 PFsection5. (******************************************************************************) @@ -245,7 +256,7 @@ have piU0p: p \in \pi(U0). rewrite -pi_of_exponent expU0 pi_of_exponent (pi_of_dvd _ _ piW1p) //=. rewrite -(@dvdn_pmul2l #|H|) ?cardG_gt0 // (sdprod_card defM_K). rewrite -(sdprod_card defM) dvdn_pmul2r ?cardSg //. - by case/sdprodP: defM' => _ <- _ _; exact: mulG_subl. + by case/sdprodP: defM' => _ <- _ _; apply: mulG_subl. have [|X EpX]:= @p_rank_geP _ p 1 U0 _; first by rewrite p_rank_gt0. have [ntX [sXU0 abelX _]] := (nt_pnElem EpX isT, pnElemP EpX). have piW1_X: \pi(W1).-group X by apply: pi_pgroup piW1p; case/andP: abelX. @@ -307,7 +318,7 @@ have ctiW: cyclicTI_hypothesis G defW by split; rewrite ?mFT_odd. split=> //; first by rewrite dprodE ?subsetIr //= setIA tiHU setI1g. split. apply: subset_trans (_ : U :&: 'F(M) \subset _). - by rewrite subsetI der_sub (subset_trans (dergS 1 sUM')). + by rewrite subsetI gFsub (subset_trans (dergS 1 sUM')). by rewrite -defF -group_modr ?subsetIl // setIC tiHU mul1g subsetIr. apply: contra => cHU; rewrite -subG1 -tiHU subsetIidr (subset_trans sUM') //. by rewrite (Fcore_max hallM') ?der_normal // -mulHU mulg_nil ?Fcore_nil. @@ -512,7 +523,7 @@ split=> // [|S /SylowP[p _ sylS] ntS]. exact: pHall_subl (pcore_sub _ M) (Fcore_Hall M). have s_p: p \in \sigma(M). by rewrite (pnatPpi sMs) // -p_rank_gt0 -(rank_Sylow sylS) rank_gt0. -by apply: (norm_sigma_Sylow s_p); exact: (subHall_Sylow (Msigma_Hall maxM)). +by apply: (norm_sigma_Sylow s_p); apply: (subHall_Sylow (Msigma_Hall maxM)). Qed. (* This is Peterfalvi (8.12). *) @@ -601,7 +612,7 @@ Let is_FTsignalizer : is_Dade_signalizer G M 'A0(M) 'R_M. Proof. rewrite /'R_M => x A0x /=; rewrite setTI. case: ifPn => [sCxM | not_sCxM]; first by rewrite sdprod1g (setIidPr sCxM). -by have [_ _ /(_ x)[| [] //]] := FTsupport_facts; exact/setIdP. +by have [_ _ /(_ x)[| [] //]] := FTsupport_facts; apply/setIdP. Qed. (* This is Peterfalvi (8.15), second assertion. *) @@ -957,15 +968,15 @@ have{PG} ->: PG = [set class_support M^~~ G | M : {group gT} in 'M]. have [c1 c2] := mFT_partition gT. split=> [M H maxM maxH eq_MH | Gtype1 | S T W W1 W2 defW VG pairST]. - apply: injMG => //; move/sMG_M in maxM; move/sMG_M in maxH. - apply/orbit_transl/idPn => not_HG_M. + apply/orbit_eqP/idPn => not_HG_M. have /negP[]: ~~ [disjoint 'A1~(M) & 'A1~(H)]. rewrite eq_MH -setI_eq0 setIid -defDsup //. - by apply: contraNneq not_PG_set0 => <-; exact: mem_imset. + by apply: contraNneq not_PG_set0 => <-; apply: mem_imset. rewrite -!defDsup // -setI_eq0 class_supportEr big_distrl -subset0. apply/bigcupsP=> x /class_supportGidr <- /=; rewrite -conjIg sub_conjg conj0g. rewrite class_supportEr big_distrr /=; apply/bigcupsP=> {x}x _. rewrite subset0 setI_eq0 -sigma_supportJ sigma_support_disjoint ?mmaxJ //. - by rewrite (orbit_transr _ (mem_orbit _ _ _)) ?in_setT // orbit_sym. + by rewrite (orbit_transl _ (mem_orbit _ _ _)) ?in_setT // orbit_sym. - rewrite c1 // setD_eq0; apply/subsetP=> M maxM. by rewrite FTtype_Fmax ?(forall_inP Gtype1). have [[[cycW maxS _] _ _ _ _] [U_S StypeP]] := (pairST, typeP_pairW pairST). @@ -1035,16 +1046,16 @@ have part_a2 S T (maxS : S \in 'M) (maxT : T \in 'M) (ncST : NC S T) : by rewrite sub1set !inE ntx. by apply/trivgPn; exists z; rewrite //= -defTs inE Ts_z cent_set1 cent1C. split; last by case/mem_uniq_mmax: uniqCx. - by apply: contra ncST => /(eq_uniq_mmax uniqCx maxS)->; exact: orbit_refl. + by apply: contra ncST => /(eq_uniq_mmax uniqCx maxS)->; apply: orbit_refl. have part_a1 S T (maxS : S \in 'M) (maxT : T \in 'M) (ncST : NC S T) : FTsupports S T = ~~ [disjoint 'A1(S) & 'A(T)]. - apply/existsP/pred0Pn=> [[x /and3P[ASx not_sCxS sCxT]] | [x /andP[A1Sx Atx]]]. have [_ [/subsetP]] := FTsupport_facts maxS; set D := finset _. have Dx: x \in D by rewrite !inE ASx. move=> /(_ x Dx) A1x /(_ x Dx)uniqCx /(_ x Dx)[_ _ /setDP[ATx _] _]. - by rewrite (eq_uniq_mmax uniqCx maxT sCxT); exists x; exact/andP. + by rewrite (eq_uniq_mmax uniqCx maxT sCxT); exists x; apply/andP. exists x; rewrite (subsetP (FTsupp1_sub maxS)) //=. - by apply/andP/part_a2=> //; exact/setIP. + by apply/andP/part_a2=> //; apply/setIP. have part_b S T (maxS : S \in 'M) (maxT : T \in 'M) (ncST : NC S T) : [exists x, FTsupports S (T :^ x)] = ~~ [disjoint 'A1~(S) & 'A~(T)]. - apply/existsP/pred0Pn=> [[x] | [y /andP[/= A1GSy AGTy]]]. @@ -1074,7 +1085,7 @@ have part_b S T (maxS : S \in 'M) (maxT : T \in 'M) (ncST : NC S T) : rewrite class_supportGidr ?inE {z}//. case/imset2P=> _ z /rcosetP[y Hy ->] _ def_x2. exists z^-1%g; rewrite part_a1 ?mmaxJ //; last first. - by rewrite /NC (orbit_transr _ (mem_orbit _ _ _)) ?inE. + by rewrite /NC (orbit_transl _ (mem_orbit _ _ _)) ?inE. apply/pred0Pn; exists x1; rewrite /= A1Sx1 FTsuppJ mem_conjgV; apply/bigcupP. pose ddS := FT_Dade1_hyp maxS; have [/andP[sA1S _] _ notA1_1 _ _] := ddS. have [ntx1 Sx1] := (memPn notA1_1 _ A1Sx1, subsetP sA1S _ A1Sx1). @@ -1093,7 +1104,7 @@ move=> S T maxS maxT ncST; split; first split; auto. apply/orP/idPn; rewrite negb_or -part_b // => /andP[suppST /negP[]]. without loss{suppST} suppST: T maxT ncST / FTsupports S T. move=> IH; case/existsP: suppST => x /IH {IH}. - rewrite FT_Dade1_supportJ (orbit_transr _ (mem_orbit _ _ _)) ?in_setT //. + rewrite FT_Dade1_supportJ (orbit_transl _ (mem_orbit _ _ _)) ?in_setT //. by rewrite mmaxJ => ->. have{suppST} [y /and3P[ASy not_sCyS sCyT]] := existsP suppST. have Dy: y \in [set z in 'A0(S) | ~~ ('C[z] \subset S)] by rewrite !inE ASy. @@ -1101,7 +1112,7 @@ have [_ [_ /(_ y Dy) uCy] /(_ y Dy)[_ coTcS _ typeT]] := FTsupport_facts maxS. rewrite -mem_iota -(eq_uniq_mmax uCy maxT sCyT) !inE in coTcS typeT. apply/negbNE; rewrite -part_b /NC 1?orbit_sym // negb_exists. apply/forallP=> x; rewrite part_a1 ?mmaxJ ?negbK //; last first. - by rewrite /NC (orbit_transr _ (mem_orbit _ _ _)) ?in_setT // orbit_sym. + by rewrite /NC (orbit_transl _ (mem_orbit _ _ _)) ?in_setT // orbit_sym. rewrite -setI_eq0 -subset0 FTsuppJ -bigcupJ big_distrr; apply/bigcupsP=> z Sxz. rewrite conjD1g /= -setDIl coprime_TIg ?setDv //= cardJg. rewrite -(Fcore_eq_FTcore maxT _) ?inE ?orbA; last by have [->] := typeT. diff --git a/mathcomp/odd_order/PFsection9.v b/mathcomp/odd_order/PFsection9.v index 361d5fe..c2e28bb 100644 --- a/mathcomp/odd_order/PFsection9.v +++ b/mathcomp/odd_order/PFsection9.v @@ -1,13 +1,24 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div choice. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp Require Import fintype tuple finfun bigop prime binomial ssralg poly finset. +From mathcomp Require Import fingroup morphism perm automorphism quotient action finalg zmodp. +From mathcomp Require Import gfunctor gproduct cyclic commutator center gseries nilpotent. +From mathcomp Require Import pgroup sylow hall abelian maximal frobenius. +From mathcomp Require Import matrix mxalgebra mxrepresentation mxabelem vector. +From mathcomp Require Import BGsection1 BGsection3 BGsection7 BGsection15 BGsection16. +From mathcomp Require Import algC classfun character inertia vcharacter. +From mathcomp Require Import PFsection1 PFsection2 PFsection3 PFsection4. +From mathcomp Require Import PFsection5 PFsection6 PFsection8. (******************************************************************************) @@ -245,14 +256,11 @@ have nsH0C: H0C <| M. rewrite -quotientYK // -{1}(quotientGK nsH0H) morphpre_norms //= [C]unlock. by rewrite cents_norm // centsC -quotient_astabQ quotientS ?subsetIr. split=> //; first by rewrite /= -{1}(joing_idPl sH0H) -joingA normalY ?gFnormal. - rewrite normalY // /normal (subset_trans (der_sub 1 U)) //=. - rewrite -{1}defM sdprodEY //= -defHU sdprodEY //=. - rewrite !join_subG gFnorm cents_norm 1?centsC //=. - by rewrite (char_norm_trans (der_char _ _)). -suffices ->: H0C' :=: H0 <*> H0C^`(1). - by rewrite normalY ?(char_normal_trans (der_char _ _)). -rewrite /= -?quotientYK ?(subset_trans (der_sub _ _)) ?subsetIl //=. -by rewrite !quotient_der ?cosetpreK ?subsetIl. + rewrite normalY // /normal gFsub_trans //=. + rewrite -defM sdprodEY //= -defHU sdprodEY //=. + by rewrite !join_subG gFnorm cents_norm ?gFnorm_trans // centsC. +suffices ->: H0C' :=: H0 <*> H0C^`(1) by rewrite normalY ?gFnormal_trans. +by rewrite /= -!quotientYK ?gFsub_trans ?quotient_der ?subsetIl //= cosetpreK. Qed. Local Notation nsH0xx_M := Ptype_Fcore_extensions_normal. @@ -537,9 +545,9 @@ have{cEE} [F [outF [inF outFK inFK] E_F]]: have outK: cancel outF inF by move=> a; apply: outI; rewrite inK ?E_F. pose one := inF 1%R; pose mul a b := inF (outF a * outF b)%R. have outM: {morph outF: a b / mul a b >-> a * b}%R. - by move=> a b; rewrite inK //; apply: envelop_mxM; exact: E_F. + by move=> a b; rewrite inK //; apply: envelop_mxM; apply: E_F. have out0: outF 0%R = 0%R by apply: raddf0. - have out1: outF one = 1%R by rewrite inK //; exact: envelop_mx1. + have out1: outF one = 1%R by rewrite inK //; apply: envelop_mx1. have nzFone: one != 0%R by rewrite -(inj_eq outI) out1 out0 oner_eq0. have mulA: associative mul by move=> *; apply: outI; rewrite !{1}outM mulrA. have mulC: commutative mul. @@ -589,7 +597,7 @@ have inj_phi': injective phi'. move=> a b /rVabelem_inj eq_sab; apply: contraNeq nz_sb. rewrite -[sb]mulmx1 idmxE -(rmorph1 outF) -subr_eq0 => /divff <-. by rewrite rmorphM mulmxA !raddfB /= eq_sab subrr mul0mx. -have injm_phi': 'injm (Morphism phi'D) by apply/injmP; exact: in2W. +have injm_phi': 'injm (Morphism phi'D) by apply/injmP; apply: in2W. have Dphi: 'dom (invm injm_phi') = Hbar. apply/setP=> h; apply/morphimP/idP=> [[a _ _ ->] // | Hh]. have /cyclic_mxP[A E_A def_h]: (outHb h <= cyclic_mx rU sb)%MS. @@ -598,7 +606,7 @@ have Dphi: 'dom (invm injm_phi') = Hbar. have [phi [def_phi Kphi _ im_phi]] := domP _ Dphi. have{Kphi} inj_phi: 'injm phi by rewrite Kphi injm_invm. have{im_phi} im_phi: phi @* Hbar = setT by rewrite im_phi -Dphi im_invm. -have phiK: {in Hbar, cancel phi phi'} by rewrite def_phi -Dphi; exact: invmK. +have phiK: {in Hbar, cancel phi phi'} by rewrite def_phi -Dphi; apply: invmK. have{def_phi Dphi injm_phi'} phi'K: cancel phi' phi. by move=> a; rewrite def_phi /= invmE ?inE. have phi'1: phi' 1%R = s by rewrite /phi' rmorph1 mulmx1 [inHb _]abelem_rV_K. @@ -704,15 +712,15 @@ exists F. move=> fRM; suff <-: map_poly (RMorphism fRM) P = P by apply: rmorph_root. apply/polyP=> i; rewrite coef_map. have [/(nth_default _)-> | lt_i_P] := leqP (size P) i; first exact: rmorph0. - by have /cycleP[n ->] := all_nthP 0%R nP i lt_i_P; exact: rmorph_nat. + by have /cycleP[n ->] := all_nthP 0%R nP i lt_i_P; apply: rmorph_nat. apply: (iffP morphimP) => [[w _ Ww ->] | alphaRM]; first exact: etaRM. suffices /setP/(_ (alpha r)): [set (eta w) r | w in W1] = [set t | root P t]. rewrite inE fPr0 // => /imsetP[w Ww def_wr]; exists w => //. - by apply: prim_r => //; exact: etaRM. + by apply: prim_r => //; apply: etaRM. apply/eqP; rewrite eqEcard; apply/andP; split. - by apply/subsetP=> _ /imsetP[w Ww ->]; rewrite inE fPr0 //; exact: etaRM. + by apply/subsetP=> _ /imsetP[w Ww ->]; rewrite inE fPr0 //; apply: etaRM. rewrite (@cardsE F) card_in_imset // => w1 w2 Ww1 Ww2 /= /prim_r eq_w12. - by apply: (injmP inj_eta) => //; apply: eq_w12; exact: etaRM. + by apply: (injmP inj_eta) => //; apply: eq_w12; apply: etaRM. have isoUb: isog Ubar (psi @* U) by rewrite /Ubar -Kpsi first_isog. pose unF := [set in_uF a | a in nF^#]. have unF_E: {in nF^#, cancel in_uF val} by move=> a /setD1P[/in_uF_E]. @@ -780,13 +788,11 @@ have nb_redM K: case/seqIndP=> s /setDP[kerK ker'H] Dphi; rewrite !inE in kerK ker'H. pose s1 := quo_Iirr K s; have Ds: s = mod_Iirr s1 by rewrite quo_IirrK. rewrite {phi}Dphi Ds mod_IirrE ?cfIndMod // in kerK ker'H red_phi *. - have [[j Ds1] | [/idPn[]]] := prTIres_irr_cases ptiWMb s1. - rewrite Ds1 cfInd_prTIres -/(muK j) in ker'H *; exists j => //. - by apply: contraNneq ker'H => ->; rewrite prTIres0 rmorph1 cfker_cfun1. - by apply: contra red_phi => /cfMod_irr/= ->. - have red_j: redM (muK j). - apply: contra (prTIred_not_irr ptiWMb j) => /(cfQuo_irr nsKM). - by rewrite cfker_mod ?cfModK // => ->. + have{red_phi} red_s1: 'Ind 'chi_s1 \notin irr (M / K) by rewrite -cfMod_irr. + have [[j Ds1] | [/idPn//]] := prTIres_irr_cases ptiWMb s1. + rewrite Ds1 cfInd_prTIres -/(muK j) in ker'H *; exists j => //. + by apply: contraNneq ker'H => ->; rewrite prTIres0 rmorph1 cfker_cfun1. + have red_j: redM (muK j) by rewrite /redM /= cfMod_irr // prTIred_not_irr. have [s DmuKj]: exists s, muK j = 'Ind[M, HU] 'chi_s. exists (mod_Iirr (primeTI_Ires ptiWMb j)). by rewrite mod_IirrE // cfIndMod // cfInd_prTIres. @@ -901,7 +907,7 @@ have Dtheta f: {in W1bar & H1, forall w xb, theta f (xb ^ w) = 'chi_(f w) xb}. transitivity ('Res[H1 :^ w] ('Res[Hbar] (theta f)) (xb ^ w)); last first. by rewrite cfDprodlK cfBigdprodKabelian // isom_IirrE cfIsomE. by rewrite cfResRes ?sH1wH // cfResE ?memJ_conjg ?(subset_trans (sH1wH w _)). -have lin_theta f: theta f \is a linear_char by apply: cfDprodl_lin_char. +have lin_theta f: theta f \is a linear_char by rewrite cfDprodl_lin_char. pose Ftheta := pffun_on (0 : Iirr H1) W1bar (predC1 0). have inj_theta: {in Ftheta &, injective theta}. move=> f1 f2 /pffun_onP[/supportP W1f1 _] /pffun_onP[/supportP W1f2 _] eq_f12. @@ -1199,7 +1205,7 @@ have ->: #|Mtheta| = (#|Xtheta| * a)%N. by rewrite cfConjgMod_norm ?(subsetP _ _ Uyb) // quotient_norms ?gFnorm. rewrite leq_pmul2r ?indexg_gt0 // cardE -(size_map (fun s => 'Ind[M] 'chi_s)). have kerH1c s: s \in Xtheta -> H1c \subset (cfker 'chi_s / H0)%g. - case/imsetP=> r Mr ->; have [i j _ _ Dr] := imset2P Mr. + case/imsetP=> r Mr {s}->; have [i j _ _ Dr] := imset2P Mr. rewrite -(setIidPr (normal_sub nsH1cHCH1)) -morphim_setIpre quotientS //. rewrite cfIirrE ?irr_Xtheta ?sub_cfker_Ind_irr //; last first. by rewrite normsI ?normal_norm // -(quotientGK nsH0_HU) cosetpre_normal. @@ -1328,7 +1334,7 @@ have Part_a': part_a'. rewrite inE Ds sub_cfker_Ind_irr // in KsH0C'. by rewrite (subset_trans sHUM) ?normal_norm. rewrite lin_irr_der1 (subset_trans _ KrH0C') //= (norm_joinEr nH0C'). - rewrite -quotientSK ?(subset_trans (der_sub 1 _)) ?quotient_der //= -/C. + rewrite -quotientSK ?gFsub_trans ?quotient_der //= -/C. by rewrite -(der_dprod 1 defHCbar) (derG1P abHbar) dprod1g. split=> // [s /Part_a[r ->] | | {Part_a' part_a'}red_H0C']. - by rewrite Du cfInd1 // dvdC_mulr // Cint_Cnat ?Cnat_irr1. @@ -1632,14 +1638,14 @@ without loss [[eqS12 irrS1 H0C_S1] [Da_p defC] [S3qu ne_qa_qu] [oS1 oS1ua]]: exists chi => //; have /hasP[xi S1xi _]: has predT S1 by rewrite has_predT. have xi1: xi 1%g = (q * a)%:R. by rewrite mem_filter in S1xi; have [/eqP] := andP S1xi. - apply: ((extend_coherent scohS0) _ xi) => //; first by rewrite S0chi sS12. + apply: (extend_coherent scohS0 _ (sS12 _ S1xi)) => //. split=> //; last by rewrite mulrAC xi1 -natrM mulnA. rewrite xi1 Dchi1 irr1_degree -natrM dvdC_nat dvdn_pmul2l ?cardG_gt0 //. rewrite -dvdC_nat /= !nCdivE -irr1_degree a_dv_XH0 //. by rewrite (subsetP (Iirr_kerDS _ _ _) _ X0C's) ?joing_subl. have lb1S2 := lerif_trans lb12 (lerif_trans lb23 (lerif_trans lb3S1' lbS1'2)). rewrite ltr_neqAle !(lerif_trans lb01 lb1S2) andbT has_predC !negbK. - case/and5P=> /eqP chi1qu /eqP Da_p /eqP defC /eqP sz_S1' /allP sS21'. + case/and5P=> /eqP-chi1qu /eqP-Da_p /eqP-defC /eqP-sz_S1' /allP/=sS21'. have defS1': S1' = S1. apply/eqP; rewrite -(geq_leqif (size_subseq_leqif (filter_subseq _ S1))). by rewrite uniq_leq_size // => psi /sS12/sS21'. @@ -2112,9 +2118,8 @@ have [Gamma [S4_Gamma normGamma [b Dbeta]]]: by have [_ _ ->] := sS10. rewrite Dbeta -Dtau3 //; apply: contraNneq => ->. rewrite add0r raddfB cfdotBr !(orthoPl oG'4) ?map_f ?subr0 //. - rewrite mem_filter /= negbK /= S3lam1s irr_aut. - move: S4lam1; rewrite mem_filter /= negbK /= -andbA => /and3P[-> H0Clam1 _]. - by rewrite cfAut_seqInd. + move: S4lam1; rewrite ![_ \in S4]mem_filter /= !negbK /= cfAut_irr S3lam1s. + by case/andP=> /andP[-> /cfAut_seqInd->]. have ubG: '[G] + (b ^+ 2 - b) * (u %/ a).*2%:R + '[Delta] = 1. apply: (addrI ((u %/ a) ^ 2)%:R); transitivity '[beta^\tau]. rewrite -!addrA addrCA Dbeta cfnormDd; last first. diff --git a/mathcomp/odd_order/all.v b/mathcomp/odd_order/all.v deleted file mode 100644 index 613acb0..0000000 --- a/mathcomp/odd_order/all.v +++ /dev/null @@ -1,33 +0,0 @@ -Require Export BGappendixAB. -Require Export BGappendixC. -Require Export BGsection10. -Require Export BGsection11. -Require Export BGsection12. -Require Export BGsection13. -Require Export BGsection14. -Require Export BGsection15. -Require Export BGsection16. -Require Export BGsection1. -Require Export BGsection2. -Require Export BGsection3. -Require Export BGsection4. -Require Export BGsection5. -Require Export BGsection6. -Require Export BGsection7. -Require Export BGsection8. -Require Export BGsection9. -Require Export PFsection10. -Require Export PFsection11. -Require Export PFsection12. -Require Export PFsection13. -Require Export PFsection14. -Require Export PFsection1. -Require Export PFsection2. -Require Export PFsection3. -Require Export PFsection4. -Require Export PFsection5. -Require Export PFsection6. -Require Export PFsection7. -Require Export PFsection8. -Require Export PFsection9. -Require Export stripped_odd_order_theorem. diff --git a/mathcomp/odd_order/wielandt_fixpoint.v b/mathcomp/odd_order/wielandt_fixpoint.v index beebc3d..d91210b 100644 --- a/mathcomp/odd_order/wielandt_fixpoint.v +++ b/mathcomp/odd_order/wielandt_fixpoint.v @@ -1,9 +1,16 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq path div. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div. +From mathcomp Require Import fintype bigop prime binomial finset ssralg fingroup finalg. -Require Import morphism perm automorphism quotient action commutator gproduct. -Require Import zmodp cyclic center pgroup gseries nilpotent sylow finalg. -Require Import finmodule abelian frobenius maximal extremal hall finmodule. +From mathcomp +Require Import morphism perm automorphism quotient action gfunctor commutator. +From mathcomp +Require Import gproduct zmodp cyclic center pgroup gseries nilpotent sylow. +From mathcomp +Require Import finalg finmodule abelian frobenius maximal extremal hall. +From mathcomp Require Import matrix mxalgebra mxrepresentation mxabelem BGsection1. (******************************************************************************) @@ -36,12 +43,12 @@ have [-> | ntA] := eqsVneq A 1. by exists set0 => [|B]; rewrite ?big_set0 ?inE. have [p_pr _ _] := pgroup_pdiv pA ntA; have p_gt1 := prime_gt1 p_pr. case: n1 => [|n] in eA; first by rewrite trivg_exponent eA in ntA. -have nA1X: X \subset 'N('Ohm_1(A)) := char_norm_trans (Ohm_char 1 A) nAX. +have nA1X: X \subset 'N('Ohm_1(A)) := gFnorm_trans _ nAX. have sAnA1: 'Mho^n(A) \subset 'Ohm_1(A). rewrite (MhoE n pA) (OhmE 1 pA) genS //. - apply/subsetP=> xpn; case/imsetP=> x Ax ->{xpn}; rewrite !inE groupX //. + apply/subsetP=> _ /imsetP[x Ax ->]; rewrite !inE groupX //. by rewrite -expgM -expnSr -eA -order_dvdn dvdn_exponent. -have nAnX: X \subset 'N('Mho^n(A)) := char_norm_trans (Mho_char n A) nAX. +have nAnX: X \subset 'N('Mho^n(A)) := gFnorm_trans _ nAX. have [B minB sBAn]: {B : {group gT} | minnormal B X & B \subset 'Mho^n(A)}. apply: mingroup_exists; rewrite nAnX andbT; apply/trivgPn. have [x Ax ox] := exponent_witness (abelian_nil cAA). @@ -49,7 +56,7 @@ have [B minB sBAn]: {B : {group gT} | minnormal B X & B \subset 'Mho^n(A)}. by rewrite -order_dvdn -ox eA dvdn_Pexp2l ?ltnn. have abelA1: p.-abelem 'Ohm_1(A) by rewrite Ohm1_abelem. have sBA1: B \subset 'Ohm_1(A) := subset_trans sBAn sAnA1. -case/mingroupP: minB; case/andP=> ntB nBX minB. +have{minB} [/andP[ntB nBX] minB] := mingroupP minB. have{nBX sBA1} [U defA1 nUX] := Maschke_abelem abelA1 p'X sBA1 nA1X nBX. have [_ mulBU _ tiBU] := dprodP defA1; have{mulBU} [_ sUA1] := mulG_sub mulBU. have sUA: U \subset A := subset_trans sUA1 (Ohm_sub 1 _). @@ -73,14 +80,14 @@ have [U1 | {defA1 minB}ntU] := eqsVneq U 1. have im_g: g @* (A / 'Phi(A)) = B by rewrite def_g // defA1 OhmMho subn1. have defAb: A / 'Phi(A) = g @*^-1 B by rewrite -im_g injmK. have nsPhiA: 'Phi(A) <| A := Phi_normal A. - have nPhiX: X \subset 'N('Phi(A)) := char_norm_trans (Phi_char A) nAX. + have nPhiX: X \subset 'N('Phi(A)) := gFnorm_trans _ nAX. rewrite defAb; apply/mingroupP; split=> [|Hb]. by rewrite -(morphim_injm_eq1 injg) ?morphpreK /= -?defAb ?im_g ?ntB ?actsQ. case/andP=> ntHb actsXHb /= sgHbB; have [sHbA _] := subsetIP sgHbB. rewrite -sub_morphim_pre // in sgHbB; rewrite -(minB _ _ sgHbB) ?injmK //. rewrite morphim_injm_eq1 // {}ntHb {actsXHb}(subset_trans actsXHb) //=. have{sHbA} [H defHb sPhiH sHA] := inv_quotientS nsPhiA sHbA. - rewrite defHb def_g // (char_norm_trans (Mho_char n H)) //. + rewrite defHb def_g // gFnorm_trans //=. by rewrite astabsQ ?subsetIr ?(normalS sPhiH sHA). have nsUA: U <| A by rewrite -sub_abelian_normal. have nUA: A \subset 'N(U) by case/andP: nsUA. @@ -114,7 +121,7 @@ have nsUK: U <| K := normalS sUK sKA nsUA; have [_ nUK] := andP nsUK. have nKX: X \subset 'N(K). by rewrite -(quotientSGK nUX) ?normsG ?quotient_normG // -defKu. pose K1 := 'Mho^1(K); have nsK1K: K1 <| K := Mho_normal 1 K. -have nXKb: X / K1 \subset 'N(K / K1) by exact: quotient_norms. +have nXKb: X / K1 \subset 'N(K / K1) by apply: quotient_norms. pose K'u := \big[dprod/1]_(Bu in S :\ Ku) Bu. have{S_Ku} defAu_K: K / U \x K'u = A / U by rewrite -defKu -big_setD1. have [[_ Pu _ defK'u]] := dprodP defAu_K; rewrite defK'u => mulKPu _ tiKPu. @@ -146,7 +153,7 @@ have tiUK1: U :&: K1 = 1. have [Db defKb nDXb] := Maschke_abelem abelKb p'Xb sUKb nXKb nUXb. have [_ mulUDb _ tiUDb] := dprodP defKb; have [_ sDKb] := mulG_sub mulUDb. have [D defDb sK1D sDK] := inv_quotientS (Mho_normal 1 K) sDKb. -have nK1X: X \subset 'N(K1) := char_norm_trans (Mho_char 1 K) nKX. +have nK1X: X \subset 'N(K1) := gFnorm_trans _ nKX. have [cDU [sK1K nK1K]] := (centSS sUK sDK cKK, andP nsK1K). have nDX: X \subset 'N(D). rewrite -(quotientSGK nK1X) ?normsG // quotient_normG ?(normalS _ sDK) //. @@ -185,7 +192,7 @@ case/setU1P=> [-> {B S simS} | ]; last exact: simS. have [[pD cDD] nUD] := (pgroupS sDA pA, abelianS sDA cAA, subset_trans sDA nUA). have isoD: D \isog Ku by rewrite defKu -mulUD quotientMidl quotient_isog. rewrite {isoD}(isog_homocyclic isoD); split=> //. -have nPhiDX: X \subset 'N('Phi(D)) := char_norm_trans (Phi_char D) nDX. +have nPhiDX: X \subset 'N('Phi(D)) := gFnorm_trans _ nDX. have [f [injf im_f act_f]]: exists f : {morphism D / 'Phi(D) >-> coset_of 'Phi(Ku)}, [/\ 'injm f, f @* (D / 'Phi(D)) = Ku / 'Phi(Ku) @@ -205,17 +212,17 @@ have [f [injf im_f act_f]]: - by rewrite (subsetP (morphim_norm _ _)) ?mem_morphim. rewrite morphim_restrm (setIidPr (Phi_sub _)). by rewrite (subsetP (morphim_norm _ _)) ?mem_quotient. -apply/mingroupP; split=> [|Y]. - rewrite -subG1 quotient_sub1 ?(normal_norm (Phi_normal _)) //. +apply/mingroupP; split=> [|Y /andP[ntY actsXY] sYD]. + rewrite -subG1 quotient_sub1 ?gFnorm //. by rewrite proper_subn ?Phi_proper // actsQ. -case/andP=> ntY actsXY sYD; have{minKu} [_ minKu] := mingroupP minKu. +have{minKu} [_ minKu] := mingroupP minKu. apply: (injm_morphim_inj injf); rewrite // im_f. apply: minKu; last by rewrite /= -im_f morphimS. rewrite morphim_injm_eq1 // ntY. -apply/subsetP=> xb; case/morphimP=> x Nx Xx ->{xb}. +apply/subsetP=> _ /morphimP[x Nx Xx ->]. rewrite 2!inE /= qact_domE ?subsetT // astabsJ. -rewrite (subsetP (char_norm_trans (Phi_char _) nKuX)) ?mem_quotient //=. -apply/subsetP=> fy; case/morphimP=> y Dy Yy ->{fy}. +rewrite (subsetP (gFnorm_trans _ nKuX)) ?mem_quotient //=. +apply/subsetP=> _ /morphimP[y Dy Yy ->]. by rewrite inE /= -act_f // morphimEsub // mem_imset // (acts_act actsXY). Qed. @@ -243,7 +250,7 @@ have [wT [fL injL [fX injX fJ]]]: exists wT : finGroupType, {in setT & G, morph_act MR_G 'J fL fX}. - exists (sdprod_groupType MR_G). exists (sdpair1_morphism MR_G); first exact: injm_sdpair1. - by exists (sdpair2_morphism MR_G); [exact: injm_sdpair2 | exact: sdpair_act]. + by exists (sdpair2_morphism MR_G); [apply: injm_sdpair2 | apply: sdpair_act]. move imfL: (fL @* [set: _])%G => L; move imfX: (fX @* G)%G => X. have cLL: abelian L by rewrite -imfL morphim_abelian // zmod_abelian. have pL: p.-group L. @@ -327,7 +334,7 @@ have simS: forall U, U \in S -> mxsimple aG (gMx U). rewrite -trivg_rowg -subG1 (subset_trans s_U2_1) //. rewrite -(morphim_ker (Morphism gM)) morphimS // kerg. by rewrite subIset ?(PhiS pL) ?orbT. - rewrite actsQ //; first by rewrite (char_norm_trans (Phi_char U)). + rewrite actsQ //; first by rewrite gFnorm_trans. rewrite normsI //; apply/subsetP=> x Xx; rewrite inE. apply/subsetP=> _ /imsetP[u g'U2u ->]. have [Lu U2gu] := morphpreP g'U2u; rewrite mem_rowg in U2gu. @@ -353,12 +360,12 @@ have Fp'G: [char 'F_p]^'.-group G. have [VK [modVK defVK]] := rsim_regular_submod mx_irrV Fp'G. have [U S_U isoUV]: {U | U \in S & mx_iso (regular_repr _ G) (gMx U) VK}. apply: hom_mxsemisimple_iso (scalar_mx_hom _ 1 _) _ => [|U S_U _|]; auto. - by apply/(submod_mx_irr modVK); exact: (mx_rsim_irr defVK). + by apply/(submod_mx_irr modVK); apply: (mx_rsim_irr defVK). by rewrite mulmx1 sumS submx1. have simU := simS U S_U; have [modU _ _] := simU. pose rV := abelem_repr abelV ntV nVG. have{VK modVK defVK isoUV} [h dimU h_free hJ]: mx_rsim (submod_repr modU) rV. - by apply: mx_rsim_trans (mx_rsim_sym defVK); exact/mx_rsim_iso. + by apply: mx_rsim_trans (mx_rsim_sym defVK); apply/mx_rsim_iso. have sUL : U \subset L. by move: defL; rewrite (big_setD1 U) //= => /dprodP[[_ U1 _ ->] /mulG_sub[]]. pose W := [set: 'rV['Z_(p ^ m)](V)]%G. @@ -574,7 +581,7 @@ have{tiRCW} rCW : 'r('C_W(Ai1)) = rC i. suffices /card_isog ->: 'C_V(A i) \isog 'C_W(Ai1) / 'Mho^1(W). by rewrite card_quotient // subIset // normal_norm ?Mho_normal. rewrite coprime_quotient_cent ?Mho_sub ?abelian_sol //= -/Ai1; last first. - by rewrite (subset_trans sAiG1) // (char_norm_trans _ nWG1) ?Mho_char. + by rewrite (subset_trans sAiG1) // gFnorm_trans. have ->: A i :=: fG @* Ai1. by rewrite /Ai1 morphim_invmE morphpreK // im_restrm imfG1. rewrite -imfW morphim_restrm (setIidPr sAiG1). diff --git a/mathcomp/real_closed/Make b/mathcomp/real_closed/Make index 08eedc2..1e013d3 100644 --- a/mathcomp/real_closed/Make +++ b/mathcomp/real_closed/Make @@ -1,6 +1,4 @@ --R . mathcomp.real_closed - -all.v +all_real_closed.v bigenough.v cauchyreals.v complex.v @@ -10,4 +8,6 @@ polyrcf.v qe_rcf_th.v qe_rcf.v realalg.v -mxtens.v
\ No newline at end of file +mxtens.v + +-R . mathcomp.real_closed
\ No newline at end of file diff --git a/mathcomp/real_closed/all.v b/mathcomp/real_closed/all_real_closed.v index 184ee4a..184ee4a 100644 --- a/mathcomp/real_closed/all.v +++ b/mathcomp/real_closed/all_real_closed.v diff --git a/mathcomp/real_closed/bigenough.v b/mathcomp/real_closed/bigenough.v index d9c89ca..819f8d9 100644 --- a/mathcomp/real_closed/bigenough.v +++ b/mathcomp/real_closed/bigenough.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. (****************************************************************************) (* This is a small library to do epsilon - N reasonning. *) diff --git a/mathcomp/real_closed/cauchyreals.v b/mathcomp/real_closed/cauchyreals.v index 83504be..977fbe7 100644 --- a/mathcomp/real_closed/cauchyreals.v +++ b/mathcomp/real_closed/cauchyreals.v @@ -1,6 +1,10 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp Require Import bigop ssralg ssrnum ssrint rat poly polydiv polyorder. +From mathcomp Require Import perm matrix mxpoly polyXY binomial bigenough. (***************************************************************************) diff --git a/mathcomp/real_closed/complex.v b/mathcomp/real_closed/complex.v index 1c26a9d..23c0301 100644 --- a/mathcomp/real_closed/complex.v +++ b/mathcomp/real_closed/complex.v @@ -1,6 +1,10 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp Require Import bigop ssralg ssrint div ssrnum rat poly closed_field polyrcf. +From mathcomp Require Import matrix mxalgebra tuple mxpoly zmodp binomial realalg. (**********************************************************************) diff --git a/mathcomp/real_closed/mxtens.v b/mathcomp/real_closed/mxtens.v new file mode 100644 index 0000000..48e5c10 --- /dev/null +++ b/mathcomp/real_closed/mxtens.v @@ -0,0 +1,315 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp +Require Import bigop ssralg matrix zmodp div. + +Set Implicit Arguments. +Unset Strict Implicit. +Unset Printing Implicit Defensive. + +Import GRing.Theory. +Local Open Scope nat_scope. +Local Open Scope ring_scope. + +Section ExtraBigOp. + +Lemma sumr_add : forall (R : ringType) m n (F : 'I_(m + n) -> R), + \sum_(i < m + n) F i = \sum_(i < m) F (lshift _ i) + + \sum_(i < n) F (rshift _ i). +Proof. +move=> R; elim=> [|m ihm] n F. + rewrite !big_ord0 add0r; apply: congr_big=> // [[i hi]] _. + by rewrite /rshift /=; congr F; apply: val_inj. +rewrite !big_ord_recl ihm -addrA. +congr (_ + _); first by congr F; apply: val_inj. +congr (_ + _); by apply: congr_big=> // i _ /=; congr F; apply: val_inj. +Qed. + +Lemma mxtens_index_proof m n (ij : 'I_m * 'I_n) : ij.1 * n + ij.2 < m * n. +Proof. +case: m ij=> [[[] //]|] m ij; rewrite mulSn addnC -addSn leq_add //. +by rewrite leq_mul2r; case: n ij=> // n ij; rewrite leq_ord orbT. +Qed. + +Definition mxtens_index m n ij := Ordinal (@mxtens_index_proof m n ij). + +Lemma mxtens_index_proof1 m n (k : 'I_(m * n)) : k %/ n < m. +Proof. by move: m n k=> [_ [] //|m] [|n] k; rewrite ?divn0 // ltn_divLR. Qed. +Lemma mxtens_index_proof2 m n (k : 'I_(m * n)) : k %% n < n. +Proof. by rewrite ltn_mod; case: n k=> //; rewrite muln0=> [] []. Qed. + +Definition mxtens_unindex m n k := + (Ordinal (@mxtens_index_proof1 m n k), Ordinal (@mxtens_index_proof2 m n k)). + +Implicit Arguments mxtens_index [[m] [n]]. +Implicit Arguments mxtens_unindex [[m] [n]]. + +Lemma mxtens_indexK m n : cancel (@mxtens_index m n) (@mxtens_unindex m n). +Proof. +case: m=> [[[] //]|m]; case: n=> [[_ [] //]|n]. +move=> [i j]; congr (_, _); apply: val_inj=> /=. + by rewrite divnMDl // divn_small. +by rewrite modnMDl // modn_small. +Qed. + +Lemma mxtens_unindexK m n : cancel (@mxtens_unindex m n) (@mxtens_index m n). +Proof. +case: m=> [[[] //]|m]. case: n=> [|n] k. + by suff: False by []; move: k; rewrite muln0=> [] []. +by apply: val_inj=> /=; rewrite -divn_eq. +Qed. + +CoInductive is_mxtens_index (m n : nat) : 'I_(m * n) -> Type := + IsMxtensIndex : forall (i : 'I_m) (j : 'I_n), + is_mxtens_index (mxtens_index (i, j)). + +Lemma mxtens_indexP (m n : nat) (k : 'I_(m * n)) : is_mxtens_index k. +Proof. by rewrite -[k]mxtens_unindexK; constructor. Qed. + +Lemma mulr_sum (R : ringType) m n (Fm : 'I_m -> R) (Fn : 'I_n -> R) : + (\sum_(i < m) Fm i) * (\sum_(i < n) Fn i) + = \sum_(i < m * n) ((Fm (mxtens_unindex i).1) * (Fn (mxtens_unindex i).2)). +Proof. +rewrite mulr_suml; transitivity (\sum_i (\sum_(j < n) Fm i * Fn j)). + by apply: eq_big=> //= i _; rewrite -mulr_sumr. +rewrite pair_big; apply: reindex=> //=. +by exists mxtens_index=> i; rewrite (mxtens_indexK, mxtens_unindexK). +Qed. + +End ExtraBigOp. + +Section ExtraMx. + +Lemma castmx_mul (R : ringType) + (m m' n p p': nat) (em : m = m') (ep : p = p') + (M : 'M[R]_(m, n)) (N : 'M[R]_(n, p)) : + castmx (em, ep) (M *m N) = castmx (em, erefl _) M *m castmx (erefl _, ep) N. +Proof. by case: m' / em; case: p' / ep. Qed. + +Lemma mulmx_cast (R : ringType) + (m n n' p p' : nat) (en : n' = n) (ep : p' = p) + (M : 'M[R]_(m, n)) (N : 'M[R]_(n', p')) : + M *m (castmx (en, ep) N) = + (castmx (erefl _, (esym en)) M) *m (castmx (erefl _, ep) N). +Proof. by case: n / en in M *; case: p / ep in N *. Qed. + +Lemma castmx_row (R : Type) (m m' n1 n2 n1' n2' : nat) + (eq_n1 : n1 = n1') (eq_n2 : n2 = n2') (eq_n12 : (n1 + n2 = n1' + n2')%N) + (eq_m : m = m') (A1 : 'M[R]_(m, n1)) (A2 : 'M_(m, n2)) : + castmx (eq_m, eq_n12) (row_mx A1 A2) = + row_mx (castmx (eq_m, eq_n1) A1) (castmx (eq_m, eq_n2) A2). +Proof. +case: _ / eq_n1 in eq_n12 *; case: _ / eq_n2 in eq_n12 *. +by case: _ / eq_m; rewrite castmx_id. +Qed. + +Lemma castmx_col (R : Type) (m m' n1 n2 n1' n2' : nat) + (eq_n1 : n1 = n1') (eq_n2 : n2 = n2') (eq_n12 : (n1 + n2 = n1' + n2')%N) + (eq_m : m = m') (A1 : 'M[R]_(n1, m)) (A2 : 'M_(n2, m)) : + castmx (eq_n12, eq_m) (col_mx A1 A2) = + col_mx (castmx (eq_n1, eq_m) A1) (castmx (eq_n2, eq_m) A2). +Proof. +case: _ / eq_n1 in eq_n12 *; case: _ / eq_n2 in eq_n12 *. +by case: _ / eq_m; rewrite castmx_id. +Qed. + +Lemma castmx_block (R : Type) (m1 m1' m2 m2' n1 n2 n1' n2' : nat) + (eq_m1 : m1 = m1') (eq_n1 : n1 = n1') (eq_m2 : m2 = m2') (eq_n2 : n2 = n2') + (eq_m12 : (m1 + m2 = m1' + m2')%N) (eq_n12 : (n1 + n2 = n1' + n2')%N) + (ul : 'M[R]_(m1, n1)) (ur : 'M[R]_(m1, n2)) + (dl : 'M[R]_(m2, n1)) (dr : 'M[R]_(m2, n2)) : + castmx (eq_m12, eq_n12) (block_mx ul ur dl dr) = + block_mx (castmx (eq_m1, eq_n1) ul) (castmx (eq_m1, eq_n2) ur) + (castmx (eq_m2, eq_n1) dl) (castmx (eq_m2, eq_n2) dr). +Proof. +case: _ / eq_m1 in eq_m12 *; case: _ / eq_m2 in eq_m12 *. +case: _ / eq_n1 in eq_n12 *; case: _ / eq_n2 in eq_n12 *. +by rewrite !castmx_id. +Qed. + +End ExtraMx. + +Section MxTens. + +Variable R : ringType. + +Definition tensmx {m n p q : nat} + (A : 'M_(m, n)) (B : 'M_(p, q)) : 'M[R]_(_,_) := nosimpl + (\matrix_(i, j) (A (mxtens_unindex i).1 (mxtens_unindex j).1 + * B (mxtens_unindex i).2 (mxtens_unindex j).2)). + +Notation "A *t B" := (tensmx A B) + (at level 40, left associativity, format "A *t B"). + +Lemma tensmxE {m n p q} (A : 'M_(m, n)) (B : 'M_(p, q)) i j k l : + (A *t B) (mxtens_index (i, j)) (mxtens_index (k, l)) = A i k * B j l. +Proof. by rewrite !mxE !mxtens_indexK. Qed. + +Lemma tens0mx {m n p q} (M : 'M[R]_(p,q)) : (0 : 'M_(m,n)) *t M = 0. +Proof. by apply/matrixP=> i j; rewrite !mxE mul0r. Qed. + +Lemma tensmx0 {m n p q} (M : 'M[R]_(m,n)) : M *t (0 : 'M_(p,q)) = 0. +Proof. by apply/matrixP=> i j; rewrite !mxE mulr0. Qed. + +Lemma tens_scalar_mx (m n : nat) (c : R) (M : 'M_(m,n)): + c%:M *t M = castmx (esym (mul1n _), esym (mul1n _)) (c *: M). +Proof. +apply/matrixP=> i j. +case: (mxtens_indexP i)=> i0 i1; case: (mxtens_indexP j)=> j0 j1. +rewrite tensmxE [i0]ord1 [j0]ord1 !castmxE !mxE /= mulr1n. +by congr (_ * M _ _); apply: val_inj. +Qed. + +Lemma tens_scalar1mx (m n : nat) (M : 'M_(m,n)) : + 1 *t M = castmx (esym (mul1n _), esym (mul1n _)) M. +Proof. by rewrite tens_scalar_mx scale1r. Qed. + +Lemma tens_scalarN1mx (m n : nat) (M : 'M_(m,n)) : + (-1) *t M = castmx (esym (mul1n _), esym (mul1n _)) (-M). +Proof. by rewrite [-1]mx11_scalar /= tens_scalar_mx !mxE scaleNr scale1r. Qed. + +Lemma trmx_tens {m n p q} (M :'M[R]_(m,n)) (N : 'M[R]_(p,q)) : + (M *t N)^T = M^T *t N^T. +Proof. by apply/matrixP=> i j; rewrite !mxE. Qed. + +Lemma tens_col_mx {m n p q} (r : 'rV[R]_n) + (M :'M[R]_(m, n)) (N : 'M[R]_(p, q)) : + (col_mx r M) *t N = + castmx (esym (mulnDl _ _ _), erefl _) (col_mx (r *t N) (M *t N)). +Proof. +apply/matrixP=> i j. +case: (mxtens_indexP i)=> i0 i1; case: (mxtens_indexP j)=> j0 j1. +rewrite !tensmxE castmxE /= cast_ord_id esymK !mxE /=. +case: splitP=> i0' /= hi0'; case: splitP=> k /= hk. ++ case: (mxtens_indexP k) hk=> k0 k1 /=; rewrite tensmxE. + move=> /(f_equal (edivn^~ p)); rewrite !edivn_eq // => [] [h0 h1]. + by congr (r _ _ * N _ _); apply:val_inj; rewrite /= -?h0 ?h1. ++ move: hk (ltn_ord i1); rewrite hi0'. + by rewrite [i0']ord1 mul0n mul1n add0n ltnNge=> ->; rewrite leq_addr. ++ move: (ltn_ord k); rewrite -hk hi0' ltnNge {1}mul1n. + by rewrite mulnDl {1}mul1n -addnA leq_addr. +case: (mxtens_indexP k) hk=> k0 k1 /=; rewrite tensmxE. +rewrite hi0' mulnDl -addnA=> /addnI. + move=> /(f_equal (edivn^~ p)); rewrite !edivn_eq // => [] [h0 h1]. +by congr (M _ _ * N _ _); apply:val_inj; rewrite /= -?h0 ?h1. +Qed. + +Lemma tens_row_mx {m n p q} (r : 'cV[R]_m) (M :'M[R]_(m,n)) (N : 'M[R]_(p,q)) : + (row_mx r M) *t N = + castmx (erefl _, esym (mulnDl _ _ _)) (row_mx (r *t N) (M *t N)). +Proof. +rewrite -[_ *t _]trmxK trmx_tens tr_row_mx tens_col_mx. +apply/eqP; rewrite -(can2_eq (castmxKV _ _) (castmxK _ _)); apply/eqP. +by rewrite trmx_cast castmx_comp castmx_id tr_col_mx -!trmx_tens !trmxK. +Qed. + +Lemma tens_block_mx {m n p q} + (ul : 'M[R]_1) (ur : 'rV[R]_n) (dl : 'cV[R]_m) + (M :'M[R]_(m,n)) (N : 'M[R]_(p,q)) : + (block_mx ul ur dl M) *t N = + castmx (esym (mulnDl _ _ _), esym (mulnDl _ _ _)) + (block_mx (ul *t N) (ur *t N) (dl *t N) (M *t N)). +Proof. +rewrite !block_mxEv tens_col_mx !tens_row_mx -!cast_col_mx castmx_comp. +by congr (castmx (_,_)); apply nat_irrelevance. +Qed. + + +Fixpoint ntensmx_rec {m n} (A : 'M_(m,n)) k : 'M_(m ^ k.+1,n ^ k.+1) := + if k is k'.+1 then (A *t (ntensmx_rec A k')) else A. + +Definition ntensmx {m n} (A : 'M_(m, n)) k := nosimpl + (if k is k'.+1 return 'M[R]_(m ^ k,n ^ k) then ntensmx_rec A k' else 1). + +Notation "A ^t k" := (ntensmx A k) + (at level 39, left associativity, format "A ^t k"). + +Lemma ntensmx0 : forall {m n} (A : 'M_(m,n)) , A ^t 0 = 1. +Proof. by []. Qed. + +Lemma ntensmx1 : forall {m n} (A : 'M_(m,n)) , A ^t 1 = A. +Proof. by []. Qed. + +Lemma ntensmx2 : forall {m n} (A : 'M_(m,n)) , A ^t 2 = A *t A. +Proof. by []. Qed. + +Lemma ntensmxSS : forall {m n} (A : 'M_(m,n)) k, A ^t k.+2 = A *t A ^t k.+1. +Proof. by []. Qed. + +Definition ntensmxS := (@ntensmx1, @ntensmx2, @ntensmxSS). + +End MxTens. + +Notation "A *t B" := (tensmx A B) + (at level 40, left associativity, format "A *t B"). + +Notation "A ^t k" := (ntensmx A k) + (at level 39, left associativity, format "A ^t k"). + +Section MapMx. +Variables (aR rR : ringType). +Hypothesis f : {rmorphism aR -> rR}. +Local Notation "A ^f" := (map_mx f A) : ring_scope. + +Variables m n p q: nat. +Implicit Type A : 'M[aR]_(m, n). +Implicit Type B : 'M[aR]_(p, q). + +Lemma map_mxT A B : (A *t B)^f = A^f *t B^f :> 'M_(m*p, n*q). +Proof. by apply/matrixP=> i j; rewrite !mxE /= rmorphM. Qed. + +End MapMx. + +Section Misc. + +Lemma tensmx_mul (R : comRingType) m n p q r s + (A : 'M[R]_(m,n)) (B : 'M[R]_(p,q)) (C : 'M[R]_(n, r)) (D : 'M[R]_(q, s)) : + (A *t B) *m (C *t D) = (A *m C) *t (B *m D). +Proof. +apply/matrixP=> /= i j. +case (mxtens_indexP i)=> [im ip] {i}; case (mxtens_indexP j)=> [jr js] {j}. +rewrite !mxE !mxtens_indexK mulr_sum; apply: congr_big=> // k _. +by rewrite !mxE !mxtens_indexK mulrCA !mulrA [C _ _ * A _ _]mulrC. +Qed. + +(* Todo : move to div ? *) +Lemma eq_addl_mul q q' m m' d : m < d -> m' < d -> + (q * d + m == q' * d + m')%N = ((q, m) == (q', m')). +Proof. +move=> lt_md lt_m'd; apply/eqP/eqP; last by move=> [-> ->]. +by move=> /(f_equal (edivn^~ d)); rewrite !edivn_eq. +Qed. + +Lemma tensmx_unit (R : fieldType) m n (A : 'M[R]_m%N) (B : 'M[R]_n%N) : + m != 0%N -> n != 0%N -> A \in unitmx -> B \in unitmx -> (A *t B) \in unitmx. +Proof. +move: m n A B => [|m] [|n] // A B _ _ uA uB. +suff : (A^-1 *t B^-1) *m (A *t B) = 1 by case/mulmx1_unit. +rewrite tensmx_mul !mulVmx //; apply/matrixP=> /= i j. +rewrite !mxE /=; symmetry; rewrite -natrM -!val_eqE /=. +rewrite {1}(divn_eq i n.+1) {1}(divn_eq j n.+1). +by rewrite eq_addl_mul ?ltn_mod // xpair_eqE mulnb. +Qed. + + +Lemma tens_mx_scalar : forall (R : comRingType) + (m n : nat) (c : R) (M : 'M[R]_(m,n)), + M *t c%:M = castmx (esym (muln1 _), esym (muln1 _)) (c *: M). +Proof. +move=> R0 m n c M; apply/matrixP=> i j. +case: (mxtens_indexP i)=> i0 i1; case: (mxtens_indexP j)=> j0 j1. +rewrite tensmxE [i1]ord1 [j1]ord1 !castmxE !mxE /= mulr1n mulrC. +by congr (_ * M _ _); apply: val_inj=> /=; rewrite muln1 addn0. +Qed. + +Lemma tensmx_decr : forall (R : comRingType) m n (M :'M[R]_m) (N : 'M[R]_n), + M *t N = (M *t 1%:M) *m (1%:M *t N). +Proof. by move=> R0 m n M N; rewrite tensmx_mul mul1mx mulmx1. Qed. + +Lemma tensmx_decl : forall (R : comRingType) m n (M :'M[R]_m) (N : 'M[R]_n), + M *t N = (1%:M *t N) *m (M *t 1%:M). +Proof. by move=> R0 m n M N; rewrite tensmx_mul mul1mx mulmx1. Qed. + +End Misc. diff --git a/mathcomp/real_closed/ordered_qelim.v b/mathcomp/real_closed/ordered_qelim.v index c718d74..6ec0cf7 100644 --- a/mathcomp/real_closed/ordered_qelim.v +++ b/mathcomp/real_closed/ordered_qelim.v @@ -1,6 +1,10 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq div choice fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq div choice fintype. +From mathcomp Require Import bigop ssralg finset fingroup zmodp. +From mathcomp Require Import poly ssrnum. @@ -360,7 +364,7 @@ split=> t1. set tr := _ m. suffices: all (@rterm R) (tr.1 :: tr.2)%PAIR. case: tr => {t1} t1 r /= /andP[t1_r]. - by elim: r m => [|t r IHr] m; rewrite /= ?andbT // => /andP[->]; exact: IHr. + by elim: r m => [|t r IHr] m; rewrite /= ?andbT // => /andP[->]; apply: IHr. have: all (@rterm R) [::] by []. rewrite {}/tr; elim: t1 [::] => //=. + move=> t1 IHt1 t2 IHt2 r. @@ -379,7 +383,7 @@ split=> t1. - rewrite /lt0_rform; move: (ub_var t1) => m; set tr := _ m. suffices: all (@rterm R) (tr.1 :: tr.2)%PAIR. case: tr => {t1} t1 r /= /andP[t1_r]. - by elim: r m => [|t r IHr] m; rewrite /= ?andbT // => /andP[->]; exact: IHr. + by elim: r m => [|t r IHr] m; rewrite /= ?andbT // => /andP[->]; apply: IHr. have: all (@rterm R) [::] by []. rewrite {}/tr; elim: t1 [::] => //=. + move=> t1 IHt1 t2 IHt2 r. @@ -398,7 +402,7 @@ split=> t1. - rewrite /leq0_rform; move: (ub_var t1) => m; set tr := _ m. suffices: all (@rterm R) (tr.1 :: tr.2)%PAIR. case: tr => {t1} t1 r /= /andP[t1_r]. - by elim: r m => [|t r IHr] m; rewrite /= ?andbT // => /andP[->]; exact: IHr. + by elim: r m => [|t r IHr] m; rewrite /= ?andbT // => /andP[->]; apply: IHr. have: all (@rterm R) [::] by []. rewrite {}/tr; elim: t1 [::] => //=. + move=> t1 IHt1 t2 IHt2 r. @@ -430,11 +434,11 @@ suffices{e f} [equal0_equiv lt0_equiv le0_equiv]: by split; [move/equal0_equiv/eqP | move/eqP/equal0_equiv]. + by move=> t1 t2 e; split; move/lt0_equiv. + by move=> t1 t2 e; split; move/le0_equiv. - + move=> t1 e; rewrite unitrE; exact: equal0_equiv. - + move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. - + move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. - + move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. - + move=> f1 IHf1 e; move: (IHf1 e); tauto. + + by move=> t1 e; rewrite unitrE; apply: equal0_equiv. + + by move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. + + by move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. + + by move=> f1 IHf1 f2 IHf2 e; move: (IHf1 e) (IHf2 e); tauto. + + by move=> f1 IHf1 e; move: (IHf1 e); tauto. + by move=> n f1 IHf1 e; split=> [] [x] /IHf1; exists x. + by move=> n f1 IHf1 e; split=> Hx x; apply/IHf1. suffices h e t1 t2 : @@ -443,7 +447,7 @@ suffices h e t1 t2 : holds e (leq0_rform (t1 - t2)) <-> (eval e t1 <= eval e t2)]. by split => e t1 t2; case: (h e t1 t2). rewrite -{1}(add0r (eval e t2)) -(can2_eq (subrK _) (addrK _)). -rewrite -subr_lt0 -subr_le0 -/(eval e (t1 - t2)); move: (t1 - t2)%T => {t1 t2} t. +rewrite -subr_lt0 -subr_le0 -/(eval e (t1 - t2)); move: {t1 t2}(t1 - t2)%T => t. have sub_var_tsubst s t0: (s.1%PAIR >= ub_var t0)%N -> tsubst t0 s = t0. elim: t0 {t} => //=. - by move=> n; case: ltngtP. @@ -586,8 +590,8 @@ Definition qf_eval e := fix loop (f : formula R) : bool := (* qf_eval is equivalent to holds *) Lemma qf_evalP e f : qf_form f -> reflect (holds e f) (qf_eval e f). Proof. -elim: f => //=; try by move=> *; exact: idP. -- move=> t1 t2 _; exact: eqP. +elim: f => //=; try by move=> *; apply: idP. +- by move=> t1 t2 _; apply: eqP. - move=> f1 IHf1 f2 IHf2 /= /andP[/IHf1[] f1T]; last by right; case. by case/IHf2; [left | right; case]. - move=> f1 IHf1 f2 IHf2 /= /andP[/IHf1[] f1F]; first by do 2 left. @@ -704,7 +708,7 @@ Lemma qf_to_dnf_rterm f b : rformula f -> all dnf_rterm (qf_to_odnf f b). Proof. set ok := all dnf_rterm. have cat_ok bcs1 bcs2: ok bcs1 -> ok bcs2 -> ok (bcs1 ++ bcs2). - by move=> ok1 ok2; rewrite [ok _]all_cat; exact/andP. + by move=> ok1 ok2; rewrite [ok _]all_cat; apply/andP. have and_ok bcs1 bcs2: ok bcs1 -> ok bcs2 -> ok (and_odnf bcs1 bcs2). rewrite /and_odnf unlock; elim: bcs1 => //= cl1 bcs1 IH1; rewrite -andbA. case/and3P=> ok11 ok12 ok1 ok2; rewrite cat_ok ?{}IH1 {bcs1 ok1}//. @@ -1120,7 +1124,7 @@ suffices or_wf fs : let ofs := foldr Or False fs in - apply: or_wf. suffices map_proj_wf bcs: let mbcs := map (proj n) bcs in all dnf_rterm bcs -> all qf_form mbcs && all rformula mbcs. - by apply: map_proj_wf; exact: qf_to_dnf_rterm. + by apply: map_proj_wf; apply: qf_to_dnf_rterm. elim: bcs => [|bc bcs ihb] bcsr //= /andP[rbc rbcs]. by rewrite andbAC andbA wf_QE_proj //= andbC ihb. elim: fs => //= g gs ihg; rewrite -andbA => /and4P[-> qgs -> rgs] /=. @@ -1137,7 +1141,7 @@ have auxP f0 e0 n0: qf_form f0 && rformula f0 -> apply: (@iffP (rc e0 n0 (odnf_to_oform bcs))); last first. - by case=> x; rewrite -qf_to_dnfP //; exists x. - by case=> x; rewrite qf_to_dnfP //; exists x. - have: all dnf_rterm bcs by case/andP: cf => _; exact: qf_to_dnf_rterm. + have: all dnf_rterm bcs by case/andP: cf => _; apply: qf_to_dnf_rterm. elim: {f0 cf}bcs => [|bc bcs IHbcs] /=; first by right; case. case/andP=> r_bc /IHbcs {IHbcs}bcsP. have f_qf := dnf_to_form_qf [:: bc]. @@ -1149,10 +1153,10 @@ have auxP f0 e0 n0: qf_form f0 && rformula f0 -> case/orP => [bc_x|]; last by exists x. by case: no_x; exists x; apply/(qf_evalP _ f_qf); rewrite /= bc_x. elim: f e => //. -- move=> b e _; exact: idP. -- move=> t1 t2 e _; exact: eqP. -- move=> t1 t2 e _; exact: idP. -- move=> t1 t2 e _; exact: idP. +- by move=> b e _; apply: idP. +- by move=> t1 t2 e _; apply: eqP. +- by move=> t1 t2 e _; apply: idP. +- by move=> t1 t2 e _; apply: idP. - move=> f1 IH1 f2 IH2 e /= /andP[/IH1[] f1e]; last by right; case. by case/IH2; [left | right; case]. - move=> f1 IH1 f2 IH2 e /= /andP[/IH1[] f1e]; first by do 2!left. @@ -1161,7 +1165,7 @@ elim: f e => //. by case/IH2; [left | right; move/(_ f1e)]. - by move=> f IHf e /= /IHf[]; [right | left]. - move=> n f IHf e /= rf; have rqf := quantifier_elim_wf rf. - by apply: (iffP (auxP _ _ _ rqf)) => [] [x]; exists x; exact/IHf. + by apply: (iffP (auxP _ _ _ rqf)) => [] [x]; exists x; apply/IHf. move=> n f IHf e /= rf; have rqf := quantifier_elim_wf rf. case: auxP => // [f_x|no_x]; first by right=> no_x; case: f_x => x /IHf[]. by left=> x; apply/IHf=> //; apply/idPn=> f_x; case: no_x; exists x. diff --git a/mathcomp/real_closed/polyorder.v b/mathcomp/real_closed/polyorder.v index be4b7cc..b7b3695 100644 --- a/mathcomp/real_closed/polyorder.v +++ b/mathcomp/real_closed/polyorder.v @@ -1,5 +1,8 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp Require Import ssralg poly ssrnum zmodp polydiv interval. Import GRing.Theory. @@ -207,7 +210,7 @@ move=> hn. case p0: (p == 0); first by rewrite (eqP p0) div0p mu0 sub0n. case: (@mu_spec p x); rewrite ?p0 // => q hq hp. rewrite {1}hp -{1}(subnK hn) exprD mulrA. -rewrite Pdiv.IdomainMonic.mulpK; last by apply: monic_exp; exact: monicXsubC. +rewrite Pdiv.IdomainMonic.mulpK; last by apply: monic_exp; apply: monicXsubC. rewrite mu_mul ?mulf_eq0 ?expf_eq0 ?polyXsubC_eq0 ?andbF ?orbF; last first. by apply: contra hq; move/eqP->; rewrite root0. by rewrite mu_exp muNroot // add0n mu_XsubC mul1n. diff --git a/mathcomp/real_closed/polyrcf.v b/mathcomp/real_closed/polyrcf.v index b49e729..3ab6a62 100644 --- a/mathcomp/real_closed/polyrcf.v +++ b/mathcomp/real_closed/polyrcf.v @@ -1,6 +1,10 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp Require Import bigop ssralg poly polydiv ssrnum zmodp. +From mathcomp Require Import polyorder path interval ssrint. (****************************************************************************) @@ -46,9 +50,9 @@ Local Notation mid x y := ((x + y) / 2%:R). Section more. Section SeqR. -Lemma last1_neq0 : forall (R : ringType) (s: seq R) (c:R), c != 0 -> - (last c s != 0) = (last 1 s != 0). -Proof. by move=> R'; elim=> [|t s ihs] c cn0 //; rewrite oner_eq0 cn0. Qed. +Lemma last1_neq0 (R : ringType) (s : seq R) (c : R) : + c != 0 -> (last c s != 0) = (last 1 s != 0). +Proof. by elim: s c => [|t s ihs] c cn0 //; rewrite oner_eq0 cn0. Qed. End SeqR. @@ -66,8 +70,8 @@ Proof. by move/lead_coefDl<-; rewrite addrC. Qed. Lemma leq1_size_polyC (c : R) : (size c%:P <= 1)%N. Proof. by rewrite size_polyC; case: (c == 0). Qed. -Lemma my_size_exp p n : p != 0 -> - (size (p ^+ n)) = ((size p).-1 * n).+1%N. +Lemma my_size_exp p n : + p != 0 -> (size (p ^+ n)) = ((size p).-1 * n).+1%N. Proof. by move=> hp; rewrite -size_exp prednK // lt0n size_poly_eq0 expf_neq0. Qed. @@ -84,7 +88,8 @@ Proof. by move=> h; rewrite -size_poly_eq0 lt0n_neq0 //; apply: leq_ltn_trans h. Qed. -Lemma lead_coef_comp_poly p q : (size q > 1)%N -> +Lemma lead_coef_comp_poly p q : + (size q > 1)%N -> lead_coef (p \Po q) = (lead_coef p) * (lead_coef q) ^+ (size p).-1. Proof. move=> sq; rewrite !lead_coefE coef_comp_poly size_comp_poly. @@ -130,7 +135,8 @@ rewrite /sgp_pinfty /sgp_minfty lead_coef_comp_poly ?size_opp ?size_polyX //. by rewrite lead_coef_opp lead_coefX mulrC. Qed. -Lemma poly_pinfty_gt_lc p : lead_coef p > 0 -> +Lemma poly_pinfty_gt_lc p : + lead_coef p > 0 -> exists n, forall x, x >= n -> p.[x] >= lead_coef p. Proof. elim/poly_ind: p => [| q c IHq]; first by rewrite lead_coef0 ltrr. @@ -151,7 +157,8 @@ by rewrite -[c]opprK subr_ge0 normrN ler_norm. Qed. (* :REMARK: not necessary here ! *) -Lemma poly_lim_infty p m : lead_coef p > 0 -> (size p > 1)%N -> +Lemma poly_lim_infty p m : + lead_coef p > 0 -> (size p > 1)%N -> exists n, forall x, x >= n -> p.[x] >= m. Proof. elim/poly_ind: p m => [| q c _] m; first by rewrite lead_coef0 ltrr. @@ -181,10 +188,10 @@ Definition cauchy_bound (p : {poly R}) := 1 + `|lead_coef p|^-1 * \sum_(i < size p) `|p`_i|. (* Could be a sharp bound, and proof should shrink... *) -Lemma cauchy_boundP : forall (p : {poly R}) x, +Lemma cauchy_boundP (p : {poly R}) x : p != 0 -> p.[x] = 0 -> `| x | < cauchy_bound p. Proof. -move=> p x np0 rpx; rewrite ltr_spaddl ?ltr01 //. +move=> np0 rpx; rewrite ltr_spaddl ?ltr01 //. case e: (size p) => [|n]; first by move: np0; rewrite -size_poly_eq0 e eqxx. have lcp : `|lead_coef p| > 0 by move: np0; rewrite -lead_coef_eq0 -normr_gt0. have lcn0 : `|lead_coef p| != 0 by rewrite normr_eq0 -normr_gt0. @@ -200,8 +207,8 @@ have {h1} h1 : `|p`_n| * `|x| ^+ n <= \sum_(i < n) `|p`_i * x ^+ i|. rewrite -normrX -normrM -normrN h1. by rewrite (ler_trans (ler_norm_sum _ _ _)) // lerr. have xp : `| x | > 0 by rewrite (ltr_trans _ cx1) ?ltr01. -move: h1; rewrite {-1}es exprS mulrA -ler_pdivl_mulr ?exprn_gt0 // big_distrl /=. -rewrite big_ord_recr /= normrM normrX -mulrA es mulfV; last first. +move: h1; rewrite {-1}es exprS mulrA -ler_pdivl_mulr ?exprn_gt0 //. +rewrite big_distrl /= big_ord_recr /= normrM normrX -mulrA es mulfV; last first. by rewrite expf_eq0 negb_and eq_sym (ltr_eqF xp) orbT. have pnp : 0 < `|p`_n| by move: lcp; rewrite /lead_coef e es. rewrite mulr1 -es mulrC -ler_pdivl_mulr //. @@ -263,10 +270,11 @@ Implicit Types p q r : {poly R}. Lemma poly_ivt (p : {poly R}) (a b : R) : a <= b -> 0 \in `[p.[a], p.[b]] -> { x : R | x \in `[a, b] & root p x }. -Proof. by move=> leab root_p_ab; exact/sig2W/poly_ivt. Qed. +Proof. by move=> leab root_p_ab; apply/sig2W/poly_ivt. Qed. -Lemma polyrN0_itv (i : interval R) (p : {poly R}) : {in i, noroot p} - -> forall y x : R, y \in i -> x \in i -> sgr p.[x] = sgr p.[y]. +Lemma polyrN0_itv (i : interval R) (p : {poly R}) : + {in i, noroot p} -> + forall y x : R, y \in i -> x \in i -> sgr p.[x] = sgr p.[y]. Proof. move=> hi y x hy hx; wlog xy: x y hx hy / x <= y=> [hwlog|]. by case/orP: (ler_total x y)=> xy; [|symmetry]; apply: hwlog. @@ -281,20 +289,21 @@ move=> /ltrW p0y /ltrW px0; case: (@poly_ivt p x y); rewrite ?inE ?px0 //. by move=> z hz; rewrite (negPf (hi z _)) // hxyi. Qed. -Lemma poly_div_factor : forall (a:R) (P : {poly R} -> Prop), - (forall k, P k%:P) -> - (forall p n k, p.[a] != 0 -> P p -> - P (p * ('X - a%:P)^+(n.+1) + k%:P)%R) +Lemma poly_div_factor (a : R) (P : {poly R} -> Prop) : + (forall k, P k%:P) -> + (forall p n k, p.[a] != 0 -> P p -> P (p * ('X - a%:P)^+(n.+1) + k%:P)%R) -> forall p, P p. Proof. -move=> a P Pk Pq p. +move=> Pk Pq p. move: {-2}p (leqnn (size p)); elim: (size p)=> {p} [|n ihn] p spn. move: spn; rewrite leqn0 size_poly_eq0; move/eqP->; rewrite -polyC0. exact: Pk. case: (leqP (size p) 1)=> sp1; first by rewrite [p]size1_polyC ?sp1//. -rewrite (Pdiv.IdomainMonic.divp_eq (monicXsubC a) p) [_ %% _]size1_polyC; last first. +rewrite (Pdiv.IdomainMonic.divp_eq (monicXsubC a) p). +rewrite [_ %% _]size1_polyC; last first. rewrite -ltnS. - by rewrite (@leq_trans (size ('X - a%:P))) // ?ltn_modp ?polyXsubC_eq0 ?size_XsubC. + by rewrite (@leq_trans (size ('X - a%:P))) // + ?ltn_modp ?polyXsubC_eq0 ?size_XsubC. have [n' [q hqa hp]] := multiplicity_XsubC (p %/ ('X - a%:P)) a. rewrite divpN0 ?size_XsubC ?polyXsubC_eq0 ?sp1 //= in hqa. rewrite hp -mulrA -exprSr; apply: Pq=> //; apply: ihn. @@ -325,8 +334,8 @@ rewrite andbT; apply/eqP=> y0; move: hyx; rewrite y0. by rewrite exprS mul0r=> x0; move: l0x; rewrite -x0 ltrr. Qed. -Lemma poly_cont x p e : e > 0 -> exists2 d, - d > 0 & forall y, `|y - x| < d -> `|p.[y] - p.[x]| < e. +Lemma poly_cont x p e : + e > 0 -> exists2 d, d > 0 & forall y, `|y - x| < d -> `|p.[y] - p.[x]| < e. Proof. elim/(@poly_div_factor x): p e. move=> e ep; exists 1; rewrite ?ltr01// => y hy. @@ -351,9 +360,9 @@ by rewrite (ler_lt_trans _ (He' y _)) // ler_sub_dist. Qed. (* Todo : orderedpoly !! *) -(* Lemma deriv_expz_nat : forall (n : nat) p, (p ^ n)^`() = (p^`() * p ^ (n.-1)) *~ n. *) +(* Lemma deriv_expz_nat (n : nat) p : (p ^ n)^`() = (p^`() * p ^ (n.-1)) *~ n. *) (* Proof. *) -(* elim=> [|n ihn] p /=; first by rewrite expr0z derivC mul0zr. *) +(* elim: n => [|n ihn] /= in p *; first by rewrite expr0z derivC mul0zr. *) (* rewrite exprSz_nat derivM ihn mulzrAr mulrCA -exprSz_nat. *) (* by case: n {ihn}=> [|n] //; rewrite mul0zr addr0 mul1zr. *) (* Qed. *) @@ -392,31 +401,32 @@ Qed. (* by move:spn; rewrite leq_eqVlt spSn /= ltnS; by move/ihn. *) (* Qed. *) -Lemma poly_ltsp_roots : forall p (rs : seq R), +Lemma poly_ltsp_roots p (rs : seq R) : (size rs >= size p)%N -> uniq rs -> all (root p) rs -> p = 0. Proof. -move=> p rs hrs urs rrs; apply/eqP; apply: contraLR hrs=> np0. +move=> hrs urs rrs; apply/eqP; apply: contraLR hrs=> np0. by rewrite -ltnNge; apply: max_poly_roots. Qed. -Lemma ivt_sign : forall (p : {poly R}) (a b : R), +Lemma ivt_sign (p : {poly R}) (a b : R) : a <= b -> sgr p.[a] * sgr p.[b] = -1 -> { x : R | x \in `]a, b[ & root p x}. Proof. -move=> p a b hab /eqP; rewrite mulrC mulr_sg_eqN1=> /andP [spb0 /eqP spb]. +move=> hab /eqP; rewrite mulrC mulr_sg_eqN1=> /andP [spb0 /eqP spb]. case: (@poly_ivt (sgr p.[b] *: p) a b)=> //. by rewrite !hornerZ {1}spb mulNr -!normrEsg inE /= oppr_cp0 !normrE. move=> c hc; rewrite rootZ ?sgr_eq0 // => rpc; exists c=> //. (* need for a lemma reditvP *) -rewrite inE /= !ltr_neqAle andbCA -!andbA [_ && (_ <= _)]hc andbT eq_sym -negb_or. +rewrite inE /= !ltr_neqAle andbCA -!andbA [_ && (_ <= _)]hc andbT. +rewrite eq_sym -negb_or. apply/negP=> /orP [] /eqP ec; move: rpc; rewrite -ec /root ?(negPf spb0) //. by rewrite -sgr_cp0 -[sgr _]opprK -spb eqr_oppLR oppr0 sgr_cp0 (negPf spb0). Qed. -Let rolle_weak : forall a b p, a < b -> - p.[a] = 0 -> p.[b] = 0 -> +Let rolle_weak a b p : + a < b -> p.[a] = 0 -> p.[b] = 0 -> {c | (c \in `]a, b[) & (p^`().[c] == 0) || (p.[c] == 0)}. Proof. -move=> a b p lab pa0 pb0; apply/sig2W. +move=> lab pa0 pb0; apply/sig2W. case p0: (p == 0). rewrite (eqP p0); exists (mid a b); first by rewrite mid_in_itv. by rewrite derivC horner0 eqxx. @@ -457,11 +467,11 @@ rewrite sgr_cp0 (rootPf qb0) orFb=> rc0. by exists c=> //; rewrite !hornerM !mulf_eq0 rc0. Qed. -Theorem rolle : forall a b p, a < b -> - p.[a] = p.[b] -> {c | c \in `]a, b[ & p^`().[c] = 0}. +Theorem rolle a b p : + a < b -> p.[a] = p.[b] -> {c | c \in `]a, b[ & p^`().[c] = 0}. Proof. -move=> a b p lab pab. -wlog pb0 : p pab / p.[b] = 0=> [hwlog|]. +move=> lab pab. +wlog pb0 : p pab / p.[b] = 0 => [hwlog|]. case: (hwlog (p - p.[b]%:P)); rewrite ?hornerE ?pab ?subrr //. by move=> c acb; rewrite derivE derivC subr0=> hc; exists c. move: pab; rewrite pb0=> pa0. @@ -485,10 +495,10 @@ move=> rs hrs srs urs rrs; apply: (max_roots (c :: rs))=> //=; last exact/andP. by rewrite urs andbT; apply/negP; move/hrs; rewrite bound_in_itv. Qed. -Theorem mvt : forall a b p, a < b -> - {c | c \in `]a, b[ & p.[b] - p.[a] = p^`().[c] * (b - a)}. +Theorem mvt a b p : + a < b -> {c | c \in `]a, b[ & p.[b] - p.[a] = p^`().[c] * (b - a)}. Proof. -move=> a b p lab. +move=> lab. pose q := (p.[b] - p.[a])%:P * ('X - a%:P) - (b - a)%:P * (p - p.[a]%:P). case: (@rolle a b q)=> //. by rewrite /q !hornerE !(subrr,mulr0) mulrC subrr. @@ -497,12 +507,12 @@ move: q'x0; rewrite /q !derivE !(mul0r,add0r,subr0,mulr1). by move/eqP; rewrite !hornerE mulrC subr_eq0; move/eqP. Qed. -Lemma deriv_sign : forall a b p, +Lemma deriv_sign a b p : (forall x, x \in `]a, b[ -> p^`().[x] >= 0) -> (forall x y, (x \in `]a, b[) && (y \in `]a, b[) -> x < y -> p.[x] <= p.[y] ). Proof. -move=> a b p Pab x y; case/andP=> hx hy xy. +move=> Pab x y; case/andP=> hx hy xy. rewrite -subr_gte0; case: (@mvt x y p)=> //. move=> c hc ->; rewrite pmulr_lge0 ?subr_gt0 ?Pab //. by apply: subitvP hc; rewrite //= ?(itvP hx) ?(itvP hy). @@ -605,7 +615,8 @@ move=> pa0 x hx; rewrite hsgpN // (@derp0l _ a b) //; first exact: derq_pos. by rewrite hornerZ pa0 mulr0. Qed. -Lemma derspl : sgr p.[b] = -sp'c -> forall x, x \in `[a, b] -> sgr p.[x] = -sp'c. +Lemma derspl : + sgr p.[b] = -sp'c -> forall x, x \in `[a, b] -> sgr p.[x] = -sp'c. Proof. move=> spb x hx; rewrite hsgpN // (@derpnl _ a b) //; first exact: derq_pos. by rewrite -sgr_cp0 hornerZ sgr_smul spb mulrN -expr2 sqr_sg derp_neq0. @@ -678,33 +689,29 @@ Variable T : predType R. Definition roots_on (p : {poly R}) (i : T) (s : seq R) := forall x, (x \in i) && (root p x) = (x \in s). -Lemma roots_onP : forall p i s, roots_on p i s -> {in i, root p =1 mem s}. -Proof. by move=> p i s hp x hx; move: (hp x); rewrite hx. Qed. +Lemma roots_onP p i s : roots_on p i s -> {in i, root p =1 mem s}. +Proof. by move=> hp x hx; move: (hp x); rewrite hx. Qed. -Lemma roots_on_in : forall p i s, - roots_on p i s -> forall x, x \in s -> x \in i. -Proof. by move=> p i s hp x; rewrite -hp; case/andP. Qed. +Lemma roots_on_in p i s : + roots_on p i s -> forall x, x \in s -> x \in i. +Proof. by move=> hp x; rewrite -hp; case/andP. Qed. -Lemma roots_on_root : forall p i s, +Lemma roots_on_root p i s : roots_on p i s -> forall x, x \in s -> root p x. -Proof. by move=> p i s hp x; rewrite -hp; case/andP. Qed. +Proof. by move=> hp x; rewrite -hp; case/andP. Qed. -Lemma root_roots_on : forall p i s, +Lemma root_roots_on p i s : roots_on p i s -> forall x, x \in i -> root p x -> x \in s. -Proof. by move=> p i s hp x; rewrite -hp=> ->. Qed. +Proof. by move=> hp x; rewrite -hp=> ->. Qed. -Lemma roots_on_opp : forall p i s, - roots_on (- p) i s -> roots_on p i s. -Proof. by move=> p i s hp x; rewrite -hp rootN. Qed. +Lemma roots_on_opp p i s : roots_on (- p) i s -> roots_on p i s. +Proof. by move=> hp x; rewrite -hp rootN. Qed. -Lemma roots_on_nil : forall p i, roots_on p i [::] -> {in i, noroot p}. -Proof. -by move=> p i hp x hx; move: (hp x); rewrite in_nil hx /=; move->. -Qed. +Lemma roots_on_nil p i : roots_on p i [::] -> {in i, noroot p}. +Proof. by move=> hp x hx; move: (hp x); rewrite in_nil hx /=; move->. Qed. -Lemma roots_on_same : forall s' p i s, - s =i s' -> (roots_on p i s <-> roots_on p i s'). -Proof. by move=> s' p i s hss'; split=> hs x; rewrite (hss', (I, hss')). Qed. +Lemma roots_on_same s' p i s : s =i s' -> (roots_on p i s <-> roots_on p i s'). +Proof. by move=> hss'; split=> hs x; rewrite (hss', (I, hss')). Qed. End RootsOn. @@ -728,23 +735,21 @@ End RootsOn. (* by rewrite lter_subrA lter_add2r -lter_addlA lter_add2l. *) (* Qed. *) -Lemma roots_on_comp : forall p a b s, - roots_on (p \Po (-'X)) `](-b), (-a)[ - (map (-%R) s) <-> roots_on p `]a, b[ s. +Lemma roots_on_comp p a b s : + roots_on (p \Po (-'X)) `](-b), (-a)[ (map (-%R) s) <-> roots_on p `]a, b[ s. Proof. -move=> p a b /= s; split=> hs x; rewrite ?root_comp ?hornerE. +split=> /= hs x; rewrite ?root_comp ?hornerE. move: (hs (-x)); rewrite mem_map; last exact: (inv_inj (@opprK _)). by rewrite root_comp ?hornerE oppr_itv !opprK. rewrite -[x]opprK oppr_itv /= mem_map; last exact: (inv_inj (@opprK _)). by move: (hs (-x)); rewrite !opprK. Qed. -Lemma min_roots_on : forall p a b x s, - all (> x) s -> roots_on p `]a, b[ (x :: s) - -> [/\ x \in `]a, b[, (roots_on p `]a, x[ [::]), - (root p x) & (roots_on p `]x, b[ s)]. +Lemma min_roots_on p a b x s : + all (> x) s -> roots_on p `]a, b[ (x :: s) -> + [/\ x \in `]a, b[, roots_on p `]a, x[ [::], root p x & roots_on p `]x, b[ s]. Proof. -move=> p a b x s lxs hxs. +move=> lxs hxs. have hx: x \in `]a, b[ by rewrite (roots_on_in hxs) ?mem_head. rewrite hx (roots_on_root hxs) ?mem_head //. split=> // y; move: (hxs y); rewrite ?in_nil ?in_cons /=. @@ -762,13 +767,11 @@ move/negP; move/negP=> nhy; apply: negbTE; apply: contra nhy. by apply: subitvPl; rewrite //= ?(itvP hx). Qed. -Lemma max_roots_on : forall p a b x s, - all (< x) s -> roots_on p `]a, b[ (x :: s) - -> [/\ x \in `]a, b[, (roots_on p `]x, b[ [::]), - (root p x) & (roots_on p `]a, x[ s)]. +Lemma max_roots_on p a b x s : + all (< x) s -> roots_on p `]a, b[ (x :: s) -> + [/\ x \in `]a, b[, roots_on p `]x, b[ [::], root p x & roots_on p `]a, x[ s]. Proof. -move=> p a b x s; move/allP=> lsx; move/roots_on_comp=> /=. -move/min_roots_on; case. +move/allP=> lsx /roots_on_comp/=/min_roots_on[]. apply/allP=> y; rewrite -[y]opprK mem_map. by move/lsx; rewrite ltr_oppr opprK. exact: (inv_inj (@opprK _)). @@ -776,10 +779,10 @@ rewrite oppr_itv root_comp !hornerE !opprK=> -> rxb -> rax. by split=> //; apply/roots_on_comp. Qed. -Lemma roots_on_cons : forall p a b r s, +Lemma roots_on_cons p a b r s : sorted <%R (r :: s) -> roots_on p `]a, b[ (r :: s) -> roots_on p `]r, b[ s. Proof. -move=> p a b r s /= hrs hr. +move=> /= hrs hr. have:= (order_path_min (@ltr_trans _) hrs)=> allrs. by case: (min_roots_on allrs hr). Qed. @@ -823,18 +826,18 @@ Qed. (* by move/(_ x)=> -> //; rewrite mem_rev. *) (* Qed. *) -Lemma no_roots_on : forall (p : {poly R}) a b, +Lemma no_roots_on (p : {poly R}) a b : {in `]a, b[, noroot p} -> roots_on p `]a, b[ [::]. Proof. -move=> p a b hr x; rewrite in_nil; case hx: (x \in _) => //=. +move=> hr x; rewrite in_nil; case hx: (x \in _) => //=. by apply: negPf; apply: hr hx. Qed. -Lemma monotonic_rootN : forall (p : {poly R}) (a b : R), - {in `]a, b[, noroot p^`()} -> - ((roots_on p `]a, b[ [::]) + ({r : R | roots_on p `]a, b[ [:: r]}))%type. +Lemma monotonic_rootN (p : {poly R}) (a b : R) : + {in `]a, b[, noroot p^`()} -> + ((roots_on p `]a, b[ [::]) + ({r : R | roots_on p `]a, b[ [:: r]}))%type. Proof. -move=> p a b hp'; case: (ltrP a b); last first => leab. +move=> hp'; case: (ltrP a b); last first => leab. by left => x; rewrite in_nil itv_gte. wlog {hp'} hp'sg: p / forall x, x \in `]a, b[ -> sgr (p^`()).[x] = 1. move=> hwlog. have := (polyrN0_itv hp'). @@ -880,15 +883,13 @@ Qed. (* Todo : Lemmas about operations of intervall : itversection, restriction and splitting *) -Lemma cat_roots_on : forall (p : {poly R}) a b x, - x \in `]a, b[ -> ~~ (root p x) - -> forall s s', - sorted <%R s -> sorted <%R s' - -> roots_on p `]a, x[ s -> roots_on p `]x, b[ s' - -> roots_on p `]a, b[ (s ++ s'). -Proof. -move=> p a b x hx /= npx0 s. -elim: s a hx => [|y s ihs] a hx s' //= ss ss'. +Lemma cat_roots_on (p : {poly R}) a b x : + x \in `]a, b[ -> ~~ (root p x) -> + forall s s', sorted <%R s -> sorted <%R s' -> + roots_on p `]a, x[ s -> roots_on p `]x, b[ s' -> + roots_on p `]a, b[ (s ++ s'). +Proof. +move=> hx /= npx0 s; elim: s a hx => [|y s ihs] a hx s' //= ss ss'. move/roots_on_nil=> hax hs' y. rewrite -hs'; case py0: root; rewrite ?(andbT, andbF) //. rewrite (itv_splitU2 hx); case: (y \in `]x, b[); rewrite ?orbF ?orbT //=. @@ -913,10 +914,10 @@ CoInductive roots_spec (p : {poly R}) (i : pred R) (s : seq R) : & sorted <%R s : roots_spec p i s p false s. (* still much too long *) -Lemma itv_roots : forall (p :{poly R}) (a b : R), +Lemma itv_roots (p :{poly R}) (a b : R) : {s : seq R & roots_spec p (topred `]a, b[) s p (p == 0) s}. Proof. -move=> p a b; case p0: (_ == 0). +case p0: (_ == 0). by rewrite (eqP p0); exists [::]; constructor. elim: (size p) {-2}p (leqnn (size p)) p0 a b => {p} [| n ihn] p sp p0 a b. by exists [::]; move: p0; rewrite -size_poly_eq0 -leqn0 sp. @@ -985,65 +986,61 @@ Qed. Definition roots (p : {poly R}) a b := projT1 (itv_roots p a b). -Lemma rootsP : forall p a b, +Lemma rootsP p a b : roots_spec p (topred `]a, b[) (roots p a b) p (p == 0) (roots p a b). -Proof. by move=> p a b; rewrite /roots; case hp: itv_roots. Qed. +Proof. by rewrite /roots; case hp: itv_roots. Qed. -Lemma roots0 : forall a b, roots 0 a b = [::]. -Proof. by move=> a b; case: rootsP=> //=; rewrite eqxx. Qed. +Lemma roots0 a b : roots 0 a b = [::]. +Proof. by case: rootsP=> //=; rewrite eqxx. Qed. Lemma roots_on_roots : forall p a b, p != 0 -> roots_on p `]a, b[ (roots p a b). Proof. by move=> a b p; case:rootsP. Qed. Hint Resolve roots_on_roots. -Lemma sorted_roots : forall a b p, sorted <%R (roots p a b). -Proof. by move=> p a b; case: rootsP. Qed. +Lemma sorted_roots a b p : sorted <%R (roots p a b). +Proof. by case: rootsP. Qed. Hint Resolve sorted_roots. -Lemma path_roots : forall p a b, path <%R a (roots p a b). +Lemma path_roots p a b : path <%R a (roots p a b). Proof. -move=> p a b; case: rootsP=> //= p0 hp sp; rewrite path_min_sorted //. +case: rootsP=> //= p0 hp sp; rewrite path_min_sorted //. by move=> y; rewrite -hp; case/andP; move/itvP->. Qed. Hint Resolve path_roots. -Lemma root_is_roots : - forall (p : {poly R}) (a b : R), p != 0 -> - forall x, x \in `]a, b[ -> root p x = (x \in roots p a b). -Proof. -by move=> p a b; case: rootsP=> // p0 hs ps _ y hy /=; rewrite -hs hy. -Qed. +Lemma root_is_roots (p : {poly R}) (a b : R) : + p != 0 -> forall x, x \in `]a, b[ -> root p x = (x \in roots p a b). +Proof. by case: rootsP=> // p0 hs ps _ y hy /=; rewrite -hs hy. Qed. -Lemma root_in_roots : forall (p : {poly R}) a b, p != 0 -> - forall x, x \in `]a, b[ -> root p x -> x \in (roots p a b). -Proof. by move=> p a b p0 x axb rpx; rewrite -root_is_roots. Qed. +Lemma root_in_roots (p : {poly R}) a b : + p != 0 -> forall x, x \in `]a, b[ -> root p x -> x \in (roots p a b). +Proof. by move=> p0 x axb rpx; rewrite -root_is_roots. Qed. -Lemma root_roots : forall p a b x, x \in roots p a b -> root p x. -Proof. by move=> p a b x; case: rootsP=> // p0 <- _; case/andP. Qed. +Lemma root_roots p a b x : x \in roots p a b -> root p x. +Proof. by case: rootsP=> // p0 <- _; case/andP. Qed. -Lemma roots_nil : forall p a b, p != 0 -> +Lemma roots_nil p a b : p != 0 -> roots p a b = [::] -> {in `]a, b[, noroot p}. Proof. -move=> p a b; case: rootsP=> // p0 hs ps _ s0 x axb. +case: rootsP => // p0 hs ps _ s0 x axb. by move: (hs x); rewrite s0 in_nil !axb /= => ->. Qed. Lemma roots_in p a b x : x \in roots p a b -> x \in `]a, b[. -Proof. by case: rootsP=> //= np0 ron_p *; exact: (roots_on_in ron_p). Qed. +Proof. by case: rootsP=> //= np0 ron_p *; apply: (roots_on_in ron_p). Qed. -Lemma rootsEba : forall p a b, b <= a -> roots p a b = [::]. +Lemma rootsEba p a b : b <= a -> roots p a b = [::]. Proof. -move=> p a b; case: rootsP=> // p0; case: (roots _ _ _)=> [|x s] hs ps ba //; +case: rootsP=> // p0; case: (roots _ _ _) => [|x s] hs ps ba //; by move: (hs x); rewrite itv_gte //= mem_head. Qed. -Lemma roots_on_uniq : forall p a b s1 s2, - sorted <%R s1 -> sorted <%R s2 -> +Lemma roots_on_uniq p a b s1 s2 : + sorted <%R s1 -> sorted <%R s2 -> roots_on p `]a, b[ s1 -> roots_on p `]a, b[ s2 -> s1 = s2. Proof. -move=> p a b s1. -elim: s1 p a b => [| r1 s1 ih] p a b [| r2 s2] ps1 ps2 rs1 rs2 //. +elim: s1 p a b s2 => [| r1 s1 ih] p a b [| r2 s2] ps1 ps2 rs1 rs2 //. - have rpr2 : root p r2 by apply: (roots_on_root rs2); rewrite mem_head. have abr2 : r2 \in `]a, b[ by apply: (roots_on_in rs2); rewrite mem_head. by have:= (rs1 r2); rewrite rpr2 !abr2 in_nil. @@ -1068,12 +1065,11 @@ move: ps1 ps2=> /=; move/path_sorted=> hs1; move/path_sorted=> hs2. exact: (ih p _ b _ hs1 hs2 h3 h4). Qed. -Lemma roots_eq : forall (p q : {poly R}) (a b : R), - p != 0 -> q != 0 -> - ({in `]a, b[, root p =1 root q} - <-> roots p a b = roots q a b). +Lemma roots_eq (p q : {poly R}) (a b : R) : + p != 0 -> q != 0 -> + ({in `]a, b[, root p =1 root q} <-> roots p a b = roots q a b). Proof. -move=> p q a b p0 q0. +move=> p0 q0. case hab : (a < b); last first. split; first by rewrite !rootsEba // lerNgt hab. move=> _ x. rewrite !inE; case/andP=> ax xb. @@ -1088,24 +1084,24 @@ move=> x axb /=. by rewrite (@root_is_roots q a b) // (@root_is_roots p a b) // hpq. Qed. -Lemma roots_opp : forall p, roots (- p) =2 roots p. +Lemma roots_opp p : roots (- p) =2 roots p. Proof. -move=> p a b; case p0 : (p == 0); first by rewrite (eqP p0) oppr0. +move=> a b; case p0 : (p == 0); first by rewrite (eqP p0) oppr0. by apply/roots_eq=> [||x]; rewrite ?oppr_eq0 ?p0 ?rootN. Qed. -Lemma no_root_roots : forall (p : {poly R}) a b, +Lemma no_root_roots (p : {poly R}) a b : {in `]a, b[ , noroot p} -> roots p a b = [::]. Proof. -move=> p a b hr; case: rootsP=> // p0 hs ps. +move=> hr; case: rootsP => // p0 hs ps. apply: (@roots_on_uniq p a b)=> // x; rewrite in_nil. by apply/negP; case/andP; move/hr; move/negPf->. Qed. -Lemma head_roots_on_ge : forall p a b s, a < b -> - roots_on p `]a, b[ s -> a < head b s. +Lemma head_roots_on_ge p a b s : + a < b -> roots_on p `]a, b[ s -> a < head b s. Proof. -move=> p a b [|x s] ab //; move/(_ x). +case: s => [|x s] ab // /(_ x). by rewrite in_cons eqxx; case/andP; case/andP. Qed. @@ -1114,32 +1110,29 @@ Proof. by move=> p a b; case: rootsP=> // *; apply: head_roots_on_ge. Qed. -Lemma last_roots_on_le : forall p a b s, a < b -> - roots_on p `]a, b[ s -> last a s < b. +Lemma last_roots_on_le p a b s : + a < b -> roots_on p `]a, b[ s -> last a s < b. Proof. -move=> p a b [|x s] ab rs //. +case: s => [|x s] ab rs //. by rewrite (itvP (roots_on_in rs _)) //= mem_last. Qed. -Lemma last_roots_le : forall p a b, a < b -> last a (roots p a b) < b. -Proof. -by move=> p a b; case: rootsP=> // *; apply: last_roots_on_le. -Qed. +Lemma last_roots_le p a b : a < b -> last a (roots p a b) < b. +Proof. by case: rootsP=> // *; apply: last_roots_on_le. Qed. -Lemma roots_uniq : forall p a b s, p != 0 -> - roots_on p `]a, b[ s -> sorted <%R s -> s = roots p a b. +Lemma roots_uniq p a b s : + p != 0 -> roots_on p `]a, b[ s -> sorted <%R s -> s = roots p a b. Proof. -move=> p a b s; case: rootsP=> // p0 hs' ps' _ hs ss. +case: rootsP=> // p0 hs' ps' _ hs ss. exact: (@roots_on_uniq p a b)=> //. Qed. -Lemma roots_cons : forall p a b x s, +Lemma roots_cons p a b x s : (roots p a b == x :: s) - = [&& p != 0, x \in `]a, b[, - (roots p a x == [::]), - (root p x) & (roots p x b == s)]. + = [&& p != 0, x \in `]a, b[, + roots p a x == [::], root p x & roots p x b == s]. Proof. -move=> p a b x s; case: rootsP=> // p0 hs' ps' /=. +case: rootsP=> // p0 hs' ps' /=. apply/idP/idP. move/eqP=> es'; move: ps' hs'; rewrite es' /= => sxs. case/min_roots_on; first by apply: order_path_min=> //; apply: ltr_trans. @@ -1158,13 +1151,12 @@ case hy: (y == x); first by rewrite (eqP hy) px0 orbT. by rewrite andFb orFb rax rxb in_nil. Qed. -Lemma roots_rcons : forall p a b x s, - (roots p a b == rcons s x) - = [&& p != 0, x \in `]a , b[, - (roots p x b == [::]), - (root p x) & (roots p a x == s)]. +Lemma roots_rcons p a b x s : + (roots p a b == rcons s x) = + [&& p != 0, x \in `]a , b[, + roots p x b == [::], root p x & roots p a x == s]. Proof. -move=> p a b x s; case: rootsP; first by case: s. +case: rootsP; first by case: s. move=> // p0 hs' ps' /=. apply/idP/idP. move/eqP=> es'; move: ps' hs'; rewrite es' /= => sxs. @@ -1199,10 +1191,11 @@ Implicit Types a b : R. Implicit Types p : {poly R}. -Definition next_root (p : {poly R}) x b := if p == 0 then x else head (maxr b x) (roots p x b). +Definition next_root (p : {poly R}) x b := + if p == 0 then x else head (maxr b x) (roots p x b). -Lemma next_root0 : forall a b, next_root 0 a b = a. -Proof. by move=> *; rewrite /next_root eqxx. Qed. +Lemma next_root0 a b : next_root 0 a b = a. +Proof. by rewrite /next_root eqxx. Qed. CoInductive next_root_spec (p : {poly R}) x b : bool -> R -> Type := | NextRootSpec0 of p = 0 : next_root_spec p x b true x @@ -1211,10 +1204,10 @@ CoInductive next_root_spec (p : {poly R}) x b : bool -> R -> Type := | NextRootSpecNoRoot c of p != 0 & c = maxr b x & {in `]x, b[, forall z, ~~(root p z)} : next_root_spec p x b (p.[c] == 0) c. -Lemma next_rootP : forall (p : {poly R}) a b, next_root_spec p a b (p.[next_root p a b] == 0) (next_root p a b). +Lemma next_rootP (p : {poly R}) a b : + next_root_spec p a b (p.[next_root p a b] == 0) (next_root p a b). Proof. -move=> p a b; rewrite /next_root /=. -case hs: roots=> [|x s] /=. +rewrite /next_root /=; case hs: roots => [|x s] /=. case: (altP (p =P 0))=> p0. by rewrite {2}p0 hornerC eqxx; constructor; rewrite p0. by constructor=> // y hy; apply: (roots_nil p0 hs). @@ -1223,33 +1216,34 @@ rewrite (negPf p0) (rootPt rpx); constructor=> //; first by move/eqP: rpx. by move=> y hy /=; move/(roots_nil p0): (rap); apply. Qed. -Lemma next_root_in : forall p a b, next_root p a b \in `[a, maxr b a]. +Lemma next_root_in p a b : next_root p a b \in `[a, maxr b a]. Proof. -move=> p a b; case: next_rootP=> [p0|y np0 py0 hy _|c np0 hc _]. +case: next_rootP => [p0|y np0 py0 hy _|c np0 hc _]. * by rewrite bound_in_itv /= ler_maxr lerr orbT. * by apply: subitvP hy=> /=; rewrite ler_maxr !lerr. * by rewrite hc bound_in_itv /= ler_maxr lerr orbT. Qed. -Lemma next_root_gt : forall p a b, a < b -> p != 0 -> next_root p a b > a. +Lemma next_root_gt p a b : a < b -> p != 0 -> next_root p a b > a. Proof. -move=> p a b ab np0; case: next_rootP=> [p0|y _ py0 hy _|c _ -> _]. +move=> ab np0; case: next_rootP=> [p0|y _ py0 hy _|c _ -> _]. * by rewrite p0 eqxx in np0. * by rewrite (itvP hy). * by rewrite maxr_l // ltrW. Qed. -Lemma next_noroot : forall p a b, {in `]a, (next_root p a b)[, noroot p}. +Lemma next_noroot p a b : {in `]a, (next_root p a b)[, noroot p}. Proof. -move=> p a b z; case: next_rootP; first by rewrite itv_xx. +move=> z; case: next_rootP; first by rewrite itv_xx. by move=> y np0 py0 hy hp hz; rewrite (negPf (hp _ _)). move=> c p0 -> hp hz; rewrite (negPf (hp _ _)) //. by case: maxrP hz; last by rewrite itv_xx. Qed. -Lemma is_next_root : forall p a b x, next_root_spec p a b (root p x) x -> x = next_root p a b. +Lemma is_next_root p a b x : + next_root_spec p a b (root p x) x -> x = next_root p a b. Proof. -move=> p a b x []; first by move->; rewrite /next_root eqxx. +case; first by move->; rewrite /next_root eqxx. move=> y; case: next_rootP; first by move->; rewrite eqxx. move=> y' np0 py'0 hy' hp' _ py0 hy hp. wlog: y y' hy' hy hp' hp py0 py'0 / y <= y'. @@ -1258,15 +1252,16 @@ move=> p a b x []; first by move->; rewrite /next_root eqxx. by rewrite rootE py0 eqxx inE /= (itvP hy) hyy'; move/(_ isT). move=> c p0 ->; case: maxrP=> hab; last by rewrite itv_gte //= ltrW. by move=> hpz _ py0 hy; move/hpz:hy; rewrite rootE py0 eqxx. -case: next_rootP=> //; first by move->; rewrite eqxx. +case: next_rootP => //; first by move->; rewrite eqxx. by move=> y np0 py0 hy _ c _ _; move/(_ _ hy); rewrite rootE py0 eqxx. by move=> c _ -> _ c' _ ->. Qed. -Definition prev_root (p : {poly R}) a x := if p == 0 then x else last (minr a x) (roots p a x). +Definition prev_root (p : {poly R}) a x := + if p == 0 then x else last (minr a x) (roots p a x). -Lemma prev_root0 : forall a b, prev_root 0 a b = b. -Proof. by move=> *; rewrite /prev_root eqxx. Qed. +Lemma prev_root0 a b : prev_root 0 a b = b. +Proof. by rewrite /prev_root eqxx. Qed. CoInductive prev_root_spec (p : {poly R}) a x : bool -> R -> Type := | PrevRootSpec0 of p = 0 : prev_root_spec p a x true x @@ -1275,10 +1270,10 @@ CoInductive prev_root_spec (p : {poly R}) a x : bool -> R -> Type := | PrevRootSpecNoRoot c of p != 0 & c = minr a x & {in `]a, x[, forall z, ~~(root p z)} : prev_root_spec p a x (p.[c] == 0) c. -Lemma prev_rootP : forall (p : {poly R}) a b, prev_root_spec p a b (p.[prev_root p a b] == 0) (prev_root p a b). +Lemma prev_rootP (p : {poly R}) a b : + prev_root_spec p a b (p.[prev_root p a b] == 0) (prev_root p a b). Proof. -move=> p a b; rewrite /prev_root /=. -move hs: (roots _ _ _)=> s. +rewrite /prev_root /=; move hs: (roots _ _ _)=> s. case: (lastP s) hs=> {s} [|s x] hs /=. case: (altP (p =P 0))=> p0. by rewrite {2}p0 hornerC eqxx; constructor; rewrite p0. @@ -1289,33 +1284,34 @@ rewrite (negPf p0) (rootPt rpx); constructor=> //; first by move/eqP: rpx. by move=> y hy /=; move/(roots_nil p0): (rap); apply. Qed. -Lemma prev_root_in : forall p a b, prev_root p a b \in `[minr a b, b]. +Lemma prev_root_in p a b : prev_root p a b \in `[minr a b, b]. Proof. -move=> p a b; case: prev_rootP=> [p0|y np0 py0 hy _|c np0 hc _]. +case: prev_rootP => [p0|y np0 py0 hy _|c np0 hc _]. * by rewrite bound_in_itv /= ler_minl lerr orbT. * by apply: subitvP hy=> /=; rewrite ler_minl !lerr. * by rewrite hc bound_in_itv /= ler_minl lerr orbT. Qed. -Lemma prev_noroot : forall p a b, {in `](prev_root p a b), b[, noroot p}. +Lemma prev_noroot p a b : {in `](prev_root p a b), b[, noroot p}. Proof. -move=> p a b z; case: prev_rootP; first by rewrite itv_xx. +move=> z; case: prev_rootP; first by rewrite itv_xx. by move=> y np0 py0 hy hp hz; rewrite (negPf (hp _ _)). move=> c np0 ->; case: minrP=> hab; last by rewrite itv_xx. by move=> hp hz; rewrite (negPf (hp _ _)). Qed. -Lemma prev_root_lt : forall p a b, a < b -> p != 0 -> prev_root p a b < b. +Lemma prev_root_lt p a b : a < b -> p != 0 -> prev_root p a b < b. Proof. -move=> p a b ab np0; case: prev_rootP=> [p0|y _ py0 hy _|c _ -> _]. +move=> ab np0; case: prev_rootP=> [p0|y _ py0 hy _|c _ -> _]. * by rewrite p0 eqxx in np0. * by rewrite (itvP hy). * by rewrite minr_l // ltrW. Qed. -Lemma is_prev_root : forall p a b x, prev_root_spec p a b (root p x) x -> x = prev_root p a b. +Lemma is_prev_root p a b x : + prev_root_spec p a b (root p x) x -> x = prev_root p a b. Proof. -move=> p a b x []; first by move->; rewrite /prev_root eqxx. +case; first by move->; rewrite /prev_root eqxx. move=> y; case: prev_rootP; first by move->; rewrite eqxx. move=> y' np0 py'0 hy' hp' _ py0 hy hp. wlog: y y' hy' hy hp' hp py0 py'0 / y <= y'. @@ -1334,13 +1330,13 @@ Definition neighpr p a b := `]a, (next_root p a b)[. Definition neighpl p a b := `](prev_root p a b), b[. -Lemma neighpl_root : forall p a x, {in neighpl p a x, noroot p}. +Lemma neighpl_root p a x : {in neighpl p a x, noroot p}. Proof. exact: prev_noroot. Qed. -Lemma sgr_neighplN : forall p a x, ~~root p x -> - {in neighpl p a x, forall y, (sgr p.[y] = sgr p.[x])}. +Lemma sgr_neighplN p a x : + ~~ root p x -> {in neighpl p a x, forall y, (sgr p.[y] = sgr p.[x])}. Proof. -rewrite /neighpl=> p a x nrpx /= y hy. +rewrite /neighpl=> nrpx /= y hy. apply: (@polyrN0_itv `[y, x]); do ?by rewrite bound_in_itv /= (itvP hy). move=> z; rewrite (@itv_splitU _ x false) ?itv_xx /=; last first. (* Todo : Lemma itv_splitP *) @@ -1350,19 +1346,19 @@ rewrite (@prev_noroot _ a x) //. by apply: subitvPl hz; rewrite /= (itvP hy). Qed. -Lemma sgr_neighpl_same : forall p a x, +Lemma sgr_neighpl_same p a x : {in neighpl p a x &, forall y z, (sgr p.[y] = sgr p.[z])}. Proof. -by rewrite /neighpl=> p x b y z *; apply: (polyrN0_itv (@prev_noroot p x b)). +by rewrite /neighpl=> y z *; apply: (polyrN0_itv (@prev_noroot p a x)). Qed. -Lemma neighpr_root : forall p x b, {in neighpr p x b, noroot p}. +Lemma neighpr_root p x b : {in neighpr p x b, noroot p}. Proof. exact: next_noroot. Qed. -Lemma sgr_neighprN : forall p x b, p.[x] != 0 -> - {in neighpr p x b, forall y, (sgr p.[y] = sgr p.[x])}. +Lemma sgr_neighprN p x b : + p.[x] != 0 -> {in neighpr p x b, forall y, (sgr p.[y] = sgr p.[x])}. Proof. -rewrite /neighpr=> p x b nrpx /= y hy; symmetry. +rewrite /neighpr=> nrpx /= y hy; symmetry. apply: (@polyrN0_itv `[x, y]); do ?by rewrite bound_in_itv /= (itvP hy). move=> z; rewrite (@itv_splitU _ x true) ?itv_xx /=; last first. (* Todo : Lemma itv_splitP *) @@ -1372,39 +1368,39 @@ rewrite (@next_noroot _ x b) //. by apply: subitvPr hz; rewrite /= (itvP hy). Qed. -Lemma sgr_neighpr_same : forall p x b, +Lemma sgr_neighpr_same p x b : {in neighpr p x b &, forall y z, (sgr p.[y] = sgr p.[z])}. Proof. -by rewrite /neighpl=> p x b y z *; apply: (polyrN0_itv (@next_noroot p x b)). +by rewrite /neighpl=> y z *; apply: (polyrN0_itv (@next_noroot p x b)). Qed. -Lemma uniq_roots : forall a b p, uniq (roots p a b). +Lemma uniq_roots a b p : uniq (roots p a b). Proof. -move=> a b p; case p0: (p == 0); first by rewrite (eqP p0) roots0. -by apply: (@sorted_uniq _ <%R); [exact: ltr_trans | exact: ltrr|]. +case p0: (p == 0); first by rewrite (eqP p0) roots0. +by apply: (@sorted_uniq _ <%R); [apply: ltr_trans | apply: ltrr|]. Qed. Hint Resolve uniq_roots. -Lemma in_roots : forall p a b, forall x : R, +Lemma in_roots p (a b x : R) : (x \in roots p a b) = [&& root p x, x \in `]a, b[ & p != 0]. Proof. -move=> p a b x; case: rootsP=> //=; first by rewrite in_nil !andbF. +case: rootsP=> //=; first by rewrite in_nil !andbF. by move=> p0 hr sr; rewrite andbT -hr andbC. Qed. (* Todo : move to polyorder => need char 0 *) -Lemma gdcop_eq0 : forall p q, (gdcop p q == 0) = (q == 0) && (p != 0). +Lemma gdcop_eq0 p q : (gdcop p q == 0) = (q == 0) && (p != 0). Proof. -move=> p q; case: (eqVneq q 0) => [-> | q0]. +case: (eqVneq q 0) => [-> | q0]. rewrite gdcop0 /= eqxx /=. by case: (eqVneq p 0) => [-> | pn0]; rewrite ?(negPf pn0) eqxx ?oner_eq0. rewrite /gdcop; move: {-1}(size q) (leqnn (size q))=> k hk. case: (eqVneq p 0) => [-> | p0]. rewrite eqxx andbF; apply: negPf. elim: k q q0 {hk} => [|k ihk] q q0 /=; first by rewrite eqxx oner_eq0. - case: ifP=> _ //. - apply: ihk; rewrite gcdp0 divpp ?q0 // polyC_eq0; exact: lc_expn_scalp_neq0. + case: ifP => _ //. + by apply: ihk; rewrite gcdp0 divpp ?q0 // polyC_eq0; apply/lc_expn_scalp_neq0. rewrite p0 (negPf q0) /=; apply: negPf. elim: k q q0 hk => [|k ihk] /= q q0 hk. by move: hk q0; rewrite leqn0 size_poly_eq0; move->. @@ -1426,11 +1422,11 @@ case: n sgcd => [|n]; first by move/eqP; rewrite size_poly_eq1 gcdp_eqp1 cpq. by rewrite addnS /= -{1}[size (_ %/ _)]addn0 ltn_add2l. Qed. -Lemma roots_mul : forall a b, a < b -> forall p q, +Lemma roots_mul a b : a < b -> forall p q, p != 0 -> q != 0 -> perm_eq (roots (p*q) a b) (roots p a b ++ roots ((gdcop p q)) a b). Proof. -move=> a b hab p q np0 nq0. +move=> hab p q np0 nq0. apply: uniq_perm_eq; first exact: uniq_roots. rewrite cat_uniq ?uniq_roots andbT /=; apply/hasPn=> x /=. move/root_roots; rewrite root_gdco //; case/andP=> _. @@ -1441,12 +1437,13 @@ case: (x \in `]_, _[); last by rewrite !andbF. by rewrite negb_or !np0 !nq0 !andbT /=; do 2?case: root=> //=. Qed. -Lemma roots_mul_coprime : forall a b, a < b -> forall p q, - p != 0 -> q != 0 -> coprimep p q -> +Lemma roots_mul_coprime a b : + a < b -> + forall p q, p != 0 -> q != 0 -> coprimep p q -> perm_eq (roots (p * q) a b) (roots p a b ++ roots q a b). Proof. -move=> a b hab p q np0 nq0 cpq. +move=> hab p q np0 nq0 cpq. rewrite (perm_eq_trans (roots_mul hab np0 nq0)) //. suff ->: roots (gdcop p q) a b = roots q a b by apply: perm_eq_refl. case: gdcopP=> r rq hrp; move/(_ q (dvdpp _)). @@ -1458,10 +1455,10 @@ by rewrite -size_poly_eq0 (eqp_size erq) size_poly_eq0. Qed. -Lemma next_rootM : forall a b (p q : {poly R}), +Lemma next_rootM a b (p q : {poly R}) : next_root (p * q) a b = minr (next_root p a b) (next_root q a b). Proof. -move=> a b p q; symmetry; apply: is_next_root. +symmetry; apply: is_next_root. wlog: p q / next_root p a b <= next_root q a b. case: minrP=> hpq; first by move/(_ _ _ hpq); case: minrP hpq. by move/(_ _ _ (ltrW hpq)); rewrite mulrC minrC; case: minrP hpq. @@ -1481,17 +1478,17 @@ case: minrP=> //; case: next_rootP=> [|y np0 py0 hy|c np0 ->] hp hpq _. by move: hpq; rewrite ler_maxl; case/andP. Qed. -Lemma neighpr_mul : forall a b p q, +Lemma neighpr_mul a b p q : (neighpr (p * q) a b) =i [predI (neighpr p a b) & (neighpr q a b)]. Proof. -move=> a b p q x; rewrite inE /= !inE /= next_rootM. +move=> x; rewrite inE /= !inE /= next_rootM. by case: (a < x); rewrite // ltr_minr. Qed. -Lemma prev_rootM : forall a b (p q : {poly R}), +Lemma prev_rootM a b (p q : {poly R}) : prev_root (p * q) a b = maxr (prev_root p a b) (prev_root q a b). Proof. -move=> a b p q; symmetry; apply: is_prev_root. +symmetry; apply: is_prev_root. wlog: p q / prev_root p a b >= prev_root q a b. case: maxrP=> hpq; first by move/(_ _ _ hpq); case: maxrP hpq. by move/(_ _ _ (ltrW hpq)); rewrite mulrC maxrC; case: maxrP hpq. @@ -1511,22 +1508,22 @@ case: maxrP=> //; case: (@prev_rootP p)=> [|y np0 py0 hy|c np0 ->] hp hpq _. by move: hpq; rewrite ler_minr; case/andP. Qed. -Lemma neighpl_mul : forall a b p q, +Lemma neighpl_mul a b p q : (neighpl (p * q) a b) =i [predI (neighpl p a b) & (neighpl q a b)]. Proof. -move=> a b p q x; rewrite !inE /= prev_rootM. +move=> x; rewrite !inE /= prev_rootM. by case: (x < b); rewrite // ltr_maxl !(andbT, andbF). Qed. -Lemma neighpr_wit : forall p x b, x < b -> p != 0 -> {y | y \in neighpr p x b}. +Lemma neighpr_wit p x b : x < b -> p != 0 -> {y | y \in neighpr p x b}. Proof. -move=> p x b xb; exists (mid x (next_root p x b)). +move=> xb; exists (mid x (next_root p x b)). by rewrite mid_in_itv //= next_root_gt. Qed. -Lemma neighpl_wit : forall p a x, a < x -> p != 0 -> {y | y \in neighpl p a x}. +Lemma neighpl_wit p a x : a < x -> p != 0 -> {y | y \in neighpl p a x}. Proof. -move=> p a x xb; exists (mid (prev_root p a x) x). +move=> xb; exists (mid (prev_root p a x) x). by rewrite mid_in_itv //= prev_root_lt. Qed. @@ -1543,13 +1540,12 @@ Definition sgp_right (p : {poly R}) x := else 0 in aux p (size p). -Lemma sgp_right0 : forall x, sgp_right 0 x = 0. -Proof. by move=> x; rewrite /sgp_right size_poly0. Qed. +Lemma sgp_right0 x : sgp_right 0 x = 0. +Proof. by rewrite /sgp_right size_poly0. Qed. -Lemma sgr_neighpr : forall b p x, +Lemma sgr_neighpr b p x : {in neighpr p x b, forall y, (sgr p.[y] = sgp_right p x)}. Proof. -move=> b p x. elim: (size p) {-2}p (leqnn (size p))=> [|n ihn] {p} p. rewrite leqn0 size_poly_eq0 /neighpr; move/eqP=> -> /=. by move=>y; rewrite next_root0 itv_xx. @@ -1585,12 +1581,11 @@ case: (@neighpr_wit (p * p^`()) x b)=> [||m hm]. by rewrite (subitvP _ (@mid_in_itv _ true true _ _ _)) //= ?lerr (itvP hmp'). Qed. -Lemma sgr_neighpl : forall a p x, +Lemma sgr_neighpl a p x : {in neighpl p a x, forall y, (sgr p.[y] = (-1) ^+ (odd (\mu_x p)) * sgp_right p x) }. Proof. -move=> a p x. elim: (size p) {-2}p (leqnn (size p))=> [|n ihn] {p} p. rewrite leqn0 size_poly_eq0 /neighpl; move/eqP=> -> /=. by move=>y; rewrite prev_root0 itv_xx. @@ -1627,31 +1622,30 @@ case: (@neighpl_wit (p * p^`()) a x)=> [||m hm]. move=> u hu /=; rewrite (@prev_noroot _ a x) //. by apply: subitvPl hu; rewrite /= (itvP hmp'). rewrite ihn ?size_deriv ?sp /neighpl //; last first. - by rewrite (subitvP _ (@mid_in_itv _ true true _ _ _)) //= ?lerr (itvP hmp'). + by rewrite (subitvP _ (@mid_in_itv _ true true _ _ _)) //= + ?lerr (itvP hmp'). rewrite mu_deriv // odd_sub ?mu_gt0 //=; last by rewrite -size_poly_eq0 sp. by rewrite signr_addb /= mulrN1 mulNr opprK. Qed. -Lemma sgp_right_deriv : forall (p : {poly R}) x, root p x -> - sgp_right p x = sgp_right (p^`()) x. +Lemma sgp_right_deriv (p : {poly R}) x : + root p x -> sgp_right p x = sgp_right (p^`()) x. Proof. -move=> p; elim: (size p) {-2}p (erefl (size p)) => {p} [p|sp hp p hsp x]. +elim: (size p) {-2}p (erefl (size p)) x => {p} [p|sp hp p hsp x]. by move/eqP; rewrite size_poly_eq0; move/eqP=> -> x _; rewrite derivC. by rewrite /sgp_right size_deriv hsp /= => ->. Qed. -Lemma sgp_rightNroot : forall (p : {poly R}) x, +Lemma sgp_rightNroot (p : {poly R}) x : ~~ root p x -> sgp_right p x = sgr p.[x]. Proof. -move=> p x nrpx; rewrite /sgp_right; case hsp: (size _)=> [|sp]. +move=> nrpx; rewrite /sgp_right; case hsp: (size _)=> [|sp]. by move/eqP:hsp; rewrite size_poly_eq0; move/eqP->; rewrite hornerC sgr0. by rewrite nrpx. Qed. -Lemma sgp_right_mul : forall p q x, - sgp_right (p * q) x = sgp_right p x * sgp_right q x. +Lemma sgp_right_mul p q x : sgp_right (p * q) x = sgp_right p x * sgp_right q x. Proof. -move=> p q x. case: (altP (q =P 0))=> q0; first by rewrite q0 /sgp_right !(size_poly0,mulr0). case: (altP (p =P 0))=> p0; first by rewrite p0 /sgp_right !(size_poly0,mul0r). case: (@neighpr_wit (p * q) x (1 + x))=> [||m hpq]; do ?by rewrite mulf_neq0. @@ -1667,27 +1661,26 @@ case c0: (c == 0); first by rewrite (eqP c0) scale0r sgr0 mul0r sgp_right0. by rewrite -mul_polyC sgp_right_mul sgp_rightNroot ?hornerC ?rootC ?c0. Qed. -Lemma sgp_right_square : forall p x, p != 0 -> sgp_right p x * sgp_right p x = 1. +Lemma sgp_right_square p x : p != 0 -> sgp_right p x * sgp_right p x = 1. Proof. -move=> p x np0; case: (@neighpr_wit p x (1 + x))=> [||m hpq] //. +move=> np0; case: (@neighpr_wit p x (1 + x))=> [||m hpq] //. by rewrite ltr_spaddl ?ltr01. rewrite -(@sgr_neighpr (1 + x) _ _ m) //. by rewrite -expr2 sqr_sg (@next_noroot _ x (1 + x)). Qed. -Lemma sgp_right_rec p x : sgp_right p x = - if p == 0 then 0 else - if ~~ root p x then sgr p.[x] - else sgp_right p^`() x. +Lemma sgp_right_rec p x : + sgp_right p x = + (if p == 0 then 0 else if ~~ root p x then sgr p.[x] else sgp_right p^`() x). Proof. -rewrite /sgp_right; case hs: size=> [|s]; rewrite -size_poly_eq0 hs //=. +rewrite /sgp_right; case hs: size => [|s]; rewrite -size_poly_eq0 hs //=. by rewrite size_deriv hs. Qed. -Lemma sgp_right_addp0 : forall (p q : {poly R}) x, q != 0 -> - (\mu_x p > \mu_x q)%N -> sgp_right (p + q) x = sgp_right q x. +Lemma sgp_right_addp0 (p q : {poly R}) x : + q != 0 -> (\mu_x p > \mu_x q)%N -> sgp_right (p + q) x = sgp_right q x. Proof. -move=> p q x nq0; move hm: (\mu_x q)=> m. +move=> nq0; move hm: (\mu_x q)=> m. elim: m p q nq0 hm => [|mq ihmq] p q nq0 hmq; case hmp: (\mu_x p)=> // [mp]; do[ rewrite ltnS=> hm; @@ -1754,10 +1747,10 @@ by rewrite /rscalp unlock /= eqxx /= expr0 oner_neq0. Qed. Notation lcn_neq0 := lc_expn_rscalp_neq0. -Lemma sgp_right_mod : forall p q x, (\mu_x p < \mu_x q)%N -> +Lemma sgp_right_mod p q x : (\mu_x p < \mu_x q)%N -> sgp_right (rmodp p q) x = (sgr (lead_coef q)) ^+ (rscalp p q) * sgp_right p x. Proof. -move=> p q x mupq; case p0: (p == 0). +move=> mupq; case p0: (p == 0). by rewrite (eqP p0) rmod0p !sgp_right0 mulr0. have qn0 : q != 0. by apply/negP; move/eqP=> q0; rewrite q0 mu0 ltn0 in mupq. @@ -1804,7 +1797,7 @@ Lemma rootsR0 : rootsR 0 = [::]. Proof. exact: roots0. Qed. Lemma rootsRC c : rootsR c%:P = [::]. Proof. exact: rootsC. Qed. Lemma rootsRP p a b : - {in `]-oo, a], noroot p} -> {in `[b , +oo[, noroot p} -> + {in `]-oo, a], noroot p} -> {in `[b , +oo[, noroot p} -> roots p a b = rootsR p. Proof. move=> rpa rpb. @@ -1816,7 +1809,8 @@ apply: contraLR rpx; rewrite inE negb_and -!lerNgt. by move=> /orP[/rpa //|xb]; rewrite rpb // inE andbT. Qed. -Lemma sgp_pinftyP x (p : {poly R}) : {in `[x , +oo[, noroot p} -> +Lemma sgp_pinftyP x (p : {poly R}) : + {in `[x , +oo[, noroot p} -> {in `[x, +oo[, forall y, sgr p.[y] = sgp_pinfty p}. Proof. rewrite /sgp_pinfty; wlog lp_gt0 : x p / lead_coef p > 0 => [hwlog|rpx y Hy]. @@ -1831,7 +1825,8 @@ rewrite (@polyrN0_itv _ _ rpx (maxr y z)) ?inE ?ler_maxr ?(itvP Hy) //. by rewrite Hz ?gtr0_sg // inE ler_maxr lerr orbT. Qed. -Lemma sgp_minftyP x (p : {poly R}) : {in `]-oo, x], noroot p} -> +Lemma sgp_minftyP x (p : {poly R}) : + {in `]-oo, x], noroot p} -> {in `]-oo, x], forall y, sgr p.[y] = sgp_minfty p}. Proof. move=> rpx y Hy; rewrite -sgp_pinfty_sym. diff --git a/mathcomp/real_closed/qe_rcf.v b/mathcomp/real_closed/qe_rcf.v index 1791aca..92ceed2 100644 --- a/mathcomp/real_closed/qe_rcf.v +++ b/mathcomp/real_closed/qe_rcf.v @@ -1,8 +1,14 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp Require Import finfun path matrix. +From mathcomp Require Import bigop ssralg poly polydiv ssrnum zmodp div ssrint. +From mathcomp Require Import polyorder polyrcf interval polyXY. +From mathcomp Require Import qe_rcf_th ordered_qelim mxtens. Set Implicit Arguments. @@ -14,6 +20,20 @@ Import GRing.Theory Num.Theory. Local Open Scope nat_scope. Local Open Scope ring_scope. +Definition grab (X Y : Type) (pattern : Y -> Prop) (P : Prop -> Prop) + (y : X) (f : X -> Y) : + (let F := f in P (forall x, y = x -> pattern (F x))) + -> P (forall x : X, y = x -> pattern (f x)) := id. + +Definition grab_eq X Y u := @grab X Y (fun v => u = v :> Y). + +Tactic Notation "grab_eq" ident(f) open_constr(PAT1) := + let Edef := fresh "Edef" in + let E := fresh "E" in + move Edef: PAT1 => E; + move: E Edef; + elim/grab_eq: _ => f _ <-. + Import ord. Section QF. @@ -599,8 +619,7 @@ Lemma eval_SeqPInfty e ps k k' : = k' (map lead_coef (map (eval_poly e) ps)). Proof. elim: ps k k' => [|p ps ihps] k k' Pk /=; first by rewrite Pk. -rewrite (eval_LeadCoef (fun lp => - k' (lp :: [seq lead_coef i |i <- [seq eval_poly e i | i <- ps]]))) => // lp. +set X := lead_coef _; grab_eq k'' X; apply: (eval_LeadCoef k'') => lp {X}. rewrite (ihps _ (fun ps => k' (eval e lp :: ps))) => //= lps. by rewrite Pk. Qed. @@ -615,13 +634,9 @@ Lemma eval_SeqMInfty e ps k k' : (map (eval_poly e) ps)). Proof. elim: ps k k' => [|p ps ihps] k k' Pk /=; first by rewrite Pk. -rewrite (eval_LeadCoef (fun lp => - k' ((-1) ^+ (~~ odd (size (eval_poly e p))) * lp - :: [seq (-1) ^+ (~~ odd (size p)) * lead_coef p - | p : {poly _} <- [seq eval_poly e i | i <- ps]]))) => // lp. -rewrite eval_Size /= (ihps _ (fun ps => - k' (((-1) ^+ (~~ odd (size (eval_poly e p))) * eval e lp) :: ps))) => //= lps. -by rewrite Pk. +set X := lead_coef _; grab_eq k'' X; apply: eval_LeadCoef => lp {X}. +rewrite eval_Size /= /k'' {k''}. +by set X := map _ _; grab_eq k'' X; apply: ihps => {X} lps; rewrite Pk. Qed. Implicit Arguments eval_SeqMInfty [e ps k]. @@ -659,18 +674,10 @@ Lemma eval_Rediv_rec_loop e q sq cq c qq r n k k' Proof. move=> Pk; elim: n c qq r k Pk @d=> [|n ihn] c qq r k Pk /=. rewrite eval_Size /=; have [//=|gtq] := ltnP. - rewrite (eval_LeadCoef (fun lr => - let m := lr *: 'X^(size (eval_poly e r) - sq) in - let qq1 := (eval_poly e qq) * (eval e cq)%:P + m in - let r1 := (eval_poly e r) * (eval e cq)%:P - m * (eval_poly e q) in - k' (c.+1, qq1, r1))) //. - by move=> x /=; rewrite Pk /= !eval_OpPoly /= !mul_polyC. + set X := lead_coef _; grab_eq k'' X; apply: eval_LeadCoef => {X}. + by move=> x /=; rewrite Pk /= !eval_OpPoly /= !mul_polyC. rewrite eval_Size /=; have [//=|gtq] := ltnP. -rewrite (eval_LeadCoef (fun lr => - let m := lr *: 'X^(size (eval_poly e r) - sq) in - let qq1 := (eval_poly e qq) * (eval e cq)%:P + m in - let r1 := (eval_poly e r) * (eval e cq)%:P - m * (eval_poly e q) in - k' (redivp_rec_loop (eval_poly e q) sq (eval e cq) c.+1 qq1 r1 n))) //=. +set X := lead_coef _; grab_eq k'' X; apply: eval_LeadCoef => {X}. by move=> x; rewrite ihn // !eval_OpPoly /= !mul_polyC. Qed. @@ -988,7 +995,8 @@ End ProjCorrect. (* Section Example. *) (* no chances it computes *) -(* Require Import rat. *) +(* From mathcomp +Require Import rat. *) (* Eval vm_compute in (54%:R / 289%:R + 2%:R^-1 :rat). *) diff --git a/mathcomp/real_closed/qe_rcf_th.v b/mathcomp/real_closed/qe_rcf_th.v index f1e5a61..71c6e45 100644 --- a/mathcomp/real_closed/qe_rcf_th.v +++ b/mathcomp/real_closed/qe_rcf_th.v @@ -1,6 +1,10 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice path fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice path fintype. +From mathcomp Require Import div bigop ssralg poly polydiv ssrnum perm zmodp ssrint. +From mathcomp Require Import polyorder polyrcf interval matrix mxtens. Set Implicit Arguments. @@ -1234,7 +1238,7 @@ case: (altP ((\prod_(q0 <- sq) q0) =P 0)) => [ | pn0]. by rewrite in_cons ht orbT. rewrite big_cons size_mul // (eqP (ih _)) //; last first. by move=> t ht; apply: hs; rewrite in_cons ht orbT. -rewrite addnS addn0; apply/eqP; apply: hs; exact: mem_head. +by rewrite addnS addn0; apply/eqP; apply: hs; apply: mem_head. Qed. Definition ccount_gt0 (sp sq : seq {poly R}) := diff --git a/mathcomp/real_closed/realalg.v b/mathcomp/real_closed/realalg.v index 88d656a..f425876 100644 --- a/mathcomp/real_closed/realalg.v +++ b/mathcomp/real_closed/realalg.v @@ -1,7 +1,12 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. +From mathcomp Require Import bigop ssralg ssrnum ssrint rat poly polydiv polyorder. +From mathcomp Require Import perm matrix mxpoly polyXY binomial generic_quotient. +From mathcomp Require Import cauchyreals separable zmodp bigenough. (*************************************************************************) @@ -1520,7 +1525,8 @@ End RatRealAlg. Canonical RatRealAlg.realalg_countType. -(* Require Import countalg. *) +(* From mathcomp +Require Import countalg. *) (* Canonical realalg_countZmodType := [countZmodType of realalg]. *) (* Canonical realalg_countRingType := [countRingType of realalg]. *) (* Canonical realalg_countComRingType := [countComRingType of realalg]. *) diff --git a/mathcomp/solvable/Make b/mathcomp/solvable/Make index 41bd3e0..d89d213 100644 --- a/mathcomp/solvable/Make +++ b/mathcomp/solvable/Make @@ -1,9 +1,10 @@ abelian.v -all.v +all_solvable.v alt.v burnside_app.v center.v commutator.v +cyclic.v extraspecial.v extremal.v finmodule.v diff --git a/mathcomp/solvable/abelian.v b/mathcomp/solvable/abelian.v index 52f92ec..0d87755 100644 --- a/mathcomp/solvable/abelian.v +++ b/mathcomp/solvable/abelian.v @@ -1,14 +1,13 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div fintype finfun bigop finset prime binomial. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism action quotient gproduct. -From mathcomp.algebra -Require Import cyclic zmodp. -Require Import gfunctor pgroup gseries nilpotent sylow. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div fintype. +From mathcomp +Require Import finfun bigop finset prime binomial fingroup morphism perm. +From mathcomp +Require Import automorphism action quotient gfunctor gproduct zmodp cyclic. +From mathcomp +Require Import pgroup gseries nilpotent sylow. (******************************************************************************) (* Constructions based on abelian groups and their structure, with some *) @@ -245,7 +244,7 @@ Lemma exponent1 : exponent [1 gT] = 1%N. Proof. by apply/eqP; rewrite -dvdn1 -trivg_exponent eqxx. Qed. Lemma exponent_dvdn G : exponent G %| #|G|. -Proof. by apply/dvdn_biglcmP=> x Gx; exact: order_dvdG. Qed. +Proof. by apply/dvdn_biglcmP=> x Gx; apply: order_dvdG. Qed. Lemma exponent_gt0 G : 0 < exponent G. Proof. exact: dvdn_gt0 (exponent_dvdn G). Qed. @@ -257,7 +256,7 @@ congr (_ && _); first by rewrite cardG_gt0 exponent_gt0. apply: eq_all_r => p; rewrite !mem_primes cardG_gt0 exponent_gt0 /=. apply: andb_id2l => p_pr; apply/idP/idP=> pG. exact: dvdn_trans pG (exponent_dvdn G). -by case/Cauchy: pG => // x Gx <-; exact: dvdn_exponent. +by case/Cauchy: pG => // x Gx <-; apply: dvdn_exponent. Qed. Lemma exponentJ A x : exponent (A :^ x) = exponent A. @@ -284,7 +283,7 @@ Lemma exponent_cycle x : exponent <[x]> = #[x]. Proof. by apply/eqP; rewrite eqn_dvd exponent_dvdn dvdn_exponent ?cycle_id. Qed. Lemma exponent_cyclic X : cyclic X -> exponent X = #|X|. -Proof. by case/cyclicP=> x ->; exact: exponent_cycle. Qed. +Proof. by case/cyclicP=> x ->; apply: exponent_cycle. Qed. Lemma primes_exponent G : primes (exponent G) = primes (#|G|). Proof. @@ -320,7 +319,7 @@ Lemma exponent_Hall pi G H : pi.-Hall(G) H -> exponent H = (exponent G)`_pi. Proof. move=> hallH; have [sHG piH _] := and3P hallH. rewrite -(partn_exponentS sHG) -?(card_Hall hallH) ?part_pnat_id //. -by apply: pnat_dvd piH; exact: exponent_dvdn. +by apply: pnat_dvd piH; apply: exponent_dvdn. Qed. Lemma exponent_Zgroup G : Zgroup G -> exponent G = #|G|. @@ -371,7 +370,7 @@ Proof. move=> cGG; apply/group_setP. split=> [|x y]; rewrite !inE ?group1 ?expg1n //=. case/andP=> Gx /eqP xn /andP[Gy /eqP yn]. -rewrite groupM //= expgMn ?xn ?yn ?mulg1 //; exact: (centsP cGG). +by rewrite groupM //= expgMn ?xn ?yn ?mulg1 //; apply: (centsP cGG). Qed. Lemma abelian_exponent_gen A : abelian A -> exponent <<A>> = exponent A. @@ -414,7 +413,7 @@ Qed. Lemma cyclic_abelem_prime p X : p.-abelem X -> cyclic X -> X :!=: 1 -> #|X| = p. Proof. move=> abelX cycX; case/cyclicP: cycX => x -> in abelX *. -by rewrite cycle_eq1; exact: abelem_order_p abelX (cycle_id x). +by rewrite cycle_eq1; apply: abelem_order_p abelX (cycle_id x). Qed. Lemma cycle_abelem p x : p.-elt x || prime p -> p.-abelem <[x]> = (#[x] %| p). @@ -423,7 +422,7 @@ move=> p_xVpr; rewrite /abelem cycle_abelian /=. apply/andP/idP=> [[_ xp1] | x_dvd_p]. by rewrite order_dvdn (exponentP xp1) ?cycle_id. split; last exact: dvdn_trans (exponent_dvdn _) x_dvd_p. -by case/orP: p_xVpr => // /pnat_id; exact: pnat_dvd. +by case/orP: p_xVpr => // /pnat_id; apply: pnat_dvd. Qed. Lemma exponent2_abelem G : exponent G %| 2 -> 2.-abelem G. @@ -494,7 +493,7 @@ by rewrite (is_abelem_pgroup (abelem_pgroup abelG)). Qed. Lemma pElemP p A E : reflect (E \subset A /\ p.-abelem E) (E \in 'E_p(A)). -Proof. by rewrite inE; exact: andP. Qed. +Proof. by rewrite inE; apply: andP. Qed. Implicit Arguments pElemP [p A E]. Lemma pElemS p A B : A \subset B -> 'E_p(A) \subset 'E_p(B). @@ -602,7 +601,7 @@ apply/and3P; split; last 1 first. - apply/imsetP=> [[X /card_p1Elem oX X'0]]. by rewrite -oX (cardsD1 1) -X'0 group1 cards0 in p_pr. - rewrite eqEsubset; apply/andP; split. - by apply/bigcupsP=> _ /imsetP[X /pnElemP[sXE _ _] ->]; exact: setSD. + by apply/bigcupsP=> _ /imsetP[X /pnElemP[sXE _ _] ->]; apply: setSD. apply/subsetP=> x /setD1P[ntx Ex]. apply/bigcupP; exists <[x]>^#; last by rewrite !inE ntx cycle_id. apply/imsetP; exists <[x]>%G; rewrite ?p1ElemE // !inE cycle_subG Ex /=. @@ -680,7 +679,7 @@ Qed. Lemma pmaxElemP p A E : reflect (E \in 'E_p(A) /\ forall H, H \in 'E_p(A) -> E \subset H -> H :=: E) (E \in 'E*_p(A)). -Proof. by rewrite [E \in 'E*_p(A)]inE; exact: (iffP maxgroupP). Qed. +Proof. by rewrite [E \in 'E*_p(A)]inE; apply: (iffP maxgroupP). Qed. Lemma pmaxElem_exists p A D : D \in 'E_p(A) -> {E | E \in 'E*_p(A) & D \subset E}. @@ -716,7 +715,7 @@ Proof. move=> sAB; apply/subsetP=> E; rewrite !inE. case/andP=> /maxgroupP[/pElemP[_ abelE] maxE] sEA. apply/maxgroupP; rewrite inE sEA; split=> // D EpD. -by apply: maxE; apply: subsetP EpD; exact: pElemS. +by apply: maxE; apply: subsetP EpD; apply: pElemS. Qed. Lemma pmaxElemJ p A E x : ((E :^ x)%G \in 'E*_p(A :^ x)) = (E \in 'E*_p(A)). @@ -736,7 +735,7 @@ Qed. Lemma grank_witness G : {B | <<B>> = G & #|B| = 'm(G)}. Proof. rewrite /gen_rank; case: arg_minP => [|B defG _]; first by rewrite genGid. -by exists B; first exact/eqP. +by exists B; first apply/eqP. Qed. Lemma p_rank_witness p G : {E | E \in 'E_p^('r_p(G))(G)}. @@ -751,7 +750,7 @@ Proof. apply: (iffP idP) => [|[E]]; last first. by rewrite inE => /andP[Ep_E /eqP <-]; rewrite (bigmax_sup E). have [D /pnElemP[sDG abelD <-]] := p_rank_witness p G. -by case/abelem_pnElem=> // E; exists E; exact: (subsetP (pnElemS _ _ sDG)). +by case/abelem_pnElem=> // E; exists E; apply: (subsetP (pnElemS _ _ sDG)). Qed. Lemma p_rank_gt0 p H : ('r_p(H) > 0) = (p \in \pi(H)). @@ -772,7 +771,7 @@ Proof. by move=> EpA_E; rewrite (bigmax_sup E). Qed. Lemma p_rank_le_logn p G : 'r_p(G) <= logn p #|G|. Proof. have [E EpE] := p_rank_witness p G. -by have [sEG _ <-] := pnElemP EpE; exact: lognSg. +by have [sEG _ <-] := pnElemP EpE; apply: lognSg. Qed. Lemma p_rank_abelem p G : p.-abelem G -> 'r_p(G) = logn p #|G|. @@ -931,7 +930,7 @@ Qed. Lemma morphim_Ldiv n A : f @* 'Ldiv_n(A) \subset 'Ldiv_n(f @* A). Proof. -by apply: subset_trans (morphimI f A _) (setIS _ _); exact: morphim_LdivT. +by apply: subset_trans (morphimI f A _) (setIS _ _); apply: morphim_LdivT. Qed. Lemma morphim_abelem p G : p.-abelem G -> p.-abelem (f @* G). @@ -984,7 +983,7 @@ by rewrite sub_morphim_pre ?morphim_sub // morphpre_invm. Qed. Lemma injm_abelem p : p.-abelem (f @* G) = p.-abelem G. -Proof. by apply/idP/idP; first rewrite -{2}defG; exact: morphim_abelem. Qed. +Proof. by apply/idP/idP; first rewrite -{2}defG; apply: morphim_abelem. Qed. Lemma injm_pElem p (E : {group aT}) : E \subset D -> ((f @* E)%G \in 'E_p(f @* G)) = (E \in 'E_p(G)). @@ -1255,7 +1254,7 @@ move=> pG; apply/eqP; rewrite eqEsubset !gen_subG; apply/andP. do [split; apply/subsetP=> xpn; case/imsetP=> x] => [|Gx ->]; last first. by rewrite Mho_p_elt ?(mem_p_elt pG). case/setIdP=> Gx _ ->; have [-> | ntx] := eqVneq x 1; first by rewrite expg1n. -by rewrite (pdiv_p_elt (mem_p_elt pG Gx) ntx) mem_gen //; exact: mem_imset. +by rewrite (pdiv_p_elt (mem_p_elt pG Gx) ntx) mem_gen //; apply: mem_imset. Qed. Lemma MhoEabelian p G : @@ -1264,7 +1263,7 @@ Proof. move=> pG cGG; rewrite (MhoE pG); rewrite gen_set_id //; apply/group_setP. split=> [|xn yn]; first by apply/imsetP; exists 1; rewrite ?expg1n. case/imsetP=> x Gx ->; case/imsetP=> y Gy ->. -by rewrite -expgMn; [apply: mem_imset; rewrite groupM | exact: (centsP cGG)]. +by rewrite -expgMn; [apply: mem_imset; rewrite groupM | apply: (centsP cGG)]. Qed. Lemma trivg_Mho G : 'Mho^n(G) == 1 -> 'Ohm_n(G) == G. @@ -1274,14 +1273,14 @@ rewrite -{1}(Sylow_gen G) genS //; apply/bigcupsP=> P. case/SylowP=> p p_pr /and3P[sPG pP _]; apply/subsetP=> x Px. have Gx := subsetP sPG x Px; rewrite inE Gx //=. rewrite (sameP eqP set1P) (subsetP Gp1) ?mem_gen //; apply: mem_imset. -by rewrite inE Gx; exact: pgroup_p (mem_p_elt pP Px). +by rewrite inE Gx; apply: pgroup_p (mem_p_elt pP Px). Qed. Lemma Mho_p_cycle p x : p.-elt x -> 'Mho^n(<[x]>) = <[x ^+ (p ^ n)]>. Proof. move=> p_x. apply/eqP; rewrite (MhoE p_x) eqEsubset cycle_subG mem_gen; last first. - by apply: mem_imset; exact: cycle_id. + by apply: mem_imset; apply: cycle_id. rewrite gen_subG andbT; apply/subsetP=> _ /imsetP[_ /cycleP[k ->] ->]. by rewrite -expgM mulnC expgM mem_cycle. Qed. @@ -1332,7 +1331,7 @@ Proof. exact: morphimF. Qed. Lemma injm_Ohm (f : {morphism D >-> rT}) : 'injm f -> G \subset D -> f @* 'Ohm_n(G) = 'Ohm_n(f @* G). -Proof. by move=> injf; exact: injmF. Qed. +Proof. by move=> injf; apply: injmF. Qed. Lemma isog_Ohm (H : {group rT}) : G \isog H -> 'Ohm_n(G) \isog 'Ohm_n(H). Proof. exact: gFisog. Qed. @@ -1425,7 +1424,7 @@ Lemma abelem_Ohm1P p G : abelian G -> p.-group G -> reflect ('Ohm_1(G) = G) (p.-abelem G). Proof. move=> cGG pG. -apply: (iffP idP) => [| <-]; [exact: Ohm1_id | exact: Ohm1_abelem]. +by apply: (iffP idP) => [| <-]; [apply: Ohm1_id | apply: Ohm1_abelem]. Qed. Lemma TI_Ohm1 G H : H :&: 'Ohm_1(G) = 1 -> H :&: G = 1. @@ -1699,7 +1698,7 @@ have{lnOx} lnOy i: lnO i <[x]> = lnO i <[y]>. have co_yx': coprime #[y] #[x.`_p^'] by rewrite !order_constt coprime_partC. have defX: <[y]> \x <[x.`_p^']> = <[x]>. rewrite dprodE ?coprime_TIg //. - by rewrite -cycleM ?consttC //; apply: (centsP cXX); exact: mem_cycle. + by rewrite -cycleM ?consttC //; apply: (centsP cXX); apply: mem_cycle. by apply: (sub_abelian_cent2 cXX); rewrite cycle_subG mem_cycle. rewrite -(lnOx i _ _ _ defX) addnC {1}/lnO lognE. case: and3P => // [[p_pr _ /idPn[]]]; rewrite -p'natE //. @@ -1760,7 +1759,7 @@ apply/eqP; rewrite -def_t size_map eqn_leq andbC; apply/andP; split. case/lastP def_b: b => // [b' x]; pose p := pdiv #[x]. have p_pr: prime p. have:= abelian_type_gt1 G; rewrite -def_t def_b map_rcons -cats1 all_cat. - by rewrite /= andbT => /andP[_]; exact: pdiv_prime. + by rewrite /= andbT => /andP[_]; apply: pdiv_prime. suffices: all [pred y | logn p #[y] > 0] b. rewrite all_count (count_logn_dprod_cycle _ _ defG) -def_b; move/eqP <-. by rewrite Ohm0 indexg1 -p_rank_abelian ?p_rank_le_rank. @@ -1791,9 +1790,9 @@ case p_x: (p_group <[x]>); last first. rewrite -order_gt1 /p_group -orderE; set p := pdiv _ => ntx p'x. have def_x: <[x.`_p]> \x <[x.`_p^']> = <[x]>. have ?: coprime #[x.`_p] #[x.`_p^'] by rewrite !order_constt coprime_partC. - have ?: commute x.`_p x.`_p^' by exact: commuteX2. + have ?: commute x.`_p x.`_p^' by apply: commuteX2. rewrite dprodE ?coprime_TIg -?cycleM ?consttC //. - by rewrite cent_cycle cycle_subG; exact/cent1P. + by rewrite cent_cycle cycle_subG; apply/cent1P. rewrite -(dprod_card (Ohm_dprod n def_x)) -(dprod_card (Mho_dprod n def_x)). rewrite mulnCA -mulnA mulnCA mulnA. rewrite !{}IHm ?(dprod_card def_x) ?(leq_trans _ lexm) {m lexm}//. @@ -1801,7 +1800,7 @@ case p_x: (p_group <[x]>); last first. rewrite p_part -(expn0 p) ltn_exp2l 1?lognE ?prime_gt1 ?pdiv_prime //. by rewrite order_gt0 pdiv_dvd. rewrite proper_card // properEneq cycle_subG mem_cycle andbT. - by apply: contra (negbT p'x); move/eqP <-; exact: p_elt_constt. + by apply: contra (negbT p'x); move/eqP <-; apply: p_elt_constt. case/p_groupP: p_x => p p_pr p_x. rewrite (Ohm_p_cycle n p_x) (Mho_p_cycle n p_x) -!orderE. set k := logn p #[x]; have ox: #[x] = (p ^ k)%N by rewrite -card_pgroup. @@ -1865,7 +1864,7 @@ Lemma homocyclic_Ohm_Mho n p G : p.-group G -> homocyclic G -> 'Ohm_n(G) = 'Mho^(logn p (exponent G) - n)(G). Proof. move=> pG /andP[cGG homoG]; set e := exponent G. -have{pG} p_e: p.-nat e by apply: pnat_dvd pG; exact: exponent_dvdn. +have{pG} p_e: p.-nat e by apply: pnat_dvd pG; apply: exponent_dvdn. have{homoG}: all (pred1 e) (abelian_type G). move: homoG; rewrite /abelian_type -(prednK (cardG_gt0 G)) /=. by case: (_ && _) (tag _); rewrite //= genGid eqxx. @@ -1884,7 +1883,7 @@ Lemma Ohm_Mho_homocyclic (n p : nat) G : Proof. set e := exponent G => cGG pG /andP[n_gt0 n_lte] eq_Ohm_Mho. suffices: all (pred1 e) (abelian_type G). - by rewrite /homocyclic cGG; exact: all_pred1_constant. + by rewrite /homocyclic cGG; apply: all_pred1_constant. case/abelian_structure: cGG (abelian_type_gt1 G) => b defG <-. elim: b {-3}G defG (subxx G) eq_Ohm_Mho => //= x b IHb H. rewrite big_cons => defG; case/dprodP: defG (defG) => [[_ K _ defK]]. @@ -2081,7 +2080,7 @@ case: (ltnP e _) (b_sorted p) => [lt_e_x | le_x_e]. move: ordb; rewrite (take_nth 1 ltib) -/x rev_rcons all_rcons /= lt_e_x. case/andP=> _ /=; move/(order_path_min leq_trans); apply: contraLR. rewrite -!has_predC !has_map; case/hasP=> y b_y /= le_y_e; apply/hasP. - by exists y; rewrite ?mem_rev //=; apply: contra le_y_e; exact: leq_trans. + by exists y; rewrite ?mem_rev //=; apply: contra le_y_e; apply: leq_trans. rewrite -(cat_take_drop i b) -map_rev rev_cat !map_cat cat_path. case/andP=> ordb _; rewrite count_cat -{1}(size_takel (ltnW ltib)) ltnNge. rewrite addnC ((count _ _ =P 0) _) ?count_size //. diff --git a/mathcomp/solvable/all.v b/mathcomp/solvable/all_solvable.v index 8c2ea8f..b05a3c4 100644 --- a/mathcomp/solvable/all.v +++ b/mathcomp/solvable/all_solvable.v @@ -3,6 +3,7 @@ Require Export alt. Require Export burnside_app. Require Export center. Require Export commutator. +Require Export cyclic. Require Export extraspecial. Require Export extremal. Require Export finmodule. diff --git a/mathcomp/solvable/alt.v b/mathcomp/solvable/alt.v index 3ee2526..18a6588 100644 --- a/mathcomp/solvable/alt.v +++ b/mathcomp/solvable/alt.v @@ -1,14 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div fintype tuple finfun bigop prime finset. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism quotient action. -From mathcomp.algebra -Require Import cyclic. -Require Import pgroup gseries sylow primitive_action. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div fintype tuple. +From mathcomp +Require Import tuple bigop prime finset fingroup morphism perm automorphism. +From mathcomp +Require Import quotient action cyclic pgroup gseries sylow primitive_action. (******************************************************************************) (* Definitions of the symmetric and alternate groups, and some properties. *) @@ -189,10 +186,10 @@ have FF (H : {group {perm T}}): H <| 'Alt_T -> H :<>: 1 -> 20 %| #|H|. - move=> Hh1 Hh3. have [x _]: exists x, x \in T by apply/existsP/eqP; rewrite oT. have F2 := Alt_trans T; rewrite oT /= in F2. - have F3: [transitive 'Alt_T, on setT | 'P] by exact: ntransitive1 F2. - have F4: [primitive 'Alt_T, on setT | 'P] by exact: ntransitive_primitive F2. + have F3: [transitive 'Alt_T, on setT | 'P] by apply: ntransitive1 F2. + have F4: [primitive 'Alt_T, on setT | 'P] by apply: ntransitive_primitive F2. case: (prim_trans_norm F4 Hh1) => F5. - case: Hh3; apply/trivgP; exact: subset_trans F5 (aperm_faithful _). + by case: Hh3; apply/trivgP; apply: subset_trans F5 (aperm_faithful _). have F6: 5 %| #|H| by rewrite -oT -cardsT (atrans_dvd F5). have F7: 4 %| #|H|. have F7: #|[set~ x]| = 4 by rewrite cardsC1 oT. @@ -206,10 +203,10 @@ have FF (H : {group {perm T}}): H <| 'Alt_T -> H :<>: 1 -> 20 %| #|H|. exact: ntransitive1 F9. have F11: [primitive Gx, on [set~ x] | 'P]. exact: ntransitive_primitive F9. - have F12: K \subset Gx by apply: setSI; exact: normal_sub. + have F12: K \subset Gx by apply: setSI; apply: normal_sub. have F13: K <| Gx by rewrite /(K <| _) F12 normsIG // normal_norm. case: (prim_trans_norm F11 F13) => Ksub; last first. - apply: dvdn_trans (cardSg F8); rewrite -F7; exact: atrans_dvd Ksub. + by apply: dvdn_trans (cardSg F8); rewrite -F7; apply: atrans_dvd Ksub. have F14: [faithful Gx, on [set~ x] | 'P]. apply/subsetP=> g; do 2![case/setIP] => Altg cgx cgx'. apply: (subsetP (aperm_faithful 'Alt_T)). @@ -293,7 +290,7 @@ Lemma rfd_funP (p : {perm T}) (u : T') : let p1 := if p x == x then p else 1 in p1 (val u) != x. Proof. case: (p x =P x) => /= [pxx|_]; last by rewrite perm1 (valP u). -by rewrite -{2}pxx (inj_eq (@perm_inj _ p)); exact: (valP u). +by rewrite -{2}pxx (inj_eq (@perm_inj _ p)); apply: (valP u). Qed. Definition rfd_fun p := [fun u => Sub ((_ : {perm T}) _) (rfd_funP p u) : T']. @@ -366,7 +363,7 @@ have Hcp1: #|[set x | p1 x != x]| <= n. by move: Hp; rewrite (cardD1 x1) inE Hx1. have ->: p = p1 * tperm x1 (p x1) by rewrite -mulgA tperm2 mulg1. rewrite odd_permM odd_tperm eq_sym Hx1 morphM; last 2 first. -- by rewrite 2!inE; exact/astab1P. +- by rewrite 2!inE; apply/astab1P. - by rewrite 2!inE; apply/astab1P; rewrite -{1}Hpx /= /aperm -permM. rewrite odd_permM Hrec //=; congr (_ (+) _). pose x2 : T' := Sub x1 nx1x; pose px2 : T' := Sub (p x1) npx1x. @@ -441,7 +438,7 @@ have F11: [primitive Gx, on [set~ x] | 'P]. have F12: K \subset Gx by rewrite setSI // normal_sub. have F13: K <| Gx by apply/andP; rewrite normsIG. have:= prim_trans_norm F11; case/(_ K) => //= => Ksub; last first. - have F14: Gx * H = 'Alt_T by exact/(subgroup_transitiveP _ _ F3). + have F14: Gx * H = 'Alt_T by apply/(subgroup_transitiveP _ _ F3). have: simple Gx. by rewrite (isog_simple (rfd_iso x)) Hrec //= card_sig cardC1 Hde. case/simpleP=> _ simGx; case/simGx: F13 => /= HH2. @@ -470,7 +467,7 @@ have Hreg g z: g \in H -> g z = z -> g = 1. by rewrite memJ_norm ?(subsetP nH). clear K F8 F12 F13 Ksub F14. have Hcard: 5 < #|H|. - apply: (leq_trans oT); apply dvdn_leq; first by exact: cardG_gt0. + apply: (leq_trans oT); apply dvdn_leq; first by apply: cardG_gt0. by rewrite -cardsT (atrans_dvd F5). case Eh: (pred0b [predD1 H & 1]). by move: Hcard; rewrite /pred0b in Eh; rewrite (cardD1 1) group1 (eqP Eh). @@ -491,7 +488,7 @@ case diff_gx_hx: (g x == h x). by rewrite permM -(eqP diff_gx_hx) -permM mulgV perm1. case diff_hgx_x: ((h * g) x == x). case/negP: diff_h1_g; apply/eqP; apply: (mulgI h); rewrite !gsimp. - by apply: (Hreg _ x); [exact: groupM | apply/eqP]. + by apply: (Hreg _ x); [apply: groupM | apply/eqP]. case diff_hgx_hx: ((h * g) x == h x). case/negP: diff_1_g; apply/eqP. by apply: (Hreg _ (h x)) => //; apply/eqP; rewrite -permM. diff --git a/mathcomp/solvable/burnside_app.v b/mathcomp/solvable/burnside_app.v index 635fd84..e8fe7dc 100644 --- a/mathcomp/solvable/burnside_app.v +++ b/mathcomp/solvable/burnside_app.v @@ -1,12 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype tuple finfun bigop finset. -From mathcomp.fingroup -Require Import fingroup action perm. -Require Import primitive_action. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div choice fintype. +From mathcomp +Require Import tuple finfun bigop finset fingroup action perm primitive_action. (* Application of the Burside formula to count the number of distinct *) (* colorings of the vertices of a square and a cube. *) @@ -120,7 +117,7 @@ Qed. Lemma rot_r1 : forall r, is_rot r -> r = r1 ^+ (r c0). Proof. move=> r hr;apply: rot_eq_c0 => //;apply/eqP. - by symmetry; exact: commuteX. + by symmetry; apply: commuteX. by case: (r c0); do 4?case => //=; rewrite ?permM !permE /=. Qed. @@ -304,12 +301,12 @@ Definition square_coloring_number4 := #|orbit to rotations @: setT|. Definition square_coloring_number8 := #|orbit to isometries @: setT|. Lemma Fid : 'Fix_to(1) = setT. -Proof. apply/setP=> x /=; rewrite in_setT; apply/afix1P; exact: act1. Qed. +Proof. by apply/setP=> x /=; rewrite in_setT; apply/afix1P; apply: act1. Qed. Lemma card_Fid : #|'Fix_to(1)| = (n ^ 4)%N. Proof. rewrite -[4]card_ord -[n]card_ord -card_ffun_on Fid cardsE. -by symmetry; apply: eq_card => f; exact/ffun_onP. +by symmetry; apply: eq_card => f; apply/ffun_onP. Qed. Definition coin0 (sc : col_squares) : colors := sc c0. @@ -691,7 +688,7 @@ Qed. Definition sop (p : {perm cube}) : seq cube := val (val (val p)). Lemma sop_inj : injective sop. -Proof. do 2!apply: (inj_comp val_inj); exact: val_inj. Qed. +Proof. by do 2!apply: (inj_comp val_inj); apply: val_inj. Qed. Definition prod_tuple (t1 t2 : seq cube) := map (fun n : 'I_6 => nth F0 t2 n) t1. @@ -728,7 +725,7 @@ Definition seq_iso_L := [:: Lemma seqs1 : forall f injf, sop (@perm _ f injf) = map f ecubes. Proof. move=> f ?; rewrite ecubes_def /sop /= -codom_ffun pvalE. -apply: eq_codom; exact: permE. +by apply: eq_codom; apply: permE. Qed. Lemma Lcorrect : seq_iso_L == map sop [:: id3; s05; s14; s23; r05; r14; r23; @@ -741,7 +738,7 @@ Proof. by move=> p; rewrite /= !inE /= -!(eq_sym p). Qed. Lemma L_iso : forall p, (p \in dir_iso3) = (sop p \in seq_iso_L). Proof. -move=> p; rewrite (eqP Lcorrect) mem_map ?iso0_1 //; exact: sop_inj. +by move=> p; rewrite (eqP Lcorrect) mem_map ?iso0_1 //; apply: sop_inj. Qed. Lemma stable : forall x y, @@ -831,7 +828,7 @@ Canonical iso_group3 := Group group_set_iso3. Lemma group_set_diso3 : group_set dir_iso3. Proof. apply/group_setP;split;first by rewrite inE eqxx /=. -by exact:stable. +by apply:stable. Qed. Canonical diso_group3 := Group group_set_diso3. @@ -862,8 +859,8 @@ have -> : s3 = r05 * r14 * r05 by iso_tac. have -> : s4 = r05 * r14 * r14 * r14 * r05 by iso_tac. have -> : s5 = r14 * r05 * r05 by iso_tac. have -> : s6 = r05 * r05 * r14 by iso_tac. -do ?case/predU1P=> [<-|]; first exact: group1; last (move/eqP => <-); - by rewrite ?groupMl ?mem_gen // !inE eqxx ?orbT. +by do ?case/predU1P=> [<-|]; first exact: group1; last (move/eqP => <-); + rewrite ?groupMl ?mem_gen // !inE eqxx ?orbT. Qed. Notation col_cubes := {ffun cube -> colors}. @@ -887,7 +884,7 @@ Proof. by apply/setP=> x /=; rewrite (sameP afix1P eqP) !inE act1 eqxx. Qed. Lemma card_Fid3 : #|'Fix_to_g[1]| = (n ^ 6)%N. Proof. rewrite -[6]card_ord -[n]card_ord -card_ffun_on Fid3 cardsT. -by symmetry; apply: eq_card => ff; exact/ffun_onP. +by symmetry; apply: eq_card => ff; apply/ffun_onP. Qed. Definition col0 (sc : col_cubes) : colors := sc F0. diff --git a/mathcomp/solvable/center.v b/mathcomp/solvable/center.v index 6226861..c0c0451 100644 --- a/mathcomp/solvable/center.v +++ b/mathcomp/solvable/center.v @@ -1,14 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div fintype bigop finset. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism quotient action gproduct. -From mathcomp.algebra -Require Import cyclic. -Require Import gfunctor. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div fintype bigop. +From mathcomp +Require Import finset fingroup morphism perm automorphism quotient action. +From mathcomp +Require Import gproduct gfunctor cyclic. (******************************************************************************) (* Definition of the center of a group and of external central products: *) @@ -63,7 +60,7 @@ Notation "''Z' ( A )" := (center A) : group_scope. Notation "''Z' ( H )" := (center_group H) : Group_scope. Lemma morphim_center : GFunctor.pcontinuous center. -Proof. move=> gT rT G D f; exact: morphim_subcent. Qed. +Proof. by move=> gT rT G D f; apply: morphim_subcent. Qed. Canonical center_igFun := [igFun by fun _ _ => subsetIl _ _ & morphim_center]. Canonical center_gFun := [gFun by morphim_center]. @@ -103,8 +100,8 @@ Proof. exact: subcentP. Qed. Lemma center_sub A : 'Z(A) \subset A. Proof. exact: subsetIl. Qed. -Lemma center1 : 'Z(1) = [1 gT]. -Proof. by apply/eqP; rewrite eqEsubset center_sub sub1G. Qed. +Lemma center1 : 'Z(1) = 1 :> {set gT}. +Proof. exact: gF1. Qed. Lemma centerC A : {in A, centralised 'Z(A)}. Proof. by apply/centsP; rewrite centsC subsetIr. Qed. @@ -137,13 +134,13 @@ by apply: (iffP cent1P) => [|[]]. Qed. Lemma subcent1_id x G : x \in G -> x \in 'C_G[x]. -Proof. move=> Gx; rewrite inE Gx; exact/cent1P. Qed. +Proof. by move=> Gx; rewrite inE Gx; apply/cent1P. Qed. Lemma subcent1_sub x G : 'C_G[x] \subset G. Proof. exact: subsetIl. Qed. Lemma subcent1C x y G : x \in G -> y \in 'C_G[x] -> x \in 'C_G[y]. -Proof. by move=> Gx /subcent1P[_ cxy]; exact/subcent1P. Qed. +Proof. by move=> Gx /subcent1P[_ cxy]; apply/subcent1P. Qed. Lemma subcent1_cycle_sub x G : x \in G -> <[x]> \subset 'C_G[x]. Proof. by move=> Gx; rewrite cycle_subG ?subcent1_id. Qed. @@ -171,9 +168,9 @@ Lemma cyclic_factor_abelian H G : H \subset 'Z(G) -> cyclic (G / H) -> abelian G. Proof. move=> sHZ cycGH; apply: cyclic_center_factor_abelian. -have nG: G \subset 'N(_) := normal_norm (sub_center_normal _). -have [f <-]:= homgP (homg_quotientS (nG _ sHZ) (nG _ (subxx _)) sHZ). -exact: morphim_cyclic. +have /andP[_ nHG]: H <| G := sub_center_normal sHZ. +have [f <-]:= homgP (homg_quotientS nHG (gFnorm _ G) sHZ). +exact: morphim_cyclic cycGH. Qed. Section Injm. @@ -293,7 +290,7 @@ Let kerHK := ker_cprod_by isoZ. Let injgz : 'injm gz. Proof. by case/isomP: isoZ. Qed. Let gzZ : gz @* 'Z(H) = 'Z(K). Proof. by case/isomP: isoZ. Qed. -Let gzZchar : gz @* 'Z(H) \char 'Z(K). Proof. by rewrite gzZ char_refl. Qed. +Let gzZchar : gz @* 'Z(H) \char 'Z(K). Proof. by rewrite gzZ. Qed. Let sgzZZ : gz @* 'Z(H) \subset 'Z(K) := char_sub gzZchar. Let sZH := center_sub H. Let sZK := center_sub K. @@ -557,7 +554,7 @@ by split=> //; apply/isomP; rewrite ker_f' injm_invm im_f' -fC im_invm. Qed. Lemma isog_cprod_by : G \isog C. -Proof. by have [f [isoG _ _]] := cprod_by_uniq; exact: isom_isog isoG. Qed. +Proof. by have [f [isoG _ _]] := cprod_by_uniq; apply: isom_isog isoG. Qed. End Isomorphism. @@ -651,7 +648,7 @@ Lemma Aut_ncprod_full n : Proof. move=> AutZinG; elim: n => [|n IHn]. by rewrite center_ncprod0; apply/Aut_sub_fullP=> // g injg gG0; exists g. -by case: ncprodS => gz isoZ; exact: Aut_cprod_by_full. +by case: ncprodS => gz isoZ; apply: Aut_cprod_by_full. Qed. End IterCprod. diff --git a/mathcomp/solvable/commutator.v b/mathcomp/solvable/commutator.v index 6f6a2b5..7a42a9a 100644 --- a/mathcomp/solvable/commutator.v +++ b/mathcomp/solvable/commutator.v @@ -1,12 +1,9 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat. -From mathcomp.discrete -Require Import fintype bigop finset prime binomial. -From mathcomp.fingroup -Require Import fingroup morphism automorphism quotient. -Require Import gfunctor. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat fintype bigop finset. +From mathcomp +Require Import binomial fingroup morphism automorphism quotient gfunctor. (******************************************************************************) (* This files contains the proofs of several key properties of commutators, *) @@ -17,8 +14,8 @@ Require Import gfunctor. (* G^`(0) == G *) (* G^`(n.+1) == [~: G^`(n), G^`(n)] *) (* as several classical results involve the (first) derived group G^`(1), *) -(* such as the equivalence H <| G /\ G / H abelian <-> G^`(1) \subset H. The *) -(* connection between the derived series and solvable groups will only be *) +(* such as the equivalence H <| G /\ G / H abelian <-> G^`(1) \subset H. *) +(* The connection between the derived series and solvable groups will only be *) (* established in nilpotent.v, however. *) (******************************************************************************) @@ -49,7 +46,7 @@ Lemma derg1 A : A^`(1) = [~: A, A]. Proof. by []. Qed. Lemma dergSn n A : A^`(n.+1) = [~: A^`(n), A^`(n)]. Proof. by []. Qed. Lemma der_group_set G n : group_set G^`(n). -Proof. by case: n => [|n]; exact: groupP. Qed. +Proof. by case: n => [|n]; apply: groupP. Qed. Canonical derived_at_group G n := Group (der_group_set G n). @@ -130,7 +127,7 @@ Variables (i j : nat) (x y : gT). Hypotheses (cxz : commute x [~ x, y]) (cyz : commute y [~ x, y]). Lemma commXXg : [~ x ^+ i, y ^+ j] = [~ x, y] ^+ (i * j). -Proof. rewrite expgM commgX commXg //; exact: commuteX. Qed. +Proof. by rewrite expgM commgX commXg //; apply: commuteX. Qed. Lemma expMg_Rmul : (y * x) ^+ i = y ^+ i * x ^+ i * [~ x, y] ^+ 'C(i, 2). Proof. @@ -176,10 +173,10 @@ Lemma commg_normal G H : [~: G, H] <| G <*> H. Proof. by rewrite /(_ <| _) commg_sub commg_norm. Qed. Lemma normsRl A G B : A \subset G -> A \subset 'N([~: G, B]). -Proof. by move=> sAG; exact: subset_trans (commg_norml G B). Qed. +Proof. by move=> sAG; apply: subset_trans (commg_norml G B). Qed. Lemma normsRr A G B : A \subset G -> A \subset 'N([~: B, G]). -Proof. by move=> sAG; exact: subset_trans (commg_normr G B). Qed. +Proof. by move=> sAG; apply: subset_trans (commg_normr G B). Qed. Lemma commg_subr G H : ([~: G, H] \subset H) = (G \subset 'N(H)). Proof. @@ -225,7 +222,7 @@ Lemma sub_der1_normal G H : G^`(1) \subset H -> H \subset G -> H <| G. Proof. by move=> sG'H sHG; rewrite /(H <| G) sHG sub_der1_norm. Qed. Lemma sub_der1_abelian G H : G^`(1) \subset H -> abelian (G / H). -Proof. by move=> sG'H; exact: quotient_cents2r. Qed. +Proof. by move=> sG'H; apply: quotient_cents2r. Qed. Lemma der1_min G H : G \subset 'N(H) -> abelian (G / H) -> G^`(1) \subset H. Proof. by move=> nHG abGH; rewrite -quotient_cents2. Qed. @@ -341,7 +338,7 @@ by rewrite !dergSn -IHn morphimR ?(subset_trans (der_sub n G)). Qed. Lemma dergS n G H : G \subset H -> G^`(n) \subset H^`(n). -Proof. by move=> sGH; elim: n => // n IHn; exact: commgSS. Qed. +Proof. by move=> sGH; elim: n => // n IHn; apply: commgSS. Qed. Lemma quotient_der n G H : G \subset 'N(H) -> G^`(n) / H = (G / H)^`(n). Proof. exact: morphim_der. Qed. diff --git a/mathcomp/algebra/cyclic.v b/mathcomp/solvable/cyclic.v index 4365406..be77dd9 100644 --- a/mathcomp/algebra/cyclic.v +++ b/mathcomp/solvable/cyclic.v @@ -1,12 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div fintype bigop prime finset. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism quotient gproduct. -Require Import ssralg finalg zmodp poly. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div fintype bigop. +From mathcomp +Require Import prime finset fingroup morphism perm automorphism quotient. +From mathcomp +Require Import gproduct ssralg finalg zmodp poly. (******************************************************************************) (* Properties of cyclic groups. *) @@ -118,7 +117,7 @@ End Zpm. (***********************************************************************) Lemma cyclic_abelian A : cyclic A -> abelian A. -Proof. by case/cyclicP=> a ->; exact: cycle_abelian. Qed. +Proof. by case/cyclicP=> a ->; apply: cycle_abelian. Qed. Lemma cycleMsub a b : commute a b -> coprime #[a] #[b] -> <[a]> \subset <[a * b]>. @@ -142,7 +141,7 @@ Lemma cyclicM A B : cyclic (A * B). Proof. move=> /cyclicP[a ->] /cyclicP[b ->]; rewrite cent_cycle cycle_subG => cab coab. -by rewrite -cycleM ?cycle_cyclic //; exact/esym/cent1P. +by rewrite -cycleM ?cycle_cyclic //; apply/esym/cent1P. Qed. Lemma cyclicY K H : @@ -158,7 +157,7 @@ Lemma order_dvdn a n : #[a] %| n = (a ^+ n == 1). Proof. by rewrite (eq_expg_mod_order a n 0) mod0n. Qed. Lemma order_inf a n : a ^+ n.+1 == 1 -> #[a] <= n.+1. -Proof. by rewrite -order_dvdn; exact: dvdn_leq. Qed. +Proof. by rewrite -order_dvdn; apply: dvdn_leq. Qed. Lemma order_dvdG G a : a \in G -> #[a] %| #|G|. Proof. by move=> Ga; apply: cardSg; rewrite cycle_subG. Qed. @@ -193,7 +192,7 @@ by rewrite order_dvdn xp. Qed. Lemma orderXdvd a n : #[a ^+ n] %| #[a]. -Proof. by apply: order_dvdG; exact: mem_cycle. Qed. +Proof. by apply: order_dvdG; apply: mem_cycle. Qed. Lemma orderXgcd a n : #[a ^+ n] = #[a] %/ gcdn #[a] n. Proof. @@ -282,7 +281,7 @@ Lemma generator_cycle a : generator <[a]> a. Proof. exact: eqxx. Qed. Lemma cycle_generator a x : generator <[a]> x -> x \in <[a]>. -Proof. by move/(<[a]> =P _)->; exact: cycle_id. Qed. +Proof. by move/(<[a]> =P _)->; apply: cycle_id. Qed. Lemma generator_order a b : generator <[a]> b -> #[a] = #[b]. Proof. by rewrite /order => /(<[a]> =P _)->. Qed. @@ -374,7 +373,7 @@ Qed. Lemma cycle_subgroup_char a (H : {group gT}) : H \subset <[a]> -> H \char <[a]>. Proof. move=> sHa; apply: lone_subgroup_char => // J sJa isoJH. -have dvHa: #|H| %| #[a] by exact: cardSg. +have dvHa: #|H| %| #[a] by apply: cardSg. have{dvHa} /setP Huniq := esym (cycle_sub_group dvHa). move: (Huniq H) (Huniq J); rewrite !inE /=. by rewrite sHa sJa (card_isog isoJH) eqxx => /eqP<- /eqP<-. @@ -433,7 +432,7 @@ Lemma cardSg_cyclic G H K : Proof. move=> cycG sHG sKG; apply/idP/idP; last exact: cardSg. case/cyclicP: (cyclicS sKG cycG) => x defK; rewrite {K}defK in sKG *. -case/dvdnP=> k ox; suffices ->: H :=: <[x ^+ k]> by exact: cycleX. +case/dvdnP=> k ox; suffices ->: H :=: <[x ^+ k]> by apply: cycleX. apply/eqP; rewrite (eq_subG_cyclic cycG) ?(subset_trans (cycleX _ _)) //. rewrite -orderE orderXdiv orderE ox ?dvdn_mulr ?mulKn //. by have:= order_gt0 x; rewrite orderE ox; case k. @@ -496,7 +495,7 @@ Lemma injm_cyclic G H (f : {morphism G >-> rT}) : 'injm f -> H \subset G -> cyclic (f @* H) = cyclic H. Proof. move=> injf sHG; apply/idP/idP; last exact: morphim_cyclic. -rewrite -{2}(morphim_invm injf sHG); exact: morphim_cyclic. +by rewrite -{2}(morphim_invm injf sHG); apply: morphim_cyclic. Qed. Lemma isog_cyclic G M : G \isog M -> cyclic G = cyclic M. @@ -516,7 +515,7 @@ Lemma injm_generator G H (f : {morphism G >-> rT}) x : Proof. move=> injf Gx sHG; apply/idP/idP; last exact: morph_generator. rewrite -{2}(morphim_invm injf sHG) -{2}(invmE injf Gx). -by apply: morph_generator; exact: mem_morphim. +by apply: morph_generator; apply: mem_morphim. Qed. End IsoCyclic. @@ -577,7 +576,7 @@ Definition cyclem of gT := fun x : gT => x ^+ n. Lemma cyclemM : {in <[a]> & , {morph cyclem a : x y / x * y}}. Proof. -by move=> x y ax ay; apply: expgMn; exact: (centsP (cycle_abelian a)). +by move=> x y ax ay; apply: expgMn; apply: (centsP (cycle_abelian a)). Qed. Canonical cyclem_morphism := Morphism cyclemM. @@ -647,7 +646,7 @@ have a_fa: <[a]> = <[f a]>. have def_n: a ^+ n = f a. by rewrite -/(Zpm n) invmK // im_Zpm a_fa cycle_id. have co_a_n: coprime #[a].-2.+2 n. - by rewrite {1}Zp_cast ?order_gt1 // -generator_coprime def_n; exact/eqP. + by rewrite {1}Zp_cast ?order_gt1 // -generator_coprime def_n; apply/eqP. exists (FinRing.unit 'Z_#[a] co_a_n); rewrite ?inE //. apply: eq_Aut (Af) (Aut_aut _ _) _ => x ax. rewrite autE //= /cyclem; case/cycleP: ax => k ->{x}. @@ -669,7 +668,7 @@ have [lea1 | lt1a] := leqP #[a] 1. rewrite /order card_le1_trivg // cards1 (@eq_card1 _ 1) // => x. by rewrite !inE -cycle_eq1 eq_sym. rewrite -(card_injm (injm_invm (injm_Zpm a))) /= ?im_Zpm; last first. - by apply/subsetP=> x; rewrite inE; exact: cycle_generator. + by apply/subsetP=> x; rewrite inE; apply: cycle_generator. rewrite -card_units_Zp // cardsE card_sub morphim_invmE; apply: eq_card => /= d. by rewrite !inE /= qualifE /= /Zp lt1a inE /= generator_coprime {1}Zp_cast. Qed. @@ -682,10 +681,10 @@ End CycleAutomorphism. Variable G : {group gT}. Lemma Aut_cyclic_abelian : cyclic G -> abelian (Aut G). -Proof. by case/cyclicP=> x ->; exact: Aut_cycle_abelian. Qed. +Proof. by case/cyclicP=> x ->; apply: Aut_cycle_abelian. Qed. Lemma card_Aut_cyclic : cyclic G -> #|Aut G| = totient #|G|. -Proof. by case/cyclicP=> x ->; exact: card_Aut_cycle. Qed. +Proof. by case/cyclicP=> x ->; apply: card_Aut_cycle. Qed. Lemma sum_ncycle_totient : \sum_(d < #|G|.+1) #|[set <[x]> | x in G & #[x] == d]| * totient d = #|G|. @@ -765,13 +764,13 @@ apply/eqP; rewrite eq_sym eqEcard -[#|_|]/n yn leqnn andbT cycle_subG /=. suff{y Gy yn} ->: <[x]> = G :&: [set z | #[z] %| n] by rewrite !inE Gy yn /=. apply/eqP; rewrite eqEcard subsetI cycle_subG {}Gx /= cardE; set rs := enum _. apply/andP; split; first by apply/subsetP=> y xy; rewrite inE order_dvdG. -pose P : {poly R} := ('X^n - 1)%R; have n_gt0: n > 0 by exact: order_gt0. +pose P : {poly R} := ('X^n - 1)%R; have n_gt0: n > 0 by apply: order_gt0. have szP: size P = n.+1 by rewrite size_addl size_polyXn ?size_opp ?size_poly1. rewrite -ltnS -szP -(size_map f) max_ring_poly_roots -?size_poly_eq0 ?{}szP //. apply/allP=> fy /mapP[y]; rewrite mem_enum !inE order_dvdn => /andP[Gy]. move/eqP=> yn1 ->{fy}; apply/eqP. by rewrite !(hornerE, hornerXn) -fX // yn1 f1 subrr. -have: uniq rs by exact: enum_uniq. +have: uniq rs by apply: enum_uniq. have: all (mem G) rs by apply/allP=> y; rewrite mem_enum; case/setIP. elim: rs => //= y rs IHrs /andP[Gy Grs] /andP[y_rs]; rewrite andbC. move/IHrs=> -> {IHrs}//; apply/allP=> _ /mapP[z rs_z ->]. @@ -786,7 +785,7 @@ Lemma field_mul_group_cyclic (F : fieldType) (f : gT -> F) : {in G, forall x, f x = 1%R <-> x = 1} -> cyclic G. Proof. -move=> fM f1P; have f1 : f 1 = 1%R by exact/f1P. +move=> fM f1P; have f1 : f 1 = 1%R by apply/f1P. apply: (div_ring_mul_group_cyclic f1 fM) => [x|]. case/setD1P=> x1 Gx; rewrite unitfE; apply: contra x1. by rewrite subr_eq0 => /eqP/f1P->. @@ -801,7 +800,7 @@ Lemma field_unit_group_cyclic (F : finFieldType) (G : {group {unit F}}) : cyclic G. Proof. apply: field_mul_group_cyclic FinRing.uval _ _ => // u _. -by split=> /eqP ?; exact/eqP. +by split=> /eqP ?; apply/eqP. Qed. Section PrimitiveRoots. @@ -832,7 +831,7 @@ pose sgT := BaseFinGroupType _ (FinGroup.Mixin sG_Ag sG_1g sG_Vg). pose gT := @FinGroupType sgT sG_Vg. have /cyclicP[x gen_x]: @cyclic gT setT. apply: (@field_mul_group_cyclic gT [set: _] F r) => // x _. - by split=> [ri1 | ->]; first exact: val_inj. + by split=> [ri1 | ->]; first apply: val_inj. apply/hasP; exists (r x); first exact: (valP x). have [m prim_x dvdmn] := prim_order_exists n_gt0 (rn1 x). rewrite -((m =P n) _) // eqn_dvd {}dvdmn -sz_rs -(card_seq_sub Urs) -cardsT. diff --git a/mathcomp/solvable/extraspecial.v b/mathcomp/solvable/extraspecial.v index 737838b..137518b 100644 --- a/mathcomp/solvable/extraspecial.v +++ b/mathcomp/solvable/extraspecial.v @@ -1,17 +1,15 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype bigop finset prime binomial. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism quotient action gproduct. -From mathcomp.fingroup -Require Import presentation. -From mathcomp.algebra -Require Import ssralg finalg zmodp matrix cyclic. -Require Import pgroup center gseries commutator gfunctor. -Require Import nilpotent sylow abelian finmodule maximal extremal. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div choice fintype. +From mathcomp +Require Import bigop finset prime binomial fingroup morphism perm automorphism. +From mathcomp +Require Import presentation quotient action commutator gproduct gfunctor. +From mathcomp +Require Import ssralg finalg zmodp cyclic pgroup center gseries. +From mathcomp +Require Import nilpotent sylow abelian finmodule matrix maximal extremal. (******************************************************************************) (* This file contains the fine structure thorems for extraspecial p-groups. *) @@ -332,7 +330,7 @@ Lemma Ohm1_extraspecial_odd (gT : finGroupType) (G : {group gT}) : Proof. move=> pG esG oddG Y; have [spG _] := esG. have [defPhiG defG'] := spG; set Z := 'Z(G) in defPhiG defG'. -have{spG} expG: exponent G %| p ^ 2 by exact: exponent_special. +have{spG} expG: exponent G %| p ^ 2 by apply: exponent_special. have p_pr := extraspecial_prime pG esG. have p_gt1 := prime_gt1 p_pr; have p_gt0 := ltnW p_gt1. have oZ: #|Z| = p := card_center_extraspecial pG esG. @@ -341,7 +339,7 @@ have nsYG: Y <| G := Ohm_normal 1 G; have [sYG nYG] := andP nsYG. have ntZ: Z != 1 by rewrite -cardG_gt1 oZ. have sZY: Z \subset Y. by apply: contraR ntZ => ?; rewrite -(setIidPl sZG) TI_Ohm1 ?prime_TIg ?oZ. -have ntY: Y != 1 by apply: contra ntZ; rewrite -!subG1; exact: subset_trans. +have ntY: Y != 1 by apply: subG1_contra ntZ. have p_odd: odd p by rewrite -oZ (oddSg sZG). have expY: exponent Y %| p by rewrite exponent_Ohm1_class2 // nil_class2 defG'. rewrite (prime_nt_dvdP p_pr _ expY) -?dvdn1 -?trivg_exponent //. @@ -372,7 +370,7 @@ have iYG: #|G : Y| = p. have [sVG [u Gu notVu]] := properP (maxgroupp maxV). without loss [v Vv notYv]: / exists2 v, v \in V & v \notin Y. have [->| ] := eqVneq Y V; first by rewrite (p_maximal_index pG). - by rewrite eqEsubset sYV => not_sVY; apply; exact/subsetPn. + by rewrite eqEsubset sYV => not_sVY; apply; apply/subsetPn. pose U := <[u]> <*> <[v]>; have Gv := subsetP sVG v Vv. have sUG: U \subset G by rewrite join_subG !cycle_subG Gu. have Uu: u \in U by rewrite -cycle_subG joing_subl. @@ -418,7 +416,7 @@ have iC1U (U : {group gT}) x: rewrite -(@dvdn_pmul2l (#|U| * #|'C_G[x]|)) ?muln_gt0 ?cardG_gt0 //. have maxCx: maximal 'C_G[x] G. rewrite cent1_extraspecial_maximal //; apply: contra not_cUx. - by rewrite inE Gx; exact: subsetP (centS sUG) _. + by rewrite inE Gx; apply: subsetP (centS sUG) _. rewrite {1}mul_cardG setIA (setIidPl sUG) -(p_maximal_index pG maxCx) -!mulnA. rewrite !Lagrange ?subsetIl // mulnC dvdn_pmul2l //. have [sCxG nCxG] := andP (p_maximal_normal pG maxCx). @@ -432,7 +430,7 @@ have oCG (U : {group gT}): by apply/subsetPn; rewrite eqEsubset sZU subsetI sUG centsC in neZU. pose W := 'C_U[x]; have iWU: #|U : W| = p by rewrite iC1U // inE not_cUx. have maxW: maximal W U by rewrite p_index_maximal ?subsetIl ?iWU. - have ltWU: W \proper U by exact: maxgroupp maxW. + have ltWU: W \proper U by apply: maxgroupp maxW. have [sWU [u Uu notWu]] := properP ltWU. have defU: W * <[u]> = U. have nsWU: W <| U := p_maximal_normal (pgroupS sUG pG) maxW. @@ -548,8 +546,8 @@ have not_cEE: ~~ abelian E. by apply: contra not_le_p3_p => cEE; rewrite -oE -oZ -defZE (center_idP _). have sES: E \subset S by rewrite -defET mulG_subl. have sTS: T \subset S by rewrite -defET mulG_subr. -have expE: exponent E %| p by exact: dvdn_trans (exponentS sES) expG. -have expT: exponent T %| p by exact: dvdn_trans (exponentS sTS) expG. +have expE: exponent E %| p by apply: dvdn_trans (exponentS sES) expG. +have expT: exponent T %| p by apply: dvdn_trans (exponentS sTS) expG. have{expE not_cEE} isoE: E \isog p^{1+2}. apply: isog_pX1p2 => //. by apply: card_p3group_extraspecial p_pr oE _; rewrite defZE. @@ -589,7 +587,7 @@ Qed. Lemma Q8_extraspecial : extraspecial 'Q_8. Proof. -have gt32: 3 > 2 by []; have isoQ: 'Q_8 \isog 'Q_(2 ^ 3) by exact: isog_refl. +have gt32: 3 > 2 by []; have isoQ: 'Q_8 \isog 'Q_(2 ^ 3) by apply: isog_refl. have [[x y] genQ _] := generators_quaternion gt32 isoQ. have [_ [defQ' defPhiQ _ _]] := quaternion_structure gt32 genQ isoQ. case=> defZ oZ _ _ _ _ _; split; last by rewrite oZ. @@ -673,9 +671,9 @@ rewrite subEproper; case/predU1P=> [defG1 | ltZG1]. have [n' n'_gt2 isoG]: exists2 n', n' > 2 & G \isog 'Q_(2 ^ n'). apply/quaternion_classP; apply/eqP. have not_cycG: ~~ cyclic G. - by apply: contra (extraspecial_nonabelian esG); exact: cyclic_abelian. + by apply: contra (extraspecial_nonabelian esG); apply: cyclic_abelian. move: oZ; rewrite defG1; move/prime_Ohm1P; rewrite (negPf not_cycG) /=. - by apply=> //; apply: contra not_cycG; move/eqP->; exact: cyclic1. + by apply=> //; apply: contra not_cycG; move/eqP->; apply: cyclic1. have [n0 n'3]: n = 0%N /\ n' = 3. have [[x y] genG _] := generators_quaternion n'_gt2 isoG. have n'3: n' = 3. @@ -689,7 +687,7 @@ rewrite subEproper; case/predU1P=> [defG1 | ltZG1]. by rewrite center_ncprod0. case/andP: ltZG1 => _; rewrite (OhmE _ pG) gen_subG. case/subsetPn=> x; case/LdivP=> Gx x2 notZx. -have ox: #[x] = 2 by exact: nt_prime_order (group1_contra notZx). +have ox: #[x] = 2 by apply: nt_prime_order (group1_contra notZx). have Z'x: x \in G :\: 'Z(G) by rewrite inE notZx. have [E [R [[oE oR] [defG ziER]]]] := split1_extraspecial pG esG Z'x. case=> defZE defZR [esE Ex] esR. @@ -744,8 +742,8 @@ Qed. Lemma rank_Dn n : 'r_2('D^n) = n.+1. Proof. elim: n => [|n IHn]; first by rewrite p_rank_abelem ?prime_abelem ?card_pX1p2n. -have oDDn: #|'D^n.+1| = (2 ^ n.+1.*2.+1)%N by exact: card_pX1p2n. -have esDDn: extraspecial 'D^n.+1 by exact: pX1p2n_extraspecial. +have oDDn: #|'D^n.+1| = (2 ^ n.+1.*2.+1)%N by apply: card_pX1p2n. +have esDDn: extraspecial 'D^n.+1 by apply: pX1p2n_extraspecial. do [case: pX1p2S => gz isoZ; set DDn := [set: _]] in oDDn esDDn *. have pDDn: 2.-group DDn by rewrite /pgroup oDDn pnat_exp. apply/eqP; rewrite eqn_leq; apply/andP; split. @@ -787,7 +785,7 @@ have esDnQ: extraspecial 'D^n*Q := DnQ_extraspecial n. do [case: DnQ_P => gz isoZ; set DnQ := setT] in pDnQ esDnQ *. suffices [E]: exists2 E, E \in 'E*_2(DnQ) & logn 2 #|E| = n.+1. by rewrite (pmaxElem_extraspecial pDnQ esDnQ); case/pnElemP=> _ _ <-. -have oZ: #|'Z(DnQ)| = 2 by exact: card_center_extraspecial. +have oZ: #|'Z(DnQ)| = 2 by apply: card_center_extraspecial. pose Dn := cpairg1 isoZ @* 'D^n; pose Q := cpair1g isoZ @* 'Q_8. have [injDn injQ] := (injm_cpairg1 isoZ, injm_cpair1g isoZ). have [E EprE]:= p_rank_witness 2 [group of Dn]. diff --git a/mathcomp/solvable/extremal.v b/mathcomp/solvable/extremal.v index e233a86..83a9fb8 100644 --- a/mathcomp/solvable/extremal.v +++ b/mathcomp/solvable/extremal.v @@ -1,17 +1,15 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype bigop finset prime binomial. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism quotient action gproduct. -From mathcomp.fingroup -Require Import presentation. -From mathcomp.algebra -Require Import ssralg finalg zmodp cyclic matrix. -Require Import pgroup center gseries commutator gfunctor. -Require Import nilpotent sylow abelian finmodule maximal. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div choice fintype. +From mathcomp +Require Import bigop finset prime binomial fingroup morphism perm automorphism. +From mathcomp +Require Import presentation quotient action commutator gproduct gfunctor. +From mathcomp +Require Import ssralg finalg zmodp cyclic pgroup center gseries. +From mathcomp +Require Import nilpotent sylow abelian finmodule matrix maximal. (******************************************************************************) (* This file contains the definition and properties of extremal p-groups; *) @@ -253,10 +251,10 @@ have mV: {in A, {morph m : a / a^-1 >-> (a^-1)%R}}. by rewrite -mM ?groupV ?mulVg. have inv_m (u : 'Z_q) : coprime q u -> {a | a \in A & m a = u}. rewrite -?unitZpE // natr_Zp -im_m => m_u. - by exists (iinv m_u); [exact: mem_iinv | rewrite f_iinv]. + by exists (iinv m_u); [apply: mem_iinv | rewrite f_iinv]. have [cycF ffulF]: cyclic F /\ [faithful F, on 'Ohm_1(G) | [Aut G]]. have Um0 a: ((m a)%:R : 'F_p) \in GRing.unit. - have: m a \in GRing.unit by exact: valP. + have: m a \in GRing.unit by apply: valP. by rewrite -{1}[m a]natr_Zp unitFpE ?unitZpE // {1}/q oG coprime_pexpl. pose fm0 a := FinRing.unit 'F_p (Um0 a). have natZqp u: (u%:R : 'Z_q)%:R = u %:R :> 'F_p. @@ -271,7 +269,7 @@ have [cycF ffulF]: cyclic F /\ [faithful F, on 'Ohm_1(G) | [Aut G]]. by rewrite {1}[q]oG coprime_pexpl // -unitFpE // natZqp natr_Zp. by exists a => //; apply: val_inj; rewrite /= m_a natZqp natr_Zp. have [x1 defG1]: exists x1, 'Ohm_1(G) = <[x1]>. - by apply/cyclicP; exact: cyclicS (Ohm_sub _ _) cycG. + by apply/cyclicP; apply: cyclicS (Ohm_sub _ _) cycG. have ox1: #[x1] = p by rewrite orderE -defG1 (Ohm1_cyclic_pgroup_prime _ pG). have Gx1: x1 \in G by rewrite -cycle_subG -defG1 Ohm_sub. have ker_m0: 'ker m0 = 'C('Ohm_1(G) | [Aut G]). @@ -293,7 +291,7 @@ have [cycF ffulF]: cyclic F /\ [faithful F, on 'Ohm_1(G) | [Aut G]]. have inj_m0: 'ker_F m0 \subset [1] by rewrite setIC ker_m0_P tiPF. split; last by rewrite /faithful -ker_m0. have isogF: F \isog [set: {unit 'F_p}]. - have sFA: F \subset A by exact: pcore_sub. + have sFA: F \subset A by apply: pcore_sub. apply/isogP; exists (restrm_morphism sFA m0); first by rewrite ker_restrm. apply/eqP; rewrite eqEcard subsetT card_injm ?ker_restrm //= oF. by rewrite card_units_Fp. @@ -441,7 +439,7 @@ Lemma Grp_modular_group : 'Mod_(p ^ n) \isog Grp (x : y : (x ^+ q, y ^+ p, x ^ y = x ^+ r.+1)). Proof. rewrite /modular_gtype def_p def_q def_r; apply: Extremal.Grp => //. -set B := <[_]>; have Bb: Zp1 \in B by exact: cycle_id. +set B := <[_]>; have Bb: Zp1 \in B by apply: cycle_id. have oB: #|B| = q by rewrite -orderE order_Zp1 Zp_cast. have cycB: cyclic B by rewrite cycle_cyclic. have pB: p.-group B by rewrite /pgroup oB pnat_exp ?pnat_id. @@ -472,7 +470,7 @@ have [Gx Gy]: x \in G /\ y \in G. have notXy: y \notin <[x]>. apply: contraL ltqm; rewrite -cycle_subG -oG -defG; move/mulGidPl->. by rewrite -leqNgt dvdn_leq ?(ltnW q_gt1) // order_dvdn xq. -have oy: #[y] = p by exact: nt_prime_order (group1_contra notXy). +have oy: #[y] = p by apply: nt_prime_order (group1_contra notXy). exists (x, y) => //=; split; rewrite ?inE ?notXy //. apply/eqP; rewrite -(eqn_pmul2r p_gt0) -expnSr -{1}oy (ltn_predK n_gt2) -/m. by rewrite -TI_cardMg ?defG ?oG // setIC prime_TIg ?cycle_subG // -orderE oy. @@ -528,7 +526,7 @@ have defZ: 'Z(G) = <[x ^+ p]>. rewrite (dvdn_prime_cyclic p_pr) // card_quotient //. rewrite -(dvdn_pmul2l (cardG_gt0 'Z(G))) Lagrange // oG -def_m dvdn_pmul2r //. case/p_natP: (pgroupS sZG pG) gtZr => k ->. - by rewrite ltn_exp2l // def_n1; exact: dvdn_exp2l. + by rewrite ltn_exp2l // def_n1; apply: dvdn_exp2l. have Zxr: x ^+ r \in 'Z(G) by rewrite /r def_n expnS expgM defZ mem_cycle. have rxy: [~ x, y] = x ^+ r by rewrite /commg xy expgS mulKg. have defG': G^`(1) = <[x ^+ r]>. @@ -542,7 +540,7 @@ have nil2_G: nil_class G = 2. have XYp: {in X & <[y]>, forall z t, (z * t) ^+ p \in z ^+ p *: <[x ^+ r ^+ 'C(p, 2)]>}. - move=> z t Xz Yt; have Gz := subsetP sXG z Xz; have Gt := subsetP sYG t Yt. - have Rtz: [~ t, z] \in G^`(1) by exact: mem_commg. + have Rtz: [~ t, z] \in G^`(1) by apply: mem_commg. have cGtz: [~ t, z] \in 'C(G) by case/setIP: (subsetP sG'Z _ Rtz). rewrite expMg_Rmul /commute ?(centP cGtz) //. have ->: t ^+ p = 1 by apply/eqP; rewrite -order_dvdn -oy order_dvdG. @@ -646,7 +644,7 @@ Lemma card_dihedral : #|'D_m| = m. Proof. by rewrite /('D_m)%type def_q card_ext_dihedral ?mul1n. Qed. Lemma Grp_dihedral : 'D_m \isog Grp (x : y : (x ^+ q, y ^+ 2, x ^ y = x^-1)). -Proof. by rewrite /('D_m)%type def_q; exact: Grp_ext_dihedral. Qed. +Proof. by rewrite /('D_m)%type def_q; apply: Grp_ext_dihedral. Qed. Lemma Grp'_dihedral : 'D_m \isog Grp (x : y : (x ^+ 2, y ^+ 2, (x * y) ^+ q)). Proof. @@ -873,7 +871,7 @@ have{defG} defG: <[x]> * <[y]> = G. have notXy: y \notin <[x]>. apply: contraL ltqm => Xy; rewrite -leqNgt -oG -defG mulGSid ?cycle_subG //. by rewrite dvdn_leq // order_dvdn xq. -have oy: #[y] = 2 by exact: nt_prime_order (group1_contra notXy). +have oy: #[y] = 2 by apply: nt_prime_order (group1_contra notXy). have ox: #[x] = q. apply: double_inj; rewrite -muln2 -oy -mul2n def2q -oG -defG TI_cardMg //. by rewrite setIC prime_TIg ?cycle_subG // -orderE oy. @@ -895,7 +893,7 @@ have{defG} defG: <[x]> * <[y]> = G. have notXy: y \notin <[x]>. apply: contraL ltqm => Xy; rewrite -leqNgt -oG -defG mulGSid ?cycle_subG //. by rewrite dvdn_leq // order_dvdn xq. -have oy: #[y] = 2 by exact: nt_prime_order (group1_contra notXy). +have oy: #[y] = 2 by apply: nt_prime_order (group1_contra notXy). have ox: #[x] = q. apply: double_inj; rewrite -muln2 -oy -mul2n def2q -oG -defG TI_cardMg //. by rewrite setIC prime_TIg ?cycle_subG // -orderE oy. @@ -1042,7 +1040,7 @@ have maxMt: {in G :\: X, forall t, maximal <<t ^: G>> G}. have X'xy: x * y \in G :\: X by rewrite !inE !groupMl ?cycle_id ?notXy. have ti_yG_xyG: [disjoint yG & xyG]. apply/pred0P=> t; rewrite /= /yG /xyG !def_tG //; apply/andP=> [[yGt]]. - rewrite rcoset_sym (rcoset_transl yGt) mem_rcoset mulgK; move/order_dvdG. + rewrite rcoset_sym (rcoset_eqP yGt) mem_rcoset mulgK; move/order_dvdG. by rewrite -orderE ox2 ox gtnNdvd. have s_tG_X': {in G :\: X, forall t, t ^: G \subset G :\: X}. by move=> t X't /=; rewrite class_sub_norm // normsD ?normG. @@ -1066,7 +1064,7 @@ split. apply/eqP; rewrite eqEsubset Ohm_sub -{1}defXY mulG_subG !cycle_subG. by rewrite -(groupMr _ (sX'G1 y X'y)) !sX'G1. - split=> //= H; apply/idP/idP=> [maxH |]; last first. - by case/or3P; move/eqP->; rewrite ?maxMt. + by case/or3P=> /eqP->; rewrite ?maxMt. have [sHG nHG]:= andP (p_maximal_normal pG maxH). have oH: #|H| = q. apply: double_inj; rewrite -muln2 -(p_maximal_index pG maxH) Lagrange //. @@ -1075,7 +1073,7 @@ split. case sHX: (H \subset X) => //=; case/subsetPn: sHX => t Ht notXt. have: t \in yG :|: xyG by rewrite defX' inE notXt (subsetP sHG). rewrite !andbT !gen_subG /yG /xyG. - by case/setUP; move/class_transr <-; rewrite !class_sub_norm ?Ht ?orbT. + by case/setUP; move/class_eqP <-; rewrite !class_sub_norm ?Ht ?orbT. rewrite eqn_leq n_gt1; case: leqP n2_abelG => //= n_gt2 _. have ->: 'Z(G) = <[x ^+ r]>. apply/eqP; rewrite eqEcard andbC -orderE oxr -{1}(setIidPr (center_sub G)). @@ -1240,14 +1238,14 @@ have oMt: {in G :\: X, forall t, #|<<t ^: G>>| = q}. rewrite quotient_cycle ?(subsetP (nXiG 2)) //= -defPhi. rewrite -orderE (abelem_order_p (Phi_quotient_abelem pG)) ?mem_quotient //. apply: contraNneq notXt; move/coset_idr; move/implyP=> /=. - by rewrite defPhi ?(subsetP (nXiG 2)) //; apply: subsetP; exact: cycleX. + by rewrite defPhi ?(subsetP (nXiG 2)) //; apply: subsetP; apply: cycleX. have maxMt: {in G :\: X, forall t, maximal <<t ^: G>> G}. move=> t X't; rewrite /= p_index_maximal -?divgS ?sMtG ?oMt //. by rewrite oG -def2q mulnK. have X'xy: x * y \in G :\: X by rewrite !inE !groupMl ?cycle_id ?notXy. have ti_yG_xyG: [disjoint yG & xyG]. apply/pred0P=> t; rewrite /= /yG /xyG !def_tG //; apply/andP=> [[yGt]]. - rewrite rcoset_sym (rcoset_transl yGt) mem_rcoset mulgK; move/order_dvdG. + rewrite rcoset_sym (rcoset_eqP yGt) mem_rcoset mulgK; move/order_dvdG. by rewrite -orderE ox2 ox gtnNdvd. have s_tG_X': {in G :\: X, forall t, t ^: G \subset G :\: X}. by move=> t X't /=; rewrite class_sub_norm // normsD ?normG. @@ -1286,7 +1284,7 @@ rewrite pprodE //; split=> // [|||n_gt3]. case sHX: (H \subset X) => //=; case/subsetPn: sHX => z Hz notXz. have: z \in yG :|: xyG by rewrite defX' inE notXz (subsetP sHG). rewrite !andbT !gen_subG /yG /xyG. - by case/setUP=> /class_transr <-; rewrite !class_sub_norm ?Hz ?orbT. + by case/setUP=> /class_eqP <-; rewrite !class_sub_norm ?Hz ?orbT. have isoMt: {in G :\: X, forall z, <<z ^: G>> \isog 'Q_q}. have n1_gt2: n.-1 > 2 by rewrite -(subnKC n_gt3). move=> z X'z /=; rewrite isogEcard card_quaternion ?oMt // leqnn andbT. @@ -1385,7 +1383,7 @@ have defG1: 'Mho^1(G) = <[x ^+ 2]>. by case/cycleP: Xz => i ->; rewrite -expgM mulnC expgM mem_cycle. have{Xz Gz} [xi Xxi ->{z}]: exists2 xi, xi \in X & z = xi * y. have Uvy: y \in <[u]> :* v by rewrite defUv -(defU x). - apply/rcosetP; rewrite /X defU // (rcoset_transl Uvy) defUv. + apply/rcosetP; rewrite /X defU // (rcoset_eqP Uvy) defUv. by rewrite inE -(defU x) ?Xz. rewrite expn1 expgS {2}(conjgC xi) -{2}[y]/(y ^+ 2.-1) -{1}oy -invg_expg. rewrite mulgA mulgK invXX' // -expgS prednK // /r -(subnKC n_gt3) expnS. @@ -1427,14 +1425,14 @@ have oMt: {in G :\: X, forall t, #|<<t ^: G>>| = q}. rewrite quotient_cycle ?(subsetP (nXiG 2)) //= -defPhi -orderE. rewrite (abelem_order_p (Phi_quotient_abelem pG)) ?mem_quotient //. apply: contraNneq notXt; move/coset_idr; move/implyP=> /=. - by rewrite /= defPhi (subsetP (nXiG 2)) //; apply: subsetP; exact: cycleX. + by rewrite /= defPhi (subsetP (nXiG 2)) //; apply: subsetP; apply: cycleX. have maxMt: {in G :\: X, forall t, maximal <<t ^: G>> G}. move=> t X't /=; rewrite p_index_maximal -?divgS ?sMtG ?oMt //. by rewrite oG -def2q mulnK. have X'xy: x * y \in G :\: X by rewrite !inE !groupMl ?cycle_id ?notXy. have ti_yG_xyG: [disjoint yG & xyG]. apply/pred0P=> t; rewrite /= /yG /xyG !def_tG //; apply/andP=> [[yGt]]. - rewrite rcoset_sym (rcoset_transl yGt) mem_rcoset mulgK; move/order_dvdG. + rewrite rcoset_sym (rcoset_eqP yGt) mem_rcoset mulgK; move/order_dvdG. by rewrite -orderE ox2 ox gtnNdvd. have s_tG_X': {in G :\: X, forall t, t ^: G \subset G :\: X}. by move=> t X't /=; rewrite class_sub_norm // normsD ?normG. @@ -1481,7 +1479,7 @@ split. case sHX: (H \subset X) => //=; case/subsetPn: sHX => t Ht notXt. have: t \in yG :|: xyG by rewrite defX' inE notXt (subsetP sHG). rewrite !andbT !gen_subG /yG /xyG. - by case/setUP=> /class_transr <-; rewrite !class_sub_norm ?Ht ?orbT. + by case/setUP=> /class_eqP <-; rewrite !class_sub_norm ?Ht ?orbT. have n1_gt2: n.-1 > 2 by [rewrite -(subnKC n_gt3)]; have n1_gt1 := ltnW n1_gt2. rewrite !isogEcard card_2dihedral ?card_quaternion ?oMt // leqnn !andbT. have invX2X': {in G :\: X, forall t, x ^+ 2 ^ t == x ^- 2}. @@ -1686,7 +1684,7 @@ have:= semidihedral_structure n_gt3 genG isoG oy. rewrite -/X -/q -/r -/yG -/xyG -/My -/Mxy. case=> [[defG oxy invXX'] nilG [defZ oZ [-> ->] defMho] [[defX' tiX'] maxG]]. case=> isoMy isoMxy defX; do 2!split=> //. - by apply/dihedral_classP; exists n.-1; first exact: ltnW. + by apply/dihedral_classP; exists n.-1; first apply: ltnW. by apply/quaternion_classP; exists n.-1. Qed. @@ -1724,7 +1722,7 @@ pose Z := <[x ^+ r]>. have defZ: Z = 'Ohm_1(X) by rewrite defX (Ohm_p_cycle _ p_x) ox subn1 pfactorK. have oZ: #|Z| = p by rewrite -orderE orderXdiv ox -def_pr ?dvdn_mull ?mulnK. have cGZ: Z \subset 'C(G). - have nsZG: Z <| G by rewrite defZ (char_normal_trans (Ohm_char 1 _)). + have nsZG: Z <| G by rewrite defZ gFnormal_trans. move/implyP: (meet_center_nil (pgroup_nil pG) nsZG). rewrite -cardG_gt1 oZ p_gt1 setIA (setIidPl (normal_sub nsZG)). by apply: contraR; move/prime_TIg=> -> //; rewrite oZ. @@ -1867,9 +1865,9 @@ Proof. move=> pG /SCN_P[nsUG scUG] not_cGG cycU; have [sUG nUG] := andP nsUG. have [cUU pU] := (cyclic_abelian cycU, pgroupS sUG pG). have ltUG: ~~ (G \subset U). - by apply: contra not_cGG => sGU; exact: abelianS cUU. + by apply: contra not_cGG => sGU; apply: abelianS cUU. have ntU: U :!=: 1. - by apply: contra ltUG; move/eqP=> U1; rewrite -(setIidPl (cents1 G)) -U1 scUG. + by apply: contraNneq ltUG => U1; rewrite -scUG subsetIidl U1 cents1. have [p_pr _ [n oU]] := pgroup_pdiv pU ntU. have p_gt1 := prime_gt1 p_pr; have p_gt0 := ltnW p_gt1. have [u defU] := cyclicP cycU; have Uu: u \in U by rewrite defU cycle_id. @@ -1879,7 +1877,7 @@ have modM1 (M : {group gT}): [/\ U \subset M, #|M : U| = p & extremal_class M = ModularGroup] -> M :=: 'C_M('Mho^1(U)) /\ 'Ohm_1(M)%G \in 'E_p^2(M). - case=> sUM iUM /modular_group_classP[q q_pr {n oU}[n n_gt23 isoM]]. - have n_gt2: n > 2 by exact: leq_trans (leq_addl _ _) n_gt23. + have n_gt2: n > 2 by apply: leq_trans (leq_addl _ _) n_gt23. have def_n: n = (n - 3).+3 by rewrite -{1}(subnKC n_gt2). have oM: #|M| = (q ^ n)%N by rewrite (card_isog isoM) card_modular_group. have pM: q.-group M by rewrite /pgroup oM pnat_exp pnat_id. @@ -1905,7 +1903,7 @@ have [_ _ [[|k] oGs]] := pgroup_pdiv pGs ntGs. by right; exists G; case: (modM1 G) => // <- ->; rewrite Ohm_char. by left; case: eqP ext2G => // <-. pose M := 'C_G('Mho^1(U)); right; exists [group of M]. -have sMG: M \subset G by exact: subsetIl. +have sMG: M \subset G by apply: subsetIl. have [pM nUM] := (pgroupS sMG pG, subset_trans sMG nUG). have sUM: U \subset M by rewrite subsetI sUG sub_abelian_cent ?Mho_sub. pose A := Aut U; have cAA: abelian A by rewrite Aut_cyclic_abelian. @@ -1986,9 +1984,9 @@ have not_cMM: ~~ abelian M. by rewrite (setIidPl _) ?indexgg // -scUG subsetI sMG sub_abelian_cent. have modM: extremal_class M = ModularGroup. have sU1Z: 'Mho^1(U) \subset 'Z(M). - by rewrite subsetI (subset_trans (Mho_sub 1 U)) // centsC subsetIr. - case: (predU1P (maximal_cycle_extremal _ _ _ _ iUM)) => //=; rewrite -/M. - case/andP; move/eqP=> p2 ext2M; rewrite p2 add1n in n_gt01. + by rewrite subsetI gFsub_trans // centsC subsetIr. + have /maximal_cycle_extremal/predU1P[] //= := iUM; rewrite -/M. + case/andP=> /eqP-p2 ext2M; rewrite p2 add1n in n_gt01. suffices{sU1Z}: #|'Z(M)| = 2. move/eqP; rewrite eqn_leq leqNgt (leq_trans _ (subset_leq_card sU1Z)) //. by rewrite defU1 -orderE (orderXexp 1 ou) subn1 p2 (ltn_exp2l 1). @@ -2005,9 +2003,9 @@ have modM: extremal_class M = ModularGroup. by case: (m == 2) => [|[]//]; move/abelem_abelian->. split=> //. have [//|_] := modM1 [group of M]; rewrite !inE -andbA /=. - by case/andP; move/subset_trans->. + by case/andP=> /subset_trans->. have{cycGs} [cycGs | [p2 [c fGs_c u_c]]] := cycGs. - suffices ->: 'Ohm_1(M) = 'Ohm_1(G) by exact: Ohm_char. + suffices ->: 'Ohm_1(M) = 'Ohm_1(G) by apply: Ohm_char. suffices sG1M: 'Ohm_1(G) \subset M. by apply/eqP; rewrite eqEsubset -{2}(Ohm_id 1 G) !OhmS. rewrite -(quotientSGK _ sUM) ?(subset_trans (Ohm_sub _ G)) //= defMs. @@ -2017,13 +2015,12 @@ have{cycGs} [cycGs | [p2 [c fGs_c u_c]]] := cycGs. rewrite -(part_pnat_id (pgroupS (Ohm_sub _ _) pGs)) p_part (leq_exp2l _ 1) //. by rewrite -p_rank_abelian -?rank_pgroup -?abelian_rank1_cyclic. suffices charU1: 'Mho^1(U) \char G^`(1). - rewrite (char_trans (Ohm_char _ _)) // subcent_char ?char_refl //. - exact: char_trans (der_char 1 G). + by rewrite gFchar_trans // subcent_char ?(char_trans charU1) ?gFchar. suffices sUiG': 'Mho^1(U) \subset G^`(1). - have cycG': cyclic G^`(1) by rewrite (cyclicS _ cycU) // der1_min. - by case/cyclicP: cycG' sUiG' => zs ->; exact: cycle_subgroup_char. + have /cyclicP[zs cycG']: cyclic G^`(1) by rewrite (cyclicS _ cycU) ?der1_min. + by rewrite cycG' in sUiG' *; apply: cycle_subgroup_char. rewrite defU1 cycle_subG p2 -groupV invMg -{2}u_c. -by case/morphimP: fGs_c => _ _ /morphimP[z _ Gz ->] ->; rewrite fG ?mem_commg. +by have [_ _ /morphimP[z _ Gz ->] ->] := morphimP fGs_c; rewrite fG ?mem_commg. Qed. (* This is Aschbacher, exercise (8.4) *) @@ -2038,7 +2035,7 @@ have [X maxX]: {X | [max X | X <| G & abelian X]}. have cycX: cyclic X by rewrite dn_G_1; case/andP: (maxgroupp maxX). have scX: X \in 'SCN(G) := max_SCN pG maxX. have [[p2 _ cG] | [M [_ _ _]]] := cyclic_SCN pG scX not_cGG cycX; last first. - rewrite 2!inE -andbA; case/and3P=> sEG abelE dimE_2 charE. + rewrite 2!inE -andbA => /and3P[sEG abelE dimE_2] charE. have:= dn_G_1 _ (char_normal charE) (abelem_abelian abelE). by rewrite (abelem_cyclic abelE) (eqP dimE_2). have [n oG] := p_natP pG; right; rewrite p2 cG /= in oG *. @@ -2113,9 +2110,9 @@ move=> pG sympG; have [H [charH]] := Thompson_critical pG. have sHG := char_sub charH; have pH := pgroupS sHG pG. set U := 'Z(H) => sPhiH_U sHG_U defU; set Z := 'Ohm_1(U). have sZU: Z \subset U by rewrite Ohm_sub. -have charU: U \char G := char_trans (center_char H) charH. +have charU: U \char G := gFchar_trans _ charH. have cUU: abelian U := center_abelian H. -have cycU: cyclic U by exact: sympG. +have cycU: cyclic U by apply: sympG. have pU: p.-group U := pgroupS (char_sub charU) pG. have cHU: U \subset 'C(H) by rewrite subsetIr. have cHsHs: abelian (H / Z). @@ -2130,11 +2127,10 @@ have [K /=] := inv_quotientS nsZH (Ohm_sub 1 (H / Z)); fold Z => defKs sZK sKH. have nsZK: Z <| K := normalS sZK sKH nsZH; have [_ nZK] := andP nsZK. have abelKs: p.-abelem (K / Z) by rewrite -defKs Ohm1_abelem ?quotient_pgroup. have charK: K \char G. - have charZ: Z \char H := char_trans (Ohm_char _ _) (center_char H). + have charZ: Z \char H := gFchar_trans _ (center_char H). rewrite (char_trans _ charH) // (char_from_quotient nsZK) //. by rewrite -defKs Ohm_char. -have cycZK: cyclic 'Z(K). - by rewrite sympG ?center_abelian ?(char_trans (center_char _)). +have cycZK: cyclic 'Z(K) by rewrite sympG ?center_abelian ?gFchar_trans. have [cKK | not_cKK] := orP (orbN (abelian K)). have defH: U = H. apply: center_idP; apply: cyclic_factor_abelian (Ohm_sub 1 _) _. @@ -2159,14 +2155,14 @@ have [cKK | not_cKK] := orP (orbN (abelian K)). by apply: (p3group_extraspecial pG); rewrite // p2 oG pfactorK. exists G; [by right | exists ('Z(G))%G; rewrite cprod_center_id setIA setIid]. by split=> //; left; rewrite prime_cyclic; case: esG. -have ntK: K :!=: 1 by apply: contra not_cKK; move/eqP->; exact: abelian1. +have ntK: K :!=: 1 by apply: contra not_cKK => /eqP->; apply: abelian1. have [p_pr _ _] := pgroup_pdiv (pgroupS sKH pH) ntK. have p_gt1 := prime_gt1 p_pr; have p_gt0 := ltnW p_gt1. have oZ: #|Z| = p. apply: Ohm1_cyclic_pgroup_prime => //=; apply: contra ntK; move/eqP. by move/(trivg_center_pgroup pH)=> GH; rewrite -subG1 -GH. have sZ_ZK: Z \subset 'Z(K). - by rewrite subsetI sZK (subset_trans (Ohm_sub _ _ )) // subIset ?centS ?orbT. + by rewrite subsetI sZK gFsub_trans // subIset ?centS ?orbT. have sZsKs: 'Z(K) / Z \subset K / Z by rewrite quotientS ?center_sub. have [Es /= splitKs] := abelem_split_dprod abelKs sZsKs. have [_ /= defEsZs cEsZs tiEsZs] := dprodP splitKs. @@ -2187,7 +2183,7 @@ have sEH := subset_trans sEK sKH; have pE := pgroupS sEH pH. have esE: extraspecial E. split; last by rewrite defZE oZ. have sPhiZ: 'Phi(E) \subset Z. - rewrite -quotient_sub1 ?(subset_trans (Phi_sub _)) ?(quotient_Phi pE) //. + rewrite -quotient_sub1 ?gFsub_trans ?(quotient_Phi pE) //. rewrite subG1 (trivg_Phi (quotient_pgroup _ pE)) /= -defEs. by rewrite (abelemS sEsKs) //= -defKs Ohm1_abelem ?quotient_pgroup. have sE'Phi: E^`(1) \subset 'Phi(E) by rewrite (Phi_joing pE) joing_subl. @@ -2207,8 +2203,8 @@ have{defE'} sEG_E': [~: E, G] \subset E^`(1). have Uge: [~ g, e] \in U by rewrite (subsetP sHG_U) ?mem_commg. rewrite inE Uge inE -commgX; last by red; rewrite -(centsP cHU). have sZ_ZG: Z \subset 'Z(G). - have charZ: Z \char G := char_trans (Ohm_char _ _) charU. - move/implyP: (meet_center_nil (pgroup_nil pG) (char_normal charZ)). + have charZ: Z \char G := gFchar_trans _ charU. + have/implyP:= meet_center_nil (pgroup_nil pG) (char_normal charZ). rewrite -cardG_gt1 oZ prime_gt1 //=; apply: contraR => not_sZ_ZG. by rewrite prime_TIg ?oZ. have: e ^+ p \in 'Z(G). @@ -2217,7 +2213,7 @@ have{defE'} sEG_E': [~: E, G] \subset E^`(1). by case/setIP=> _ /centP cGep; apply/commgP; red; rewrite cGep. have sEG: E \subset G := subset_trans sEK (char_sub charK). set R := 'C_G(E). -have{sEG_E'} defG: E \* R = G by exact: (critical_extraspecial pG). +have{sEG_E'} defG: E \* R = G by apply: (critical_extraspecial pG). have [_ defER cRE] := cprodP defG. have defH: E \* 'C_H(E) = H by rewrite -(setIidPr sHG) setIAC (cprod_modl defG). have{defH} [_ defH cRH_E] := cprodP defH. @@ -2232,7 +2228,7 @@ have cRH_RH: abelian 'C_H(E). have:= Ohm_dprod 1 defHs; rewrite /= defKs (Ohm1_id (abelemS sEsKs abelKs)). rewrite dprodC; case/dprodP=> _ defEsRHs1 cRHs1Es tiRHs1Es. have sRHsHs: 'C_H(E) / Z \subset H / Z by rewrite quotientS ?subsetIl. - have cRHsRHs: abelian ('C_H(E) / Z) by exact: abelianS cHsHs. + have cRHsRHs: abelian ('C_H(E) / Z) by apply: abelianS cHsHs. have pHs: p.-group (H / Z) by rewrite quotient_pgroup. rewrite abelian_rank1_cyclic // (rank_pgroup (pgroupS sRHsHs pHs)). rewrite p_rank_abelian // -(leq_add2r (logn p #|Es|)) -lognM ?cardG_gt0 //. @@ -2259,14 +2255,12 @@ have [[p2 _ ext2R] | [M []]] := cyclic_SCN pR scUR not_cRR cycU. exists E; [by right | exists [group of R]; split=> //; right]. by rewrite dvdn_leq ?(@pfactor_dvdn 2 4) ?cardG_gt0 // -{2}p2. rewrite /= -/R => defM iUM modM _ _; pose N := 'C_G('Mho^1(U)). -have charZN2: 'Z('Ohm_2(N)) \char G. - rewrite (char_trans (center_char _)) // (char_trans (Ohm_char _ _)) //. - by rewrite subcent_char ?char_refl // (char_trans (Mho_char _ _)). +have charZN2: 'Z('Ohm_2(N)) \char G by rewrite !(gFchar_trans, subcent_char). have:= sympG _ charZN2 (center_abelian _). rewrite abelian_rank1_cyclic ?center_abelian // leqNgt; case/negP. have defN: E \* M = N. - rewrite defM (cprod_modl defG) // centsC (subset_trans (Mho_sub 1 _)) //. - by rewrite /= -/U -defRH subsetIr. + rewrite defM (cprod_modl defG) // centsC gFsub_trans //= -/U. + by rewrite -defRH subsetIr. case/modular_group_classP: modM => q q_pr [n n_gt23 isoM]. have{n_gt23} n_gt2 := leq_trans (leq_addl _ _) n_gt23. have n_gt1 := ltnW n_gt2; have n_gt0 := ltnW n_gt1. @@ -2276,14 +2270,14 @@ have{q_pr} defq: q = p; last rewrite {q}defq in genM modM isoM. by have [-> _ _ _] := genM; rewrite Euclid_dvdX // dvdn_prime2 //; case: eqP. have [oM Mx ox X'y] := genM; have [My _] := setDP X'y; have [oy _] := modM. have [sUM sMR]: U \subset M /\ M \subset R. - by rewrite defM subsetI sUR subsetIl centsC (subset_trans (Mho_sub _ _)). + by rewrite defM subsetI sUR subsetIl centsC gFsub_trans. have oU1: #|'Mho^1(U)| = (p ^ n.-2)%N. have oU: #|U| = (p ^ n.-1)%N. by rewrite -(divg_indexS sUM) iUM oM -subn1 expnB. case/cyclicP: cycU pU oU => u -> p_u ou. by rewrite (Mho_p_cycle 1 p_u) -orderE (orderXexp 1 ou) subn1. have sZU1: Z \subset 'Mho^1(U). - rewrite -(cardSg_cyclic cycU) ?Ohm_sub ?Mho_sub // oZ oU1. + rewrite -(cardSg_cyclic cycU) ?gFsub // oZ oU1. by rewrite -(subnKC n_gt2) expnS dvdn_mulr. case/modular_group_structure: genM => // _ [defZM _ oZM] _ _. have:= n_gt2; rewrite leq_eqVlt eq_sym !xpair_eqE andbC. @@ -2319,7 +2313,7 @@ have defN2: (E <*> X2) \x <[y]> = 'Ohm_2(N). rewrite /= (OhmE 2 pN) sub_gen /=; last 1 first. by rewrite subsetI -defN mulG_subl sub_LdivT expE. rewrite -cent_joinEl // -genM_join genS // -defN. - apply/subsetP=> ez; case/setIP; case/imset2P=> e z Ee Mz ->{ez}. + apply/subsetP=> _ /setIP[/imset2P[e z Ee Mz ->]]. rewrite inE expgMn; last by red; rewrite -(centsP cME). rewrite (exponentP expE) // mul1g => zp2; rewrite mem_mulg //=. by rewrite (OhmE 2 (pgroupS sMR pR)) mem_gen // !inE Mz. diff --git a/mathcomp/solvable/finmodule.v b/mathcomp/solvable/finmodule.v index dd84def..2b2aa5f 100644 --- a/mathcomp/solvable/finmodule.v +++ b/mathcomp/solvable/finmodule.v @@ -1,14 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype bigop prime finset. -From mathcomp.fingroup -Require Import fingroup morphism perm action gproduct. -From mathcomp.algebra -Require Import ssralg finalg cyclic. -Require Import commutator. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path div choice. +From mathcomp +Require Import fintype bigop ssralg finset fingroup morphism perm. +From mathcomp +Require Import finalg action gproduct commutator cyclic. (******************************************************************************) (* This file regroups constructions and results that are based on the most *) @@ -17,7 +14,7 @@ Require Import commutator. (* splitting and transitivity theorem, from which we will later derive the *) (* Schur-Zassenhaus theorem and the elementary abelian special case of *) (* Maschke's theorem, the coprime abelian centraliser/commutator trivial *) -(* intersection theorem, from which we will derive that p-groups under coprime*) +(* intersection theorem, which is used to show that p-groups under coprime *) (* action factor into special groups, and the construction of the transfer *) (* homomorphism and its expansion relative to a cycle, from which we derive *) (* the Higman Focal Subgroup and the Burnside Normal Complement theorems. *) @@ -26,7 +23,7 @@ Require Import commutator. (* needed much outside this file, which contains all the results that exploit *) (* this construction. *) (* FiniteModule defines the Z[N(A)]-module associated with a finite abelian *) -(* abelian group A, given a proof abelA : abelian A) : *) +(* abelian group A, given a proof (abelA : abelian A) : *) (* fmod_of abelA == the type of elements of the module (similar to but *) (* distinct from [subg A]). *) (* fmod abelA x == the injection of x into fmod_of abelA if x \in A, else 0 *) @@ -95,16 +92,16 @@ Definition fmod_opp u := sub2f u^-1. Definition fmod_add u v := sub2f (u * v). Fact fmod_add0r : left_id (sub2f 1) fmod_add. -Proof. move=> u; apply: val_inj; exact: mul1g. Qed. +Proof. by move=> u; apply: val_inj; apply: mul1g. Qed. Fact fmod_addrA : associative fmod_add. -Proof. move=> u v w; apply: val_inj; exact: mulgA. Qed. +Proof. by move=> u v w; apply: val_inj; apply: mulgA. Qed. Fact fmod_addNr : left_inverse (sub2f 1) fmod_opp fmod_add. -Proof. move=> u; apply: val_inj; exact: mulVg. Qed. +Proof. by move=> u; apply: val_inj; apply: mulVg. Qed. Fact fmod_addrC : commutative fmod_add. -Proof. case=> x Ax [y Ay]; apply: val_inj; exact: (centsP abelA). Qed. +Proof. by case=> x Ax [y Ay]; apply: val_inj; apply: (centsP abelA). Qed. Definition fmod_zmodMixin := ZmodMixin fmod_addrA fmod_addrC fmod_add0r fmod_addNr. @@ -149,7 +146,7 @@ Qed. Lemma injm_fmod : 'injm fmod. Proof. -apply/injmP=> x y Ax Ay []; move/val_inj; exact: (injmP (injm_subg A)). +by apply/injmP=> x y Ax Ay []; move/val_inj; apply: (injmP (injm_subg A)). Qed. Notation "u ^@ x" := (actr u x) : ring_scope. @@ -275,9 +272,9 @@ have PpP x: pP x \in P by rewrite -mem_rcoset rcoset_repr rcoset_refl. have rPmul x y: x \in P -> rP (x * y) = rP y. by move=> Px; rewrite /rP rcosetM rcoset_id. pose pQ x := remgr H Q x; pose rH x := pQ (pP x) * rP x. -have pQhq: {in H & Q, forall h q, pQ (h * q) = q} by exact: remgrMid. +have pQhq: {in H & Q, forall h q, pQ (h * q) = q} by apply: remgrMid. have pQmul: {in P &, {morph pQ : x y / x * y}}. - apply: remgrM; [exact/complP | exact: normalS (nsHG)]. + by apply: remgrM; [apply/complP | apply: normalS (nsHG)]. have HrH x: rH x \in H :* x. by rewrite rcoset_sym mem_rcoset invMg mulgA mem_divgr // eqHQ PpP. have GrH x: x \in G -> rH x \in G. @@ -292,7 +289,7 @@ pose nu y := (\sum_(Px in rcosets P G) mu (repr Px) y)%R. have rHmul: {in G &, forall x y, rH (x * y) = rH x * rH y * val (mu x y)}. move=> x y Gx Gy; rewrite /= fmodK ?mulKVg // -mem_lcoset lcoset_sym. rewrite -norm_rlcoset; last by rewrite nHG ?GrH ?groupM. - by rewrite (rcoset_transl (HrH _)) -rcoset_mul ?nHG ?GrH // mem_mulg. + by rewrite (rcoset_eqP (HrH _)) -rcoset_mul ?nHG ?GrH // mem_mulg. have actrH a x: x \in G -> (a ^@ rH x = a ^@ x)%R. move=> Gx; apply: val_inj; rewrite /= !fmvalJ ?nHG ?GrH //. case/rcosetP: (HrH x) => b /(fmodK abelH) <- ->; rewrite conjgM. @@ -338,7 +335,7 @@ exists (Morphism fM @* G)%G; apply/complP; split. apply/trivgP/subsetP=> x /setIP[Hx /morphimP[y _ Gy eq_x]]. apply/set1P; move: Hx; rewrite {x}eq_x /= groupMr ?subgP //. rewrite -{1}(mulgKV y (rH y)) groupMl -?mem_rcoset // => Hy. - by rewrite -(mulg1 y) /f nu_Hmul // rH_Hmul //; exact: (morph1 (Morphism fM)). + by rewrite -(mulg1 y) /f nu_Hmul // rH_Hmul //; apply: (morph1 (Morphism fM)). apply/setP=> x; apply/mulsgP/idP=> [[h y Hh fy ->{x}] | Gx]. rewrite groupMl; last exact: (subsetP sHG). case/morphimP: fy => z _ Gz ->{h Hh y}. @@ -519,7 +516,7 @@ have HGgHzg: Hzg \in HG :* <[g]>. by rewrite mem_mulg ?set11 // -rcosetE mem_imset. have Hzg_x: x \in Hzg by rewrite (repr_mem_pblock trX). exists x; first by rewrite (repr_mem_transversal trX). -case/mulsgP: Hzg_x => y u /rcoset_transl <- /(orbit_act 'Rs) <- -> /=. +case/mulsgP: Hzg_x => y u /rcoset_eqP <- /(orbit_act 'Rs) <- -> /=. by rewrite rcosetE -rcosetM. Qed. @@ -543,7 +540,7 @@ pose pcyc x := pcycle (actperm 'Rs g) (H :* x). pose traj x := traject (actperm 'Rs g) (H :* x) #|pcyc x|. have Hgr_eq x: H_g_rcosets x = pcyc x. by rewrite /H_g_rcosets -orbitRs -pcycle_actperm ?inE. -have pcyc_eq x: pcyc x =i traj x by exact: pcycle_traject. +have pcyc_eq x: pcyc x =i traj x by apply: pcycle_traject. have uniq_traj x: uniq (traj x) by apply: uniq_traject_pcycle. have n_eq x: n_ x = #|pcyc x| by rewrite -Hgr_eq. have size_traj x: size (traj x) = n_ x by rewrite n_eq size_traject. @@ -572,7 +569,7 @@ have trY: is_transversal Y HG G. by rewrite -[x](mulgK (g ^+ i)) mem_mulg ?rcoset_refl // groupV mem_cycle. apply/set1P; rewrite /y eq_xx'; congr (_ * _ ^+ _) => //; apply/eqP. rewrite -(@nth_uniq _ (H :* x) (traj x)) ?size_traj // ?eq_xx' //. - by rewrite !nth_traj ?(rcoset_transl Hy_x'gj) // -eq_xx'. + by rewrite !nth_traj ?(rcoset_eqP Hy_x'gj) // -eq_xx'. have rYE x i : x \in X -> i < n_ x -> rY (H :* x :* g ^+ i) = x * g ^+ i. move=> Xx lt_i_x; rewrite -rcosetM; apply: (canLR_in (pblockK trY 1)). by apply/bigcupP; exists x => //; apply/imsetP; exists (Ordinal lt_i_x). diff --git a/mathcomp/solvable/frobenius.v b/mathcomp/solvable/frobenius.v index 64fe7e6..cc589a0 100644 --- a/mathcomp/solvable/frobenius.v +++ b/mathcomp/solvable/frobenius.v @@ -1,14 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat. -From mathcomp.discrete -Require Import div fintype bigop prime finset. -From mathcomp.fingroup -Require Import fingroup morphism perm action quotient gproduct. -From mathcomp.algebra -Require Import cyclic. -Require Import center pgroup nilpotent sylow hall abelian. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat div fintype bigop prime. +From mathcomp +Require Import finset fingroup morphism perm action quotient gproduct. +From mathcomp +Require Import cyclic center pgroup nilpotent sylow hall abelian. (******************************************************************************) (* Definition of Frobenius groups, some basic results, and the Frobenius *) @@ -19,8 +16,7 @@ Require Import center pgroup nilpotent sylow hall abelian. (* condition. *) (* semiprime K H <-> *) (* the internal action of H on K is "prime", i.e., an element of K that *) -(* centralises a nontrivial element of H must actually centralise all *) -(* of H. *) +(* centralises a nontrivial element of H must centralise all of H. *) (* normedTI A G L <=> *) (* A is nonempty, strictly disjoint from its conjugates in G, and has *) (* normaliser L in G. *) @@ -36,8 +32,8 @@ Require Import center pgroup nilpotent sylow hall abelian. (* G is (isomorphic to) a Frobenius group with complement H; same as *) (* above, but without the semi-direct product. The proof that this form *) (* is equivalent to the above (i.e., the existence of Frobenius *) -(* kernels) requires chareacter theory and will only be proved in the *) -(* vcharacter module. *) +(* kernels) requires character theory and will only be proved in the *) +(* vcharacter.v file. *) (* [Frobenius G] <=> G is a Frobenius group. *) (* Frobenius_action G H S to <-> *) (* The action to of G on S defines an isomorphism of G with a *) @@ -306,7 +302,7 @@ Lemma Frobenius_actionP G H : Proof. apply: (iffP andP) => [[neqHG] | [sT S to [ffulG transG regG ntH [u Su defH]]]]. case/normedTI_P=> nzH /subsetIP[sHG _] tiHG. - suffices: Frobenius_action G H (rcosets H G) 'Rs by exact: HasFrobeniusAction. + suffices: Frobenius_action G H (rcosets H G) 'Rs by apply: HasFrobeniusAction. pose Hfix x := 'Fix_(rcosets H G | 'Rs)[x]. have regG: {in G^#, forall x, #|Hfix x| <= 1}. move=> x /setD1P[ntx Gx]. @@ -314,12 +310,12 @@ apply: (iffP andP) => [[neqHG] | [sT S to [ffulG transG regG ntH [u Su defH]]]]. rewrite -(cards1 Hy) => /setIP[/imsetP[y Gy ->{Hy}] cHyx]. apply/subset_leq_card/subsetP=> _ /setIP[/imsetP[z Gz ->] cHzx]. rewrite -!sub_astab1 !astab1_act !sub1set astab1Rs in cHyx cHzx *. - rewrite !rcosetE; apply/set1P/rcoset_transl; rewrite mem_rcoset. + rewrite !rcosetE; apply/set1P/rcoset_eqP; rewrite mem_rcoset. apply: tiHG; [by rewrite !in_group | apply/pred0Pn; exists (x ^ y^-1)]. by rewrite conjD1g !inE conjg_eq1 ntx -mem_conjg cHyx conjsgM memJ_conjg. have ntH: H :!=: 1 by rewrite -subG1 -setD_eq0. split=> //; first 1 last; first exact: transRs_rcosets. - by exists (H : {set gT}); rewrite ?orbit_refl // astab1Rs (setIidPr sHG). + by exists (val H); rewrite ?orbit_refl // astab1Rs (setIidPr sHG). apply/subsetP=> y /setIP[Gy cHy]; apply: contraR neqHG => nt_y. rewrite (index1g sHG) //; apply/eqP; rewrite eqn_leq indexg_gt0 andbT. apply: leq_trans (regG y _); last by rewrite setDE 2!inE Gy nt_y /=. @@ -351,7 +347,7 @@ Lemma FrobeniusWcompl : [Frobenius G with complement H]. Proof. by case/andP: frobG. Qed. Lemma FrobeniusW : [Frobenius G]. -Proof. by apply/existsP; exists H; exact: FrobeniusWcompl. Qed. +Proof. by apply/existsP; exists H; apply: FrobeniusWcompl. Qed. Lemma Frobenius_context : [/\ K ><| H = G, K :!=: 1, H :!=: 1, K \proper G & H \proper G]. @@ -408,7 +404,7 @@ by rewrite inE (subsetP (Frobenius_cent1_ker K1y)) // inE cent1C (subsetP sHG). Qed. Lemma Frobenius_reg_compl : semiregular H K. -Proof. by apply: semiregular_sym; exact: Frobenius_reg_ker. Qed. +Proof. by apply: semiregular_sym; apply: Frobenius_reg_ker. Qed. Lemma Frobenius_dvd_ker1 : #|H| %| #|K|.-1. Proof. @@ -490,11 +486,11 @@ Qed. Lemma Frobenius_ker_dvd_ker1 G K : [Frobenius G with kernel K] -> #|G : K| %| #|K|.-1. -Proof. case/existsP=> H; exact: Frobenius_index_dvd_ker1. Qed. +Proof. by case/existsP=> H; apply: Frobenius_index_dvd_ker1. Qed. Lemma Frobenius_ker_coprime G K : [Frobenius G with kernel K] -> coprime #|K| #|G : K|. -Proof. case/existsP=> H; exact: Frobenius_index_coprime. Qed. +Proof. by case/existsP=> H; apply: Frobenius_index_coprime. Qed. Lemma Frobenius_semiregularP G K H : K ><| H = G -> K :!=: 1 -> H :!=: 1 -> @@ -547,7 +543,7 @@ move=> ntH1 sH1H frobG; have [defG ntK _ _ _] := Frobenius_context frobG. apply/Frobenius_semiregularP=> //. have [_ _ /(subset_trans sH1H) nH1K tiHK] := sdprodP defG. by rewrite sdprodEY //; apply/trivgP; rewrite -tiHK setIS. -by apply: sub_in1 (Frobenius_reg_ker frobG); exact/subsetP/setSD. +by apply: sub_in1 (Frobenius_reg_ker frobG); apply/subsetP/setSD. Qed. Lemma Frobenius_kerP G K : @@ -557,7 +553,7 @@ Lemma Frobenius_kerP G K : Proof. apply: (iffP existsP) => [[H frobG] | [ntK ltKG nsKG regK]]. have [/sdprod_context[nsKG _ _ _ _] ntK _ ltKG _] := Frobenius_context frobG. - by split=> //; exact: Frobenius_cent1_ker frobG. + by split=> //; apply: Frobenius_cent1_ker frobG. have /andP[sKG nKG] := nsKG. have hallK: Hall G K. rewrite /Hall sKG //= coprime_sym coprime_pi' //. @@ -568,7 +564,7 @@ have hallK: Hall G K. have /trivgPn[z]: P :&: K :&: 'Z(P) != 1. by rewrite meet_center_nil ?(pgroup_nil pP) ?(normalGI sPG nsKG). rewrite !inE -andbA -sub_cent1=> /and4P[_ Kz _ cPz] ntz. - by apply: subset_trans (regK z _); [exact/subsetIP | exact/setD1P]. + by apply: subset_trans (regK z _); [apply/subsetIP | apply/setD1P]. have /splitsP[H /complP[tiKH defG]] := SchurZassenhaus_split hallK nsKG. have [_ sHG] := mulG_sub defG; have nKH := subset_trans sHG nKG. exists H; apply/Frobenius_semiregularP; rewrite ?sdprodE //. @@ -604,7 +600,7 @@ Proof. move=> defG FrobG. have partG: partition (gval K |: (H^# :^: K)) G. apply: Frobenius_partition; apply/andP; rewrite defG; split=> //. - by apply/Frobenius_actionP; exact: HasFrobeniusAction FrobG. + by apply/Frobenius_actionP; apply: HasFrobeniusAction FrobG. have{FrobG} [ffulG transG regG ntH [u Su defH]]:= FrobG. apply/setP=> x; rewrite !inE; have [-> | ntx] := altP eqP; first exact: group1. rewrite /= -(cover_partition partG) /cover. @@ -677,7 +673,8 @@ Qed. Lemma injm_Frobenius_ker K sGD injf : [Frobenius G with kernel K] -> [Frobenius f @* G with kernel f @* K]. Proof. -case/existsP=> H frobG; apply/existsP; exists (f @* H)%G; exact: injm_Frobenius. +case/existsP=> H frobG; apply/existsP. +by exists (f @* H)%G; apply: injm_Frobenius. Qed. Lemma injm_Frobenius_group sGD injf : [Frobenius G] -> [Frobenius f @* G]. @@ -698,7 +695,7 @@ rewrite ltnS oG mulnK // => leqm. have:= q_gt0; rewrite leq_eqVlt => /predU1P[q1 | lt1q]. rewrite -(mul1n n) q1 -oG (setIidPl _) //. by apply/subsetP=> x Gx; rewrite inE -order_dvdn order_dvdG. -pose p := pdiv q; have pr_p: prime p by exact: pdiv_prime. +pose p := pdiv q; have pr_p: prime p by apply: pdiv_prime. have lt1p: 1 < p := prime_gt1 pr_p; have p_gt0 := ltnW lt1p. have{leqm} lt_qp_mq: q %/ p < mq by apply: leq_trans leqm; rewrite ltn_Pdiv. have: n %| #|'Ldiv_(p * n)(G)|. diff --git a/mathcomp/solvable/gfunctor.v b/mathcomp/solvable/gfunctor.v index 25295ff..517010f 100644 --- a/mathcomp/solvable/gfunctor.v +++ b/mathcomp/solvable/gfunctor.v @@ -1,10 +1,8 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat. -From mathcomp.discrete -Require Import fintype bigop finset. -From mathcomp.fingroup +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat fintype bigop finset. +From mathcomp Require Import fingroup morphism automorphism quotient gproduct. (******************************************************************************) @@ -121,7 +119,7 @@ Definition iso_continuous := 'injm phi -> phi @* F G \subset F (phi @* G). Lemma continuous_is_iso_continuous : continuous -> iso_continuous. -Proof. by move=> Fcont gT hT G phi inj_phi; exact: Fcont. Qed. +Proof. by move=> Fcont gT hT G phi inj_phi; apply: Fcont. Qed. (* Functoriality on Grp with partial morphisms. *) Definition pcontinuous := @@ -129,7 +127,7 @@ Definition pcontinuous := phi @* F G \subset F (phi @* G). Lemma pcontinuous_is_continuous : pcontinuous -> continuous. -Proof. by move=> Fcont gT hT G; exact: Fcont. Qed. +Proof. by move=> Fcont gT hT G; apply: Fcont. Qed. (* Heredity with respect to inclusion *) Definition hereditary := @@ -269,6 +267,10 @@ Variable F : GFunctor.iso_map. Lemma gFsub gT (G : {group gT}) : F gT G \subset G. Proof. by case: F gT G. Qed. +Lemma gFsub_trans gT (G : {group gT}) (A : pred_class) : + G \subset A -> F gT G \subset A. +Proof. exact/subset_trans/gFsub. Qed. + Lemma gF1 gT : F gT 1 = 1. Proof. exact/trivgP/gFsub. Qed. Lemma gFiso_cont : GFunctor.iso_continuous F. @@ -282,32 +284,44 @@ by rewrite -morphimEsub ?gFsub ?gFiso_cont ?injm_autm. Qed. Lemma gFnorm gT (G : {group gT}) : G \subset 'N(F gT G). -Proof. by rewrite char_norm ?gFchar. Qed. +Proof. exact/char_norm/gFchar. Qed. + +Lemma gFnorms gT (G : {group gT}) : 'N(G) \subset 'N(F gT G). +Proof. exact/char_norms/gFchar. Qed. Lemma gFnormal gT (G : {group gT}) : F gT G <| G. -Proof. by rewrite char_normal ?gFchar. Qed. +Proof. exact/char_normal/gFchar. Qed. + +Lemma gFchar_trans gT (G H : {group gT}) : H \char G -> F gT H \char G. +Proof. exact/char_trans/gFchar. Qed. + +Lemma gFnormal_trans gT (G H : {group gT}) : H <| G -> F gT H <| G. +Proof. exact/char_normal_trans/gFchar. Qed. + +Lemma gFnorm_trans gT (A : pred_class) (G : {group gT}) : + A \subset 'N(G) -> A \subset 'N(F gT G). +Proof. by move/subset_trans/(_ (gFnorms G)). Qed. Lemma injmF_sub gT rT (G D : {group gT}) (f : {morphism D >-> rT}) : 'injm f -> G \subset D -> f @* (F gT G) \subset F rT (f @* G). Proof. -move=> injf sGD; apply/eqP; rewrite -(setIidPr (gFsub G)). -by rewrite-{3}(setIid G) -!(morphim_restrm sGD) gFiso_cont // injm_restrm. +move=> injf sGD; have:= gFiso_cont (injm_restrm sGD injf). +by rewrite im_restrm morphim_restrm (setIidPr _) ?gFsub. Qed. Lemma injmF gT rT (G D : {group gT}) (f : {morphism D >-> rT}) : 'injm f -> G \subset D -> f @* (F gT G) = F rT (f @* G). Proof. -move=> injf sGD; apply/eqP; rewrite eqEsubset injmF_sub //=. -rewrite -{2}(morphim_invm injf sGD) -[f @* F _ _](morphpre_invm injf). -have Fsubs := subset_trans (gFsub _). -by rewrite -sub_morphim_pre (injmF_sub, Fsubs) ?morphimS ?injm_invm. +move=> injf sGD; have [sfGD injf'] := (morphimS f sGD, injm_invm injf). +apply/esym/eqP; rewrite eqEsubset -(injmSK injf') ?gFsub_trans //. +by rewrite !(subset_trans (injmF_sub _ _)) ?morphim_invm // gFsub_trans. Qed. Lemma gFisom gT rT (G D : {group gT}) R (f : {morphism D >-> rT}) : G \subset D -> isom G (gval R) f -> isom (F gT G) (F rT R) f. Proof. -case/(restrmP f)=> g [gf _ _ _]; rewrite -{f}gf. -by case/isomP=> injg <-; rewrite sub_isom ?gFsub ?injmF. +case/(restrmP f)=> g [gf _ _ _]; rewrite -{f}gf => /isomP[injg <-]. +by rewrite sub_isom ?gFsub ?injmF. Qed. Lemma gFisog gT rT (G : {group gT}) (R : {group rT}) : @@ -440,7 +454,7 @@ Section Composition. Variables (F1 : GFunctor.mono_map) (F2 : GFunctor.map). Lemma gFcomp_closed : GFunctor.closed (F1 \o F2). -Proof. by move=> gT G; rewrite (subset_trans (gFsub _ _)) ?gFsub. Qed. +Proof. by move=> gT G; rewrite !gFsub_trans. Qed. Lemma gFcomp_cont : GFunctor.continuous (F1 \o F2). Proof. diff --git a/mathcomp/solvable/gseries.v b/mathcomp/solvable/gseries.v index 718f074..3878413 100644 --- a/mathcomp/solvable/gseries.v +++ b/mathcomp/solvable/gseries.v @@ -1,13 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype bigop finset. -From mathcomp.fingroup -Require Import fingroup morphism automorphism quotient action. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path fintype bigop. +From mathcomp +Require Import finset fingroup morphism automorphism quotient action. +From mathcomp Require Import commutator center. - (******************************************************************************) (* H <|<| G <=> H is subnormal in G, i.e., H <| ... <| G. *) (* invariant_factor A H G <=> A normalises both H and G, and H <| G. *) @@ -173,7 +171,7 @@ Lemma invariant_subnormal A G H : Proof. move=> nGA nHA /andP[]; move: #|G| => m. elim: m => [|m IHm] in G nGA * => sHG. - by rewrite eq_sym; exists [::]; last exact/eqP. + by rewrite eq_sym; exists [::]; last apply/eqP. rewrite iterSr; set K := <<_>>. have nKA: A \subset 'N(K) by rewrite norms_gen ?norms_class_support. have sHK: H \subset K by rewrite sub_gen ?sub_class_support. @@ -297,7 +295,7 @@ apply/maxgroupP/maxgroupP; rewrite morphpre_proper //= => [] [ltMG maxM]. split=> // H ltHG sMH. have dH: H \subset D := subset_trans (proper_sub ltHG) (subsetIl D _). have defH: f @*^-1 (f @* H) = H. - by apply: morphimGK dH; apply: subset_trans sMH; exact: ker_sub_pre. + by apply: morphimGK dH; apply: subset_trans sMH; apply: ker_sub_pre. rewrite -defH morphpre_proper ?morphimS // in ltHG. by rewrite -defH [f @* H]maxM // -(morphpreK dM) morphimS. Qed. @@ -391,12 +389,12 @@ Qed. Lemma maxnormal_proper A B C : maxnormal A B C -> A \proper B. Proof. -by case/maxsetP=> /and3P[gA pAB _] _; exact: (sub_proper_trans (subset_gen A)). +by case/maxsetP=> /and3P[gA pAB _] _; apply: (sub_proper_trans (subset_gen A)). Qed. Lemma maxnormal_sub A B C : maxnormal A B C -> A \subset B. Proof. -by move=> maxA; rewrite proper_sub //; exact: (maxnormal_proper maxA). +by move=> maxA; rewrite proper_sub //; apply: (maxnormal_proper maxA). Qed. Lemma ex_maxnormal_ntrivg G : G :!=: 1-> {N : {group gT} | maxnormal N G G}. @@ -416,7 +414,7 @@ wlog suffices: H K {maxH} maxK nsHG nsKG cHK ltHK_G / H \subset K. by move=> IH; rewrite eqEsubset !IH // -cHK. have{maxK} /maxgroupP[_ maxK] := maxK. apply/joing_idPr/maxK; rewrite ?joing_subr //= comm_joingE //. -by rewrite properEneq ltHK_G; exact: normalM. +by rewrite properEneq ltHK_G; apply: normalM. Qed. Lemma maxnormal_minnormal G L M : @@ -425,7 +423,7 @@ Lemma maxnormal_minnormal G L M : Proof. move=> nMG nGL /maxgroupP[/andP[/andP[sMG ltMG] nML] maxM]; apply/mingroupP. rewrite -subG1 quotient_sub1 ?ltMG ?quotient_norms //. -split=> // Hb /andP[ntHb nHbL]; have nsMG: M <| G by exact/andP. +split=> // Hb /andP[ntHb nHbL]; have nsMG: M <| G by apply/andP. case/inv_quotientS=> // H defHb sMH sHG; rewrite defHb; congr (_ / M). apply/eqP; rewrite eqEproper sHG /=; apply: contra ntHb => ltHG. have nsMH: M <| H := normalS sMH sHG nsMG. @@ -458,14 +456,14 @@ apply: (iffP mingroupP); rewrite normG andbT => [[ntG simG]]. split=> // N /andP[sNG nNG]. by case: (eqsVneq N 1) => [|ntN]; [left | right; apply: simG; rewrite ?ntN]. split=> // N /andP[ntN nNG] sNG. -by case: (simG N) ntN => // [|->]; [exact/andP | case/eqP]. +by case: (simG N) ntN => // [|->]; [apply/andP | case/eqP]. Qed. Lemma quotient_simple gT (G H : {group gT}) : H <| G -> simple (G / H) = maxnormal H G G. Proof. move=> nsHG; have nGH := normal_norm nsHG. -by apply/idP/idP; [exact: minnormal_maxnormal | exact: maxnormal_minnormal]. +by apply/idP/idP; [apply: minnormal_maxnormal | apply: maxnormal_minnormal]. Qed. Lemma isog_simple gT rT (G : {group gT}) (M : {group rT}) : @@ -496,7 +494,7 @@ Lemma chief_factor_minnormal G V U : chief_factor G V U -> minnormal (U / V) (G / V). Proof. case/andP=> maxV /andP[sUG nUG]; apply: maxnormal_minnormal => //. -by have /andP[_ nVG] := maxgroupp maxV; exact: subset_trans sUG nVG. +by have /andP[_ nVG] := maxgroupp maxV; apply: subset_trans sUG nVG. Qed. Lemma acts_irrQ G U V : @@ -527,7 +525,7 @@ have /andP[ltVU nVG] := maxgroupp maxV. have [||s ch_s defV] := IHm V; first exact: leq_trans (proper_card ltVU) _. by rewrite /normal (subset_trans (proper_sub ltVU) (normal_sub nsUG)). exists (rcons s U); last by rewrite last_rcons. -by rewrite rcons_path defV /= ch_s /chief_factor; exact/and3P. +by rewrite rcons_path defV /= ch_s /chief_factor; apply/and3P. Qed. End Chiefs. diff --git a/mathcomp/solvable/hall.v b/mathcomp/solvable/hall.v index 90ba407..5e8a41b 100644 --- a/mathcomp/solvable/hall.v +++ b/mathcomp/solvable/hall.v @@ -1,12 +1,12 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype finset bigop prime. -From mathcomp.fingroup -Require Import fingroup morphism automorphism quotient action gproduct. -Require Import commutator center pgroup finmodule nilpotent sylow. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div fintype finset. +From mathcomp +Require Import prime fingroup morphism automorphism quotient action gproduct. +From mathcomp +Require Import gfunctor commutator center pgroup finmodule nilpotent sylow. +From mathcomp Require Import abelian maximal. (*****************************************************************************) @@ -41,7 +41,7 @@ have [-> | [p pr_p pH]] := trivgVpdiv H. have [P sylP] := Sylow_exists p H. case nPG: (P <| G); last first. pose N := ('N_G(P))%G; have sNG: N \subset G by rewrite subsetIl. - have eqHN_G: H * N = G by exact: Frattini_arg sylP. + have eqHN_G: H * N = G by apply: Frattini_arg sylP. pose H' := (H :&: N)%G. have nsH'N: H' <| N. by rewrite /normal subsetIr normsI ?normG ?(subset_trans sNG). @@ -50,7 +50,7 @@ case nPG: (P <| G); last first. by rewrite -mul_cardG (mulnC #|H'|) divnMl // cardG_gt0. have hallH': Hall N H'. rewrite /Hall -divgS subsetIr //= -eq_iH. - by case/andP: hallH => _; apply: coprimeSg; exact: subsetIl. + by case/andP: hallH => _; apply: coprimeSg; apply: subsetIl. have: [splits N, over H']. apply: IHn hallH' nsH'N; apply: {n}leq_trans Gn. rewrite proper_card // properEneq sNG andbT; apply/eqP=> eqNG. @@ -61,12 +61,12 @@ case nPG: (P <| G); last first. by rewrite /= -(setIidPr sKN) setIA tiKN. by rewrite eqEsubset -eqHN_G mulgS // -eqH'K mulGS mulSg ?subsetIl. pose Z := 'Z(P); pose Gbar := G / Z; pose Hbar := H / Z. -have sZP: Z \subset P by exact: center_sub. -have sZH: Z \subset H by exact: subset_trans (pHall_sub sylP). -have sZG: Z \subset G by exact: subset_trans sHG. -have nZG: Z <| G by apply: char_normal_trans nPG; exact: center_char. -have nZH: Z <| H by exact: normalS nZG. -have nHGbar: Hbar <| Gbar by exact: morphim_normal. +have sZP: Z \subset P by apply: center_sub. +have sZH: Z \subset H by apply: subset_trans (pHall_sub sylP). +have sZG: Z \subset G by apply: subset_trans sHG. +have nZG: Z <| G by apply: gFnormal_trans nPG. +have nZH: Z <| H by apply: normalS nZG. +have nHGbar: Hbar <| Gbar by apply: morphim_normal. have hallHbar: Hall Gbar Hbar by apply: morphim_Hall (normal_norm _) _. have: [splits Gbar, over Hbar]. apply: IHn => //; apply: {n}leq_trans Gn; rewrite ltn_quotient //. @@ -76,7 +76,7 @@ have: [splits Gbar, over Hbar]. case/splitsP=> Kbar /complP[tiHKbar eqHKbar]. have: Kbar \subset Gbar by rewrite -eqHKbar mulG_subr. case/inv_quotientS=> //= ZK quoZK sZZK sZKG. -have nZZK: Z <| ZK by exact: normalS nZG. +have nZZK: Z <| ZK by apply: normalS nZG. have cardZK: #|ZK| = (#|Z| * #|G : H|)%N. rewrite -(Lagrange sZZK); congr (_ * _)%N. rewrite -card_quotient -?quoZK; last by case/andP: nZZK. @@ -86,7 +86,7 @@ have cardZK: #|ZK| = (#|Z| * #|G : H|)%N. have: [splits ZK, over Z]. rewrite (Gaschutz_split nZZK _ sZZK) ?center_abelian //; last first. rewrite -divgS // cardZK mulKn ?cardG_gt0 //. - by case/andP: hallH => _; exact: coprimeSg. + by case/andP: hallH => _; apply: coprimeSg. by apply/splitsP; exists 1%G; rewrite inE -subG1 subsetIr mulg1 eqxx. case/splitsP=> K /complP[tiZK eqZK]. have sKZK: K \subset ZK by rewrite -(mul1g K) -eqZK mulSg ?sub1G. @@ -111,8 +111,8 @@ rewrite ltnS => leHn solH nHK; have [-> | ] := eqsVneq H 1. by apply/eqP; rewrite conjsg1 eqEcard sK1K eqK1K /=. pose G := (H <*> K)%G. have defG: G :=: H * K by rewrite -normC // -norm_joinEl // joingC. -have sHG: H \subset G by exact: joing_subl. -have sKG: K \subset G by exact: joing_subr. +have sHG: H \subset G by apply: joing_subl. +have sKG: K \subset G by apply: joing_subr. have nsHG: H <| G by rewrite /(H <| G) sHG join_subG normG. case/(solvable_norm_abelem solH nsHG)=> M [sMH nsMG ntM] /and3P[_ abelM _]. have [sMG nMG] := andP nsMG; rewrite -defG => sK1G coHK oK1K. @@ -134,7 +134,7 @@ case/morphimP=> x nMx Hx ->{xb} eqK1Kx; pose K2 := (K :^ x)%G. have{eqK1Kx} eqK12: K1 / M = K2 / M by rewrite quotientJ. suff [y My ->]: exists2 y, y \in M & K1 :=: K2 :^ y. by exists (x * y); [rewrite groupMl // (subsetP sMH) | rewrite conjsgM]. -have nMK1: K1 \subset 'N(M) by exact: nMsG. +have nMK1: K1 \subset 'N(M) by apply: nMsG. have defMK: M * K1 = M <*> K1 by rewrite -normC // -norm_joinEl // joingC. have sMKM: M \subset M <*> K1 by rewrite joing_subl. have nMKM: M <| M <*> K1 by rewrite normalYl. @@ -243,7 +243,7 @@ have{transHb} transH (K : {group gT}): have: y \in coset M y by rewrite val_coset (subsetP nMK, rcoset_refl). have: coset M y \in (H :^ x) / M. rewrite /quotient morphimJ //=. - rewrite def_x def_H in sKHxb; apply: (subsetP sKHxb); exact: mem_quotient. + by rewrite def_x def_H in sKHxb; apply/(subsetP sKHxb)/mem_quotient. case/morphimP=> z Nz Hxz ->. rewrite val_coset //; case/rcosetP=> t Mt ->; rewrite groupMl //. by rewrite mem_conjg (subsetP sMH) // -mem_conjg (normP Nx). @@ -256,7 +256,7 @@ have [pi_p | pi'p] := boolP (p \in pi). have [ltHG | leGH {n IHn leGn transH}] := ltnP #|H| #|G|. case: (IHn _ H (leq_trans ltHG leGn)) => [|H1]; first exact: solvableS solG. case/and3P=> sH1H piH1 pi'H1' transH1. - have sH1G: H1 \subset G by exact: subset_trans sHG. + have sH1G: H1 \subset G by apply: subset_trans sHG. exists H1 => [|K sKG piK]. apply/and3P; split => //. rewrite -divgS // -(Lagrange sHG) -(Lagrange sH1H) -mulnA. @@ -283,7 +283,7 @@ have nMK: K \subset 'N(M) by apply: subset_trans sKG nMG. have defG1: M * K = G1 by rewrite -normC -?norm_joinEl. have sK1G1: K1 \subset M * K by rewrite defG1 subsetIr. have coMK: coprime #|M| #|K|. - by rewrite coprime_sym (pnat_coprime piK) //; exact: (pHall_pgroup hallM). + by rewrite coprime_sym (pnat_coprime piK) //; apply: (pHall_pgroup hallM). case: (SchurZassenhaus_trans_sol _ nMK sK1G1 coMK) => [||x Mx defK1]. - exact: solvableS solG. - apply/eqP; rewrite -(eqn_pmul2l (cardG_gt0 M)) -TI_cardMg //; last first. @@ -389,7 +389,7 @@ pose N := 'N_(A <*> G)(H)%G. have nGN: N \subset 'N(G) by rewrite subIset ?nG_AG. have nGN_N: G :&: N <| N by rewrite /(_ <| N) subsetIr normsI ?normG. have NG_AG: G * N = A <*> G. - by apply: Hall_Frattini_arg hallH => //; exact/andP. + by apply: Hall_Frattini_arg hallH => //; apply/andP. have iGN_A: #|N| %/ #|G :&: N| = #|A|. rewrite setIC divgI -card_quotient // -quotientMidl NG_AG. rewrite card_quotient -?divgS //= norm_joinEl //. @@ -426,14 +426,14 @@ have nGN_N: G :&: N <| N. apply/normalP; rewrite subsetIr; split=> // y Ny. by rewrite conjIg (normP _) // (subsetP nGN, conjGid). have NG_AG : G * N = A <*> G. - by apply: Hall_Frattini_arg hallH => //; exact/andP. + by apply: Hall_Frattini_arg hallH => //; apply/andP. have iGN_A: #|N : G :&: N| = #|A|. rewrite -card_quotient //; last by case/andP: nGN_N. rewrite (card_isog (second_isog nGN)) /= -quotientMidr (normC nGN) NG_AG. rewrite card_quotient // -divgS //= joingC norm_joinEr //. by rewrite coprime_cardMg // mulnC mulnK. -have solGN: solvable (G :&: N) by apply: solvableS solG; exact: subsetIl. -have oAxA: #|A :^ x^-1| = #|A| by exact: cardJg. +have solGN: solvable (G :&: N) by apply: solvableS solG; apply: subsetIl. +have oAxA: #|A :^ x^-1| = #|A| by apply: cardJg. have sAN: A \subset N by rewrite subsetI -{1}genGid genS // subsetUl. have nGNA: A \subset 'N(G :&: N). by apply/normsP=> y ?; rewrite conjIg (normsP nGA) ?(conjGid, subsetP sAN). @@ -471,7 +471,7 @@ apply/eqP; rewrite eqEsubset subsetI morphimS ?subsetIl //=. rewrite (subset_trans _ (morphim_cent _ _)) ?morphimS ?subsetIr //=. apply/subsetP=> _ /setIP[/morphimP[x Nx Gx ->] cAHx]. have{cAHx} cAxR y: y \in A -> [~ x, y] \in R. - move=> Ay; have Ny: y \in 'N(H) by exact: subsetP Ay. + move=> Ay; have Ny: y \in 'N(H) by apply: subsetP Ay. rewrite inE mem_commg // andbT coset_idr ?groupR // morphR //=. by apply/eqP; apply/commgP; apply: (centP cAHx); rewrite mem_quotient. have AxRA: A :^ x \subset R * A. @@ -566,11 +566,11 @@ have defG: G :=: K * H. rewrite -{1}(mulGid G) mulgSS //= (card_Hall hallH) (card_Hall hallK). by rewrite mulnC partnC. have sGA_H: [~: G, A] \subset H. - rewrite gen_subG defG; apply/subsetP=> xya; case/imset2P=> xy a. - case/imset2P=> x y Kx Hy -> Aa -> {xya xy}. + rewrite gen_subG defG. + apply/subsetP=> _ /imset2P[_ a /imset2P[x y Kx Hy ->] Aa ->]. rewrite commMgJ (([~ x, a] =P 1) _) ?(conj1g, mul1g). by rewrite groupMl ?groupV // memJ_norm ?(subsetP nHA). - rewrite subsetI sKG in cKA; apply/commgP; exact: (centsP cKA). + by rewrite subsetI sKG in cKA; apply/commgP/(centsP cKA). apply: pcore_max; last first. by rewrite /(_ <| G) /= commg_norml commGC commg_subr nGA. by case/and3P: hallH => _ piH _; apply: pgroupS piH. @@ -592,11 +592,11 @@ have [G1 | ntG] := eqsVneq G 1. have sG_AG: G \subset A <*> G by rewrite joing_subr. have sA_AG: A \subset A <*> G by rewrite joing_subl. have nG_AG: A <*> G \subset 'N(G) by rewrite join_subG nGA normG. -have nsG_AG: G <| A <*> G by exact/andP. +have nsG_AG: G <| A <*> G by apply/andP. case: (solvable_norm_abelem solG nsG_AG) => // M [sMG nsMAG ntM]. have{nsMAG} [nMA nMG]: A \subset 'N(M) /\ G \subset 'N(M). by apply/andP; rewrite -join_subG normal_norm. -have nMX: X \subset 'N(M) by exact: subset_trans nMG. +have nMX: X \subset 'N(M) by apply: subset_trans nMG. case/is_abelemP=> p pr_p; case/and3P=> pM cMM _. have: #|G / M| < n by rewrite (leq_trans (ltn_quotient _ _)). move/(IHn _ (A / M)%G _ (X / M)%G); rewrite !(quotient_norms, quotientS) //. @@ -609,7 +609,7 @@ have{pi'Hq' sHGq} pi'HM': pi^'.-nat #|G : HM|. move: pi'Hq'; rewrite -!divgS // defHM !card_quotient //. by rewrite -(divnMl (cardG_gt0 M)) !Lagrange. have{nHAq} nHMA: A \subset 'N(HM). - by rewrite -(quotientSGK nMA) ?normsG ?quotient_normG -?defHM //; exact/andP. + by rewrite -(quotientSGK nMA) ?normsG ?quotient_normG -?defHM //; apply/andP. case/orP: (orbN (p \in pi)) => pi_p. exists HM; split=> //; apply/and3P; split; rewrite /pgroup //. by rewrite -(Lagrange sMHM) pnat_mul -card_quotient // -defHM (pi_pnat pM). @@ -618,7 +618,7 @@ case: (ltnP #|HM| #|G|) => [ltHG | leGHM {n IHn leGn}]. - exact: coprimeSg coGA. - exact: solvableS solG. case/and3P: hallH => sHHM piH pi'H'. - have sHG: H \subset G by exact: subset_trans sHMG. + have sHG: H \subset G by apply: subset_trans sHMG. exists H; split=> //; apply/and3P; split=> //. rewrite -divgS // -(Lagrange sHMG) -(Lagrange sHHM) -mulnA mulKn //. by rewrite pnat_mul pi'H'. @@ -640,7 +640,7 @@ have hallX: pi.-Hall(XM) X. have sXMG: XM \subset G by rewrite join_subG sXG. have hallY: pi.-Hall(XM) Y. have sYXM: Y \subset XM by rewrite subsetIr. - have piY: pi.-group Y by apply: pgroupS piH; exact: subsetIl. + have piY: pi.-group Y by apply: pgroupS piH; apply: subsetIl. rewrite /pHall sYXM piY -divgS // -(_ : Y * M = XM). by rewrite coprime_cardMg ?co_pi_M // mulKn //. rewrite /= setIC group_modr ?joing_subr //=; apply/setIidPl. @@ -746,7 +746,7 @@ Proof. move=> sHG nHA coHA solH; pose N := 'N_G(H). have nsHN: H <| N by rewrite normal_subnorm. have [sHN nHn] := andP nsHN. -have sNG: N \subset G by exact: subsetIl. +have sNG: N \subset G by apply: subsetIl. have nNA: {acts A, on group N | to}. split; rewrite // actsEsd // injm_subnorm ?injm_sdpair1 //=. by rewrite normsI ?norms_norm ?im_sdpair_norm -?actsEsd. @@ -882,12 +882,12 @@ Lemma sol_coprime_Sylow_subset A G X : Proof. move=> nGA coGA solA sXG pX nXA. pose nAp (Q : {group gT}) := [&& p.-group Q, Q \subset G & A \subset 'N(Q)]. -have: nAp X by exact/and3P. +have: nAp X by apply/and3P. case/maxgroup_exists=> R; case/maxgroupP; case/and3P=> pR sRG nRA maxR sXR. have [P sylP sRP]:= Sylow_superset sRG pR. suffices defP: P :=: R by exists P; rewrite sylP defP. case/and3P: sylP => sPG pP _; apply: (nilpotent_sub_norm (pgroup_nil pP)) => //. -pose N := 'N_G(R); have{sPG} sPN_N: 'N_P(R) \subset N by exact: setSI. +pose N := 'N_G(R); have{sPG} sPN_N: 'N_P(R) \subset N by apply: setSI. apply: norm_sub_max_pgroup (pgroupS (subsetIl _ _) pP) sPN_N (subsetIr _ _). have nNA: A \subset 'N(N) by rewrite normsI ?norms_norm. have coNA: coprime #|N| #|A| by apply: coprimeSg coGA; rewrite subsetIl. diff --git a/mathcomp/solvable/jordanholder.v b/mathcomp/solvable/jordanholder.v index 82d9c8b..f315efa 100644 --- a/mathcomp/solvable/jordanholder.v +++ b/mathcomp/solvable/jordanholder.v @@ -1,11 +1,10 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path div choice fintype tuple finfun bigop finset. -From mathcomp.fingroup -Require Import fingroup morphism automorphism quotient action. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq path choice fintype. +From mathcomp +Require Import bigop finset fingroup morphism automorphism quotient action. +From mathcomp Require Import gseries. (******************************************************************************) @@ -113,7 +112,7 @@ Qed. Lemma section_repr_isog s1 s2 : s1 \isog s2 -> section_repr s1 = section_repr s2. Proof. -by move=> iso12; congr (odflt _ _); apply: eq_pick => s; exact: isog_transr. +by move=> iso12; congr (odflt _ _); apply: eq_pick => s; apply: isog_transr. Qed. Definition mkfactors (G : {group gT}) (s : seq {group gT}) := @@ -167,7 +166,7 @@ have [-> | ntG] := eqVneq G 1%G; first by exists [::]; rewrite /comps eqxx. have [N maxN] := ex_maxnormal_ntrivg ntG. have [|s /andP[ls cs]] := IHn N. by rewrite -ltnS (leq_trans _ cG) // proper_card // (maxnormal_proper maxN). -by exists (N :: s); exact/and3P. +by exists (N :: s); apply/and3P. Qed. (******************************************************************************) @@ -234,7 +233,7 @@ have i3 : perm_eq fG1 fG2. rewrite -(@section_repr_isog _ (mkSec _ _) (mkSec _ _) iso2). exact: perm_eq_refl. apply: (perm_eq_trans i1); apply: (perm_eq_trans i3); rewrite perm_eq_sym. -apply: perm_eq_trans i2; exact: perm_eq_refl. +by apply: perm_eq_trans i2; apply: perm_eq_refl. Qed. End CompositionSeries. @@ -361,7 +360,7 @@ by move/maxgroupp; case/andP; rewrite properE; move/normal_sub->; case/andP. Qed. Lemma maxainv_sub : maxainv K N -> N \subset K. -Proof. move=> h; apply: proper_sub; exact: maxainv_proper. Qed. +Proof. by move=> h; apply: proper_sub; apply: maxainv_proper. Qed. Lemma maxainv_ainvar : maxainv K N -> A \subset 'N(N | to). Proof. by move/maxgroupp; case/and3P. Qed. @@ -557,7 +556,7 @@ have nKQ1 : K <| N2 / N1. have sqA : qact_dom to N1 \subset A. by apply/subsetP=> t; rewrite qact_domE // inE; case/andP. have nNN2 : (N2 :&: N1) <| N2. - rewrite /normal subsetIl; apply: normsI => //; exact: normG. + by rewrite /normal subsetIl; apply: normsI => //; apply: normG. have aKQ1 : [acts qact_dom to N1, on K | to / N1]. pose H':= coset (N2 :&: N1)@*^-1 H. have eHH' : H :=: H' / (N2 :&: N1) by rewrite cosetpreK. @@ -653,7 +652,7 @@ have i1 : perm_eq (mksrepr G N1 :: mkfactors N1 st1) apply: Hi=> //; rewrite /acomps ?lst1 //= lsN csN andbT /=. apply: asimple_quo_maxainv=> //; first by apply: subIset; rewrite sN1D. apply: asimpleI => //. - apply: subset_trans (normal_norm nN2G); exact: normal_sub. + by apply: subset_trans (normal_norm nN2G); apply: normal_sub. rewrite -quotientMidl (maxainvM _ _ maxN_2) //. by apply: maxainv_asimple_quo. by move=> e; apply: neN12. @@ -664,7 +663,7 @@ have i2 : perm_eq (mksrepr G N2 :: mkfactors N2 st2) apply: asimple_quo_maxainv=> //; first by apply: subIset; rewrite sN1D. have e : N1 :&: N2 :=: N2 :&: N1 by rewrite setIC. rewrite (group_inj (setIC N1 N2)); apply: asimpleI => //. - apply: subset_trans (normal_norm nN1G); exact: normal_sub. + by apply: subset_trans (normal_norm nN1G); apply: normal_sub. rewrite -quotientMidl (maxainvM _ _ maxN_1) //. exact: maxainv_asimple_quo. pose fG1 := [:: mksrepr G N1, mksrepr N1 N & mkfactors N sN]. @@ -675,7 +674,7 @@ have i3 : perm_eq fG1 fG2. rewrite -(@section_repr_isog _ (mkSec _ _) (mkSec _ _) iso2). exact: perm_eq_refl. apply: (perm_eq_trans i1); apply: (perm_eq_trans i3); rewrite perm_eq_sym. -apply: perm_eq_trans i2; exact: perm_eq_refl. +by apply: perm_eq_trans i2; apply: perm_eq_refl. Qed. End StrongJordanHolder. diff --git a/mathcomp/solvable/maximal.v b/mathcomp/solvable/maximal.v index a3cd11c..afca270 100644 --- a/mathcomp/solvable/maximal.v +++ b/mathcomp/solvable/maximal.v @@ -1,14 +1,14 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div choice fintype finfun bigop finset prime binomial. -From mathcomp.fingroup -Require Import fingroup morphism perm automorphism quotient action gproduct. -From mathcomp.algebra -Require Import ssralg finalg zmodp cyclic. -Require Import pgroup center gseries commutator gfunctor. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq choice div fintype. +From mathcomp +Require Import finfun bigop finset prime binomial fingroup morphism perm. +From mathcomp +Require Import automorphism quotient action commutator gproduct gfunctor. +From mathcomp +Require Import ssralg finalg zmodp cyclic pgroup center gseries. +From mathcomp Require Import nilpotent sylow abelian finmodule. (******************************************************************************) @@ -65,12 +65,12 @@ Definition Fitting A := \big[dprod/1]_(p <- primes #|A|) 'O_p(A). Lemma Fitting_group_set G : group_set (Fitting G). Proof. -suffices [F ->]: exists F : {group gT}, Fitting G = F by exact: groupP. +suffices [F ->]: exists F : {group gT}, Fitting G = F by apply: groupP. rewrite /Fitting; elim: primes (primes_uniq #|G|) => [_|p r IHr] /=. by exists [1 gT]%G; rewrite big_nil. case/andP=> rp /IHr[F defF]; rewrite big_cons defF. suffices{IHr} /and3P[p'F sFG nFG]: p^'.-group F && (F <| G). - have nFGp: 'O_p(G) \subset 'N(F) := subset_trans (pcore_sub p G) nFG. + have nFGp: 'O_p(G) \subset 'N(F) := gFsub_trans _ nFG. have pGp: p.-group('O_p(G)) := pcore_pgroup p G. have{pGp} tiGpF: 'O_p(G) :&: F = 1 by rewrite coprime_TIg ?(pnat_coprime pGp). exists ('O_p(G) <*> F)%G; rewrite dprodEY // (sameP commG1P trivgP) -tiGpF. @@ -78,8 +78,8 @@ suffices{IHr} /and3P[p'F sFG nFG]: p^'.-group F && (F <| G). move/bigdprodWY: defF => <- {F}; elim: r rp => [_|q r IHr] /=. by rewrite big_nil gen0 pgroup1 normal1. rewrite inE eq_sym big_cons -joingE -joing_idr => /norP[qp /IHr {IHr}]. -set F := <<_>> => /andP[p'F nsFG]; rewrite norm_joinEl /= -/F; last first. - exact: subset_trans (pcore_sub q G) (normal_norm nsFG). +set F := <<_>> => /andP[p'F nsFG]. +rewrite norm_joinEl /= -/F; last exact/gFsub_trans/normal_norm. by rewrite pgroupM p'F normalM ?pcore_normal //= (pi_pgroup (pcore_pgroup q G)). Qed. @@ -236,7 +236,7 @@ Lemma Phi_quotient_id G : 'Phi (G / 'Phi(G)) = 1. Proof. apply/trivgP; rewrite -cosetpreSK cosetpre1 /=; apply/bigcapsP=> M maxM. have nPhi := Phi_normal G; have nPhiM: 'Phi(G) <| M. - by apply: normalS nPhi; [exact: bigcap_inf | case/maximal_eqP: maxM]. + by apply: normalS nPhi; [apply: bigcap_inf | case/maximal_eqP: maxM]. by rewrite sub_cosetpre_quo ?bigcap_inf // quotient_maximal_eq. Qed. @@ -265,12 +265,12 @@ apply/eqP/idP=> [trPhi | abP]. apply/set1gP; rewrite -trPhi; apply/bigcapP=> M. case/predU1P=> [-> | maxM]; first exact: groupX. have /andP[_ nMP] := p_maximal_normal pP maxM. - have nMx : x \in 'N(M) by exact: subsetP Px. + have nMx : x \in 'N(M) by apply: subsetP Px. apply: coset_idr; rewrite ?groupX ?morphX //=; apply/eqP. rewrite -(p_maximal_index pP maxM) -card_quotient // -order_dvdn cardSg //=. by rewrite cycle_subG mem_quotient. apply/trivgP/subsetP=> x Phi_x; rewrite -cycle_subG. -have Px: x \in P by exact: (subsetP (Phi_sub P)). +have Px: x \in P by apply: (subsetP (Phi_sub P)). have sxP: <[x]> \subset P by rewrite cycle_subG. case/splitsP: (abelem_splits abP sxP) => K /complP[tiKx defP]. have [-> | nt_x] := eqVneq x 1; first by rewrite cycle1. @@ -301,16 +301,16 @@ have [p_pr _ _] := pgroup_pdiv pP ntP. have [abP x1P] := abelemP p_pr Phi_quotient_abelem. apply/andP; split. have nMP: P \subset 'N(P^`(1) <*> 'Mho^1(P)) by rewrite normsY // !gFnorm. - rewrite -quotient_sub1 ?(subset_trans sPhiP) //=. - suffices <-: 'Phi(P / (P^`(1) <*> 'Mho^1(P))) = 1 by exact: morphimF. + rewrite -quotient_sub1 ?gFsub_trans //=. + suffices <-: 'Phi(P / (P^`(1) <*> 'Mho^1(P))) = 1 by apply: morphimF. apply/eqP; rewrite (trivg_Phi (morphim_pgroup _ pP)) /= -quotientE. apply/abelemP=> //; rewrite [abelian _]quotient_cents2 ?joing_subl //. split=> // _ /morphimP[x Nx Px ->] /=. rewrite -morphX //= coset_id // (MhoE 1 pP) joing_idr expn1. - by rewrite mem_gen //; apply/setUP; right; exact: mem_imset. + by rewrite mem_gen //; apply/setUP; right; apply: mem_imset. rewrite -quotient_cents2 // [_ \subset 'C(_)]abP (MhoE 1 pP) gen_subG /=. apply/subsetP=> _ /imsetP[x Px ->]; rewrite expn1. -have nPhi_x: x \in 'N('Phi(P)) by exact: (subsetP nPhiP). +have nPhi_x: x \in 'N('Phi(P)) by apply: (subsetP nPhiP). by rewrite coset_idr ?groupX ?morphX ?x1P ?mem_morphim. Qed. @@ -334,8 +334,8 @@ Lemma morphim_Phi rT P D (f : {morphism D >-> rT}) : p.-group P -> P \subset D -> f @* 'Phi(P) = 'Phi(f @* P). Proof. move=> pP sPD; rewrite !(@Phi_joing _ p) ?morphim_pgroup //. -rewrite morphim_gen ?(subset_trans _ sPD) ?subUset ?der_subS ?Mho_sub //. -by rewrite morphimU -joingE morphimR ?morphim_Mho. +rewrite morphim_gen ?subUset ?gFsub_trans // morphimU -joingE. +by rewrite morphimR ?morphim_Mho. Qed. Lemma quotient_Phi P H : @@ -347,7 +347,7 @@ Lemma Phi_min G H : p.-group G -> G \subset 'N(H) -> p.-abelem (G / H) -> 'Phi(G) \subset H. Proof. move=> pG nHG; rewrite -trivg_Phi ?quotient_pgroup // -subG1 /=. -by rewrite -(quotient_Phi pG) ?quotient_sub1 // (subset_trans (Phi_sub _)). +by rewrite -(quotient_Phi pG) ?quotient_sub1 // gFsub_trans. Qed. Lemma Phi_cprod G H K : @@ -366,17 +366,17 @@ Lemma Phi_mulg H K : 'Phi(H * K) = 'Phi(H) * 'Phi(K). Proof. move=> pH pK cHK; have defHK := cprodEY cHK. -have [|_ -> _] := cprodP (Phi_cprod _ defHK); rewrite /= cent_joinEr //. -by apply: pnat_dvd (dvdn_cardMg H K) _; rewrite pnat_mul; exact/andP. +have [|_ ->] /= := cprodP (Phi_cprod _ defHK); rewrite cent_joinEr //. +by rewrite pgroupM pH. Qed. Lemma charsimpleP G : reflect (G :!=: 1 /\ forall K, K :!=: 1 -> K \char G -> K :=: G) (charsimple G). Proof. -apply: (iffP mingroupP); rewrite char_refl andbT => [[ntG simG]]. +apply: (iffP mingroupP); rewrite char_refl andbT => -[ntG simG]. by split=> // K ntK chK; apply: simG; rewrite ?ntK // char_sub. -split=> // K /andP[ntK chK] _; exact: simG. +by split=> // K /andP[ntK chK] _; apply: simG. Qed. End Frattini4. @@ -405,26 +405,22 @@ Qed. Lemma Fitting_max G H : H <| G -> nilpotent H -> H \subset 'F(G). Proof. move=> nsHG nilH; rewrite -(Sylow_gen H) gen_subG. -apply/bigcupsP=> P /SylowP[p _ SylP]. +apply/bigcupsP=> P /SylowP[p _ sylP]. case Gp: (p \in \pi(G)); last first. - rewrite card1_trivg ?sub1G // (card_Hall SylP). + rewrite card1_trivg ?sub1G // (card_Hall sylP). rewrite part_p'nat // (pnat_dvd (cardSg (normal_sub nsHG))) //. by rewrite /pnat cardG_gt0 all_predC has_pred1 Gp. -move/nilpotent_Hall_pcore: SylP => ->{P} //. +have{P sylP}[-> //] := nilpotent_Hall_pcore nilH sylP. rewrite -(bigdprodWY (erefl 'F(G))) sub_gen //. rewrite -(filter_pi_of (ltnSn _)) big_filter big_mkord. have le_pG: p < #|G|.+1. by rewrite ltnS dvdn_leq //; move: Gp; rewrite mem_primes => /and3P[]. apply: (bigcup_max (Ordinal le_pG)) => //=. -apply: pcore_max (pcore_pgroup _ _) _. -exact: char_normal_trans (pcore_char p H) nsHG. +by rewrite pcore_max ?pcore_pgroup ?gFnormal_trans. Qed. Lemma pcore_Fitting pi G : 'O_pi('F(G)) \subset 'O_pi(G). -Proof. -rewrite pcore_max ?pcore_pgroup //. -exact: char_normal_trans (pcore_char _ _) (Fitting_normal _). -Qed. +Proof. by rewrite pcore_max ?pcore_pgroup ?gFnormal_trans ?Fitting_normal. Qed. Lemma p_core_Fitting p G : 'O_p('F(G)) = 'O_p(G). Proof. @@ -473,7 +469,7 @@ Qed. Lemma FittingS gT (G H : {group gT}) : H \subset G -> H :&: 'F(G) \subset 'F(H). Proof. move=> sHG; rewrite -{2}(setIidPl sHG). -do 2!rewrite -(morphim_idm (subsetIl H _)) morphimIdom; exact: morphim_Fitting. +do 2!rewrite -(morphim_idm (subsetIl H _)) morphimIdom; apply: morphim_Fitting. Qed. Lemma FittingJ gT (G : {group gT}) x : 'F(G :^ x) = 'F(G) :^ x. @@ -655,7 +651,7 @@ Qed. Lemma charsimple_solvable G : charsimple G -> solvable G -> is_abelem G. Proof. case/charsimple_dprod=> H [sHG simH [I Aut_I defG]] solG. -have p_pr: prime #|H| by exact: simple_sol_prime (solvableS sHG solG) simH. +have p_pr: prime #|H| by apply: simple_sol_prime (solvableS sHG solG) simH. set p := #|H| in p_pr; apply/is_abelemP; exists p => //. elim/big_rec: _ (G) defG => [_ <-|f B If IH_B M defM]; first exact: abelem1. have [Af [[_ K _ defB] _ _ _]] := (subsetP Aut_I f If, dprodP defM). @@ -686,8 +682,7 @@ Qed. Lemma trivg_Fitting G : solvable G -> ('F(G) == 1) = (G :==: 1). Proof. -move=> solG; apply/idP/idP=> [F1|]; last first. - by rewrite -!subG1; apply: subset_trans; exact: Fitting_sub. +move=> solG; apply/idP/idP=> [F1 | /eqP->]; last by rewrite gF1. apply/idPn=> /(solvable_norm_abelem solG (normal_refl _))[M [_ nsMG ntM]]. case/is_abelemP=> p _ /and3P[pM _ _]; case/negP: ntM. by rewrite -subG1 -(eqP F1) Fitting_max ?(pgroup_nil pM). @@ -697,11 +692,9 @@ Lemma Fitting_pcore pi G : 'F('O_pi(G)) = 'O_pi('F(G)). Proof. apply/eqP; rewrite eqEsubset. rewrite (subset_trans _ (pcoreS _ (Fitting_sub _))); last first. - rewrite subsetI Fitting_sub Fitting_max ?Fitting_nil //. - by rewrite (char_normal_trans (Fitting_char _)) ?pcore_normal. + by rewrite subsetI Fitting_sub Fitting_max ?Fitting_nil ?gFnormal_trans. rewrite (subset_trans _ (FittingS (pcore_sub _ _))) // subsetI pcore_sub. -rewrite pcore_max ?pcore_pgroup //. -by rewrite (char_normal_trans (pcore_char _ _)) ?Fitting_normal. +by rewrite pcore_max ?pcore_pgroup ?gFnormal_trans. Qed. End CharSimple. @@ -722,7 +715,7 @@ Lemma sol_prime_factor_exists : solvable G -> G :!=: 1 -> {H : {group gT} | H <| G & prime #|G : H| }. Proof. move=> solG /ex_maxnormal_ntrivg[H maxH]. -by exists H; [exact: maxnormal_normal | exact: index_maxnormal_sol_prime]. +by exists H; [apply: maxnormal_normal | apply: index_maxnormal_sol_prime]. Qed. End SolvablePrimeFactor. @@ -738,7 +731,7 @@ move=> pG [defPhi defG']. have [-> | ntG] := eqsVneq G 1; first by rewrite center1 abelem1. have [p_pr _ _] := pgroup_pdiv pG ntG. have fM: {in 'Z(G) &, {morph expgn^~ p : x y / x * y}}. - by move=> x y /setIP[_ /centP cxG] /setIP[/cxG cxy _]; exact: expgMn. + by move=> x y /setIP[_ /centP cxG] /setIP[/cxG cxy _]; apply: expgMn. rewrite abelemE //= center_abelian; apply/exponentP=> /= z Zz. apply: (@kerP _ _ _ (Morphism fM)) => //; apply: subsetP z Zz. rewrite -{1}defG' gen_subG; apply/subsetP=> _ /imset2P[x y Gx Gy ->]. @@ -786,10 +779,10 @@ have{sZ2G'_Z} sG'Z: G^`(1) \subset 'Z(G). rewrite -quotient_der ?gFnorm // -quotientGI ?ucn_subS ?quotientS1 //=. by rewrite ucn1. have sCG': 'C_G(A) \subset G^`(1). - rewrite -quotient_sub1 //; last by rewrite subIset // char_norm ?der_char. + rewrite -quotient_sub1 //; last by rewrite subIset ?gFnorm. rewrite (subset_trans (quotient_subcent _ G A)) //= -[G in G / _]defG. have nGA: A \subset 'N(G) by rewrite -commg_subl defG. - rewrite quotientR ?(char_norm_trans (der_char _ _)) ?normG //. + rewrite quotientR ?gFnorm_trans ?normG //. rewrite coprime_abel_cent_TI ?quotient_norms ?coprime_morph //. exact: sub_der1_abelian. have defZ: 'Z(G) = G^`(1) by apply/eqP; rewrite eqEsubset (subset_trans cZA). @@ -853,7 +846,7 @@ have pEq: p.-group (E / 'Z(E))%g by rewrite quotient_pgroup. have /andP[sZE nZE] := center_normal E. have oEq: #|E / 'Z(E)|%g = (p ^ 2)%N. by rewrite card_quotient -?divgS // oEp3 oZp expnS mulKn. -have cEEq: abelian (E / 'Z(E))%g by exact: card_p2group_abelian oEq. +have cEEq: abelian (E / 'Z(E))%g by apply: card_p2group_abelian oEq. have not_cEE: ~~ abelian E. have: #|'Z(E)| < #|E| by rewrite oEp3 oZp (ltn_exp2l 1) ?prime_gt1. by apply: contraL => cEE; rewrite -leqNgt subset_leq_card // subsetI subxx. @@ -863,9 +856,9 @@ have defE': E^`(1) = 'Z(E). by rewrite setIC prime_TIg ?oZp. split; [split=> // | by rewrite oZp]; apply/eqP. rewrite eqEsubset andbC -{1}defE' {1}(Phi_joing pE) joing_subl. -rewrite -quotient_sub1 ?(subset_trans (Phi_sub _)) //. -rewrite subG1 /= (quotient_Phi pE) //= (trivg_Phi pEq); apply/abelemP=> //. -split=> // Zx EqZx; apply/eqP; rewrite -order_dvdn /order. +rewrite -quotient_sub1 ?gFsub_trans ?subG1 //=. +rewrite (quotient_Phi pE) //= (trivg_Phi pEq). +apply/abelemP=> //; split=> // Zx EqZx; apply/eqP; rewrite -order_dvdn /order. rewrite (card_pgroup (mem_p_elt pEq EqZx)) (@dvdn_exp2l _ _ 1) //. rewrite leqNgt -pfactor_dvdn // -oEq; apply: contra not_cEE => sEqZx. rewrite cyclic_center_factor_abelian //; apply/cyclicP. @@ -877,7 +870,7 @@ Lemma p3group_extraspecial G : p.-group G -> ~~ abelian G -> logn p #|G| <= 3 -> extraspecial G. Proof. move=> pG not_cGG; have /andP[sZG nZG] := center_normal G. -have ntG: G :!=: 1 by apply: contraNneq not_cGG => ->; exact: abelian1. +have ntG: G :!=: 1 by apply: contraNneq not_cGG => ->; apply: abelian1. have ntZ: 'Z(G) != 1 by rewrite (center_nil_eq1 (pgroup_nil pG)). have [p_pr _ [n oG]] := pgroup_pdiv pG ntG; rewrite oG pfactorK //. have [_ _ [m oZ]] := pgroup_pdiv (pgroupS sZG pG) ntZ. @@ -924,11 +917,11 @@ Qed. Lemma isog_special G (R : {group rT}) : G \isog R -> special G -> special R. -Proof. by case/isogP=> f injf <-; exact: injm_special. Qed. +Proof. by case/isogP=> f injf <-; apply: injm_special. Qed. Lemma isog_extraspecial G (R : {group rT}) : G \isog R -> extraspecial G -> extraspecial R. -Proof. by case/isogP=> f injf <-; exact: injm_extraspecial. Qed. +Proof. by case/isogP=> f injf <-; apply: injm_extraspecial. Qed. Lemma cprod_extraspecial G H K : p.-group G -> H \* K = G -> H :&: K = 'Z(H) -> @@ -972,10 +965,10 @@ have fM: {in G &, {morph f : y z / y * z}}%g. pose fm := Morphism fM. have fmG: fm @* G = 'Z(G). have sfmG: fm @* G \subset 'Z(G). - by apply/subsetP=> _ /morphimP[z _ Gz ->]; exact: fZ. + by apply/subsetP=> _ /morphimP[z _ Gz ->]; apply: fZ. apply/eqP; rewrite eqEsubset sfmG; apply: contraR notZx => /(prime_TIg prZ). rewrite (setIidPr _) // => fmG1; rewrite inE Gx; apply/centP=> y Gy. - by apply/commgP; rewrite -in_set1 -[[set _]]fmG1; exact: mem_morphim. + by apply/commgP; rewrite -in_set1 -[[set _]]fmG1; apply: mem_morphim. have ->: 'C_G[x] = 'ker fm. apply/setP=> z; rewrite inE (sameP cent1P commgP) !inE. by rewrite -invg_comm eq_invg_mul mulg1. @@ -995,7 +988,7 @@ apply/eqP; rewrite eqEsubset sCxH subsetI sHU /= andbT. apply: contraR not_sHU => not_sHCx. have maxCx: maximal 'C_G[x] G. rewrite cent1_extraspecial_maximal //; apply: contra not_cUx. - by rewrite inE Gx; exact: subsetP (centS sUG) _. + by rewrite inE Gx; apply: subsetP (centS sUG) _. have nsCx := p_maximal_normal pG maxCx. rewrite -(setIidPl sUG) -(mulg_normal_maximal nsCx maxCx sHG) ?subsetI ?sHG //. by rewrite -group_modr //= setIA (setIidPl sUG) mul_subG. @@ -1014,7 +1007,7 @@ have{not_cUG} [x Gx not_cUx] := subsetPn not_cUG. pose W := 'C_U[x]; have sCW_G: 'C_G(W) \subset G := subsetIl G _. have maxW: maximal W U by rewrite subcent1_extraspecial_maximal // inE not_cUx. have nsWU: W <| U := p_maximal_normal (pgroupS sUG pG) maxW. -have ltWU: W \proper U by exact: maxgroupp maxW. +have ltWU: W \proper U by apply: maxgroupp maxW. have [sWU [u Uu notWu]] := properP ltWU; have sWG := subset_trans sWU sUG. have defU: W * <[u]> = U by rewrite (mulg_normal_maximal nsWU) ?cycle_subG. have iCW_CU: #|'C_G(W) : 'C_G(U)| = p. @@ -1060,7 +1053,7 @@ have sZE: 'Z(G) \subset E. by rewrite /= -defG' inE !mem_commg. have ziER: E :&: R = 'Z(E) by rewrite setIA (setIidPl sEG). have cER: R \subset 'C(E) by rewrite subsetIr. -have iCxG: #|G : 'C_G[x]| = p by exact: p_maximal_index. +have iCxG: #|G : 'C_G[x]| = p by apply: p_maximal_index. have maxR: maximal R 'C_G[x]. rewrite /R centY !cent_cycle setIA. rewrite subcent1_extraspecial_maximal ?subsetIl // inE Gy andbT -sub_cent1. @@ -1171,7 +1164,7 @@ have defS: <<X>> = S. apply: Phi_nongen; apply/eqP; rewrite eqEsubset join_subG sPS sXS -joing_idr. rewrite -genM_join sub_gen // -quotientSK ?quotient_gen // -defSb genS //. apply/subsetP=> xb Xxb; apply/imsetP; rewrite (setIidPr nPX). - by exists (repr xb); rewrite /= ?coset_reprK //; exact: mem_imset. + by exists (repr xb); rewrite /= ?coset_reprK //; apply: mem_imset. pose f (a : {perm gT}) := [ffun x => if x \in X then x^-1 * a x else 1]. have injf: {in A &, injective f}. move=> _ _ /morphimP[y nSy Ry ->] /morphimP[z nSz Rz ->]. @@ -1235,7 +1228,7 @@ move/IHs=> {IHs}IHs; case/cprodP=> [[_ U _ defU]]; rewrite defU => defT cEU. rewrite -(mulnK #|T| (cardG_gt0 (E :&: U))) -defT -mul_cardG /=. have ->: E :&: U = 'Z(S). apply/eqP; rewrite eqEsubset subsetI -{1 2}defZE subsetIl setIS //=. - by case/cprodP: defU => [[V _ -> _]] <- _; exact: mulG_subr. + by case/cprodP: defU => [[V _ -> _]] <- _; apply: mulG_subr. rewrite (IHs U) // oEp3 oZ -expnD addSn expnS mulKn ?prime_gt0 //. by rewrite pfactorK //= uphalf_double. Qed. @@ -1251,7 +1244,7 @@ rewrite -andbA big_cons -cprodA; case/and3P; move/eqP=> oE; move/eqP=> defZE. move=> es_s; case/cprodP=> [[_ U _ defU]]; rewrite defU => defT cEU. have sUT: U \subset T by rewrite -defT mulG_subr. have sZU: 'Z(T) \subset U. - by case/cprodP: defU => [[V _ -> _] <- _]; exact: mulG_subr. + by case/cprodP: defU => [[V _ -> _] <- _]; apply: mulG_subr. have defZU: 'Z(E) = 'Z(U). apply/eqP; rewrite eqEsubset defZE subsetI sZU subIset ?centS ?orbT //=. by rewrite subsetI subIset ?sUT //= -defT centM setSI. @@ -1259,7 +1252,7 @@ apply: (Aut_cprod_full _ defZU); rewrite ?cprodE //; last first. by apply: IHs; rewrite -?defZU ?defZE. have oZE: #|'Z(E)| = p by rewrite defZE. have [p2 | odd_p] := even_prime p_pr. - suffices <-: restr_perm 'Z(E) @* Aut E = Aut 'Z(E) by exact: Aut_in_isog. + suffices <-: restr_perm 'Z(E) @* Aut E = Aut 'Z(E) by apply: Aut_in_isog. apply/eqP; rewrite eqEcard restr_perm_Aut ?center_sub //=. by rewrite card_Aut_cyclic ?prime_cyclic ?oZE // {1}p2 cardG_gt0. have pE: p.-group E by rewrite /pgroup oE pnat_exp pnat_id. @@ -1298,7 +1291,7 @@ have [y Ey not_cxy]: exists2 y, y \in E & y \notin 'C[x]. have notZy: y \notin 'Z(E). apply: contra not_cxy; rewrite inE Ey; apply: subsetP. by rewrite -cent_set1 centS ?sub1set. -pose K := 'C_E[y]; have maxK: maximal K E by exact: cent1_extraspecial_maximal. +pose K := 'C_E[y]; have maxK: maximal K E by apply: cent1_extraspecial_maximal. have nsKE: K <| E := p_maximal_normal pE maxK; have [sKE nKE] := andP nsKE. have oK: #|K| = (p ^ 2)%N. by rewrite -(divg_indexS sKE) oE (p_maximal_index pE) ?mulKn. @@ -1403,7 +1396,7 @@ have{defB} defB : B :=: A * <[x]>. rewrite -quotientK ?cycle_subG ?quotient_cycle // defB quotientGK //. exact: normalS (normal_sub nBG) nsAG. apply/setIidPl; rewrite ?defB -[_ :&: _]center_prod //=. -rewrite /center !(setIidPl _) //; exact: cycle_abelian. +rewrite /center !(setIidPl _) //; apply: cycle_abelian. Qed. (* The two other assertions of Aschbacher 23.15 state properties of the *) @@ -1472,19 +1465,19 @@ have{nil_classY pY sXW sZY sZCA} defW: W = <[x]> * Z. congr (_ * _); apply/eqP; rewrite eqEsubset {1}[Z](OhmE 1 pA). rewrite subsetI setIAC subIset //; first by rewrite sZCA -[Z]Ohm_id OhmS. rewrite sub_gen ?setIS //; apply/subsetP=> w Ww; rewrite inE. - by apply/eqP; apply: exponentP w Ww; exact: exponent_Ohm1_class2. + by apply/eqP; apply: exponentP w Ww; apply: exponent_Ohm1_class2. have{sXG sAG} sXAG: XA \subset G by rewrite join_subG sXG. have{sXAG} nilXA: nilpotent XA := nilpotentS sXAG (pgroup_nil pG). have sYXA: Y \subset XA by rewrite defY defXA mulgS ?subsetIl. rewrite -[Y](nilpotent_sub_norm nilXA) {nilXA sYXA}//= -/Y -/XA. -apply: subset_trans (setIS _ (char_norm_trans (Ohm_char 1 _) (subxx _))) _. +suffices: 'N_XA('Ohm_1(Y)) \subset Y by apply/subset_trans/setIS/gFnorms. rewrite {XA}defXA -group_modl ?normsG /= -/W ?{W}defW ?mulG_subl //. rewrite {Y}defY mulgS // subsetI subsetIl {CA C}sub_astabQ subIset ?nZA //= -/Z. rewrite (subset_trans (quotient_subnorm _ _ _)) //= quotientMidr /= -/Z. rewrite -quotient_sub1 ?subIset ?cent_norm ?orbT //. rewrite (subset_trans (quotientI _ _ _)) ?coprime_TIg //. -rewrite (@pnat_coprime p) // -/(pgroup p _) ?quotient_pgroup {pA}//=. -rewrite -(setIidPr (cent_sub _)) [pnat _ _]p'group_quotient_cent_prime //. +rewrite (@pnat_coprime p) // -/(p.-group _) ?quotient_pgroup {pA}//= -pgroupE. +rewrite -(setIidPr (cent_sub _)) p'group_quotient_cent_prime //. by rewrite (dvdn_trans (dvdn_quotient _ _)) ?order_dvdn. Qed. @@ -1497,10 +1490,8 @@ Proof. move=> p_odd pG; set X := 'Ohm_1('C_G(Z)). case/maxgroupP=> /andP[nsZG abelZ] maxZ. have [sZG nZG] := andP nsZG; have [_ cZZ expZp] := and3P abelZ. -have{nZG} nsXG: X <| G. - apply: (char_normal_trans (Ohm_char 1 'C_G(Z))). - by rewrite /normal subsetIl normsI ?normG ?norms_cent. -have cZX : X \subset 'C(Z) := subset_trans (Ohm_sub _ _) (subsetIr _ _). +have{nZG} nsXG: X <| G by rewrite gFnormal_trans ?norm_normalI ?norms_cent. +have cZX : X \subset 'C(Z) by apply/gFsub_trans/subsetIr. have{sZG expZp} sZX: Z \subset X. rewrite [X](OhmE 1 (pgroupS _ pG)) ?subsetIl ?sub_gen //. apply/subsetP=> x Zx; rewrite !inE ?(subsetP sZG) ?(subsetP cZZ) //=. @@ -1513,23 +1504,23 @@ suffices{sZX} expXp: (exponent X %| p). case/andP=> /sub_center_normal nsZdG. have{nsZdG} [D defD sZD nsDG] := inv_quotientN nsZG nsZdG; rewrite defD. have sDG := normal_sub nsDG; have nsZD := normalS sZD sDG nsZG. - rewrite quotientSGK ?quotient_sub1 ?normal_norm //= -/X => sDX; case/negP. + rewrite quotientSGK ?quotient_sub1 ?normal_norm //= -/X => sDX /negP[]. rewrite (maxZ D) // nsDG andbA (pgroupS sDG) ?(dvdn_trans (exponentS sDX)) //. have sZZD: Z \subset 'Z(D) by rewrite subsetI sZD centsC (subset_trans sDX). by rewrite (cyclic_factor_abelian sZZD) //= -defD cycle_cyclic. pose normal_abelian := [pred A : {group gT} | A <| G & abelian A]. -have{nsZG cZZ} normal_abelian_Z : normal_abelian Z by exact/andP. +have{nsZG cZZ} normal_abelian_Z : normal_abelian Z by apply/andP. have{normal_abelian_Z} [A maxA sZA] := maxgroup_exists normal_abelian_Z. have SCN_A : A \in 'SCN(G) by apply: max_SCN pG maxA. move/maxgroupp: maxA => /andP[nsAG cAA] {normal_abelian}. have pA := pgroupS (normal_sub nsAG) pG. have{abelZ maxZ nsAG cAA sZA} defA1: 'Ohm_1(A) = Z. - apply: maxZ; last by rewrite -(Ohm1_id abelZ) OhmS. - by rewrite Ohm1_abelem ?(char_normal_trans (Ohm_char _ _) nsAG). + have: Z \subset 'Ohm_1(A) by rewrite -(Ohm1_id abelZ) OhmS. + by apply: maxZ; rewrite Ohm1_abelem ?gFnormal_trans. have{SCN_A} sX'A: X^`(1) \subset A. have sX_CWA1 : X \subset 'C('Ohm_1(A)) :&: 'C_G(A / 'Ohm_1(A) | 'Q). - rewrite subsetI /X -defA1 (Ohm1_stab_Ohm1_SCN_series _ p_odd) //= andbT. - exact: subset_trans (Ohm_sub _ _) (subsetIr _ _). + rewrite subsetI /X -defA1 (Ohm1_stab_Ohm1_SCN_series _ p_odd) //=. + by rewrite gFsub_trans ?subsetIr. by apply: subset_trans (der1_stab_Ohm1_SCN_series SCN_A); rewrite commgSS. pose genXp := [pred U : {group gT} | 'Ohm_1(U) == U & ~~ (exponent U %| p)]. apply/idPn=> expXp'; have genXp_X: genXp [group of X] by rewrite /= Ohm_id eqxx. @@ -1582,12 +1573,11 @@ have [|K]:= @maxgroup_exists _ qcr 1 _. case/maxgroupP; rewrite {}/qcr subUset => /and3P[chK sPhiZ sRZ] maxK _. have sKG := char_sub chK; have nKG := char_normal chK. exists K; split=> //; apply/eqP; rewrite eqEsubset andbC setSI //=. -have chZ: 'Z(K) \char G by [exact: subcent_char]; have nZG := char_norm chZ. -have chC: 'C_G(K) \char G by exact: subcent_char (char_refl G) chK. +have chZ: 'Z(K) \char G by [apply: subcent_char]; have nZG := char_norm chZ. +have chC: 'C_G(K) \char G by apply: subcent_char chK. rewrite -quotient_sub1; last by rewrite subIset // char_norm. apply/trivgP; apply: (TI_center_nil (quotient_nil _ (pgroup_nil pG))). - rewrite quotient_normal // /normal subsetIl normsI ?normG ?norms_cent //. - exact: char_norm. + by rewrite quotient_normal ?norm_normalI ?norms_cent ?normal_norm. apply: TI_Ohm1; apply/trivgP; rewrite -trivg_quotient -sub_cosetpre_quo //. rewrite morphpreI quotientGK /=; last first. by apply: normalS (char_normal chZ); rewrite ?subsetIl ?setSI. @@ -1598,16 +1588,16 @@ rewrite subsetI centsC cXK andbT -(mul1g K) -mulSG mul1g -(cent_joinEr cXK). rewrite [_ <*> K]maxK ?joing_subr //= andbC (cent_joinEr cXK). rewrite -center_prod // (subset_trans _ (mulG_subr _ _)). rewrite charM 1?charI ?(char_from_quotient (normal_cosetpre _)) //. - by rewrite cosetpreK (char_trans _ (center_char _)) ?Ohm_char. + by rewrite cosetpreK !gFchar_trans. rewrite (@Phi_mulg p) ?(pgroupS _ pG) // subUset commGC commMG; last first. by rewrite normsR ?(normsG sKG) // cents_norm // centsC. rewrite !mul_subG 1?commGC //. apply: subset_trans (commgS _ (subsetIr _ _)) _. rewrite -quotient_cents2 ?subsetIl // centsC // cosetpreK //. - by rewrite (subset_trans (Ohm_sub _ _)) // subsetIr. -have nZX := subset_trans sXG nZG; have pX : p.-group gX by exact: pgroupS pG. -rewrite -quotient_sub1 ?(subset_trans (Phi_sub _)) //=. -have pXZ: p.-group (gX / 'Z(K)) by exact: morphim_pgroup. + exact/gFsub_trans/subsetIr. +have nZX := subset_trans sXG nZG; have pX : p.-group gX by apply: pgroupS pG. +rewrite -quotient_sub1 ?gFsub_trans //=. +have pXZ: p.-group (gX / 'Z(K)) by apply: morphim_pgroup. rewrite (quotient_Phi pX nZX) subG1 (trivg_Phi pXZ). apply: (abelemS (quotientS _ (subsetIr _ _))); rewrite /= cosetpreK /=. have pZ: p.-group 'Z(G / 'Z(K)). @@ -1644,7 +1634,7 @@ have Hy: y \in H. rewrite -(injmSK inj1) ?cycle_subG // injm_subcent // subsetI. rewrite morphimS ?morphim_cycle ?cycle_subG //=. suffices: sdpair1 [Aut G] y \in [~: G', F']. - by rewrite commGC; apply: subsetP; exact/commG1P. + by rewrite commGC; apply: subsetP; apply/commG1P. rewrite morphM ?groupV ?morphV //= sdpair_act // -commgEl. by rewrite mem_commg ?mem_morphim ?cycle_id. have fy: f y = y := astabP cHFP _ Hy. diff --git a/mathcomp/solvable/nilpotent.v b/mathcomp/solvable/nilpotent.v index e1baa3f..d9a6f8b 100644 --- a/mathcomp/solvable/nilpotent.v +++ b/mathcomp/solvable/nilpotent.v @@ -1,14 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import path fintype div bigop prime finset. -From mathcomp.fingroup -Require Import fingroup morphism automorphism quotient gproduct. -From mathcomp.algebra -Require Import cyclic. -Require Import commutator gfunctor center gseries. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq path fintype div. +From mathcomp +Require Import bigop prime finset fingroup morphism automorphism quotient. +From mathcomp +Require Import commutator gproduct gfunctor center gseries cyclic. (******************************************************************************) (* This file defines nilpotent and solvable groups, and give some of their *) @@ -132,29 +129,27 @@ Lemma lcn0 A : 'L_0(A) = A. Proof. by []. Qed. Lemma lcn1 A : 'L_1(A) = A. Proof. by []. Qed. Lemma lcnSn n A : 'L_n.+2(A) = [~: 'L_n.+1(A), A]. Proof. by []. Qed. Lemma lcnSnS n G : [~: 'L_n(G), G] \subset 'L_n.+1(G). -Proof. by case: n => //; exact: der1_subG. Qed. +Proof. by case: n => //; apply: der1_subG. Qed. Lemma lcnE n A : 'L_n.+1(A) = lower_central_at_rec n A. Proof. by []. Qed. Lemma lcn2 A : 'L_2(A) = A^`(1). Proof. by []. Qed. Lemma lcn_group_set n G : group_set 'L_n(G). -Proof. by case: n => [|[|n]]; exact: groupP. Qed. +Proof. by case: n => [|[|n]]; apply: groupP. Qed. Canonical lower_central_at_group n G := Group (lcn_group_set n G). Lemma lcn_char n G : 'L_n(G) \char G. -Proof. -by case: n => [|n]; last elim: n => [|n IHn]; rewrite ?lcnSn ?charR ?char_refl. -Qed. +Proof. by case: n; last elim=> [|n IHn]; rewrite ?char_refl ?lcnSn ?charR. Qed. Lemma lcn_normal n G : 'L_n(G) <| G. -Proof. by apply: char_normal; exact: lcn_char. Qed. +Proof. exact/char_normal/lcn_char. Qed. Lemma lcn_sub n G : 'L_n(G) \subset G. -Proof. by case/andP: (lcn_normal n G). Qed. +Proof. exact/char_sub/lcn_char. Qed. Lemma lcn_norm n G : G \subset 'N('L_n(G)). -Proof. by case/andP: (lcn_normal n G). Qed. +Proof. exact/char_norm/lcn_char. Qed. Lemma lcn_subS n G : 'L_n.+1(G) \subset 'L_n(G). Proof. @@ -173,7 +168,7 @@ Qed. Lemma lcn_sub_leq m n G : n <= m -> 'L_m(G) \subset 'L_n(G). Proof. -by move/subnK <-; elim: {m}(m - n) => // m; exact: subset_trans (lcn_subS _ _). +by move/subnK <-; elim: {m}(m - n) => // m; apply: subset_trans (lcn_subS _ _). Qed. Lemma lcnS n A B : A \subset B -> 'L_n(A) \subset 'L_n(B). @@ -268,13 +263,13 @@ Qed. Lemma lcnP G : reflect (exists n, 'L_n.+1(G) = 1) (nilpotent G). Proof. apply: (iffP idP) => [nilG | [n Ln1]]. - by exists (nil_class G); exact/lcn_nil_classP. + by exists (nil_class G); apply/lcn_nil_classP. apply/forall_inP=> H /subsetIP[sHG sHR]; rewrite -subG1 -{}Ln1. by elim: n => // n IHn; rewrite (subset_trans sHR) ?commSg. Qed. Lemma abelian_nil G : abelian G -> nilpotent G. -Proof. by move=> abG; apply/lcnP; exists 1%N; exact/commG1P. Qed. +Proof. by move=> abG; apply/lcnP; exists 1%N; apply/commG1P. Qed. Lemma nil_class0 G : (nil_class G == 0) = (G :==: 1). Proof. @@ -349,15 +344,15 @@ Qed. (* Now extract all the intermediate facts of the last proof. *) Lemma ucn_group_set gT (G : {group gT}) : group_set 'Z_n(G). -Proof. by have [hZ ->] := ucn_pmap; exact: groupP. Qed. +Proof. by have [hZ ->] := ucn_pmap; apply: groupP. Qed. Canonical upper_central_at_group gT G := Group (@ucn_group_set gT G). Lemma ucn_sub gT (G : {group gT}) : 'Z_n(G) \subset G. -Proof. by have [hZ ->] := ucn_pmap; exact: gFsub. Qed. +Proof. by have [hZ ->] := ucn_pmap; apply: gFsub. Qed. Lemma morphim_ucn : GFunctor.pcontinuous (upper_central_at n). -Proof. by have [hZ ->] := ucn_pmap; exact: pmorphimF. Qed. +Proof. by have [hZ ->] := ucn_pmap; apply: pmorphimF. Qed. Canonical ucn_igFun := [igFun by ucn_sub & morphim_ucn]. Canonical ucn_gFun := [gFun by morphim_ucn]. @@ -432,7 +427,7 @@ have <-: 'Z(H / Z) * 'Z(K / Z) = 'Z(G / Z). by rewrite -mulHK quotientMl // center_prod ?quotient_cents. have ZquoZ (B A : {group gT}): B \subset 'C(A) -> 'Z_n(A) * 'Z_n(B) = Z -> 'Z(A / Z) = 'Z_n.+1(A) / Z. -- move=> cAB {defZ}defZ; have cAZnB := subset_trans (ucn_sub n B) cAB. +- move=> cAB {defZ}defZ; have cAZnB: 'Z_n(B) \subset 'C(A) := gFsub_trans _ cAB. have /second_isom[/=]: A \subset 'N(Z). by rewrite -defZ normsM ?gFnorm ?cents_norm // centsC. suffices ->: Z :&: A = 'Z_n(A). @@ -442,7 +437,7 @@ have ZquoZ (B A : {group gT}): by apply: subset_trans (ucn_sub_geq A n_gt0); rewrite /= setIC ucn1 setIS. rewrite (ZquoZ H K) 1?centsC 1?(centC cZn) // {ZquoZ}(ZquoZ K H) //. have cZn1: 'Z_n.+1(K) \subset 'C('Z_n.+1(H)) by apply: centSS cHK; apply: gFsub. -rewrite -quotientMl ?quotientK ?mul_subG ?(subset_trans (gFsub _ _)) //=. +rewrite -quotientMl ?quotientK ?mul_subG ?gFsub_trans //=. rewrite cprodE // -cent_joinEr ?mulSGid //= cent_joinEr //= -/Z. by rewrite -defZ mulgSS ?ucn_subS. Qed. @@ -475,25 +470,24 @@ Proof. rewrite !eqEsubset sub1G ucn_sub /= andbT -(ucn0 G). elim: {1 3}n 0 (addn0 n) => [j <- //|i IHi j]. rewrite addSnnS => /IHi <- {IHi}; rewrite ucnSn lcnSn. -have nZG := normal_norm (ucn_normal j G). -have nZL := subset_trans (lcn_sub _ _) nZG. -by rewrite -sub_morphim_pre // subsetI morphimS ?lcn_sub // quotient_cents2. +rewrite -sub_morphim_pre ?gFsub_trans ?gFnorm_trans // subsetI. +by rewrite morphimS ?gFsub // quotient_cents2 ?gFsub_trans ?gFnorm_trans. Qed. Lemma ucnP G : reflect (exists n, 'Z_n(G) = G) (nilpotent G). Proof. -apply: (iffP (lcnP G)) => [] [n /eqP clGn]; +apply: (iffP (lcnP G)) => -[n /eqP-clGn]; by exists n; apply/eqP; rewrite ucn_lcnP in clGn *. Qed. Lemma ucn_nil_classP n G : nilpotent G -> reflect ('Z_n(G) = G) (nil_class G <= n). Proof. -move=> nilG; rewrite (sameP (lcn_nil_classP n nilG) eqP) ucn_lcnP; exact: eqP. +move=> nilG; rewrite (sameP (lcn_nil_classP n nilG) eqP) ucn_lcnP; apply: eqP. Qed. Lemma ucn_id n G : 'Z_n('Z_n(G)) = 'Z_n(G). -Proof. by rewrite -{2}['Z_n(G)]gFid. Qed. +Proof. exact: gFid. Qed. Lemma ucn_nilpotent n G : nilpotent 'Z_n(G). Proof. by apply/ucnP; exists n; rewrite ucn_id. Qed. @@ -527,7 +521,7 @@ Lemma injm_nil G : 'injm f -> G \subset D -> nilpotent (f @* G) = nilpotent G. Proof. move=> injf sGD; apply/idP/idP; last exact: morphim_nil. case/ucnP=> n; rewrite -injm_ucn // => /injm_morphim_inj defZ. -by apply/ucnP; exists n; rewrite defZ ?(subset_trans (ucn_sub n G)). +by apply/ucnP; exists n; rewrite defZ ?gFsub_trans. Qed. Lemma nil_class_morphim G : nilpotent G -> nil_class (f @* G) <= nil_class G. @@ -542,7 +536,7 @@ Proof. move=> injf sGD; case nilG: (nilpotent G). apply/eqP; rewrite eqn_leq nil_class_morphim //. rewrite (sameP (lcn_nil_classP _ nilG) eqP) -subG1. - rewrite -(injmSK injf) ?(subset_trans (lcn_sub _ _)) // morphim1. + rewrite -(injmSK injf) ?gFsub_trans // morphim1. by rewrite morphim_lcn // (lcn_nil_classP _ _ (leqnn _)) //= injm_nil. transitivity #|G|; apply/eqP; rewrite eqn_leq. rewrite -(card_injm injf sGD) (leq_trans (index_size _ _)) ?size_mkseq //. @@ -608,21 +602,21 @@ have{nsGH} [i sZH []]: exists2 i, 'Z_i(G) \subset H & ~ 'Z_i.+1(G) \subset H. elim: n => [|i IHi] in nsGH *; first by rewrite sub1G in nsGH. by case sZH: ('Z_i(G) \subset H); [exists i | apply: IHi; rewrite sZH]. apply: subset_trans sNH; rewrite subsetI ucn_sub -commg_subr. -by apply: subset_trans sZH; apply: subset_trans (ucn_comm i G); exact: commgS. +by apply: subset_trans sZH; apply: subset_trans (ucn_comm i G); apply: commgS. Qed. Lemma nilpotent_proper_norm G H : nilpotent G -> H \proper G -> H \proper 'N_G(H). Proof. move=> nilG; rewrite properEneq properE subsetI normG => /andP[neHG sHG]. -by rewrite sHG; apply: contra neHG; move/(nilpotent_sub_norm nilG)->. +by rewrite sHG; apply: contra neHG => /(nilpotent_sub_norm nilG)->. Qed. Lemma nilpotent_subnormal G H : nilpotent G -> H \subset G -> H <|<| G. Proof. move=> nilG; elim: {H}_.+1 {-2}H (ltnSn (#|G| - #|H|)) => // m IHm H. -rewrite ltnS => leGHm sHG; have:= sHG; rewrite subEproper. -case/predU1P=> [->|]; first exact: subnormal_refl. +rewrite ltnS => leGHm sHG. +have [->|] := eqVproper sHG; first exact: subnormal_refl. move/(nilpotent_proper_norm nilG); set K := 'N_G(H) => prHK. have snHK: H <|<| K by rewrite normal_subnormal ?normalSG. have sKG: K \subset G by rewrite subsetIl. @@ -676,7 +670,7 @@ by rewrite (forall_inP nilG) // subsetI sHG (subset_trans sHH') ?commgS. Qed. Lemma abelian_sol G : abelian G -> solvable G. -Proof. by move/abelian_nil; exact: nilpotent_sol. Qed. +Proof. by move/abelian_nil/nilpotent_sol. Qed. Lemma solvable1 : solvable [1 gT]. Proof. exact: abelian_sol (abelian1 gT). Qed. @@ -705,7 +699,7 @@ elim: n => [|n IHn]; first by rewrite subn0 prednK. rewrite dergSn subnS -ltnS. have [-> | ntGn] := eqVneq G^`(n) 1; first by rewrite commG1 cards1. case: (_ - _) IHn => [|n']; first by rewrite leqNgt cardG_gt1 ntGn. -by apply: leq_trans (proper_card _); exact: sol_der1_proper (der_sub _ _) _. +by apply: leq_trans (proper_card _); apply: sol_der1_proper (der_sub _ _) _. Qed. End Solvable. @@ -725,8 +719,7 @@ Lemma injm_sol : 'injm f -> G \subset D -> solvable (f @* G) = solvable G. Proof. move=> injf sGD; apply/idP/idP; last exact: morphim_sol. case/derivedP=> n Gn1; apply/derivedP; exists n; apply/trivgP. -rewrite -(injmSK injf) ?(subset_trans (der_sub _ _)) ?morphim_der //. -by rewrite Gn1 morphim1. +by rewrite -(injmSK injf) ?gFsub_trans ?morphim_der // Gn1 morphim1. Qed. End MorphSol. diff --git a/mathcomp/solvable/pgroup.v b/mathcomp/solvable/pgroup.v index b1d2bf6..3d48a8a 100644 --- a/mathcomp/solvable/pgroup.v +++ b/mathcomp/solvable/pgroup.v @@ -1,14 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div fintype bigop finset prime. -From mathcomp.fingroup -Require Import fingroup morphism automorphism quotient action gproduct. -From mathcomp.algebra -Require Import cyclic. -Require Import gfunctor. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div. +From mathcomp +Require Import fintype bigop finset prime fingroup morphism. +From mathcomp +Require Import gfunctor automorphism quotient action gproduct cyclic. (******************************************************************************) (* Standard group notions and constructions based on the prime decomposition *) @@ -125,7 +122,7 @@ Implicit Types (x y z : gT) (A B C D : {set gT}) (G H K P Q R : {group gT}). Lemma trivgVpdiv G : G :=: 1 \/ (exists2 p, prime p & p %| #|G|). Proof. -have [leG1|lt1G] := leqP #|G| 1; first by left; exact: card_le1_trivg. +have [leG1|lt1G] := leqP #|G| 1; first by left; apply: card_le1_trivg. by right; exists (pdiv #|G|); rewrite ?pdiv_dvd ?pdiv_prime. Qed. @@ -140,13 +137,13 @@ Qed. Lemma pgroupE pi A : pi.-group A = pi.-nat #|A|. Proof. by []. Qed. Lemma sub_pgroup pi rho A : {subset pi <= rho} -> pi.-group A -> rho.-group A. -Proof. by move=> pi_sub_rho; exact: sub_in_pnat (in1W pi_sub_rho). Qed. +Proof. by move=> pi_sub_rho; apply: sub_in_pnat (in1W pi_sub_rho). Qed. Lemma eq_pgroup pi rho A : pi =i rho -> pi.-group A = rho.-group A. Proof. exact: eq_pnat. Qed. Lemma eq_p'group pi rho A : pi =i rho -> pi^'.-group A = rho^'.-group A. -Proof. by move/eq_negn; exact: eq_pnat. Qed. +Proof. by move/eq_negn; apply: eq_pnat. Qed. Lemma pgroupNK pi A : pi^'^'.-group A = pi.-group A. Proof. exact: pnatNK. Qed. @@ -164,7 +161,7 @@ Lemma p'groupEpi p G : p^'.-group G = (p \notin \pi(G)). Proof. exact: p'natEpi (cardG_gt0 G). Qed. Lemma pgroup_pi G : \pi(G).-group G. -Proof. by rewrite /=; exact: pnat_pi. Qed. +Proof. by rewrite /=; apply: pnat_pi. Qed. Lemma partG_eq1 pi G : (#|G|`_pi == 1%N) = pi^'.-group G. Proof. exact: partn_eq1 (cardG_gt0 G). Qed. @@ -178,10 +175,10 @@ Lemma pgroup1 pi : pi.-group [1 gT]. Proof. by rewrite /pgroup cards1. Qed. Lemma pgroupS pi G H : H \subset G -> pi.-group G -> pi.-group H. -Proof. by move=> sHG; exact: pnat_dvd (cardSg sHG). Qed. +Proof. by move=> sHG; apply: pnat_dvd (cardSg sHG). Qed. Lemma oddSg G H : H \subset G -> odd #|G| -> odd #|H|. -Proof. by rewrite !odd_2'nat; exact: pgroupS. Qed. +Proof. by rewrite !odd_2'nat; apply: pgroupS. Qed. Lemma odd_pgroup_odd p G : odd p -> p.-group G -> odd #|G|. Proof. @@ -212,7 +209,7 @@ Proof. by rewrite /pgroup cardJg. Qed. Lemma pgroup_p p P : p.-group P -> p_group P. Proof. -case: (leqP #|P| 1); first by move=> /card_le1_trivg-> _; exact: pgroup1. +case: (leqP #|P| 1); first by move=> /card_le1_trivg-> _; apply: pgroup1. move/pdiv_prime=> pr_q pgP; have:= pgroupP pgP _ pr_q (pdiv_dvd _). by rewrite /p_group => /eqnP->. Qed. @@ -220,7 +217,7 @@ Qed. Lemma p_groupP P : p_group P -> exists2 p, prime p & p.-group P. Proof. case: (ltnP 1 #|P|); first by move/pdiv_prime; exists (pdiv #|P|). -move/card_le1_trivg=> -> _; exists 2 => //; exact: pgroup1. +by move/card_le1_trivg=> -> _; exists 2 => //; apply: pgroup1. Qed. Lemma pgroup_pdiv p G : @@ -254,7 +251,7 @@ Proof. by case/and3P. Qed. Lemma pHallP pi G H : reflect (H \subset G /\ #|H| = #|G|`_pi) (pi.-Hall(G) H). Proof. apply: (iffP idP) => [piH | [sHG oH]]. - split; [exact: pHall_sub piH | exact: card_Hall]. + by split; [apply: pHall_sub piH | apply: card_Hall]. rewrite /pHall sHG -divgS // /pgroup oH. by rewrite -{2}(@partnC pi #|G|) ?mulKn ?part_pnat. Qed. @@ -268,7 +265,7 @@ Lemma coprime_mulpG_Hall pi G K R : Proof. move=> defG piK pi'R; apply/andP. rewrite /pHall piK -!divgS /= -defG ?mulG_subl ?mulg_subr //= pnatNK. -by rewrite coprime_cardMg ?(pnat_coprime piK) // mulKn ?mulnK //; exact/and3P. +by rewrite coprime_cardMg ?(pnat_coprime piK) // mulKn ?mulnK //; apply/and3P. Qed. Lemma coprime_mulGp_Hall pi G K R : @@ -284,15 +281,15 @@ Lemma eq_in_pHall pi rho G H : Proof. move=> eq_pi_rho; apply: andb_id2l => sHG. congr (_ && _); apply: eq_in_pnat => p piHp. - by apply: eq_pi_rho; exact: (piSg sHG). + by apply: eq_pi_rho; apply: (piSg sHG). by congr (~~ _); apply: eq_pi_rho; apply: (pi_of_dvd (dvdn_indexg G H)). Qed. Lemma eq_pHall pi rho G H : pi =i rho -> pi.-Hall(G) H = rho.-Hall(G) H. -Proof. by move=> eq_pi_rho; exact: eq_in_pHall (in1W eq_pi_rho). Qed. +Proof. by move=> eq_pi_rho; apply: eq_in_pHall (in1W eq_pi_rho). Qed. Lemma eq_p'Hall pi rho G H : pi =i rho -> pi^'.-Hall(G) H = rho^'.-Hall(G) H. -Proof. by move=> eq_pi_rho; exact: eq_pHall (eq_negn _). Qed. +Proof. by move=> eq_pi_rho; apply: eq_pHall (eq_negn _). Qed. Lemma pHallNK pi G H : pi^'^'.-Hall(G) H = pi.-Hall(G) H. Proof. exact: eq_pHall (negnK _). Qed. @@ -322,7 +319,7 @@ by case/andP=> sHG coHG /=; rewrite /pHall sHG /pgroup pnat_pi -?coprime_pi'. Qed. Lemma HallP G H : Hall G H -> exists pi, pi.-Hall(G) H. -Proof. by exists \pi(H); exact: Hall_pi. Qed. +Proof. by exists \pi(H); apply: Hall_pi. Qed. Lemma sdprod_Hall G K H : K ><| H = G -> Hall G K = Hall G H. Proof. @@ -355,7 +352,7 @@ Qed. Lemma compl_p'Hall pi K H G : pi^'.-Hall(G) K -> (H \in [complements to K in G]) = pi.-Hall(G) H. -Proof. by move/compl_pHall->; exact: eq_pHall (negnK pi). Qed. +Proof. by move/compl_pHall->; apply: eq_pHall (negnK pi). Qed. Lemma sdprod_normal_p'HallP pi K H G : K <| G -> pi^'.-Hall(G) H -> reflect (K ><| H = G) (pi.-Hall(G) K). @@ -402,8 +399,7 @@ Qed. Lemma pHall_subl pi G K H : H \subset K -> K \subset G -> pi.-Hall(G) H -> pi.-Hall(K) H. Proof. -move=> sHK sKG; rewrite /pHall sHK; case/and3P=> _ ->. -by apply: pnat_dvd; exact: indexSg. +by move=> sHK sKG; rewrite /pHall sHK => /and3P[_ ->]; apply/pnat_dvd/indexSg. Qed. Lemma Hall1 G : Hall G 1. @@ -446,14 +442,14 @@ Lemma p_eltM_norm pi x y : x \in 'N(<[y]>) -> pi.-elt x -> pi.-elt y -> pi.-elt (x * y). Proof. move=> nyx pi_x pi_y; apply: (@mem_p_elt pi _ (<[x]> <*> <[y]>)%G). - by rewrite /= norm_joinEl ?cycle_subG // pgroupM; exact/andP. + by rewrite /= norm_joinEl ?cycle_subG // pgroupM; apply/andP. by rewrite groupM // mem_gen // inE cycle_id ?orbT. Qed. Lemma p_eltM pi x y : commute x y -> pi.-elt x -> pi.-elt y -> pi.-elt (x * y). Proof. move=> cxy; apply: p_eltM_norm; apply: (subsetP (cent_sub _)). -by rewrite cent_gen cent_set1; exact/cent1P. +by rewrite cent_gen cent_set1; apply/cent1P. Qed. Lemma p_elt1 pi : pi.-elt (1 : gT). @@ -469,10 +465,10 @@ Lemma p_eltJ pi x y : pi.-elt (x ^ y) = pi.-elt x. Proof. by congr pnat; rewrite orderJ. Qed. Lemma sub_p_elt pi1 pi2 x : {subset pi1 <= pi2} -> pi1.-elt x -> pi2.-elt x. -Proof. by move=> pi12; apply: sub_in_pnat => q _; exact: pi12. Qed. +Proof. by move=> pi12; apply: sub_in_pnat => q _; apply: pi12. Qed. Lemma eq_p_elt pi1 pi2 x : pi1 =i pi2 -> pi1.-elt x = pi2.-elt x. -Proof. by move=> pi12; exact: eq_pnat. Qed. +Proof. by move=> pi12; apply: eq_pnat. Qed. Lemma p_eltNK pi x : pi^'^'.-elt x = pi.-elt x. Proof. exact: pnatNK. Qed. @@ -480,7 +476,7 @@ Proof. exact: pnatNK. Qed. Lemma eq_constt pi1 pi2 x : pi1 =i pi2 -> x.`_pi1 = x.`_pi2. Proof. move=> pi12; congr (x ^+ (chinese _ _ 1 0)); apply: eq_partn => // a. -by congr (~~ _); exact: pi12. +by congr (~~ _); apply: pi12. Qed. Lemma consttNK pi x : x.`_pi^'^' = x.`_pi. @@ -514,8 +510,8 @@ Proof. by rewrite -{1}(consttC pi^' x) consttNK mulgK p_elt_constt. Qed. Lemma order_constt pi (x : gT) : #[x.`_pi] = #[x]`_pi. Proof. rewrite -{2}(consttC pi x) orderM; [|exact: commuteX2|]; last first. - by apply: (@pnat_coprime pi); exact: p_elt_constt. -by rewrite partnM // part_pnat_id ?part_p'nat ?muln1 //; exact: p_elt_constt. + by apply: (@pnat_coprime pi); apply: p_elt_constt. +by rewrite partnM // part_pnat_id ?part_p'nat ?muln1 //; apply: p_elt_constt. Qed. Lemma consttM pi x y : commute x y -> (x * y).`_pi = x.`_pi * y.`_pi. @@ -535,7 +531,7 @@ Qed. Lemma consttX pi x n : (x ^+ n).`_pi = x.`_pi ^+ n. Proof. elim: n => [|n IHn]; first exact: constt1. -rewrite !expgS consttM ?IHn //; exact: commuteX. +by rewrite !expgS consttM ?IHn //; apply: commuteX. Qed. Lemma constt1P pi x : reflect (x.`_pi = 1) (pi^'.-elt x). @@ -562,12 +558,12 @@ Qed. Lemma prod_constt x : \prod_(0 <= p < #[x].+1) x.`_p = x. Proof. pose lp n := [pred p | p < n]. -have: (lp #[x].+1).-elt x by apply/pnatP=> // p _; exact: dvdn_leq. +have: (lp #[x].+1).-elt x by apply/pnatP=> // p _; apply: dvdn_leq. move/constt_p_elt=> def_x; symmetry; rewrite -{1}def_x {def_x}. elim: _.+1 => [|p IHp]. - by rewrite big_nil; apply/constt1P; exact/pgroupP. + by rewrite big_nil; apply/constt1P; apply/pgroupP. rewrite big_nat_recr //= -{}IHp -(consttC (lp p) x.`__); congr (_ * _). - rewrite sub_in_constt // => q _; exact: leqW. + by rewrite sub_in_constt // => q _; apply: leqW. set y := _.`__; rewrite -(consttC p y) (constt1P p^' _ _) ?mulg1. by rewrite 2?sub_in_constt // => q _; move/eqnP->; rewrite !inE ?ltnn. rewrite /p_elt pnatNK !order_constt -partnI. @@ -604,7 +600,7 @@ Qed. Lemma norm_sub_max_pgroup pi H M G : [max M | pi.-subgroup(G) M] -> pi.-group H -> H \subset G -> H \subset 'N(M) -> H \subset M. -Proof. by move=> maxM piH sHG /normC; exact: comm_sub_max_pgroup piH sHG. Qed. +Proof. by move=> maxM piH sHG /normC; apply: comm_sub_max_pgroup piH sHG. Qed. Lemma sub_pHall pi H G K : pi.-Hall(G) H -> pi.-group K -> H \subset K -> K \subset G -> K :=: H. @@ -615,9 +611,9 @@ Qed. Lemma Hall_max pi H G : pi.-Hall(G) H -> [max H | pi.-subgroup(G) H]. Proof. -move=> hallH; apply/maxgroupP; split=> [|K]. +move=> hallH; apply/maxgroupP; split=> [|K /andP[sKG piK] sHK]. by rewrite /psubgroup; case/and3P: hallH => ->. -case/andP=> sKG piK sHK; exact: (sub_pHall hallH). +exact: (sub_pHall hallH). Qed. Lemma pHall_id pi H G : pi.-Hall(G) H -> pi.-group G -> H :=: G. @@ -644,7 +640,7 @@ have{pG n leGn IHn} pZ: p %| #|'C_G(G)|. by have /andP[] := no_x y. by apply/implyP; rewrite -index_cent1 indexgI implyNb -Euclid_dvdM ?LagrangeI. have [Q maxQ _]: {Q | [max Q | p^'.-subgroup('C_G(G)) Q] & 1%G \subset Q}. - apply: maxgroup_exists; exact: psubgroup1. + by apply: maxgroup_exists; apply: psubgroup1. case/andP: (maxgroupp maxQ) => sQC; rewrite /pgroup p'natE // => /negP[]. apply: dvdn_trans pZ (cardSg _); apply/subsetP=> x /setIP[Gx Cx]. rewrite -sub1set -gen_subG (normal_sub_max_pgroup maxQ) //; last first. @@ -670,7 +666,7 @@ Qed. Lemma mem_normal_Hall pi H G x : pi.-Hall(G) H -> H <| G -> x \in G -> (x \in H) = pi.-elt x. -Proof. by rewrite -!cycle_subG; exact: sub_normal_Hall. Qed. +Proof. by rewrite -!cycle_subG; apply: sub_normal_Hall. Qed. Lemma uniq_normal_Hall pi H G K : pi.-Hall(G) H -> H <| G -> [max K | pi.-subgroup(G) K] -> K :=: H. @@ -723,17 +719,17 @@ Variables (aT rT : finGroupType) (D : {group aT}) (f : {morphism D >-> rT}). Implicit Types (pi : nat_pred) (G H P : {group aT}). Lemma morphim_pgroup pi G : pi.-group G -> pi.-group (f @* G). -Proof. by apply: pnat_dvd; exact: dvdn_morphim. Qed. +Proof. by apply: pnat_dvd; apply: dvdn_morphim. Qed. Lemma morphim_odd G : odd #|G| -> odd #|f @* G|. -Proof. by rewrite !odd_2'nat; exact: morphim_pgroup. Qed. +Proof. by rewrite !odd_2'nat; apply: morphim_pgroup. Qed. Lemma pmorphim_pgroup pi G : pi.-group ('ker f) -> G \subset D -> pi.-group (f @* G) = pi.-group G. Proof. move=> piker sGD; apply/idP/idP=> [pifG|]; last exact: morphim_pgroup. apply: (@pgroupS _ _ (f @*^-1 (f @* G))); first by rewrite -sub_morphim_pre. -by rewrite /pgroup card_morphpre ?morphimS // pnat_mul; exact/andP. +by rewrite /pgroup card_morphpre ?morphimS // pnat_mul; apply/andP. Qed. Lemma morphim_p_index pi G H : @@ -760,7 +756,7 @@ Qed. Lemma morphim_Hall G H : H \subset D -> Hall G H -> Hall (f @* G) (f @* H). Proof. -by move=> sHD /HallP[pi piH]; apply: (@pHall_Hall _ pi); exact: morphim_pHall. +by move=> sHD /HallP[pi piH]; apply: (@pHall_Hall _ pi); apply: morphim_pHall. Qed. Lemma morphim_pSylow p G P : @@ -768,7 +764,7 @@ Lemma morphim_pSylow p G P : Proof. exact: morphim_pHall. Qed. Lemma morphim_p_group P : p_group P -> p_group (f @* P). -Proof. by move/morphim_pgroup; exact: pgroup_p. Qed. +Proof. by move/morphim_pgroup; apply: pgroup_p. Qed. Lemma morphim_Sylow G P : P \subset D -> Sylow G P -> Sylow (f @* G) (f @* P). Proof. @@ -776,12 +772,12 @@ by move=> sPD /andP[pP hallP]; rewrite /Sylow morphim_p_group // morphim_Hall. Qed. Lemma morph_p_elt pi x : x \in D -> pi.-elt x -> pi.-elt (f x). -Proof. by move=> Dx; apply: pnat_dvd; exact: morph_order. Qed. +Proof. by move=> Dx; apply: pnat_dvd; apply: morph_order. Qed. Lemma morph_constt pi x : x \in D -> f x.`_pi = (f x).`_pi. Proof. move=> Dx; rewrite -{2}(consttC pi x) morphM ?groupX //. -rewrite consttM; last by rewrite !morphX //; exact: commuteX2. +rewrite consttM; last by rewrite !morphX //; apply: commuteX2. have: pi.-elt (f x.`_pi) by rewrite morph_p_elt ?groupX ?p_elt_constt //. have: pi^'.-elt (f x.`_pi^') by rewrite morph_p_elt ?groupX ?p_elt_constt //. by move/constt1P->; move/constt_p_elt->; rewrite mulg1. @@ -817,7 +813,7 @@ Lemma ltn_log_quotient : Proof. move=> pG ntH sHG; apply: contraLR (ltn_quotient ntH sHG); rewrite -!leqNgt. rewrite {2}(card_pgroup pG) {2}(card_pgroup (morphim_pgroup _ pG)). -by case: (posnP p) => [-> //|]; exact: leq_pexp2l. +by case: (posnP p) => [-> //|]; apply: leq_pexp2l. Qed. End Pquotient. @@ -883,7 +879,7 @@ Canonical pcore_mod_group pi B : {group gT} := Definition pseries := foldr pcore_mod 1 (rev pis). Lemma pseries_group_set : group_set pseries. -Proof. rewrite /pseries; case: rev => [|pi1 pi1']; exact: groupP. Qed. +Proof. by rewrite /pseries; case: rev => [|pi1 pi1']; apply: groupP. Qed. Canonical pseries_group : {group gT} := group pseries_group_set. @@ -906,7 +902,7 @@ Lemma pcore_psubgroup G : pi.-subgroup(G) 'O_pi(G). Proof. have [M maxM _]: {M | [max M | pi.-subgroup(G) M] & 1%G \subset M}. by apply: maxgroup_exists; rewrite /psubgroup sub1G pgroup1. -have sOM: 'O_pi(G) \subset M by exact: bigcap_inf. +have sOM: 'O_pi(G) \subset M by apply: bigcap_inf. have /andP[piM sMG] := maxgroupp maxM. by rewrite /psubgroup (pgroupS sOM) // (subset_trans sOM). Qed. @@ -918,7 +914,7 @@ Lemma pcore_sub G : 'O_pi(G) \subset G. Proof. by case/andP: (pcore_psubgroup G). Qed. Lemma pcore_sub_Hall G H : pi.-Hall(G) H -> 'O_pi(G) \subset H. -Proof. by move/Hall_max=> maxH; exact: bigcap_inf. Qed. +Proof. by move/Hall_max=> maxH; apply: bigcap_inf. Qed. Lemma pcore_max G H : pi.-group H -> H <| G -> H \subset 'O_pi(G). Proof. @@ -952,11 +948,11 @@ Qed. Lemma sub_Hall_pcore G K : pi.-Hall(G) 'O_pi(G) -> K \subset G -> (K \subset 'O_pi(G)) = pi.-group K. -Proof. by move=> hallGpi; exact: sub_normal_Hall (pcore_normal G). Qed. +Proof. by move=> hallGpi; apply: sub_normal_Hall (pcore_normal G). Qed. Lemma mem_Hall_pcore G x : pi.-Hall(G) 'O_pi(G) -> x \in G -> (x \in 'O_pi(G)) = pi.-elt x. -Proof. move=> hallGpi; exact: mem_normal_Hall (pcore_normal G). Qed. +Proof. by move=> hallGpi; apply: mem_normal_Hall (pcore_normal G). Qed. Lemma sdprod_Hall_pcoreP H G : pi.-Hall(G) 'O_pi(G) -> reflect ('O_pi(G) ><| H = G) (pi^'.-Hall(G) H). @@ -986,14 +982,14 @@ Lemma morphim_pcore pi : GFunctor.pcontinuous (pcore pi). Proof. move=> gT rT D G f; apply/bigcapsP=> M /normal_sub_max_pgroup; apply. by rewrite morphim_pgroup ?pcore_pgroup. -apply: morphim_normal; exact: pcore_normal. +by apply: morphim_normal; apply: pcore_normal. Qed. Lemma pcoreS pi gT (G H : {group gT}) : H \subset G -> H :&: 'O_pi(G) \subset 'O_pi(H). Proof. move=> sHG; rewrite -{2}(setIidPl sHG). -do 2!rewrite -(morphim_idm (subsetIl H _)) morphimIdom; exact: morphim_pcore. +by do 2!rewrite -(morphim_idm (subsetIl H _)) morphimIdom; apply: morphim_pcore. Qed. Canonical pcore_igFun pi := [igFun by pcore_sub pi & morphim_pcore pi]. @@ -1009,17 +1005,12 @@ Variable F : GFunctor.pmap. Lemma pcore_mod_sub pi gT (G : {group gT}) : pcore_mod G pi (F _ G) \subset G. Proof. -have nFD := gFnorm F G; rewrite sub_morphpre_im ?pcore_sub //=. - by rewrite ker_coset_prim subIset // gen_subG gFsub. -by apply: subset_trans (pcore_sub _ _) _; apply: morphimS. +by rewrite sub_morphpre_im ?gFsub_trans ?morphimS ?gFnorm //= ker_coset gFsub. Qed. Lemma quotient_pcore_mod pi gT (G : {group gT}) (B : {set gT}) : pcore_mod G pi B / B = 'O_pi(G / B). -Proof. -apply: morphpreK; apply: subset_trans (pcore_sub _ _) _. -by rewrite /= /quotient -morphimIdom morphimS ?subsetIl. -Qed. +Proof. exact/morphpreK/gFsub_trans/morphim_sub. Qed. Lemma morphim_pcore_mod pi gT rT (D G : {group gT}) (f : {morphism D >-> rT}) : f @* pcore_mod G pi (F _ G) \subset pcore_mod (f @* G) pi (F _ (f @* G)). @@ -1111,7 +1102,7 @@ Qed. Lemma pseries_pop2 pi1 pi2 gT (G : {group gT}) : 'O_pi1(G) = 1 -> 'O_{pi1, pi2}(G) = 'O_pi2(G). -Proof. by move/pseries_pop->; exact: pseries1. Qed. +Proof. by move/pseries_pop->; apply: pseries1. Qed. Lemma pseries_sub_catl pi1s pi2s gT (G : {group gT}) : pseries pi1s G \subset pseries (pi1s ++ pi2s) G. @@ -1126,9 +1117,7 @@ Proof. by rewrite pseries_rcons quotient_pcore_mod. Qed. Lemma pseries_norm2 pi1s pi2s gT (G : {group gT}) : pseries pi2s G \subset 'N(pseries pi1s G). -Proof. -apply: subset_trans (normal_norm (pseries_normal pi1s G)); exact: pseries_sub. -Qed. +Proof. by rewrite gFsub_trans ?gFnorm. Qed. Lemma pseries_sub_catr pi1s pi2s gT (G : {group gT}) : pseries pi2s G \subset pseries (pi1s ++ pi2s) G. @@ -1177,7 +1166,7 @@ rewrite /= cat_rcons -(IHpi (pi :: pi2s)) {1}quotient_pseries IHpi. apply/eqP; rewrite quotient_pseries eqEsubset !pcore_max ?pcore_pgroup //=. rewrite -quotient_pseries morphim_normal // /(_ <| _) pseries_norm2. by rewrite -cat_rcons pseries_sub_catl. -by rewrite (char_normal_trans (pcore_char _ _)) ?quotient_normal ?gFnormal. +by rewrite gFnormal_trans ?quotient_normal ?gFnormal. Qed. Lemma pseries_char_catl pi1s pi2s gT (G : {group gT}) : @@ -1197,8 +1186,7 @@ rewrite /= -Epis {1}quotient_pseries Epis quotient_pseries. apply/eqP; rewrite eqEsubset !pcore_max ?pcore_pgroup //=. rewrite -quotient_pseries morphim_normal // /(_ <| _) pseries_norm2. by rewrite pseries_sub_catr. -apply: char_normal_trans (pcore_char pi _) (morphim_normal _ _). -exact: pseries_normal. +by rewrite gFnormal_trans ?morphim_normal ?gFnormal. Qed. Lemma pseries_char_catr pi1s pi2s gT (G : {group gT}) : @@ -1208,12 +1196,11 @@ Proof. by rewrite -(pseries_catr_id pi1s pi2s G) pseries_char. Qed. Lemma pcore_modp pi gT (G H : {group gT}) : H <| G -> pi.-group H -> pcore_mod G pi H = 'O_pi(G). Proof. -move=> nsHG piH; apply/eqP; rewrite eqEsubset andbC. -have nHG := normal_norm nsHG; have sOG := subset_trans (pcore_sub pi _). -rewrite -sub_morphim_pre ?(sOG, morphim_pcore) // pcore_max //. - rewrite -(pquotient_pgroup piH) ?subsetIl //. - by rewrite quotient_pcore_mod pcore_pgroup. -by rewrite -{2}(quotientGK nsHG) morphpre_normal ?pcore_normal ?sOG ?morphimS. +move=> nsHG piH; have nHG := normal_norm nsHG; apply/eqP. +rewrite eqEsubset andbC -sub_morphim_pre ?(gFsub_trans, morphim_pcore) //=. +rewrite -[G in 'O_pi(G)](quotientGK nsHG) pcore_max //. + by rewrite -(pquotient_pgroup piH) ?subsetIl // cosetpreK pcore_pgroup. +by rewrite morphpre_normal ?gFnormal ?gFsub_trans ?morphim_sub. Qed. Lemma pquotient_pcore pi gT (G H : {group gT}) : @@ -1221,9 +1208,7 @@ Lemma pquotient_pcore pi gT (G H : {group gT}) : Proof. by move=> nsHG piH; rewrite -quotient_pcore_mod pcore_modp. Qed. Lemma trivg_pcore_quotient pi gT (G : {group gT}) : 'O_pi(G / 'O_pi(G)) = 1. -Proof. -by rewrite pquotient_pcore ?pcore_normal ?pcore_pgroup // trivg_quotient. -Qed. +Proof. by rewrite pquotient_pcore ?gFnormal ?pcore_pgroup ?trivg_quotient. Qed. Lemma pseries_rcons_id pis pi gT (G : {group gT}) : pseries (rcons (rcons pis pi) pi) G = pseries (rcons pis pi) G. @@ -1246,11 +1231,11 @@ Lemma sub_in_pcore pi rho G : Proof. move=> pi_sub_rho; rewrite pcore_max ?pcore_normal //. apply: sub_in_pnat (pcore_pgroup _ _) => p. -move/(piSg (pcore_sub _ _)); exact: pi_sub_rho. +by move/(piSg (pcore_sub _ _)); apply: pi_sub_rho. Qed. Lemma sub_pcore pi rho G : {subset pi <= rho} -> 'O_pi(G) \subset 'O_rho(G). -Proof. by move=> pi_sub_rho; exact: sub_in_pcore (in1W pi_sub_rho). Qed. +Proof. by move=> pi_sub_rho; apply: sub_in_pcore (in1W pi_sub_rho). Qed. Lemma eq_in_pcore pi rho G : {in \pi(G), pi =i rho} -> 'O_pi(G) = 'O_rho(G). Proof. @@ -1259,31 +1244,30 @@ by rewrite !sub_in_pcore // => p /eq_pi_rho->. Qed. Lemma eq_pcore pi rho G : pi =i rho -> 'O_pi(G) = 'O_rho(G). -Proof. by move=> eq_pi_rho; exact: eq_in_pcore (in1W eq_pi_rho). Qed. +Proof. by move=> eq_pi_rho; apply: eq_in_pcore (in1W eq_pi_rho). Qed. Lemma pcoreNK pi G : 'O_pi^'^'(G) = 'O_pi(G). -Proof. by apply: eq_pcore; exact: negnK. Qed. +Proof. by apply: eq_pcore; apply: negnK. Qed. Lemma eq_p'core pi rho G : pi =i rho -> 'O_pi^'(G) = 'O_rho^'(G). -Proof. by move/eq_negn; exact: eq_pcore. Qed. +Proof. by move/eq_negn; apply: eq_pcore. Qed. Lemma sdprod_Hall_p'coreP pi H G : pi^'.-Hall(G) 'O_pi^'(G) -> reflect ('O_pi^'(G) ><| H = G) (pi.-Hall(G) H). -Proof. by rewrite -(pHallNK pi G H); exact: sdprod_Hall_pcoreP. Qed. +Proof. by rewrite -(pHallNK pi G H); apply: sdprod_Hall_pcoreP. Qed. Lemma sdprod_p'core_HallP pi H G : pi.-Hall(G) H -> reflect ('O_pi^'(G) ><| H = G) (pi^'.-Hall(G) 'O_pi^'(G)). -Proof. by rewrite -(pHallNK pi G H); exact: sdprod_pcore_HallP. Qed. +Proof. by rewrite -(pHallNK pi G H); apply: sdprod_pcore_HallP. Qed. Lemma pcoreI pi rho G : 'O_[predI pi & rho](G) = 'O_pi('O_rho(G)). Proof. apply/eqP; rewrite eqEsubset !pcore_max //. -- rewrite /pgroup pnatI [pnat _ _]pcore_pgroup. - exact: pgroupS (pcore_sub _ _) (pcore_pgroup _ _). -- exact: char_normal_trans (pcore_char _ _) (pcore_normal _ _). -- by apply: sub_in_pnat (pcore_pgroup _ _) => p _ /andP[]. +- by rewrite /pgroup pnatI -!pgroupE !(pcore_pgroup, pgroupS (pcore_sub pi _)). +- by rewrite !gFnormal_trans. +- by apply: sub_pgroup (pcore_pgroup _ _) => p /andP[]. apply/andP; split; first by apply: sub_pcore => p /andP[]. -by rewrite (subset_trans (pcore_sub _ _)) ?gFnorm. +by rewrite gFnorm_trans ?normsG ?gFsub. Qed. Lemma bigcap_p'core pi G : @@ -1298,7 +1282,7 @@ apply/eqP; rewrite eqEsubset subsetI pcore_sub pcore_max /=. rewrite (dvdn_trans qGpi') ?cardSg ?subIset //= orbC. by rewrite (bigcap_inf (Ordinal ltqG)). rewrite /normal subsetIl normsI ?normG // norms_bigcap //. -by apply/bigcapsP => p _; exact: gFnorm. +by apply/bigcapsP => p _; apply: gFnorm. Qed. Lemma coprime_pcoreC (rT : finGroupType) pi G (R : {group rT}) : @@ -1310,10 +1294,9 @@ Proof. by rewrite coprime_TIg ?coprime_pcoreC. Qed. Lemma pcore_setI_normal pi G H : H <| G -> 'O_pi(G) :&: H = 'O_pi(H). Proof. -move=> nsHG; apply/eqP; rewrite eqEsubset subsetI pcore_sub. -rewrite !pcore_max ?(pgroupS (subsetIl _ H)) ?pcore_pgroup //=. - exact: char_normal_trans (pcore_char pi H) nsHG. -by rewrite setIC (normalGI (normal_sub nsHG)) ?pcore_normal. +move=> nsHG; apply/eqP; rewrite eqEsubset subsetI pcore_sub setIC. +rewrite !pcore_max ?(pgroupS (subsetIr H _)) ?pcore_pgroup ?gFnormal_trans //=. +by rewrite norm_normalI ?gFnorm_trans ?normsG ?normal_sub. Qed. End EqPcore. diff --git a/mathcomp/solvable/primitive_action.v b/mathcomp/solvable/primitive_action.v index 06ab30e..2d07bba 100644 --- a/mathcomp/solvable/primitive_action.v +++ b/mathcomp/solvable/primitive_action.v @@ -1,12 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div fintype tuple finset. -From mathcomp.fingroup -Require Import fingroup action. -Require Import gseries. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat. +From mathcomp +Require Import div seq fintype tuple finset. +From mathcomp +Require Import fingroup action gseries. (******************************************************************************) (* n-transitive and primitive actions: *) @@ -65,7 +64,7 @@ move=> Sx trG; rewrite /primitive trG negb_exists. apply/forallP/maximal_eqP=> /= [primG | [_ maxCx] Q]. split=> [|H sCH sHG]; first exact: subsetIl. pose X := orbit to H x; pose Q := orbit (to^*)%act G X. - have Xx: x \in X by exact: orbit_refl. + have Xx: x \in X by apply: orbit_refl. have defH: 'N_(G)(X | to) = H. have trH: [transitive H, on X | to] by apply/imsetP; exists x. have sHN: H \subset 'N_G(X | to) by rewrite subsetI sHG atrans_acts. @@ -79,13 +78,13 @@ apply/forallP/maximal_eqP=> /= [primG | [_ maxCx] Q]. - rewrite card_orbit astab1_set defH -(@ltn_pmul2l #|H|) ?Lagrange // muln1. rewrite oHG -(@ltn_pmul2l #|H|) ?Lagrange // -(card_orbit_stab to G x). by rewrite -(atransP trG x Sx) mulnC card_orbit ltn_pmul2r. - - by apply/actsP=> a Ga Y; apply: orbit_transr; exact: mem_orbit. + - by apply/actsP=> a Ga Y; apply/orbit_transl/mem_orbit. apply/and3P; split; last 1 first. - rewrite orbit_sym; apply/imsetP=> [[a _]] /= defX. by rewrite defX /setact imset0 inE in Xx. - apply/eqP/setP=> y; apply/bigcupP/idP=> [[_ /imsetP[a Ga ->]] | Sy]. case/imsetP=> _ /imsetP[b Hb ->] ->. - by rewrite !(actsP (atrans_acts trG)) //; exact: subsetP Hb. + by rewrite !(actsP (atrans_acts trG)) //; apply: subsetP Hb. case: (atransP2 trG Sx Sy) => a Ga ->. by exists ((to^*)%act X a); apply: mem_imset; rewrite // orbit_refl. apply/trivIsetP=> _ _ /imsetP[a Ga ->] /imsetP[b Gb ->]. @@ -179,7 +178,7 @@ Lemma n_act_dtuple t a : Proof. move/astabsP=> toSa /dtuple_onP[t_inj St]; apply/dtuple_onP. split=> [i j | i]; rewrite !tnth_map ?[_ \in S]toSa //. -by move/act_inj; exact: t_inj. +by move/act_inj; apply: t_inj. Qed. End NTransitive. @@ -202,7 +201,7 @@ Variables (gT : finGroupType) (sT : finType). Variables (to : {action gT &-> sT}) (G : {group gT}) (S : {set sT}). Lemma card_uniq_tuple n (t : n.-tuple sT) : uniq t -> #|t| = n. -Proof. by move/card_uniqP->; exact: size_tuple. Qed. +Proof. by move/card_uniqP->; apply: size_tuple. Qed. Lemma n_act0 (t : 0.-tuple sT) a : n_act to t a = [tuple]. Proof. exact: tuple0. Qed. @@ -223,7 +222,7 @@ Qed. Lemma dtuple_on_subset n (S1 S2 : {set sT}) t : S1 \subset S2 -> t \in n.-dtuple(S1) -> t \in n.-dtuple(S2). -Proof. by move=> sS12; rewrite !inE => /andP[-> /subset_trans]; exact. Qed. +Proof. by move=> sS12; rewrite !inE => /andP[-> /subset_trans]; apply. Qed. Lemma n_act_add n x (t : n.-tuple sT) a : n_act to [tuple of x :: t] a = [tuple of to x a :: n_act to t a]. @@ -254,7 +253,7 @@ case/and3P=> Sx ntx dt; set xt := [tuple of _] => tr_xt. apply/imsetP; exists t => //. apply/setP=> u; apply/idP/imsetP=> [du | [a Ga ->{u}]]. case: (ext_t u du) => y; rewrite tr_xt. - by case/imsetP=> a Ga [_ def_u]; exists a => //; exact: val_inj. + by case/imsetP=> a Ga [_ def_u]; exists a => //; apply: val_inj. have: n_act to xt a \in dtuple_on _ S by rewrite tr_xt mem_imset. by rewrite n_act_add dtuple_on_add; case/and3P. Qed. @@ -267,14 +266,14 @@ have trdom1 x: ([tuple x] \in 1.-dtuple(S)) = (x \in S). move=> m_gt0 /(ntransitive_weak m_gt0) {m m_gt0}. case/imsetP; case/tupleP=> x t0; rewrite {t0}(tuple0 t0) trdom1 => Sx trx. apply/imsetP; exists x => //; apply/setP=> y; rewrite -trdom1 trx. -apply/imsetP/imsetP=> [[a ? [->]]|[a ? ->]]; exists a => //; exact: val_inj. +by apply/imsetP/imsetP=> [[a ? [->]]|[a ? ->]]; exists a => //; apply: val_inj. Qed. Lemma ntransitive_primitive m : 1 < m -> [transitive^m G, on S | to] -> [primitive G, on S | to]. Proof. move=> lt1m /(ntransitive_weak lt1m) {m lt1m}tr2G. -have trG: [transitive G, on S | to] by exact: ntransitive1 tr2G. +have trG: [transitive G, on S | to] by apply: ntransitive1 tr2G. have [x Sx _]:= imsetP trG; rewrite (trans_prim_astab Sx trG). apply/maximal_eqP; split=> [|H]; first exact: subsetIl; rewrite subEproper. case/predU1P; first by [left]; case/andP=> sCH /subsetPn[a Ha nCa] sHG. @@ -287,7 +286,7 @@ rewrite eqEsubset acts_sub_orbit // Sy andbT; apply/subsetP=> z Sz. have [-> | zx] := eqVneq z x; first by rewrite orbit_sym mem_orbit. pose ty := [tuple y; x]; pose tz := [tuple z; x]. have [Sty Stz]: ty \in 2.-dtuple(S) /\ tz \in 2.-dtuple(S). - rewrite !inE !memtE !subset_all /= !mem_seq1 !andbT; split; exact/and3P. + by rewrite !inE !memtE !subset_all /= !mem_seq1 !andbT; split; apply/and3P. case: (atransP2 tr2G Sty Stz) => b Gb [->] /esym/astab1P cxb. by rewrite mem_orbit // (subsetP sCH) // inE Gb. Qed. @@ -310,12 +309,12 @@ case/and3P=> Sx1 nt1x1 dt1 trt1; have Gtr1 := ntransitive1 (ltn0Sn _) Gtr. case: (atransP2 Gtr1 Sx1 Sx) => // a Ga x1ax. pose t := n_act to t1 a. have dxt: [tuple of x :: t] \in m.+1.-dtuple(S). - rewrite trt1 x1ax; apply/imsetP; exists a => //; exact: val_inj. + by rewrite trt1 x1ax; apply/imsetP; exists a => //; apply: val_inj. apply/imsetP; exists t; first by rewrite dtuple_on_add_D1 Sx in dxt. apply/setP=> t2; apply/idP/imsetP => [dt2|[b]]. have: [tuple of x :: t2] \in dtuple_on _ S by rewrite dtuple_on_add_D1 Sx. case/(atransP2 Gtr dxt)=> b Gb [xbx tbt2]. - exists b; [rewrite inE Gb; exact/astab1P | exact: val_inj]. + by exists b; [rewrite inE Gb; apply/astab1P | apply: val_inj]. case/setIP=> Gb /astab1P xbx ->{t2}. rewrite n_act_dtuple //; last by rewrite dtuple_on_add_D1 Sx in dxt. apply/astabsP=> y; rewrite !inE -{1}xbx (inj_eq (act_inj _ _)). diff --git a/mathcomp/solvable/sylow.v b/mathcomp/solvable/sylow.v index abaa2a5..e347a74 100644 --- a/mathcomp/solvable/sylow.v +++ b/mathcomp/solvable/sylow.v @@ -1,16 +1,11 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import mathcomp.ssreflect.ssreflect. -From mathcomp.ssreflect -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. -From mathcomp.discrete -Require Import div fintype finset bigop prime. -From mathcomp.fingroup -Require Import fingroup morphism automorphism quotient action gproduct. -From mathcomp.algebra -Require Import ssralg poly ssrnum ssrint rat. -From mathcomp.algebra -Require Import polydiv finalg zmodp matrix mxalgebra vector cyclic. -Require Import commutator pgroup center nilpotent. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq div fintype prime. +From mathcomp +Require Import bigop finset fingroup morphism automorphism quotient action. +From mathcomp +Require Import cyclic gproduct gfunctor commutator pgroup center nilpotent. (******************************************************************************) (* The Sylow theorem and its consequences, including the Frattini argument, *) @@ -68,11 +63,11 @@ Lemma pcore_sub_astab_irr G M : 'O_p(G) \subset 'C_G(M | to). Proof. move=> pM sMR /mingroupP[/andP[ntM nMG] minM]. -have [sGpG nGpG]:= andP (pcore_normal p G). -have sGD := acts_dom nMG; have sGpD := subset_trans sGpG sGD. +have /andP[sGpG nGpG]: 'O_p(G) <| G := gFnormal _ G. +have sGD := acts_dom nMG; have sGpD: 'O_p(G) \subset D := gFsub_trans _ sGD. rewrite subsetI sGpG -gacentC //=; apply/setIidPl; apply: minM (subsetIl _ _). rewrite nontrivial_gacent_pgroup ?pcore_pgroup //=; last first. - by split; rewrite ?(subset_trans sGpG). + by split; rewrite ?gFsub_trans. by apply: subset_trans (acts_subnorm_subgacent sGpD nMG); rewrite subsetI subxx. Qed. @@ -108,7 +103,7 @@ have S_pG P: P \in S -> P \subset G /\ p.-group P. have SmaxN P Q: Q \in S -> Q \subset 'N(P) -> maxp 'N_G(P) Q. rewrite inE => /maxgroupP[/andP[sQG pQ] maxQ] nPQ. apply/maxgroupP; rewrite /psubgroup subsetI sQG nPQ. - by split=> // R; rewrite subsetI -andbA andbCA => /andP[_]; exact: maxQ. + by split=> // R; rewrite subsetI -andbA andbCA => /andP[_]; apply: maxQ. have nrmG P: P \subset G -> P <| 'N_G(P). by move=> sPG; rewrite /normal subsetIr subsetI sPG normG. have sylS P: P \in S -> p.-Sylow('N_G(P)) P. @@ -123,7 +118,7 @@ have{defCS} oG_mod: {in S &, forall P Q, #|oG P| = (Q \in oG P) %[mod p]}. move=> P Q S_P S_Q; have [sQG pQ] := S_pG _ S_Q. have soP_S: oG P \subset S by rewrite acts_sub_orbit. have /pgroup_fix_mod-> //: [acts Q, on oG P | 'JG]. - apply/actsP=> x /(subsetP sQG) Gx R; apply: orbit_transr. + apply/actsP=> x /(subsetP sQG) Gx R; apply: orbit_transl. exact: mem_orbit. rewrite -{1}(setIidPl soP_S) -setIA defCS // (cardsD1 Q) setDE. by rewrite -setIA setICr setI0 cards0 addn0 inE set11 andbT. @@ -206,7 +201,7 @@ Lemma card_Syl P : p.-Sylow(G) P -> #|'Syl_p(G)| = #|G : 'N_G(P)|. Proof. by case: Sylow's_theorem P. Qed. Lemma card_Syl_dvd : #|'Syl_p(G)| %| #|G|. -Proof. by case Sylow_exists => P /card_Syl->; exact: dvdn_indexg. Qed. +Proof. by case Sylow_exists => P /card_Syl->; apply: dvdn_indexg. Qed. Lemma card_Syl_mod : prime p -> #|'Syl_p(G)| %% p = 1%N. Proof. by case Sylow's_theorem. Qed. @@ -262,13 +257,13 @@ Qed. Lemma p2group_abelian P : p.-group P -> logn p #|P| <= 2 -> abelian P. Proof. move=> pP lePp2; pose Z := 'Z(P); have sZP: Z \subset P := center_sub P. -case: (eqVneq Z 1); first by move/(trivg_center_pgroup pP)->; exact: abelian1. +case: (eqVneq Z 1); first by move/(trivg_center_pgroup pP)->; apply: abelian1. case/(pgroup_pdiv (pgroupS sZP pP)) => p_pr _ [k oZ]. apply: cyclic_center_factor_abelian. -case: (eqVneq (P / Z) 1) => [-> |]; first exact: cyclic1. +have [->|] := eqVneq (P / Z) 1; first exact: cyclic1. have pPq := quotient_pgroup 'Z(P) pP; case/(pgroup_pdiv pPq) => _ _ [j oPq]. rewrite prime_cyclic // oPq; case: j oPq lePp2 => //= j. -rewrite card_quotient ?gfunctor.gFnorm //. +rewrite card_quotient ?gFnorm //. by rewrite -(Lagrange sZP) lognM // => ->; rewrite oZ !pfactorK ?addnS. Qed. @@ -343,7 +338,7 @@ Proof. move=> defG nHG coKR; apply/eqP; rewrite eqEcard mulG_subG /= -defG. rewrite !setSI ?mulG_subl ?mulG_subr //=. rewrite coprime_cardMg ?(coKR, coprimeSg (subsetIl _ _), coprime_sym) //=. -pose pi := \pi(K); have piK: pi.-group K by exact: pgroup_pi. +pose pi := \pi(K); have piK: pi.-group K by apply: pgroup_pi. have pi'R: pi^'.-group R by rewrite /pgroup -coprime_pi' /=. have [hallK hallR] := coprime_mulpG_Hall defG piK pi'R. have nsHG: H :&: G <| G by rewrite /normal subsetIr normsI ?normG. @@ -372,7 +367,7 @@ by rewrite ltn_Pdiv // ltnNge -trivg_card_le1. Qed. Lemma pgroup_sol p P : p.-group P -> solvable P. -Proof. by move/pgroup_nil; exact: nilpotent_sol. Qed. +Proof. by move/pgroup_nil; apply: nilpotent_sol. Qed. Lemma small_nil_class G : nil_class G <= 5 -> nilpotent G. Proof. @@ -404,10 +399,10 @@ have nHN: H <| 'N_G(H) by rewrite normal_subnorm. have{maxH} hallH: pi.-Hall('N_G(H)) H. apply: normal_max_pgroup_Hall => //; apply/maxgroupP. rewrite /psubgroup normal_sub // piH; split=> // K. - by rewrite subsetI -andbA andbCA => /andP[_]; exact: maxH. -rewrite /normal sHG; apply/setIidPl; symmetry. -apply: nilpotent_sub_norm; rewrite ?subsetIl ?setIS //=. -by rewrite char_norms // -{1}(normal_Hall_pcore hallH) // pcore_char. + by rewrite subsetI -andbA andbCA => /andP[_ /maxH]. +rewrite /normal sHG; apply/setIidPl/esym. +apply: nilpotent_sub_norm; rewrite ?subsetIl ?setIS //= char_norms //. +by congr (_ \char _): (pcore_char pi 'N_G(H)); apply: normal_Hall_pcore. Qed. Lemma nilpotent_Hall_pcore pi G H : @@ -430,12 +425,12 @@ Qed. Lemma nilpotent_pcoreC pi G : nilpotent G -> 'O_pi(G) \x 'O_pi^'(G) = G. Proof. move=> nilG; have trO: 'O_pi(G) :&: 'O_pi^'(G) = 1. - by apply: coprime_TIg; apply: (@pnat_coprime pi); exact: pcore_pgroup. + by apply: coprime_TIg; apply: (@pnat_coprime pi); apply: pcore_pgroup. rewrite dprodE //. apply/eqP; rewrite eqEcard mul_subG ?pcore_sub // (TI_cardMg trO). by rewrite !(card_Hall (nilpotent_pcore_Hall _ _)) // partnC ?leqnn. rewrite (sameP commG1P trivgP) -trO subsetI commg_subl commg_subr. -by rewrite !(subset_trans (pcore_sub _ _)) ?normal_norm ?pcore_normal. +by rewrite !gFsub_trans ?gFnorm. Qed. Lemma sub_nilpotent_cent2 H K G : @@ -485,7 +480,7 @@ elim: c => // c IHc in gT P def_c pP *; set e := logn p _. have nilP := pgroup_nil pP; have sZP := center_sub P. have [e_le2 | e_gt2] := leqP e 2. by rewrite -def_c leq_max nil_class1 (p2group_abelian pP). -have pPq: p.-group (P / 'Z(P)) by exact: quotient_pgroup. +have pPq: p.-group (P / 'Z(P)) by apply: quotient_pgroup. rewrite -(subnKC e_gt2) ltnS (leq_trans (IHc _ _ _ pPq)) //. by rewrite nil_class_quotient_center ?def_c. rewrite geq_max /= -add1n -leq_subLR -subn1 -subnDA -subSS leq_sub2r //. @@ -534,7 +529,7 @@ have /cyclicP[x' def_p']: cyclic 'O_p^'(G). by have:= pcore_pgroup p^' G; rewrite def_p' /pgroup p'natE ?pG. apply/cyclicP; exists (x * x'); rewrite -{}defG def_p def_p' cycleM //. by red; rewrite -(centsP Cpp') // (def_p, def_p') cycle_id. -rewrite /order -def_p -def_p' (@pnat_coprime p) //; exact: pcore_pgroup. +by rewrite /order -def_p -def_p' (@pnat_coprime p) //; apply: pcore_pgroup. Qed. End Zgroups. @@ -582,8 +577,8 @@ have{pE} pE: {in E &, forall x1 x2, p.-group <<[set x1; x2]>>}. rewrite -(mulgKV y1 y2) conjgM -2!conjg_set1 -conjUg genJ pgroupJ. by rewrite pE // groupMl ?groupV. have sEG: <<E>> \subset G by rewrite gen_subG class_subG. -have nEG: G \subset 'N(E) by exact: class_norm. -have Ex: x \in E by exact: class_refl. +have nEG: G \subset 'N(E) by apply: class_norm. +have Ex: x \in E by apply: class_refl. have [P Px sylP]: exists2 P : {group gT}, x \in P & p.-Sylow(<<E>>) P. have sxxE: <<[set x; x]>> \subset <<E>> by rewrite genS // setUid sub1set. have{sxxE} [P sylP sxxP] := Sylow_superset sxxE (pE _ _ Ex Ex). @@ -600,14 +595,14 @@ have{Ex Px}: P_D [set x]. case/(@maxset_exists _ P_D)=> D /maxsetP[]; rewrite {P_yD P_D}/=. rewrite subsetI sub1set -andbA => /and3P[sDP sDE /existsP[y0]]. set B := _ |: D; rewrite inE -andbA => /and3P[Py0 Ey0 pB] maxD Dx. -have sDgE: D \subset <<E>> by exact: sub_gen. -have sDG: D \subset G by exact: subset_trans sEG. +have sDgE: D \subset <<E>> by apply: sub_gen. +have sDG: D \subset G by apply: subset_trans sEG. have sBE: B \subset E by rewrite subUset sub1set Ey0. -have sBG: <<B>> \subset G by exact: subset_trans (genS _) sEG. +have sBG: <<B>> \subset G by apply: subset_trans (genS _) sEG. have sDB: D \subset B by rewrite subsetUr. have defD: D :=: P :&: <<B>> :&: E. apply/eqP; rewrite eqEsubset ?subsetI sDP sDE sub_gen //=. - apply/setUidPl; apply: maxD; last exact: subsetUl. + apply/setUidPl; apply: maxD; last apply: subsetUl. rewrite subUset subsetI sDP sDE setIAC subsetIl. apply/existsP; exists y0; rewrite inE Py0 Ey0 /= setUA -/B. by rewrite -[<<_>>]joing_idl joingE setKI genGid. @@ -645,7 +640,7 @@ have [y1 Ny1 Py1]: exists2 y1, y1 \in 'N_E(D) & y1 \notin P. have [y2 Ny2 Dy2]: exists2 y2, y2 \in 'N_(P :&: E)(D) & y2 \notin D. case sNN: ('N_P('N_P(D)) \subset 'N_P(D)). have [z /= Ez sEzP] := Sylow_Jsub sylP (genS sBE) pB. - have Gz: z \in G by exact: subsetP Ez. + have Gz: z \in G by apply: subsetP Ez. have /subsetPn[y Bzy Dy]: ~~ (B :^ z \subset D). apply/negP; move/subset_leq_card; rewrite cardJg cardsU1. by rewrite {1}defD 2!inE (negPf Py0) ltnn. @@ -657,7 +652,7 @@ have [y2 Ny2 Dy2]: exists2 y2, y2 \in 'N_(P :&: E)(D) & y2 \notin D. case/subsetPn=> y Dzy Dy; exists y => //; apply: subsetP Dzy. rewrite -setIA setICA subsetI sub_conjg (normsP nEG) ?groupV //. by rewrite sDE -(normP NNz); rewrite conjSg subsetI sDP. - apply: subsetP Pz; exact: (subset_trans (pHall_sub sylP)). + by apply: subsetP Pz; apply: (subset_trans (pHall_sub sylP)). suff{Dy2} Dy2D: y2 |: D = D by rewrite -Dy2D setU11 in Dy2. apply: maxD; last by rewrite subsetUr. case/setIP: Ny2 => PEy2 Ny2; case/setIP: Ny1 => Ey1 Ny1. diff --git a/mathcomp/ssreflect/Make b/mathcomp/ssreflect/Make index 6c76897..39f593b 100644 --- a/mathcomp/ssreflect/Make +++ b/mathcomp/ssreflect/Make @@ -1,4 +1,4 @@ -all.v +all_ssreflect.v eqtype.v seq.v ssrbool.v diff --git a/mathcomp/ssreflect/Makefile b/mathcomp/ssreflect/Makefile index 76e2984..b97c47a 100644 --- a/mathcomp/ssreflect/Makefile +++ b/mathcomp/ssreflect/Makefile @@ -17,6 +17,12 @@ else V=$(BRANCH_coq) endif +ifeq "$V" "v8.4" +COQDEP=../../etc/utils/ssrcoqdep +else +COQDEP=$(COQBIN)/coqdep +endif + OLD_MAKEFLAGS:=$(MAKEFLAGS) MAKEFLAGS+=-B @@ -27,7 +33,7 @@ MAKEFLAGS+=-B # Override COQDEP to find only the "right" copy .ml files $(H)MAKEFLAGS=$(OLD_MAKEFLAGS) $(MAKE) --no-print-directory \ -f Makefile.coq $* \ - COQDEP='$(COQBIN)/coqdep -exclude-dir plugin -c' + COQDEP='$(COQDEP) -exclude-dir plugin -c' define coqmakefile (echo "Generating Makefile.coq for Coq $(V) with COQBIN=$(COQBIN)";\ diff --git a/mathcomp/ssreflect/all.v b/mathcomp/ssreflect/all_ssreflect.v index ce3e470..ce3e470 100644 --- a/mathcomp/ssreflect/all.v +++ b/mathcomp/ssreflect/all_ssreflect.v diff --git a/mathcomp/ssreflect/eqtype.v b/mathcomp/ssreflect/eqtype.v index 01b0f5e..1ca9891 100644 --- a/mathcomp/ssreflect/eqtype.v +++ b/mathcomp/ssreflect/eqtype.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool. (******************************************************************************) (* This file defines two "base" combinatorial interfaces: *) @@ -240,7 +242,7 @@ by case: y /; case: _ / (proj x _). Qed. Corollary eq_axiomK (T : eqType) (x : T) : all_equal_to (erefl x). -Proof. move=> eq_x_x; exact: eq_irrelevance. Qed. +Proof. by move=> eq_x_x; apply: eq_irrelevance. Qed. (* We use the module system to circumvent a silly limitation that *) (* forbids using the same constant to coerce to different targets. *) @@ -312,15 +314,15 @@ Definition predU1 (a1 : T) p := SimplPred (xpredU1 a1 p). Definition predC1 (a1 : T) := SimplPred (xpredC1 a1). Definition predD1 p (a1 : T) := SimplPred (xpredD1 p a1). -Lemma pred1E : pred1 =2 eq_op. Proof. move=> x y; exact: eq_sym. Qed. +Lemma pred1E : pred1 =2 eq_op. Proof. by move=> x y; apply: eq_sym. Qed. Variables (T2 : eqType) (x y : T) (z u : T2) (b : bool). Lemma predU1P : reflect (x = y \/ b) ((x == y) || b). -Proof. apply: (iffP orP) => [] []; by [right | move/eqP; left]. Qed. +Proof. by apply: (iffP orP); do [case=> [/eqP|]; [left | right]]. Qed. Lemma pred2P : reflect (x = y \/ z = u) ((x == y) || (z == u)). -Proof. by apply: (iffP orP) => [] [] /eqP; by [left | right]. Qed. +Proof. by apply: (iffP orP); do [case=> /eqP; [left | right]]. Qed. Lemma predD1P : reflect (x <> y /\ b) ((x != y) && b). Proof. by apply: (iffP andP)=> [] [] // /eqP. Qed. @@ -355,24 +357,24 @@ Section Exo. Variables (aT rT : eqType) (D : pred aT) (f : aT -> rT) (g : rT -> aT). Lemma inj_eq : injective f -> forall x y, (f x == f y) = (x == y). -Proof. by move=> inj_f x y; apply/eqP/eqP=> [|-> //]; exact: inj_f. Qed. +Proof. by move=> inj_f x y; apply/eqP/eqP=> [|-> //]; apply: inj_f. Qed. Lemma can_eq : cancel f g -> forall x y, (f x == f y) = (x == y). -Proof. move/can_inj; exact: inj_eq. Qed. +Proof. by move/can_inj; apply: inj_eq. Qed. Lemma bij_eq : bijective f -> forall x y, (f x == f y) = (x == y). -Proof. move/bij_inj; apply: inj_eq. Qed. +Proof. by move/bij_inj; apply: inj_eq. Qed. Lemma can2_eq : cancel f g -> cancel g f -> forall x y, (f x == y) = (x == g y). -Proof. by move=> fK gK x y; rewrite -{1}[y]gK; exact: can_eq. Qed. +Proof. by move=> fK gK x y; rewrite -{1}[y]gK; apply: can_eq. Qed. Lemma inj_in_eq : {in D &, injective f} -> {in D &, forall x y, (f x == f y) = (x == y)}. -Proof. by move=> inj_f x y Dx Dy; apply/eqP/eqP=> [|-> //]; exact: inj_f. Qed. +Proof. by move=> inj_f x y Dx Dy; apply/eqP/eqP=> [|-> //]; apply: inj_f. Qed. Lemma can_in_eq : {in D, cancel f g} -> {in D &, forall x y, (f x == f y) = (x == y)}. -Proof. by move/can_in_inj; exact: inj_in_eq. Qed. +Proof. by move/can_in_inj; apply: inj_in_eq. Qed. End Exo. @@ -383,7 +385,7 @@ Variable T : eqType. Definition frel f := [rel x y : T | f x == y]. Lemma inv_eq f : involutive f -> forall x y : T, (f x == y) = (x == f y). -Proof. by move=> fK; exact: can2_eq. Qed. +Proof. by move=> fK; apply: can2_eq. Qed. Lemma eq_frel f f' : f =1 f' -> frel f =2 frel f'. Proof. by move=> eq_f x y; rewrite /= eq_f. Qed. @@ -404,7 +406,7 @@ Lemma invariant_comp : subpred (invariant f k) (invariant f (h \o k)). Proof. by move=> x eq_kfx; rewrite /= (eqP eq_kfx). Qed. Lemma invariant_inj : injective h -> invariant f (h \o k) =1 invariant f k. -Proof. move=> inj_h x; exact: (inj_eq inj_h). Qed. +Proof. by move=> inj_h x; apply: (inj_eq inj_h). Qed. End EqFun. @@ -467,7 +469,7 @@ Hypothesis Hcompare : comparable. Definition compareb x y : bool := Hcompare x y. Lemma compareP : Equality.axiom compareb. -Proof. by move=> x y; exact: sumboolP. Qed. +Proof. by move=> x y; apply: sumboolP. Qed. Definition comparableClass := EqMixin compareP. @@ -517,14 +519,13 @@ CoInductive insub_spec x : option sT -> Type := Lemma insubP x : insub_spec x (insub x). Proof. -by rewrite /insub; case: {-}_ / idP; [left; rewrite ?SubK | right; exact/negP]. +by rewrite /insub; case: {-}_ / idP; [left; rewrite ?SubK | right; apply/negP]. Qed. Lemma insubT x Px : insub x = Some (Sub x Px). Proof. -case: insubP; last by case/negP. -case/SubP=> y Py _ def_x; rewrite -def_x SubK in Px *. -congr (Some (Sub _ _)); exact: bool_irrelevance. +do [case: insubP => [/SubP[y Py] _ <- | /negP// ]; rewrite SubK] in Px *. +by rewrite (bool_irrelevance Px Py). Qed. Lemma insubF x : P x = false -> insub x = None. @@ -543,7 +544,7 @@ Lemma valP (u : sT) : P (val u). Proof. by case/SubP: u => x Px; rewrite SubK. Qed. Lemma valK : pcancel (@val _) insub. -Proof. case/SubP=> x Px; rewrite SubK; exact: insubT. Qed. +Proof. by case/SubP=> x Px; rewrite SubK; apply: insubT. Qed. Lemma val_inj : injective (@val sT). Proof. exact: pcan_inj valK. Qed. @@ -624,7 +625,7 @@ Implicit Arguments innew [T nT]. Prenex Implicits innew. Lemma innew_val T nT : cancel val (@innew T nT). -Proof. by move=> u; apply: val_inj; exact: SubK. Qed. +Proof. by move=> u; apply: val_inj; apply: SubK. Qed. (* Prenex Implicits and renaming. *) Notation sval := (@proj1_sig _ _). @@ -680,7 +681,7 @@ Section TransferEqType. Variables (T : Type) (eT : eqType) (f : T -> eT). Lemma inj_eqAxiom : injective f -> Equality.axiom (fun x y => f x == f y). -Proof. by move=> f_inj x y; apply: (iffP eqP) => [|-> //]; exact: f_inj. Qed. +Proof. by move=> f_inj x y; apply: (iffP eqP) => [|-> //]; apply: f_inj. Qed. Definition InjEqMixin f_inj := EqMixin (inj_eqAxiom f_inj). diff --git a/mathcomp/ssreflect/opam b/mathcomp/ssreflect/opam index bef7ebd..4afccfb 100644 --- a/mathcomp/ssreflect/opam +++ b/mathcomp/ssreflect/opam @@ -9,4 +9,4 @@ license: "CeCILL-B" build: [ make "-j" "%{jobs}%" ] install: [ make "install" ] remove: [ "sh" "-c" "rm -rf '%{lib}%/coq/user-contrib/mathcomp/ssreflect'" ] -depends: [ "coq" { ( >= "8.4pl4" & < "8.5" ) | ( = "8.5beta1" ) } ] +depends: [ "coq" { ( >= "8.4pl4" & < "8.5" ) | ( = "8.5.dev" ) } ] diff --git a/mathcomp/ssreflect/plugin/v8.4/ssreflect.ml4 b/mathcomp/ssreflect/plugin/v8.4/ssreflect.ml4 index 6e2be9e..e842823 100644 --- a/mathcomp/ssreflect/plugin/v8.4/ssreflect.ml4 +++ b/mathcomp/ssreflect/plugin/v8.4/ssreflect.ml4 @@ -6055,7 +6055,7 @@ GEXTEND Gram [[ IDENT "From"; ns = Constr.global; IDENT "Require"; export = ssr_export_token; qidl = LIST1 Constr.global -> - let qidl = List.map (join_reference ns) qidl in +(* let qidl = List.map (join_reference ns) qidl in *) Vernacexpr.VernacRequire (export, None, qidl) ]]; END diff --git a/mathcomp/ssreflect/plugin/v8.5/ssreflect.ml4 b/mathcomp/ssreflect/plugin/v8.5/ssreflect.ml4 index d9a2581..1b4afb2 100644 --- a/mathcomp/ssreflect/plugin/v8.5/ssreflect.ml4 +++ b/mathcomp/ssreflect/plugin/v8.5/ssreflect.ml4 @@ -257,7 +257,7 @@ let mkRLambda n s t = GLambda (dummy_loc, n, Explicit, s, t) (** Constructors for constr *) let pf_e_type_of gl t = let sigma, env, it = project gl, pf_env gl, sig_it gl in - let sigma, ty = Typing.e_type_of env sigma t in + let sigma, ty = Typing.type_of env sigma t in re_sig it sigma, ty let mkAppRed f c = match kind_of_term f with @@ -302,6 +302,8 @@ let loc_ofCG = function let mk_term k c = k, (mkRHole, Some c) let mk_lterm c = mk_term ' ' c +let pf_type_of gl t = let sigma, ty = pf_type_of gl t in re_sig (sig_it gl) sigma, ty + let map_fold_constr g f ctx acc cstr = let array_f ctx acc x = let x, acc = f ctx acc x in acc, x in match kind_of_term cstr with @@ -2081,7 +2083,8 @@ let abs_wgen keep_let ist f gen (gl,args,c) = with NoMatch -> redex_of_pattern env cp, c in evar_closed t p; let ut = red_product_skip_id env sigma t in - pf_merge_uc ucst gl, args, mkLetIn(Name (f x), ut, pf_type_of gl t, c) + let gl, ty = pf_type_of gl t in + pf_merge_uc ucst gl, args, mkLetIn(Name (f x), ut, ty, c) | _, Some ((x, _), Some p) -> let x = hoi_id x in let cp = interp_cpattern ist gl p None in @@ -2089,7 +2092,8 @@ let abs_wgen keep_let ist f gen (gl,args,c) = try fill_occ_pattern ~raise_NoMatch:true env sigma c cp None 1 with NoMatch -> redex_of_pattern env cp, c in evar_closed t p; - pf_merge_uc ucst gl, t :: args, mkProd(Name (f x), pf_type_of gl t, c) + let gl, ty = pf_type_of gl t in + pf_merge_uc ucst gl, t :: args, mkProd(Name (f x), ty, c) | _ -> gl, args, c let clr_of_wgen gen clrs = match gen with @@ -2249,9 +2253,9 @@ ARGUMENT EXTEND ssrocc TYPED AS (bool * int list) option PRINTED BY pr_ssrocc END let pf_mkprod gl c ?(name=constr_name c) cl = - let t = pf_type_of gl c in - if name <> Anonymous || noccurn 1 cl then mkProd (name, t, cl) else - mkProd (Name (pf_type_id gl t), t, cl) + let gl, t = pf_type_of gl c in + if name <> Anonymous || noccurn 1 cl then gl, mkProd (name, t, cl) else + gl, mkProd (Name (pf_type_id gl t), t, cl) let pf_abs_prod name gl c cl = pf_mkprod gl c ~name (subst_term c cl) @@ -2457,7 +2461,8 @@ let with_view ist si env gl0 c name cl prune = let n, c', _, ucst = pf_abs_evars gl0 (sigma, c') in let c' = if not prune then c' else pf_abs_cterm gl0 n c' in let gl0 = pf_merge_uc ucst gl0 in - pf_abs_prod name gl0 c' (prod_applist cl [c]), c', pf_merge_uc_of sigma gl0 + let gl0, ap = pf_abs_prod name gl0 c' (prod_applist cl [c]) in + ap, c', pf_merge_uc_of sigma gl0 in loop let pf_with_view ist gl (prune, view) cl c = @@ -2767,11 +2772,13 @@ let injectl2rtac c = match kind_of_term c with tclTHENLIST [havetac id c; injectidl2rtac id (mkVar id, NoBindings); clear [id]] let is_injection_case c gl = - let (mind,_), _ = pf_reduce_to_quantified_ind gl (pf_type_of gl c) in + let gl, cty = pf_type_of gl c in + let (mind,_), _ = pf_reduce_to_quantified_ind gl cty in eq_gr (IndRef mind) (build_coq_eq ()) let perform_injection c gl = - let mind, t = pf_reduce_to_quantified_ind gl (pf_type_of gl c) in + let gl, cty = pf_type_of gl c in + let mind, t = pf_reduce_to_quantified_ind gl cty in let dc, eqt = decompose_prod t in if dc = [] then injectl2rtac c gl else if not (closed0 eqt) then @@ -2875,7 +2882,7 @@ let ssrmkabs id gl = Evarutil.new_evar (Environ.push_rel rd env) sigma concl in pp(lazy(pr_constr concl)); let term = mkApp (mkLambda(Name id,abstract_ty,kont) ,[|abstract_proof|]) in - let sigma, _ = Typing.e_type_of env sigma term in + let sigma, _ = Typing.type_of env sigma term in sigma, term in Proofview.V82.of_tactic (Proofview.tclTHEN @@ -3417,13 +3424,14 @@ let pf_interp_gen_aux ist gl to_ind ((oclr, occ), t) = errorstrm (str "@ can be used with variables only") else match pf_get_hyp gl (destVar c) with | _, None, _ -> errorstrm (str "@ can be used with let-ins only") - | name, Some bo, ty -> true, pat, mkLetIn (Name name,bo,ty,cl),c,clr,ucst - else false, pat, pf_mkprod gl c cl, c, clr,ucst + | name, Some bo, ty -> true, pat, mkLetIn (Name name,bo,ty,cl),c,clr,ucst,gl + else let gl, ccl = pf_mkprod gl c cl in false, pat, ccl, c, clr,ucst,gl else if to_ind && occ = None then let nv, p, _, ucst' = pf_abs_evars gl (fst pat, c) in let ucst = Evd.union_evar_universe_context ucst ucst' in if nv = 0 then anomaly "occur_existential but no evars" else - false, pat, mkProd (constr_name c, pf_type_of gl p, pf_concl gl), p, clr,ucst + let gl,pty = pf_type_of gl p in + false, pat, mkProd (constr_name c, pty, pf_concl gl), p, clr,ucst, gl else loc_error (loc_of_cpattern t) "generalized term didn't match" let genclrtac cl cs clr = @@ -3446,7 +3454,7 @@ let genclrtac cl cs clr = let gentac ist gen gl = (* pp(lazy(str"sigma@gentac=" ++ pr_evar_map None (project gl))); *) - let conv, _, cl, c, clr, ucst = pf_interp_gen_aux ist gl false gen in + let conv, _, cl, c, clr, ucst,gl = pf_interp_gen_aux ist gl false gen in pp(lazy(str"c@gentac=" ++ pr_constr c)); let gl = pf_merge_uc ucst gl in if conv @@ -3454,7 +3462,7 @@ let gentac ist gen gl = else genclrtac cl [c] clr gl let pf_interp_gen ist gl to_ind gen = - let _, _, a, b, c, ucst = pf_interp_gen_aux ist gl to_ind gen in + let _, _, a, b, c, ucst,gl = pf_interp_gen_aux ist gl to_ind gen in a, b ,c, pf_merge_uc ucst gl (** Generalization (discharge) sequence *) @@ -3612,7 +3620,8 @@ let pushcaseeqtac cl gl = let dc, cl2 = decompose_lam_n n cl1 in let _, t = List.nth dc (n - 1) in let cl3, eqc, gl = mkEq R2L cl2 args.(0) t n gl in - let prot, gl = mkProt (pf_type_of gl cl) cl3 gl in + let gl, clty = pf_type_of gl cl in + let prot, gl = mkProt clty cl3 gl in let cl4 = mkApp (compose_lam dc prot, args) in let gl, _ = pf_e_type_of gl cl4 in tclTHEN (apply_type cl4 [eqc]) @@ -3687,7 +3696,7 @@ END let viewmovetac_aux clear name_ref (_, vl as v) _ gen ist gl = let cl, c, clr, gl, gen_pat = - let _, gen_pat, a, b, c, ucst = pf_interp_gen_aux ist gl false gen in + let _, gen_pat, a, b, c, ucst,gl = pf_interp_gen_aux ist gl false gen in a, b ,c, pf_merge_uc ucst gl, gen_pat in let cl, c, gl = if vl = [] then cl, c, gl else pf_with_view ist gl v cl c in let clr = if clear then clr else [] in @@ -3825,7 +3834,7 @@ let refine_with ?(first_goes_last=false) ?beta ?(with_evars=true) oc gl = let rec mkRels = function 1 -> [] | n -> mkRel n :: mkRels (n-1) in let uct = Evd.evar_universe_context (fst oc) in let n, oc = pf_abs_evars_pirrel gl oc in - let gl = pf_merge_uc uct gl in + let gl = pf_unsafe_merge_uc uct gl in let oc = if not first_goes_last || n <= 1 then oc else let l, c = decompose_lam oc in if not (List.for_all_i (fun i (_,t) -> closedn ~-i t) (1-n) l) then oc else @@ -3921,7 +3930,7 @@ let ssrelim ?(is_case=false) ?ist deps what ?elim eqid ipats gl = let elimty = Reductionops.whd_betadeltaiota env (project gl) elimty in None, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl | None -> - let c = Option.get oc in let c_ty = pf_type_of gl c in + let c = Option.get oc in let gl, c_ty = pf_type_of gl c in let ((kn, i) as ind, _ as indu), unfolded_c_ty = pf_reduce_to_quantified_ind gl c_ty in let sort = elimination_sort_of_goal gl in @@ -3932,7 +3941,7 @@ let ssrelim ?(is_case=false) ?ist deps what ?elim eqid ipats gl = else pf_eapply (fun env sigma -> Indrec.build_case_analysis_scheme env sigma indu true) gl sort in - let elimty = pf_type_of gl elim in + let gl, elimty = pf_type_of gl elim in let pred_id,n_elim_args,is_rec,elim_is_dep,n_pred_args = analyze_eliminator elimty env (project gl) in let rctx = fst (decompose_prod_assum unfolded_c_ty) in @@ -3965,7 +3974,7 @@ let ssrelim ?(is_case=false) ?ist deps what ?elim eqid ipats gl = let res = if elim_is_dep then None else let arg = List.assoc (n_elim_args - 1) elim_args in - let arg_ty = pf_type_of gl arg in + let gl, arg_ty = pf_type_of gl arg in match saturate_until gl c c_ty (fun c c_ty gl -> pf_unify_HO (pf_unify_HO gl c_ty arg_ty) arg c) with | Some (c, _, _, gl) -> Some (false, gl) @@ -3974,7 +3983,7 @@ let ssrelim ?(is_case=false) ?ist deps what ?elim eqid ipats gl = | Some x -> x | None -> let inf_arg = List.hd inf_deps_r in - let inf_arg_ty = pf_type_of gl inf_arg in + let gl, inf_arg_ty = pf_type_of gl inf_arg in match saturate_until gl c c_ty (fun _ c_ty gl -> pf_unify_HO gl c_ty inf_arg_ty) with | Some (c, _, _,gl) -> true, gl @@ -3983,7 +3992,7 @@ let ssrelim ?(is_case=false) ?ist deps what ?elim eqid ipats gl = spc()++pr_constr c++spc()++str"or to unify it's type with"++ pr_constr inf_arg_ty) in pp(lazy(str"elim_is_dep= " ++ bool elim_is_dep)); - let predty = pf_type_of gl pred in + let gl, predty = pf_type_of gl pred in (* Patterns for the inductive types indexes to be bound in pred are computed * looking at the ones provided by the user and the inferred ones looking at * the type of the elimination principle *) @@ -4060,7 +4069,7 @@ let ssrelim ?(is_case=false) ?ist deps what ?elim eqid ipats gl = | Some (IpatId _) when not is_rec -> let k = List.length deps in let c = fire_subst gl (List.assoc (n_elim_args - k - 1) elim_args) in - let t = pf_type_of gl c in + let gl, t = pf_type_of gl c in let gen_eq_tac, gl = let refl = mkApp (eq, [|t; c; c|]) in let new_concl = mkArrow refl (lift 1 (pf_concl orig_gl)) in @@ -4078,7 +4087,8 @@ let ssrelim ?(is_case=false) ?ist deps what ?elim eqid ipats gl = let concl = List.fold_left mk_lam concl pred_rctx in let gl, concl = if eqid <> None && is_rec then - let concl, gl = mkProt (pf_type_of gl concl) concl gl in + let gl, concls = pf_type_of gl concl in + let concl, gl = mkProt concls concl gl in let gl, _ = pf_e_type_of gl concl in gl, concl else gl, concl in @@ -4145,7 +4155,7 @@ let ssrelim ?(is_case=false) ?ist deps what ?elim eqid ipats gl = let case = args.(Array.length args-1) in if not(closed0 case) then tclTHEN (introstac [IpatAnon]) gen_eq_tac gl else - let case_ty = pf_type_of gl case in + let gl, case_ty = pf_type_of gl case in let refl = mkApp (eq, [|lift 1 case_ty; mkRel 1; lift 1 case|]) in let new_concl = fire_subst gl (mkProd (Name (name gl), case_ty, mkArrow refl (lift 2 concl))) in @@ -4772,7 +4782,7 @@ let pirrel_rewrite pred rdx rdx_ty new_rdx dir (sigma, c) c_ty gl = let proof = mkApp (elim, [| rdx_ty; new_rdx; pred; p; rdx; c |]) in (* We check the proof is well typed *) let sigma, proof_ty = - try Typing.e_type_of env sigma proof with _ -> raise PRtype_error in + try Typing.type_of env sigma proof with _ -> raise PRtype_error in pp(lazy(str"pirrel_rewrite proof term of type: " ++ pr_constr proof_ty)); try refine_with ~first_goes_last:(not !ssroldreworder) ~with_evars:false (sigma, proof) gl @@ -4809,14 +4819,14 @@ let rwcltac cl rdx dir sr gl = let n, r_n,_, ucst = pf_abs_evars gl sr in let r_n' = pf_abs_cterm gl n r_n in let r' = subst_var pattern_id r_n' in - let gl = pf_merge_uc ucst gl in + let gl = pf_unsafe_merge_uc ucst gl in let rdxt = Retyping.get_type_of (pf_env gl) (fst sr) rdx in (* pp(lazy(str"sigma@rwcltac=" ++ pr_evar_map None (fst sr))); *) pp(lazy(str"r@rwcltac=" ++ pr_constr (snd sr))); let cvtac, rwtac, gl = if closed0 r' then let env, sigma, c, c_eq = pf_env gl, fst sr, snd sr, build_coq_eq () in - let sigma, c_ty = Typing.e_type_of env sigma c in + let sigma, c_ty = Typing.type_of env sigma c in pp(lazy(str"c_ty@rwcltac=" ++ pr_constr c_ty)); match kind_of_type (Reductionops.whd_betadeltaiota env sigma c_ty) with | AtomicType(e, a) when is_ind_ref e c_eq -> @@ -4824,7 +4834,7 @@ let rwcltac cl rdx dir sr gl = pirrel_rewrite cl rdx rdxt new_rdx dir (sigma,c) c_ty, tclIDTAC, gl | _ -> let cl' = mkApp (mkNamedLambda pattern_id rdxt cl, [|rdx|]) in - let sigma, _ = Typing.e_type_of env sigma cl' in + let sigma, _ = Typing.type_of env sigma cl' in let gl = pf_merge_uc_of sigma gl in Proofview.V82.of_tactic (convert_concl cl'), rewritetac dir r', gl else @@ -5534,8 +5544,8 @@ let ssrsettac ist id ((_, (pat, pty)), (_, occ)) gl = if occur_existential c then errorstrm(str"The pattern"++spc()++ pr_constr_pat c++spc()++str"did not match and has holes."++spc()++ str"Did you mean pose?") else - let c, cty = match kind_of_term c with - | Cast(t, DEFAULTcast, ty) -> t, ty + let c, (gl, cty) = match kind_of_term c with + | Cast(t, DEFAULTcast, ty) -> t, (gl, ty) | _ -> c, pf_type_of gl c in let cl' = mkLetIn (Name id, c, cty, cl) in let gl = pf_merge_uc ucst gl in @@ -5604,7 +5614,7 @@ let occur_existential_or_casted_meta c = in try occrec c; false with Not_found -> true let examine_abstract id gl = - let tid = pf_type_of gl id in + let gl, tid = pf_type_of gl id in let abstract, gl = pf_mkSsrConst "abstract" gl in if not (isApp tid) || not (Term.eq_constr (fst(destApp tid)) abstract) then errorstrm(strbrk"not an abstract constant: "++pr_constr id); @@ -5640,7 +5650,8 @@ let unfold cl = let havegentac ist t gl = let sigma, c, ucst = pf_abs_ssrterm ist gl t in let gl = pf_merge_uc ucst gl in - apply_type (mkArrow (pf_type_of gl c) (pf_concl gl)) [c] gl + let gl, cty = pf_type_of gl c in + apply_type (mkArrow cty (pf_concl gl)) [c] gl let havetac ist (transp,((((clr, pats), binders), simpl), (((fk, _), t), hint))) @@ -5695,7 +5706,7 @@ let havetac ist let cty = combineCG cty hole (mkCArrow loc) mkRArrow in let _, t, uc = interp gl false (combineCG ct cty (mkCCast loc) mkRCast) in let gl = pf_merge_uc uc gl in - let ty = pf_type_of gl t in + let gl, ty = pf_type_of gl t in let ctx, _ = decompose_prod_n 1 ty in let assert_is_conv gl = try Proofview.V82.of_tactic (convert_concl (compose_prod ctx concl)) gl diff --git a/mathcomp/ssreflect/plugin/v8.5/ssrmatching.ml4 b/mathcomp/ssreflect/plugin/v8.5/ssrmatching.ml4 index cb791ca..effd193 100644 --- a/mathcomp/ssreflect/plugin/v8.5/ssrmatching.ml4 +++ b/mathcomp/ssreflect/plugin/v8.5/ssrmatching.ml4 @@ -162,6 +162,8 @@ let loc_ofCG = function let mk_term k c = k, (mkRHole, Some c) let mk_lterm = mk_term ' ' +let pf_type_of gl t = let sigma, ty = pf_type_of gl t in re_sig (sig_it gl) sigma, ty + (* }}} *) (** Profiling {{{ *************************************************************) @@ -357,7 +359,7 @@ let pf_unif_HO gl sigma pt p c = let unify_HO env sigma0 t1 t2 = let sigma = unif_HO env sigma0 t1 t2 in let sigma, uc, _ = unif_end env sigma0 sigma t2 (fun _ -> true) in - Evd.merge_universe_context sigma uc + Evd.set_universe_context sigma uc let pf_unify_HO gl t1 t2 = let env, sigma0, si = pf_env gl, project gl, sig_it gl in @@ -990,7 +992,8 @@ let interp_pattern ist gl red redty = aux [] (Evarutil.nf_evar sigma rp) in let sigma = List.fold_left (fun sigma e -> - if Evd.is_defined sigma e then sigma else (* clear may be recursiv *) + if Evd.is_defined sigma e then sigma else (* clear may be recursive *) + if Option.is_empty !to_clean then sigma else let name = Option.get !to_clean in pp(lazy(pr_id name)); try snd(Logic.prim_refiner (Proof_type.Thin [name]) sigma e) @@ -1218,13 +1221,16 @@ END let pf_merge_uc uc gl = re_sig (sig_it gl) (Evd.merge_universe_context (project gl) uc) +let pf_unsafe_merge_uc uc gl = + re_sig (sig_it gl) (Evd.set_universe_context (project gl) uc) + let ssrpatterntac ist arg gl = let pat = interp_rpattern ist gl arg in let sigma0 = project gl in let concl0 = pf_concl gl in let (t, uc), concl_x = fill_occ_pattern (Global.env()) sigma0 concl0 pat noindex 1 in - let tty = pf_type_of gl t in + let gl, tty = pf_type_of gl t in let concl = mkLetIn (Name (id_of_string "toto"), t, tty, concl_x) in Proofview.V82.of_tactic (convert_concl concl DEFAULTcast) gl diff --git a/mathcomp/ssreflect/plugin/v8.5/ssrmatching.mli b/mathcomp/ssreflect/plugin/v8.5/ssrmatching.mli index 58ee875..0976be7 100644 --- a/mathcomp/ssreflect/plugin/v8.5/ssrmatching.mli +++ b/mathcomp/ssreflect/plugin/v8.5/ssrmatching.mli @@ -227,6 +227,7 @@ val cpattern_of_id : Names.variable -> cpattern val cpattern_of_id : Names.variable -> cpattern val pr_constr_pat : constr -> Pp.std_ppcmds val pf_merge_uc : Evd.evar_universe_context -> goal Evd.sigma -> goal Evd.sigma +val pf_unsafe_merge_uc : Evd.evar_universe_context -> goal Evd.sigma -> goal Evd.sigma (* One can also "Set SsrMatchingDebug" from a .v *) val debug : bool -> unit diff --git a/mathcomp/ssreflect/plugin/v8.5beta2/ssreflect.ml4 b/mathcomp/ssreflect/plugin/v8.5beta2/ssreflect.ml4 new file mode 100644 index 0000000..f598c21 --- /dev/null +++ b/mathcomp/ssreflect/plugin/v8.5beta2/ssreflect.ml4 @@ -0,0 +1,6164 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) + +(* This line is read by the Makefile's dist target: do not remove. *) +DECLARE PLUGIN "ssreflect" +let ssrversion = "1.5";; +let ssrAstVersion = 1;; +let () = Mltop.add_known_plugin (fun () -> + if Flags.is_verbose () && not !Flags.batch_mode then begin + Printf.printf "\nSmall Scale Reflection version %s loaded.\n" ssrversion; + Printf.printf "Copyright 2005-2014 Microsoft Corporation and INRIA.\n"; + Printf.printf "Distributed under the terms of the CeCILL-B license.\n\n" + end; + (* Disable any semantics associated with bullets *) + Goptions.set_string_option_value_gen + (Some false) ["Bullet";"Behavior"] "None") + "ssreflect" +;; + +(* Defining grammar rules with "xx" in it automatically declares keywords too, + * we thus save the lexer to restore it at the end of the file *) +let frozen_lexer = Lexer.freeze () ;; + +(*i camlp4use: "pa_extend.cmo" i*) +(*i camlp4deps: "grammar/grammar.cma" i*) + +open Names +open Pp +open Pcoq +open Genarg +open Stdarg +open Constrarg +open Term +open Vars +open Context +open Topconstr +open Libnames +open Tactics +open Tacticals +open Termops +open Namegen +open Recordops +open Tacmach +open Coqlib +open Glob_term +open Util +open Evd +open Extend +open Goptions +open Tacexpr +open Tacinterp +open Pretyping +open Constr +open Tactic +open Extraargs +open Ppconstr +open Printer + +open Globnames +open Misctypes +open Decl_kinds +open Evar_kinds +open Constrexpr +open Constrexpr_ops +open Notation_term +open Notation_ops +open Locus +open Locusops + +open Ssrmatching + + +(* Tentative patch from util.ml *) + +let array_fold_right_from n f v a = + let rec fold n = + if n >= Array.length v then a else f v.(n) (fold (succ n)) + in + fold n + +let array_app_tl v l = + if Array.length v = 0 then invalid_arg "array_app_tl"; + array_fold_right_from 1 (fun e l -> e::l) v l + +let array_list_of_tl v = + if Array.length v = 0 then invalid_arg "array_list_of_tl"; + array_fold_right_from 1 (fun e l -> e::l) v [] + +(* end patch *) + +module Intset = Evar.Set + +type loc = Loc.t +let dummy_loc = Loc.ghost +let errorstrm = Errors.errorlabstrm "ssreflect" +let loc_error loc msg = Errors.user_err_loc (loc, msg, str msg) +let anomaly s = Errors.anomaly (str s) + +(** look up a name in the ssreflect internals module *) +let ssrdirpath = make_dirpath [id_of_string "ssreflect"] +let ssrqid name = make_qualid ssrdirpath (id_of_string name) +let ssrtopqid name = make_short_qualid (id_of_string name) +let locate_reference qid = + Smartlocate.global_of_extended_global (Nametab.locate_extended qid) +let mkSsrRef name = + try locate_reference (ssrqid name) with Not_found -> + try locate_reference (ssrtopqid name) with Not_found -> + Errors.error "Small scale reflection library not loaded" +let mkSsrRRef name = GRef (dummy_loc, mkSsrRef name,None), None +let mkSsrConst name env sigma = + Evd.fresh_global env sigma (mkSsrRef name) +let pf_mkSsrConst name gl = + let sigma, env, it = project gl, pf_env gl, sig_it gl in + let sigma, t = mkSsrConst name env sigma in + t, re_sig it sigma +let pf_fresh_global name gl = + let sigma, env, it = project gl, pf_env gl, sig_it gl in + let sigma,t = Evd.fresh_global env sigma name in + t, re_sig it sigma + +(** Ssreflect load check. *) + +(* To allow ssrcoq to be fully compatible with the "plain" Coq, we only *) +(* turn on its incompatible features (the new rewrite syntax, and the *) +(* reserved identifiers) when the theory library (ssreflect.v) has *) +(* has actually been required, or is being defined. Because this check *) +(* needs to be done often (for each identifier lookup), we implement *) +(* some caching, repeating the test only when the environment changes. *) +(* We check for protect_term because it is the first constant loaded; *) +(* ssr_have would ultimately be a better choice. *) +let ssr_loaded = Summary.ref ~name:"SSR:loaded" false +let is_ssr_loaded () = + !ssr_loaded || + (if Lexer.is_keyword "SsrSyntax_is_Imported" then ssr_loaded:=true; + !ssr_loaded) + +(* 0 cost pp function. Active only if env variable SSRDEBUG is set *) +(* or if SsrDebug is Set *) +let pp_ref = ref (fun _ -> ()) +let ssr_pp s = pperrnl (str"SSR: "++Lazy.force s) +let _ = try ignore(Sys.getenv "SSRDEBUG"); pp_ref := ssr_pp with Not_found -> () +let _ = + Goptions.declare_bool_option + { Goptions.optsync = false; + Goptions.optname = "ssreflect debugging"; + Goptions.optkey = ["SsrDebug"]; + Goptions.optdepr = false; + Goptions.optread = (fun _ -> !pp_ref == ssr_pp); + Goptions.optwrite = (fun b -> + Ssrmatching.debug b; + if b then pp_ref := ssr_pp else pp_ref := fun _ -> ()) } +let pp s = !pp_ref s + +(** Utils {{{ *****************************************************************) +let env_size env = List.length (Environ.named_context env) +let safeDestApp c = + match kind_of_term c with App (f, a) -> f, a | _ -> c, [| |] +let get_index = function ArgArg i -> i | _ -> + anomaly "Uninterpreted index" +(* Toplevel constr must be globalized twice ! *) +let glob_constr ist genv = function + | _, Some ce -> + let vars = Id.Map.fold (fun x _ accu -> Id.Set.add x accu) ist.lfun Id.Set.empty in + let ltacvars = { + Constrintern.empty_ltac_sign with Constrintern.ltac_vars = vars } in + Constrintern.intern_gen WithoutTypeConstraint ~ltacvars genv ce + | rc, None -> rc + +(* Term printing utilities functions for deciding bracketing. *) +let pr_paren prx x = hov 1 (str "(" ++ prx x ++ str ")") +(* String lexing utilities *) +let skip_wschars s = + let rec loop i = match s.[i] with '\n'..' ' -> loop (i + 1) | _ -> i in loop +let skip_numchars s = + let rec loop i = match s.[i] with '0'..'9' -> loop (i + 1) | _ -> i in loop +(* We also guard characters that might interfere with the ssreflect *) +(* tactic syntax. *) +let guard_term ch1 s i = match s.[i] with + | '(' -> false + | '{' | '/' | '=' -> true + | _ -> ch1 = '(' +(* The call 'guard s i' should return true if the contents of s *) +(* starting at i need bracketing to avoid ambiguities. *) +let pr_guarded guard prc c = + msg_with Format.str_formatter (prc c); + let s = Format.flush_str_formatter () ^ "$" in + if guard s (skip_wschars s 0) then pr_paren prc c else prc c +(* More sensible names for constr printers *) +let prl_constr = pr_lconstr +let pr_constr = pr_constr +let prl_glob_constr c = pr_lglob_constr_env (Global.env ()) c +let pr_glob_constr c = pr_glob_constr_env (Global.env ()) c +let prl_constr_expr = pr_lconstr_expr +let pr_constr_expr = pr_constr_expr +let prl_glob_constr_and_expr = function + | _, Some c -> prl_constr_expr c + | c, None -> prl_glob_constr c +let pr_glob_constr_and_expr = function + | _, Some c -> pr_constr_expr c + | c, None -> pr_glob_constr c +let pr_term (k, c) = pr_guarded (guard_term k) pr_glob_constr_and_expr c +let prl_term (k, c) = pr_guarded (guard_term k) prl_glob_constr_and_expr c + +(** Adding a new uninterpreted generic argument type *) +let add_genarg tag pr = + let wit = Genarg.make0 None tag in + let glob ist x = (ist, x) in + let subst _ x = x in + let interp ist gl x = (gl.Evd.sigma, x) in + let gen_pr _ _ _ = pr in + let () = Genintern.register_intern0 wit glob in + let () = Genintern.register_subst0 wit subst in + let () = Geninterp.register_interp0 wit interp in + Pptactic.declare_extra_genarg_pprule wit gen_pr gen_pr gen_pr; + wit + +(** Constructors for cast type *) +let dC t = CastConv t + +(** Constructors for constr_expr *) +let mkCProp loc = CSort (loc, GProp) +let mkCType loc = CSort (loc, GType []) +let mkCVar loc id = CRef (Ident (loc, id),None) +let isCVar = function CRef (Ident _,_) -> true | _ -> false +let destCVar = function CRef (Ident (_, id),_) -> id | _ -> + anomaly "not a CRef" +let rec mkCHoles loc n = + if n <= 0 then [] else CHole (loc, None, IntroAnonymous, None) :: mkCHoles loc (n - 1) +let mkCHole loc = CHole (loc, None, IntroAnonymous, None) +let rec isCHoles = function CHole _ :: cl -> isCHoles cl | cl -> cl = [] +let mkCExplVar loc id n = + CAppExpl (loc, (None, Ident (loc, id), None), mkCHoles loc n) +let mkCLambda loc name ty t = + CLambdaN (loc, [[loc, name], Default Explicit, ty], t) +let mkCLetIn loc name bo t = + CLetIn (loc, (loc, name), bo, t) +let mkCArrow loc ty t = + CProdN (loc, [[dummy_loc,Anonymous], Default Explicit, ty], t) +let mkCCast loc t ty = CCast (loc,t, dC ty) +(** Constructors for rawconstr *) +let mkRHole = GHole (dummy_loc, InternalHole, IntroAnonymous, None) +let rec mkRHoles n = if n > 0 then mkRHole :: mkRHoles (n - 1) else [] +let rec isRHoles = function GHole _ :: cl -> isRHoles cl | cl -> cl = [] +let mkRApp f args = if args = [] then f else GApp (dummy_loc, f, args) +let mkRVar id = GRef (dummy_loc, VarRef id,None) +let mkRltacVar id = GVar (dummy_loc, id) +let mkRCast rc rt = GCast (dummy_loc, rc, dC rt) +let mkRType = GSort (dummy_loc, GType []) +let mkRProp = GSort (dummy_loc, GProp) +let mkRArrow rt1 rt2 = GProd (dummy_loc, Anonymous, Explicit, rt1, rt2) +let mkRConstruct c = GRef (dummy_loc, ConstructRef c,None) +let mkRInd mind = GRef (dummy_loc, IndRef mind,None) +let mkRLambda n s t = GLambda (dummy_loc, n, Explicit, s, t) + +(** Constructors for constr *) +let pf_e_type_of gl t = + let sigma, env, it = project gl, pf_env gl, sig_it gl in + let sigma, ty = Typing.e_type_of env sigma t in + re_sig it sigma, ty + +let mkAppRed f c = match kind_of_term f with +| Lambda (_, _, b) -> subst1 c b +| _ -> mkApp (f, [|c|]) + +let mkProt t c gl = + let prot, gl = pf_mkSsrConst "protect_term" gl in + mkApp (prot, [|t; c|]), gl + +let mkRefl t c gl = + let refl, gl = pf_fresh_global (build_coq_eq_data()).refl gl in + mkApp (refl, [|t; c|]), gl + + +(* Application to a sequence of n rels (for building eta-expansions). *) +(* The rel indices decrease down to imin (inclusive), unless n < 0, *) +(* in which case they're incresing (from imin). *) +let mkEtaApp c n imin = + if n = 0 then c else + let nargs, mkarg = + if n < 0 then -n, (fun i -> mkRel (imin + i)) else + let imax = imin + n - 1 in n, (fun i -> mkRel (imax - i)) in + mkApp (c, Array.init nargs mkarg) +(* Same, but optimizing head beta redexes *) +let rec whdEtaApp c n = + if n = 0 then c else match kind_of_term c with + | Lambda (_, _, c') -> whdEtaApp c' (n - 1) + | _ -> mkEtaApp (lift n c) n 1 +let mkType () = Universes.new_Type (Lib.cwd ()) + +(* ssrterm conbinators *) +let combineCG t1 t2 f g = match t1, t2 with + | (x, (t1, None)), (_, (t2, None)) -> x, (g t1 t2, None) + | (x, (_, Some t1)), (_, (_, Some t2)) -> x, (mkRHole, Some (f t1 t2)) + | _, (_, (_, None)) -> anomaly "have: mixed C-G constr" + | _ -> anomaly "have: mixed G-C constr" +let loc_ofCG = function + | (_, (s, None)) -> Glob_ops.loc_of_glob_constr s + | (_, (_, Some s)) -> Constrexpr_ops.constr_loc s + +let mk_term k c = k, (mkRHole, Some c) +let mk_lterm c = mk_term ' ' c + +let map_fold_constr g f ctx acc cstr = + let array_f ctx acc x = let x, acc = f ctx acc x in acc, x in + match kind_of_term cstr with + | (Rel _ | Meta _ | Var _ | Sort _ | Const _ | Ind _ | Construct _) -> + cstr, acc + | Proj (x,c) -> + let c', acc = f ctx acc c in + (if c == c' then cstr else mkProj (x,c')), acc + | Cast (c,k, t) -> + let c', acc = f ctx acc c in + let t', acc = f ctx acc t in + (if c==c' && t==t' then cstr else mkCast (c', k, t')), acc + | Prod (na,t,c) -> + let t', acc = f ctx acc t in + let c', acc = f (g (na,None,t) ctx) acc c in + (if t==t' && c==c' then cstr else mkProd (na, t', c')), acc + | Lambda (na,t,c) -> + let t', acc = f ctx acc t in + let c', acc = f (g (na,None,t) ctx) acc c in + (if t==t' && c==c' then cstr else mkLambda (na, t', c')), acc + | LetIn (na,b,t,c) -> + let b', acc = f ctx acc b in + let t', acc = f ctx acc t in + let c', acc = f (g (na,Some b,t) ctx) acc c in + (if b==b' && t==t' && c==c' then cstr else mkLetIn (na, b', t', c')), acc + | App (c,al) -> + let c', acc = f ctx acc c in + let acc, al' = CArray.smartfoldmap (array_f ctx) acc al in + (if c==c' && Array.for_all2 (==) al al' then cstr else mkApp (c', al')), + acc + | Evar (e,al) -> + let acc, al' = CArray.smartfoldmap (array_f ctx) acc al in + (if Array.for_all2 (==) al al' then cstr else mkEvar (e, al')), acc + | Case (ci,p,c,bl) -> + let p', acc = f ctx acc p in + let c', acc = f ctx acc c in + let acc, bl' = CArray.smartfoldmap (array_f ctx) acc bl in + (if p==p' && c==c' && Array.for_all2 (==) bl bl' then cstr else + mkCase (ci, p', c', bl')), + acc + | Fix (ln,(lna,tl,bl)) -> + let acc, tl' = CArray.smartfoldmap (array_f ctx) acc tl in + let ctx' = Array.fold_left2 (fun l na t -> g (na,None,t) l) ctx lna tl in + let acc, bl' = CArray.smartfoldmap (array_f ctx') acc bl in + (if Array.for_all2 (==) tl tl' && Array.for_all2 (==) bl bl' + then cstr + else mkFix (ln,(lna,tl',bl'))), acc + | CoFix(ln,(lna,tl,bl)) -> + let acc, tl' = CArray.smartfoldmap (array_f ctx) acc tl in + let ctx' = Array.fold_left2 (fun l na t -> g (na,None,t) l) ctx lna tl in + let acc,bl' = CArray.smartfoldmap (array_f ctx') acc bl in + (if Array.for_all2 (==) tl tl' && Array.for_all2 (==) bl bl' + then cstr + else mkCoFix (ln,(lna,tl',bl'))), acc + +let pf_merge_uc_of sigma gl = + let ucst = Evd.evar_universe_context sigma in + pf_merge_uc ucst gl + +(* }}} *) + +(** Profiling {{{ *************************************************************) +type profiler = { + profile : 'a 'b. ('a -> 'b) -> 'a -> 'b; + reset : unit -> unit; + print : unit -> unit } +let profile_now = ref false +let something_profiled = ref false +let profilers = ref [] +let add_profiler f = profilers := f :: !profilers;; +let _ = + Goptions.declare_bool_option + { Goptions.optsync = false; + Goptions.optname = "ssreflect profiling"; + Goptions.optkey = ["SsrProfiling"]; + Goptions.optread = (fun _ -> !profile_now); + Goptions.optdepr = false; + Goptions.optwrite = (fun b -> + Ssrmatching.profile b; + profile_now := b; + if b then List.iter (fun f -> f.reset ()) !profilers; + if not b then List.iter (fun f -> f.print ()) !profilers) } +let () = + let prof_total = + let init = ref 0.0 in { + profile = (fun f x -> assert false); + reset = (fun () -> init := Unix.gettimeofday ()); + print = (fun () -> if !something_profiled then + prerr_endline + (Printf.sprintf "!! %-39s %10d %9.4f %9.4f %9.4f" + "total" 0 (Unix.gettimeofday() -. !init) 0.0 0.0)) } in + let prof_legenda = { + profile = (fun f x -> assert false); + reset = (fun () -> ()); + print = (fun () -> if !something_profiled then begin + prerr_endline + (Printf.sprintf "!! %39s ---------- --------- --------- ---------" + (String.make 39 '-')); + prerr_endline + (Printf.sprintf "!! %-39s %10s %9s %9s %9s" + "function" "#calls" "total" "max" "average") end) } in + add_profiler prof_legenda; + add_profiler prof_total +;; + +let mk_profiler s = + let total, calls, max = ref 0.0, ref 0, ref 0.0 in + let reset () = total := 0.0; calls := 0; max := 0.0 in + let profile f x = + if not !profile_now then f x else + let before = Unix.gettimeofday () in + try + incr calls; + let res = f x in + let after = Unix.gettimeofday () in + let delta = after -. before in + total := !total +. delta; + if delta > !max then max := delta; + res + with exc -> + let after = Unix.gettimeofday () in + let delta = after -. before in + total := !total +. delta; + if delta > !max then max := delta; + raise exc in + let print () = + if !calls <> 0 then begin + something_profiled := true; + prerr_endline + (Printf.sprintf "!! %-39s %10d %9.4f %9.4f %9.4f" + s !calls !total !max (!total /. (float_of_int !calls))) end in + let prof = { profile = profile; reset = reset; print = print } in + add_profiler prof; + prof +;; +(* }}} *) + +let inVersion = Libobject.declare_object { + (Libobject.default_object "SSRASTVERSION") with + Libobject.load_function = (fun _ (_,v) -> + if v <> ssrAstVersion then Errors.error "Please recompile your .vo files"); + Libobject.classify_function = (fun v -> Libobject.Keep v); +} + +let _ = + Goptions.declare_bool_option + { Goptions.optsync = false; + Goptions.optname = "ssreflect version"; + Goptions.optkey = ["SsrAstVersion"]; + Goptions.optread = (fun _ -> true); + Goptions.optdepr = false; + Goptions.optwrite = (fun _ -> + Lib.add_anonymous_leaf (inVersion ssrAstVersion)) } + +let tactic_expr = Tactic.tactic_expr +let gallina_ext = Vernac_.gallina_ext +let sprintf = Printf.sprintf +let tactic_mode = G_vernac.tactic_mode + +(** 1. Utilities *) + + +let ssroldreworder = Summary.ref ~name:"SSR:oldreworder" true +let _ = + Goptions.declare_bool_option + { Goptions.optsync = false; + Goptions.optname = "ssreflect 1.3 compatibility flag"; + Goptions.optkey = ["SsrOldRewriteGoalsOrder"]; + Goptions.optread = (fun _ -> !ssroldreworder); + Goptions.optdepr = false; + Goptions.optwrite = (fun b -> ssroldreworder := b) } + +let ssrhaveNOtcresolution = ref false + +let inHaveTCResolution = Libobject.declare_object { + (Libobject.default_object "SSRHAVETCRESOLUTION") with + Libobject.cache_function = (fun (_,v) -> ssrhaveNOtcresolution := v); + Libobject.load_function = (fun _ (_,v) -> ssrhaveNOtcresolution := v); + Libobject.classify_function = (fun v -> Libobject.Keep v); +} + +let _ = + Goptions.declare_bool_option + { Goptions.optsync = false; + Goptions.optname = "have type classes"; + Goptions.optkey = ["SsrHave";"NoTCResolution"]; + Goptions.optread = (fun _ -> !ssrhaveNOtcresolution); + Goptions.optdepr = false; + Goptions.optwrite = (fun b -> + Lib.add_anonymous_leaf (inHaveTCResolution b)) } + + +(** Primitive parsing to avoid syntax conflicts with basic tactics. *) + +let accept_before_syms syms strm = + match Compat.get_tok (stream_nth 1 strm) with + | Tok.KEYWORD sym when List.mem sym syms -> () + | _ -> raise Stream.Failure + +let accept_before_syms_or_any_id syms strm = + match Compat.get_tok (stream_nth 1 strm) with + | Tok.KEYWORD sym when List.mem sym syms -> () + | Tok.IDENT _ -> () + | _ -> raise Stream.Failure + +let accept_before_syms_or_ids syms ids strm = + match Compat.get_tok (stream_nth 1 strm) with + | Tok.KEYWORD sym when List.mem sym syms -> () + | Tok.IDENT id when List.mem id ids -> () + | _ -> raise Stream.Failure + +(** Pretty-printing utilities *) + +let pr_id = Ppconstr.pr_id +let pr_name = function Name id -> pr_id id | Anonymous -> str "_" +let pr_spc () = str " " +let pr_bar () = Pp.cut() ++ str "|" +let pr_list = prlist_with_sep + +let tacltop = (5,Ppextend.E) + +(** Tactic-level diagnosis *) + +let pf_pr_constr gl = pr_constr_env (pf_env gl) + +let pf_pr_glob_constr gl = pr_glob_constr_env (pf_env gl) + +(* debug *) + +let pf_msg gl = + let ppgl = pr_lconstr_env (pf_env gl) (project gl) (pf_concl gl) in + msgnl (str "goal is " ++ ppgl) + +let msgtac gl = pf_msg gl; tclIDTAC gl + +(** Tactic utilities *) + +let last_goal gls = let sigma, gll = Refiner.unpackage gls in + Refiner.repackage sigma (List.nth gll (List.length gll - 1)) + +let pf_type_id gl t = id_of_string (hdchar (pf_env gl) t) + +let not_section_id id = not (is_section_variable id) + +let is_pf_var c = isVar c && not_section_id (destVar c) + +let pf_ids_of_proof_hyps gl = + let add_hyp (id, _, _) ids = if not_section_id id then id :: ids else ids in + Context.fold_named_context add_hyp (pf_hyps gl) ~init:[] + +let pf_nf_evar gl e = Reductionops.nf_evar (project gl) e + +let pf_partial_solution gl t evl = + let sigma, g = project gl, sig_it gl in + let sigma = Goal.V82.partial_solution sigma g t in + re_sig (List.map (fun x -> (fst (destEvar x))) evl) sigma + +let pf_new_evar gl ty = + let sigma, env, it = project gl, pf_env gl, sig_it gl in + let sigma, extra = Evarutil.new_evar env sigma ty in + re_sig it sigma, extra + +(* Basic tactics *) + +let convert_concl_no_check t = convert_concl_no_check t DEFAULTcast +let convert_concl t = convert_concl t DEFAULTcast +let reduct_in_concl t = reduct_in_concl (t, DEFAULTcast) +let havetac id c = Proofview.V82.of_tactic (pose_proof (Name id) c) +let settac id c = letin_tac None (Name id) c None +let posetac id cl = Proofview.V82.of_tactic (settac id cl nowhere) +let basecuttac name c gl = + let hd, gl = pf_mkSsrConst name gl in + let t = mkApp (hd, [|c|]) in + let gl, _ = pf_e_type_of gl t in + Proofview.V82.of_tactic (apply t) gl + +(* we reduce head beta redexes *) +let betared env = + Closure.create_clos_infos + (Closure.RedFlags.mkflags [Closure.RedFlags.fBETA]) + env +;; +let introid name = tclTHEN (fun gl -> + let g, env = pf_concl gl, pf_env gl in + match kind_of_term g with + | App (hd, _) when isLambda hd -> + let g = Closure.whd_val (betared env) (Closure.inject g) in + Proofview.V82.of_tactic (convert_concl_no_check g) gl + | _ -> tclIDTAC gl) + (Proofview.V82.of_tactic (intro_mustbe_force name)) +;; + + +(** Name generation {{{ *******************************************************) + +(* Since Coq now does repeated internal checks of its external lexical *) +(* rules, we now need to carve ssreflect reserved identifiers out of *) +(* out of the user namespace. We use identifiers of the form _id_ for *) +(* this purpose, e.g., we "anonymize" an identifier id as _id_, adding *) +(* an extra leading _ if this might clash with an internal identifier. *) +(* We check for ssreflect identifiers in the ident grammar rule; *) +(* when the ssreflect Module is present this is normally an error, *) +(* but we provide a compatibility flag to reduce this to a warning. *) + +let ssr_reserved_ids = Summary.ref ~name:"SSR:idents" true + +let _ = + Goptions.declare_bool_option + { Goptions.optsync = true; + Goptions.optname = "ssreflect identifiers"; + Goptions.optkey = ["SsrIdents"]; + Goptions.optdepr = false; + Goptions.optread = (fun _ -> !ssr_reserved_ids); + Goptions.optwrite = (fun b -> ssr_reserved_ids := b) + } + +let is_ssr_reserved s = + let n = String.length s in n > 2 && s.[0] = '_' && s.[n - 1] = '_' + +let internal_names = ref [] +let add_internal_name pt = internal_names := pt :: !internal_names +let is_internal_name s = List.exists (fun p -> p s) !internal_names + +let ssr_id_of_string loc s = + if is_ssr_reserved s && is_ssr_loaded () then begin + if !ssr_reserved_ids then + loc_error loc ("The identifier " ^ s ^ " is reserved.") + else if is_internal_name s then + msg_warning (str ("Conflict between " ^ s ^ " and ssreflect internal names.")) + else msg_warning (str ( + "The name " ^ s ^ " fits the _xxx_ format used for anonymous variables.\n" + ^ "Scripts with explicit references to anonymous variables are fragile.")) + end; id_of_string s + +let ssr_null_entry = Gram.Entry.of_parser "ssr_null" (fun _ -> ()) + +let (!@) = Compat.to_coqloc + +GEXTEND Gram + GLOBAL: Prim.ident; + Prim.ident: [[ s = IDENT; ssr_null_entry -> ssr_id_of_string !@loc s ]]; +END + +let mk_internal_id s = + let s' = sprintf "_%s_" s in + for i = 1 to String.length s do if s'.[i] = ' ' then s'.[i] <- '_' done; + add_internal_name ((=) s'); id_of_string s' + +let same_prefix s t n = + let rec loop i = i = n || s.[i] = t.[i] && loop (i + 1) in loop 0 + +let skip_digits s = + let n = String.length s in + let rec loop i = if i < n && is_digit s.[i] then loop (i + 1) else i in loop + +let mk_tagged_id t i = id_of_string (sprintf "%s%d_" t i) +let is_tagged t s = + let n = String.length s - 1 and m = String.length t in + m < n && s.[n] = '_' && same_prefix s t m && skip_digits s m = n + +let perm_tag = "_perm_Hyp_" +let _ = add_internal_name (is_tagged perm_tag) +let mk_perm_id = + let salt = ref 1 in + fun () -> salt := !salt mod 10000 + 1; mk_tagged_id perm_tag !salt + +let evar_tag = "_evar_" +let _ = add_internal_name (is_tagged evar_tag) +let mk_evar_name n = Name (mk_tagged_id evar_tag n) +let nb_evar_deps = function + | Name id -> + let s = string_of_id id in + if not (is_tagged evar_tag s) then 0 else + let m = String.length evar_tag in + (try int_of_string (String.sub s m (String.length s - 1 - m)) with _ -> 0) + | _ -> 0 + +let discharged_tag = "_discharged_" +let mk_discharged_id id = + id_of_string (sprintf "%s%s_" discharged_tag (string_of_id id)) +let has_discharged_tag s = + let m = String.length discharged_tag and n = String.length s - 1 in + m < n && s.[n] = '_' && same_prefix s discharged_tag m +let _ = add_internal_name has_discharged_tag +let is_discharged_id id = has_discharged_tag (string_of_id id) + +let wildcard_tag = "_the_" +let wildcard_post = "_wildcard_" +let mk_wildcard_id i = + id_of_string (sprintf "%s%s%s" wildcard_tag (String.ordinal i) wildcard_post) +let has_wildcard_tag s = + let n = String.length s in let m = String.length wildcard_tag in + let m' = String.length wildcard_post in + n < m + m' + 2 && same_prefix s wildcard_tag m && + String.sub s (n - m') m' = wildcard_post && + skip_digits s m = n - m' - 2 +let _ = add_internal_name has_wildcard_tag + +let max_suffix m (t, j0 as tj0) id = + let s = string_of_id id in let n = String.length s - 1 in + let dn = String.length t - 1 - n in let i0 = j0 - dn in + if not (i0 >= m && s.[n] = '_' && same_prefix s t m) then tj0 else + let rec loop i = + if i < i0 && s.[i] = '0' then loop (i + 1) else + if (if i < i0 then skip_digits s i = n else le_s_t i) then s, i else tj0 + and le_s_t i = + let ds = s.[i] and dt = t.[i + dn] in + if ds = dt then i = n || le_s_t (i + 1) else + dt < ds && skip_digits s i = n in + loop m + +let mk_anon_id t gl = + let m, si0, id0 = + let s = ref (sprintf "_%s_" t) in + if is_internal_name !s then s := "_" ^ !s; + let n = String.length !s - 1 in + let rec loop i j = + let d = !s.[i] in if not (is_digit d) then i + 1, j else + loop (i - 1) (if d = '0' then j else i) in + let m, j = loop (n - 1) n in m, (!s, j), id_of_string !s in + let gl_ids = pf_ids_of_hyps gl in + if not (List.mem id0 gl_ids) then id0 else + let s, i = List.fold_left (max_suffix m) si0 gl_ids in + let n = String.length s - 1 in + let rec loop i = + if s.[i] = '9' then (s.[i] <- '0'; loop (i - 1)) else + if i < m then (s.[n] <- '0'; s.[m] <- '1'; s ^ "_") else + (s.[i] <- Char.chr (Char.code s.[i] + 1); s) in + id_of_string (loop (n - 1)) + +(* We must not anonymize context names discharged by the "in" tactical. *) + +let ssr_anon_hyp = "Hyp" + +let anontac (x, _, _) gl = + let id = match x with + | Name id -> + if is_discharged_id id then id else mk_anon_id (string_of_id id) gl + | _ -> mk_anon_id ssr_anon_hyp gl in + introid id gl + +let rec constr_name c = match kind_of_term c with + | Var id -> Name id + | Cast (c', _, _) -> constr_name c' + | Const (cn,_) -> Name (id_of_label (con_label cn)) + | App (c', _) -> constr_name c' + | _ -> Anonymous + +(* }}} *) + +(** Open term to lambda-term coercion {{{ ************************************) + +(* This operation takes a goal gl and an open term (sigma, t), and *) +(* returns a term t' where all the new evars in sigma are abstracted *) +(* with the mkAbs argument, i.e., for mkAbs = mkLambda then there is *) +(* some duplicate-free array args of evars of sigma such that the *) +(* term mkApp (t', args) is convertible to t. *) +(* This makes a useful shorthand for local definitions in proofs, *) +(* i.e., pose succ := _ + 1 means pose succ := fun n : nat => n + 1, *) +(* and, in context of the the 4CT library, pose mid := maps id means *) +(* pose mid := fun d : detaSet => @maps d d (@id (datum d)) *) +(* Note that this facility does not extend to set, which tries *) +(* instead to fill holes by matching a goal subterm. *) +(* The argument to "have" et al. uses product abstraction, e.g. *) +(* have Hmid: forall s, (maps id s) = s. *) +(* stands for *) +(* have Hmid: forall (d : dataSet) (s : seq d), (maps id s) = s. *) +(* We also use this feature for rewrite rules, so that, e.g., *) +(* rewrite: (plus_assoc _ 3). *) +(* will execute as *) +(* rewrite (fun n => plus_assoc n 3) *) +(* i.e., it will rewrite some subterm .. + (3 + ..) to .. + 3 + ... *) +(* The convention is also used for the argument of the congr tactic, *) +(* e.g., congr (x + _ * 1). *) + +(* Replace new evars with lambda variables, retaining local dependencies *) +(* but stripping global ones. We use the variable names to encode the *) +(* the number of dependencies, so that the transformation is reversible. *) + +let pf_abs_evars gl (sigma, c0) = + let sigma0, ucst = project gl, Evd.evar_universe_context sigma in + let nenv = env_size (pf_env gl) in + let abs_evar n k = + let evi = Evd.find sigma k in + let dc = List.firstn n (evar_filtered_context evi) in + let abs_dc c = function + | x, Some b, t -> mkNamedLetIn x b t (mkArrow t c) + | x, None, t -> mkNamedProd x t c in + let t = Context.fold_named_context_reverse abs_dc ~init:evi.evar_concl dc in + Evarutil.nf_evar sigma t in + let rec put evlist c = match kind_of_term c with + | Evar (k, a) -> + if List.mem_assoc k evlist || Evd.mem sigma0 k then evlist else + let n = max 0 (Array.length a - nenv) in + let t = abs_evar n k in (k, (n, t)) :: put evlist t + | _ -> fold_constr put evlist c in + let evlist = put [] c0 in + if evlist = [] then 0, c0,[], ucst else + let rec lookup k i = function + | [] -> 0, 0 + | (k', (n, _)) :: evl -> if k = k' then i, n else lookup k (i + 1) evl in + let rec get i c = match kind_of_term c with + | Evar (ev, a) -> + let j, n = lookup ev i evlist in + if j = 0 then map_constr (get i) c else if n = 0 then mkRel j else + mkApp (mkRel j, Array.init n (fun k -> get i a.(n - 1 - k))) + | _ -> map_constr_with_binders ((+) 1) get i c in + let rec loop c i = function + | (_, (n, t)) :: evl -> + loop (mkLambda (mk_evar_name n, get (i - 1) t, c)) (i - 1) evl + | [] -> c in + List.length evlist, loop (get 1 c0) 1 evlist, List.map fst evlist, ucst + + + +(* As before but if (?i : T(?j)) and (?j : P : Prop), then the lambda for i + * looks like (fun evar_i : (forall pi : P. T(pi))) thanks to "loopP" and all + * occurrences of evar_i are replaced by (evar_i evar_j) thanks to "app". + * + * If P can be solved by ssrautoprop (that defaults to trivial), then + * the corresponding lambda looks like (fun evar_i : T(c)) where c is + * the solution found by ssrautoprop. + *) +let ssrautoprop_tac = ref (fun gl -> assert false) + +(* Thanks to Arnaud Spiwack for this snippet *) +let call_on_evar tac e s = + let { it = gs ; sigma = s } = + tac { it = e ; sigma = s; } in + gs, s + +let pf_abs_evars_pirrel gl (sigma, c0) = + pp(lazy(str"==PF_ABS_EVARS_PIRREL==")); + pp(lazy(str"c0= " ++ pr_constr c0)); + let sigma0 = project gl in + let c0 = Evarutil.nf_evar sigma0 (Evarutil.nf_evar sigma c0) in + let nenv = env_size (pf_env gl) in + let abs_evar n k = + let evi = Evd.find sigma k in + let dc = List.firstn n (evar_filtered_context evi) in + let abs_dc c = function + | x, Some b, t -> mkNamedLetIn x b t (mkArrow t c) + | x, None, t -> mkNamedProd x t c in + let t = Context.fold_named_context_reverse abs_dc ~init:evi.evar_concl dc in + Evarutil.nf_evar sigma0 (Evarutil.nf_evar sigma t) in + let rec put evlist c = match kind_of_term c with + | Evar (k, a) -> + if List.mem_assoc k evlist || Evd.mem sigma0 k then evlist else + let n = max 0 (Array.length a - nenv) in + let k_ty = + Retyping.get_sort_family_of + (pf_env gl) sigma (Evd.evar_concl (Evd.find sigma k)) in + let is_prop = k_ty = InProp in + let t = abs_evar n k in (k, (n, t, is_prop)) :: put evlist t + | _ -> fold_constr put evlist c in + let evlist = put [] c0 in + if evlist = [] then 0, c0 else + let pr_constr t = pr_constr (Reductionops.nf_beta (project gl) t) in + pp(lazy(str"evlist=" ++ pr_list (fun () -> str";") + (fun (k,_) -> str(Evd.string_of_existential k)) evlist)); + let evplist = + let depev = List.fold_left (fun evs (_,(_,t,_)) -> + Intset.union evs (Evarutil.undefined_evars_of_term sigma t)) Intset.empty evlist in + List.filter (fun i,(_,_,b) -> b && Intset.mem i depev) evlist in + let evlist, evplist, sigma = + if evplist = [] then evlist, [], sigma else + List.fold_left (fun (ev, evp, sigma) (i, (_,t,_) as p) -> + try + let ng, sigma = call_on_evar !ssrautoprop_tac i sigma in + if (ng <> []) then errorstrm (str "Should we tell the user?"); + List.filter (fun (j,_) -> j <> i) ev, evp, sigma + with _ -> ev, p::evp, sigma) (evlist, [], sigma) (List.rev evplist) in + let c0 = Evarutil.nf_evar sigma c0 in + let evlist = + List.map (fun (x,(y,t,z)) -> x,(y,Evarutil.nf_evar sigma t,z)) evlist in + let evplist = + List.map (fun (x,(y,t,z)) -> x,(y,Evarutil.nf_evar sigma t,z)) evplist in + pp(lazy(str"c0= " ++ pr_constr c0)); + let rec lookup k i = function + | [] -> 0, 0 + | (k', (n,_,_)) :: evl -> if k = k' then i,n else lookup k (i + 1) evl in + let rec get evlist i c = match kind_of_term c with + | Evar (ev, a) -> + let j, n = lookup ev i evlist in + if j = 0 then map_constr (get evlist i) c else if n = 0 then mkRel j else + mkApp (mkRel j, Array.init n (fun k -> get evlist i a.(n - 1 - k))) + | _ -> map_constr_with_binders ((+) 1) (get evlist) i c in + let rec app extra_args i c = match decompose_app c with + | hd, args when isRel hd && destRel hd = i -> + let j = destRel hd in + mkApp (mkRel j, Array.of_list (List.map (lift (i-1)) extra_args @ args)) + | _ -> map_constr_with_binders ((+) 1) (app extra_args) i c in + let rec loopP evlist c i = function + | (_, (n, t, _)) :: evl -> + let t = get evlist (i - 1) t in + let n = Name (id_of_string (ssr_anon_hyp ^ string_of_int n)) in + loopP evlist (mkProd (n, t, c)) (i - 1) evl + | [] -> c in + let rec loop c i = function + | (_, (n, t, _)) :: evl -> + let evs = Evarutil.undefined_evars_of_term sigma t in + let t_evplist = List.filter (fun (k,_) -> Intset.mem k evs) evplist in + let t = loopP t_evplist (get t_evplist 1 t) 1 t_evplist in + let t = get evlist (i - 1) t in + let extra_args = + List.map (fun (k,_) -> mkRel (fst (lookup k i evlist))) + (List.rev t_evplist) in + let c = if extra_args = [] then c else app extra_args 1 c in + loop (mkLambda (mk_evar_name n, t, c)) (i - 1) evl + | [] -> c in + let res = loop (get evlist 1 c0) 1 evlist in + pp(lazy(str"res= " ++ pr_constr res)); + List.length evlist, res + +(* Strip all non-essential dependencies from an abstracted term, generating *) +(* standard names for the abstracted holes. *) + +let pf_abs_cterm gl n c0 = + if n <= 0 then c0 else + let noargs = [|0|] in + let eva = Array.make n noargs in + let rec strip i c = match kind_of_term c with + | App (f, a) when isRel f -> + let j = i - destRel f in + if j >= n || eva.(j) = noargs then mkApp (f, Array.map (strip i) a) else + let dp = eva.(j) in + let nd = Array.length dp - 1 in + let mkarg k = strip i a.(if k < nd then dp.(k + 1) - j else k + dp.(0)) in + mkApp (f, Array.init (Array.length a - dp.(0)) mkarg) + | _ -> map_constr_with_binders ((+) 1) strip i c in + let rec strip_ndeps j i c = match kind_of_term c with + | Prod (x, t, c1) when i < j -> + let dl, c2 = strip_ndeps j (i + 1) c1 in + if noccurn 1 c2 then dl, lift (-1) c2 else + i :: dl, mkProd (x, strip i t, c2) + | LetIn (x, b, t, c1) when i < j -> + let _, _, c1' = destProd c1 in + let dl, c2 = strip_ndeps j (i + 1) c1' in + if noccurn 1 c2 then dl, lift (-1) c2 else + i :: dl, mkLetIn (x, strip i b, strip i t, c2) + | _ -> [], strip i c in + let rec strip_evars i c = match kind_of_term c with + | Lambda (x, t1, c1) when i < n -> + let na = nb_evar_deps x in + let dl, t2 = strip_ndeps (i + na) i t1 in + let na' = List.length dl in + eva.(i) <- Array.of_list (na - na' :: dl); + let x' = + if na' = 0 then Name (pf_type_id gl t2) else mk_evar_name na' in + mkLambda (x', t2, strip_evars (i + 1) c1) +(* if noccurn 1 c2 then lift (-1) c2 else + mkLambda (Name (pf_type_id gl t2), t2, c2) *) + | _ -> strip i c in + strip_evars 0 c0 + +(* Undo the evar abstractions. Also works for non-evar variables. *) + +let pf_unabs_evars gl ise n c0 = + if n = 0 then c0 else + let evv = Array.make n mkProp in + let nev = ref 0 in + let env0 = pf_env gl in + let nenv0 = env_size env0 in + let rec unabs i c = match kind_of_term c with + | Rel j when i - j < !nev -> evv.(i - j) + | App (f, a0) when isRel f -> + let a = Array.map (unabs i) a0 in + let j = i - destRel f in + if j >= !nev then mkApp (f, a) else + let ev, eva = destEvar evv.(j) in + let nd = Array.length eva - nenv0 in + if nd = 0 then mkApp (evv.(j), a) else + let evarg k = if k < nd then a.(nd - 1 - k) else eva.(k) in + let c' = mkEvar (ev, Array.init (nd + nenv0) evarg) in + let na = Array.length a - nd in + if na = 0 then c' else mkApp (c', Array.sub a nd na) + | _ -> map_constr_with_binders ((+) 1) unabs i c in + let push_rel = Environ.push_rel in + let rec mk_evar j env i c = match kind_of_term c with + | Prod (x, t, c1) when i < j -> + mk_evar j (push_rel (x, None, unabs i t) env) (i + 1) c1 + | LetIn (x, b, t, c1) when i < j -> + let _, _, c2 = destProd c1 in + mk_evar j (push_rel (x, Some (unabs i b), unabs i t) env) (i + 1) c2 + | _ -> Evarutil.e_new_evar env ise (unabs i c) in + let rec unabs_evars c = + if !nev = n then unabs n c else match kind_of_term c with + | Lambda (x, t, c1) when !nev < n -> + let i = !nev in + evv.(i) <- mk_evar (i + nb_evar_deps x) env0 i t; + incr nev; unabs_evars c1 + | _ -> unabs !nev c in + unabs_evars c0 + +(* }}} *) + +(** Tactical extensions. {{{ **************************************************) + +(* The TACTIC EXTEND facility can't be used for defining new user *) +(* tacticals, because: *) +(* - the concrete syntax must start with a fixed string *) +(* We use the following workaround: *) +(* - We use the (unparsable) "YouShouldNotTypeThis" token for tacticals that *) +(* don't start with a token, then redefine the grammar and *) +(* printer using GEXTEND and set_pr_ssrtac, respectively. *) + +type ssrargfmt = ArgSsr of string | ArgCoq of argument_type | ArgSep of string + +let ssrtac_name name = { + mltac_plugin = "ssreflect"; + mltac_tactic = "ssr" ^ name; +} + +let set_pr_ssrtac name prec afmt = + let fmt = List.map (function ArgSep s -> Some s | _ -> None) afmt in + let rec mk_akey = function + | ArgSsr s :: afmt' -> ExtraArgType ("ssr" ^ s) :: mk_akey afmt' + | ArgCoq a :: afmt' -> a :: mk_akey afmt' + | ArgSep _ :: afmt' -> mk_akey afmt' + | [] -> [] in + let tacname = ssrtac_name name in + Pptactic.declare_ml_tactic_pprule tacname + { Pptactic.pptac_args = mk_akey afmt; + Pptactic.pptac_prods = (prec, fmt) } + +let ssrtac_atom loc name args = TacML (loc, ssrtac_name name, args) +let ssrtac_expr = ssrtac_atom + + +let ssrevaltac ist gtac = + let debug = match TacStore.get ist.extra f_debug with + | None -> Tactic_debug.DebugOff | Some level -> level + in + Proofview.V82.of_tactic (interp_tac_gen ist.lfun [] debug (globTacticIn (fun _ -> gtac))) + +(* fun gl -> let lfun = [tacarg_id, val_interp ist gl gtac] in + interp_tac_gen lfun [] ist.debug tacarg_expr gl *) + +(** Generic argument-based globbing/typing utilities *) + + +let interp_wit wit ist gl x = + let globarg = in_gen (glbwit wit) x in + let sigma, arg = interp_genarg ist (pf_env gl) (project gl) (pf_concl gl) gl.Evd.it globarg in + sigma, out_gen (topwit wit) arg + +let interp_intro_pattern = interp_wit wit_intro_pattern + +let interp_constr = interp_wit wit_constr + +let interp_open_constr ist gl gc = + interp_wit wit_open_constr ist gl ((), gc) + +let interp_refine ist gl rc = + let constrvars = extract_ltac_constr_values ist (pf_env gl) in + let vars = { Pretyping.empty_lvar with + Pretyping.ltac_constrs = constrvars; ltac_genargs = ist.lfun + } in + let kind = OfType (pf_concl gl) in + let flags = { + use_typeclasses = true; + use_unif_heuristics = true; + use_hook = None; + fail_evar = false; + expand_evars = true } + in + let sigma, c = understand_ltac flags (pf_env gl) (project gl) vars kind rc in +(* pp(lazy(str"sigma@interp_refine=" ++ pr_evar_map None sigma)); *) + pp(lazy(str"c@interp_refine=" ++ pr_constr c)); + (sigma, (sigma, c)) + +let pf_match = pf_apply (fun e s c t -> understand_tcc e s ~expected_type:t c) + +(* Estimate a bound on the number of arguments of a raw constr. *) +(* This is not perfect, because the unifier may fail to *) +(* typecheck the partial application, so we use a minimum of 5. *) +(* Also, we don't handle delayed or iterated coercions to *) +(* FUNCLASS, which is probably just as well since these can *) +(* lead to infinite arities. *) + +let splay_open_constr gl (sigma, c) = + let env = pf_env gl in let t = Retyping.get_type_of env sigma c in + Reductionops.splay_prod env sigma t + +let nbargs_open_constr gl oc = + let pl, _ = splay_open_constr gl oc in List.length pl + +let interp_nbargs ist gl rc = + try + let rc6 = mkRApp rc (mkRHoles 6) in + let sigma, t = interp_open_constr ist gl (rc6, None) in + let si = sig_it gl in + let gl = re_sig si sigma in + 6 + nbargs_open_constr gl t + with _ -> 5 + +let pf_nbargs gl c = nbargs_open_constr gl (project gl, c) + +let isAppInd gl c = + try ignore (pf_reduce_to_atomic_ind gl c); true with _ -> false + +let interp_view_nbimps ist gl rc = + try + let sigma, t = interp_open_constr ist gl (rc, None) in + let si = sig_it gl in + let gl = re_sig si sigma in + let pl, c = splay_open_constr gl t in + if isAppInd gl c then List.length pl else ~-(List.length pl) + with _ -> 0 + +(* }}} *) + +(** Vernacular commands: Prenex Implicits and Search {{{ **********************) + +(* This should really be implemented as an extension to the implicit *) +(* arguments feature, but unfortuately that API is sealed. The current *) +(* workaround uses a combination of notations that works reasonably, *) +(* with the following caveats: *) +(* - The pretty-printing always elides prenex implicits, even when *) +(* they are obviously needed. *) +(* - Prenex Implicits are NEVER exported from a module, because this *) +(* would lead to faulty pretty-printing and scoping errors. *) +(* - The command "Import Prenex Implicits" can be used to reassert *) +(* Prenex Implicits for all the visible constants that had been *) +(* declared as Prenex Implicits. *) + +let declare_one_prenex_implicit locality f = + let fref = + try Smartlocate.global_with_alias f + with _ -> errorstrm (pr_reference f ++ str " is not declared") in + let rec loop = function + | a :: args' when Impargs.is_status_implicit a -> + (ExplByName (Impargs.name_of_implicit a), (true, true, true)) :: loop args' + | args' when List.exists Impargs.is_status_implicit args' -> + errorstrm (str "Expected prenex implicits for " ++ pr_reference f) + | _ -> [] in + let impls = + match Impargs.implicits_of_global fref with + | [cond,impls] -> impls + | [] -> errorstrm (str "Expected some implicits for " ++ pr_reference f) + | _ -> errorstrm (str "Multiple implicits not supported") in + match loop impls with + | [] -> + errorstrm (str "Expected some implicits for " ++ pr_reference f) + | impls -> + Impargs.declare_manual_implicits locality fref ~enriching:false [impls] + +VERNAC COMMAND EXTEND Ssrpreneximplicits CLASSIFIED AS SIDEFF + | [ "Prenex" "Implicits" ne_global_list(fl) ] + -> [ let locality = + Locality.make_section_locality (Locality.LocalityFixme.consume ()) in + List.iter (declare_one_prenex_implicit locality) fl ] +END + +(* Vernac grammar visibility patch *) + +GEXTEND Gram + GLOBAL: gallina_ext; + gallina_ext: + [ [ IDENT "Import"; IDENT "Prenex"; IDENT "Implicits" -> + Vernacexpr.VernacUnsetOption (["Printing"; "Implicit"; "Defensive"]) + ] ] + ; +END + +(** Extend Search to subsume SearchAbout, also adding hidden Type coercions. *) + +(* Main prefilter *) + +type raw_glob_search_about_item = + | RGlobSearchSubPattern of constr_expr + | RGlobSearchString of Loc.t * string * string option + +let pr_search_item = function + | RGlobSearchString (_,s,_) -> str s + | RGlobSearchSubPattern p -> pr_constr_expr p + +let wit_ssr_searchitem = add_genarg "ssr_searchitem" pr_search_item + +let interp_search_notation loc s opt_scope = + try + let interp = Notation.interp_notation_as_global_reference loc in + let ref = interp (fun _ -> true) s opt_scope in + Search.GlobSearchSubPattern (Pattern.PRef ref) + with _ -> + let diagnosis = + try + let ntns = Notation.locate_notation pr_glob_constr s opt_scope in + let ambig = "This string refers to a complex or ambiguous notation." in + str ambig ++ str "\nTry searching with one of\n" ++ ntns + with _ -> str "This string is not part of an identifier or notation." in + Errors.user_err_loc (loc, "interp_search_notation", diagnosis) + +let pr_ssr_search_item _ _ _ = pr_search_item + +(* Workaround the notation API that can only print notations *) + +let is_ident s = try Lexer.check_ident s; true with _ -> false + +let is_ident_part s = is_ident ("H" ^ s) + +let interp_search_notation loc tag okey = + let err msg = Errors.user_err_loc (loc, "interp_search_notation", msg) in + let mk_pntn s for_key = + let n = String.length s in + let s' = String.make (n + 2) ' ' in + let rec loop i i' = + if i >= n then s', i' - 2 else if s.[i] = ' ' then loop (i + 1) i' else + let j = try String.index_from s (i + 1) ' ' with _ -> n in + let m = j - i in + if s.[i] = '\'' && i < j - 2 && s.[j - 1] = '\'' then + (String.blit s (i + 1) s' i' (m - 2); loop (j + 1) (i' + m - 1)) + else if for_key && is_ident (String.sub s i m) then + (s'.[i'] <- '_'; loop (j + 1) (i' + 2)) + else (String.blit s i s' i' m; loop (j + 1) (i' + m + 1)) in + loop 0 1 in + let trim_ntn (pntn, m) = String.sub pntn 1 (max 0 m) in + let pr_ntn ntn = str "(" ++ str ntn ++ str ")" in + let pr_and_list pr = function + | [x] -> pr x + | x :: lx -> pr_list pr_comma pr lx ++ pr_comma () ++ str "and " ++ pr x + | [] -> mt () in + let pr_sc sc = str (if sc = "" then "independently" else sc) in + let pr_scs = function + | [""] -> pr_sc "" + | scs -> str "in " ++ pr_and_list pr_sc scs in + let generator, pr_tag_sc = + let ign _ = mt () in match okey with + | Some key -> + let sc = Notation.find_delimiters_scope loc key in + let pr_sc s_in = str s_in ++ spc() ++ str sc ++ pr_comma() in + Notation.pr_scope ign sc, pr_sc + | None -> Notation.pr_scopes ign, ign in + let qtag s_in = pr_tag_sc s_in ++ qstring tag ++ spc()in + let ptag, ttag = + let ptag, m = mk_pntn tag false in + if m <= 0 then err (str "empty notation fragment"); + ptag, trim_ntn (ptag, m) in + let last = ref "" and last_sc = ref "" in + let scs = ref [] and ntns = ref [] in + let push_sc sc = match !scs with + | "" :: scs' -> scs := "" :: sc :: scs' + | scs' -> scs := sc :: scs' in + let get s _ _ = match !last with + | "Scope " -> last_sc := s; last := "" + | "Lonely notation" -> last_sc := ""; last := "" + | "\"" -> + let pntn, m = mk_pntn s true in + if String.string_contains pntn ptag then begin + let ntn = trim_ntn (pntn, m) in + match !ntns with + | [] -> ntns := [ntn]; scs := [!last_sc] + | ntn' :: _ when ntn' = ntn -> push_sc !last_sc + | _ when ntn = ttag -> ntns := ntn :: !ntns; scs := [!last_sc] + | _ :: ntns' when List.mem ntn ntns' -> () + | ntn' :: ntns' -> ntns := ntn' :: ntn :: ntns' + end; + last := "" + | _ -> last := s in + pp_with (Format.make_formatter get (fun _ -> ())) generator; + let ntn = match !ntns with + | [] -> + err (hov 0 (qtag "in" ++ str "does not occur in any notation")) + | ntn :: ntns' when ntn = ttag -> + if ntns' <> [] then begin + let pr_ntns' = pr_and_list pr_ntn ntns' in + msg_warning (hov 4 (qtag "In" ++ str "also occurs in " ++ pr_ntns')) + end; ntn + | [ntn] -> + msgnl (hov 4 (qtag "In" ++ str "is part of notation " ++ pr_ntn ntn)); ntn + | ntns' -> + let e = str "occurs in" ++ spc() ++ pr_and_list pr_ntn ntns' in + err (hov 4 (str "ambiguous: " ++ qtag "in" ++ e)) in + let (nvars, body), ((_, pat), osc) = match !scs with + | [sc] -> Notation.interp_notation loc ntn (None, [sc]) + | scs' -> + try Notation.interp_notation loc ntn (None, []) with _ -> + let e = pr_ntn ntn ++ spc() ++ str "is defined " ++ pr_scs scs' in + err (hov 4 (str "ambiguous: " ++ pr_tag_sc "in" ++ e)) in + let sc = Option.default "" osc in + let _ = + let m_sc = + if osc <> None then str "In " ++ str sc ++ pr_comma() else mt() in + let ntn_pat = trim_ntn (mk_pntn pat false) in + let rbody = glob_constr_of_notation_constr loc body in + let m_body = hov 0 (Constrextern.without_symbols prl_glob_constr rbody) in + let m = m_sc ++ pr_ntn ntn_pat ++ spc () ++ str "denotes " ++ m_body in + msgnl (hov 0 m) in + if List.length !scs > 1 then + let scs' = List.remove (=) sc !scs in + let w = pr_ntn ntn ++ str " is also defined " ++ pr_scs scs' in + msg_warning (hov 4 w) + else if String.string_contains ntn " .. " then + err (pr_ntn ntn ++ str " is an n-ary notation"); + let nvars = List.filter (fun (_,(_,typ)) -> typ = NtnTypeConstr) nvars in + let rec sub () = function + | NVar x when List.mem_assoc x nvars -> GPatVar (loc, (false, x)) + | c -> + glob_constr_of_notation_constr_with_binders loc (fun _ x -> (), x) sub () c in + let _, npat = Patternops.pattern_of_glob_constr (sub () body) in + Search.GlobSearchSubPattern npat + +ARGUMENT EXTEND ssr_search_item TYPED AS ssr_searchitem + PRINTED BY pr_ssr_search_item + | [ string(s) ] -> [ RGlobSearchString (loc,s,None) ] + | [ string(s) "%" preident(key) ] -> [ RGlobSearchString (loc,s,Some key) ] + | [ constr_pattern(p) ] -> [ RGlobSearchSubPattern p ] +END + +let pr_ssr_search_arg _ _ _ = + let pr_item (b, p) = str (if b then "-" else "") ++ pr_search_item p in + pr_list spc pr_item + +ARGUMENT EXTEND ssr_search_arg TYPED AS (bool * ssr_searchitem) list + PRINTED BY pr_ssr_search_arg + | [ "-" ssr_search_item(p) ssr_search_arg(a) ] -> [ (false, p) :: a ] + | [ ssr_search_item(p) ssr_search_arg(a) ] -> [ (true, p) :: a ] + | [ ] -> [ [] ] +END + +(* Main type conclusion pattern filter *) + +let rec splay_search_pattern na = function + | Pattern.PApp (fp, args) -> splay_search_pattern (na + Array.length args) fp + | Pattern.PLetIn (_, _, bp) -> splay_search_pattern na bp + | Pattern.PRef hr -> hr, na + | _ -> Errors.error "no head constant in head search pattern" + +let coerce_search_pattern_to_sort hpat = + let env = Global.env () and sigma = Evd.empty in + let mkPApp fp n_imps args = + let args' = Array.append (Array.make n_imps (Pattern.PMeta None)) args in + Pattern.PApp (fp, args') in + let hr, na = splay_search_pattern 0 hpat in + let dc, ht = + Reductionops.splay_prod env sigma (Universes.unsafe_type_of_global hr) in + let np = List.length dc in + if np < na then Errors.error "too many arguments in head search pattern" else + let hpat' = if np = na then hpat else mkPApp hpat (np - na) [||] in + let warn () = + msg_warning (str "Listing only lemmas with conclusion matching " ++ + pr_constr_pattern hpat') in + if isSort ht then begin warn (); true, hpat' end else + let filter_head, coe_path = + try + let _, cp = + Classops.lookup_path_to_sort_from (push_rels_assum dc env) sigma ht in + warn (); + true, cp + with _ -> false, [] in + let coerce hp coe_index = + let coe = Classops.get_coercion_value coe_index in + try + let coe_ref = reference_of_constr coe in + let n_imps = Option.get (Classops.hide_coercion coe_ref) in + mkPApp (Pattern.PRef coe_ref) n_imps [|hp|] + with _ -> + errorstrm (str "need explicit coercion " ++ pr_constr coe ++ spc () + ++ str "to interpret head search pattern as type") in + filter_head, List.fold_left coerce hpat' coe_path + +let rec interp_head_pat hpat = + let filter_head, p = coerce_search_pattern_to_sort hpat in + let rec loop c = match kind_of_term c with + | Cast (c', _, _) -> loop c' + | Prod (_, _, c') -> loop c' + | LetIn (_, _, _, c') -> loop c' + | _ -> Constr_matching.is_matching (Global.env()) Evd.empty p c in + filter_head, loop + +let all_true _ = true + +let rec interp_search_about args accu = match args with +| [] -> accu +| (flag, arg) :: rem -> + fun gr env typ -> + let ans = Search.search_about_filter arg gr env typ in + (if flag then ans else not ans) && interp_search_about rem accu gr env typ + +let interp_search_arg arg = + let arg = List.map (fun (x,arg) -> x, match arg with + | RGlobSearchString (loc,s,key) -> + if is_ident_part s then Search.GlobSearchString s else + interp_search_notation loc s key + | RGlobSearchSubPattern p -> + try + let intern = Constrintern.intern_constr_pattern in + Search.GlobSearchSubPattern (snd (intern (Global.env()) p)) + with e -> let e = Errors.push e in iraise (Cerrors.process_vernac_interp_error e)) arg in + let hpat, a1 = match arg with + | (_, Search.GlobSearchSubPattern (Pattern.PMeta _)) :: a' -> all_true, a' + | (true, Search.GlobSearchSubPattern p) :: a' -> + let filter_head, p = interp_head_pat p in + if filter_head then p, a' else all_true, arg + | _ -> all_true, arg in + let is_string = + function (_, Search.GlobSearchString _) -> true | _ -> false in + let a2, a3 = List.partition is_string a1 in + interp_search_about (a2 @ a3) (fun gr env typ -> hpat typ) + +(* Module path postfilter *) + +let pr_modloc (b, m) = if b then str "-" ++ pr_reference m else pr_reference m + +let wit_ssrmodloc = add_genarg "ssrmodloc" pr_modloc + +let pr_ssr_modlocs _ _ _ ml = + if ml = [] then str "" else spc () ++ str "in " ++ pr_list spc pr_modloc ml + +ARGUMENT EXTEND ssr_modlocs TYPED AS ssrmodloc list PRINTED BY pr_ssr_modlocs + | [ ] -> [ [] ] +END + +GEXTEND Gram + GLOBAL: ssr_modlocs; + modloc: [[ "-"; m = global -> true, m | m = global -> false, m]]; + ssr_modlocs: [[ "in"; ml = LIST1 modloc -> ml ]]; +END + +let interp_modloc mr = + let interp_mod (_, mr) = + let (loc, qid) = qualid_of_reference mr in + try Nametab.full_name_module qid with Not_found -> + Errors.user_err_loc (loc, "interp_modloc", str "No Module " ++ pr_qualid qid) in + let mr_out, mr_in = List.partition fst mr in + let interp_bmod b = function + | [] -> fun _ _ _ -> true + | rmods -> Search.module_filter (List.map interp_mod rmods, b) in + let is_in = interp_bmod false mr_in and is_out = interp_bmod true mr_out in + fun gr env typ -> is_in gr env typ && is_out gr env typ + +(* The unified, extended vernacular "Search" command *) + +let ssrdisplaysearch gr env t = + let pr_res = pr_global gr ++ spc () ++ str " " ++ pr_lconstr_env env Evd.empty t in + msg (hov 2 pr_res ++ fnl ()) + +VERNAC COMMAND EXTEND SsrSearchPattern CLASSIFIED AS QUERY +| [ "Search" ssr_search_arg(a) ssr_modlocs(mr) ] -> + [ let hpat = interp_search_arg a in + let in_mod = interp_modloc mr in + let post_filter gr env typ = in_mod gr env typ && hpat gr env typ in + let display gr env typ = + if post_filter gr env typ then ssrdisplaysearch gr env typ + in + Search.generic_search None display ] +END + +(* }}} *) + +(** Alternative notations for "match" and anonymous arguments. {{{ ************) + +(* Syntax: *) +(* if <term> is <pattern> then ... else ... *) +(* if <term> is <pattern> [in ..] return ... then ... else ... *) +(* let: <pattern> := <term> in ... *) +(* let: <pattern> [in ...] := <term> return ... in ... *) +(* The scope of a top-level 'as' in the pattern extends over the *) +(* 'return' type (dependent if/let). *) +(* Note that the optional "in ..." appears next to the <pattern> *) +(* rather than the <term> in then "let:" syntax. The alternative *) +(* would lead to ambiguities in, e.g., *) +(* let: p1 := (*v---INNER LET:---v *) *) +(* let: p2 := let: p3 := e3 in k return t in k2 in k1 return t' *) +(* in b (*^--ALTERNATIVE INNER LET--------^ *) *) + +(* Caveat : There is no pretty-printing support, since this would *) +(* require a modification to the Coq kernel (adding a new match *) +(* display style -- why aren't these strings?); also, the v8.1 *) +(* pretty-printer only allows extension hooks for printing *) +(* integer or string literals. *) +(* Also note that in the v8 grammar "is" needs to be a keyword; *) +(* as this can't be done from an ML extension file, the new *) +(* syntax will only work when ssreflect.v is imported. *) + +let no_ct = None, None and no_rt = None in +let aliasvar = function + | [_, [CPatAlias (loc, _, id)]] -> Some (loc,Name id) + | _ -> None in +let mk_cnotype mp = aliasvar mp, None in +let mk_ctype mp t = aliasvar mp, Some t in +let mk_rtype t = Some t in +let mk_dthen loc (mp, ct, rt) c = (loc, mp, c), ct, rt in +let mk_let loc rt ct mp c1 = + CCases (loc, LetPatternStyle, rt, ct, [loc, mp, c1]) in +GEXTEND Gram + GLOBAL: binder_constr; + ssr_rtype: [[ "return"; t = operconstr LEVEL "100" -> mk_rtype t ]]; + ssr_mpat: [[ p = pattern -> [!@loc, [p]] ]]; + ssr_dpat: [ + [ mp = ssr_mpat; "in"; t = pattern; rt = ssr_rtype -> mp, mk_ctype mp t, rt + | mp = ssr_mpat; rt = ssr_rtype -> mp, mk_cnotype mp, rt + | mp = ssr_mpat -> mp, no_ct, no_rt + ] ]; + ssr_dthen: [[ dp = ssr_dpat; "then"; c = lconstr -> mk_dthen !@loc dp c ]]; + ssr_elsepat: [[ "else" -> [!@loc, [CPatAtom (!@loc, None)]] ]]; + ssr_else: [[ mp = ssr_elsepat; c = lconstr -> !@loc, mp, c ]]; + binder_constr: [ + [ "if"; c = operconstr LEVEL "200"; "is"; db1 = ssr_dthen; b2 = ssr_else -> + let b1, ct, rt = db1 in CCases (!@loc, MatchStyle, rt, [c, ct], [b1; b2]) + | "if"; c = operconstr LEVEL "200";"isn't";db1 = ssr_dthen; b2 = ssr_else -> + let b1, ct, rt = db1 in + let b1, b2 = + let (l1, p1, r1), (l2, p2, r2) = b1, b2 in (l1, p1, r2), (l2, p2, r1) in + CCases (!@loc, MatchStyle, rt, [c, ct], [b1; b2]) + | "let"; ":"; mp = ssr_mpat; ":="; c = lconstr; "in"; c1 = lconstr -> + mk_let (!@loc) no_rt [c, no_ct] mp c1 + | "let"; ":"; mp = ssr_mpat; ":="; c = lconstr; + rt = ssr_rtype; "in"; c1 = lconstr -> + mk_let (!@loc) rt [c, mk_cnotype mp] mp c1 + | "let"; ":"; mp = ssr_mpat; "in"; t = pattern; ":="; c = lconstr; + rt = ssr_rtype; "in"; c1 = lconstr -> + mk_let (!@loc) rt [c, mk_ctype mp t] mp c1 + ] ]; +END + +GEXTEND Gram + GLOBAL: closed_binder; + closed_binder: [ + [ ["of" | "&"]; c = operconstr LEVEL "99" -> + [LocalRawAssum ([!@loc, Anonymous], Default Explicit, c)] + ] ]; +END +(* }}} *) + +(** Tacticals (+, -, *, done, by, do, =>, first, and last). {{{ ***************) + +(** Bracketing tactical *) + +(* The tactic pretty-printer doesn't know that some extended tactics *) +(* are actually tacticals. To prevent it from improperly removing *) +(* parentheses we override the parsing rule for bracketed tactic *) +(* expressions so that the pretty-print always reflects the input. *) +(* (Removing user-specified parentheses is dubious anyway). *) + +GEXTEND Gram + GLOBAL: tactic_expr; + ssrparentacarg: [[ "("; tac = tactic_expr; ")" -> !@loc, Tacexp tac ]]; + tactic_expr: LEVEL "0" [[ arg = ssrparentacarg -> TacArg arg ]]; +END + +(** The internal "done" and "ssrautoprop" tactics. *) + +(* For additional flexibility, "done" and "ssrautoprop" are *) +(* defined in Ltac. *) +(* Although we provide a default definition in ssreflect, *) +(* we look up the definition dynamically at each call point, *) +(* to allow for user extensions. "ssrautoprop" defaults to *) +(* trivial. *) + +let donetac gl = + let tacname = + try Nametab.locate_tactic (qualid_of_ident (id_of_string "done")) + with Not_found -> try Nametab.locate_tactic (ssrqid "done") + with Not_found -> Errors.error "The ssreflect library was not loaded" in + let tacexpr = dummy_loc, Tacexpr.Reference (ArgArg (dummy_loc, tacname)) in + Proofview.V82.of_tactic (eval_tactic (Tacexpr.TacArg tacexpr)) gl + +let prof_donetac = mk_profiler "donetac";; +let donetac gl = prof_donetac.profile donetac gl;; + +let ssrautoprop gl = + try + let tacname = + try Nametab.locate_tactic (qualid_of_ident (id_of_string "ssrautoprop")) + with Not_found -> Nametab.locate_tactic (ssrqid "ssrautoprop") in + let tacexpr = dummy_loc, Tacexpr.Reference (ArgArg (dummy_loc, tacname)) in + Proofview.V82.of_tactic (eval_tactic (Tacexpr.TacArg tacexpr)) gl + with Not_found -> Proofview.V82.of_tactic (Auto.full_trivial []) gl + +let () = ssrautoprop_tac := ssrautoprop + +let tclBY tac = tclTHEN tac donetac + +(** Tactical arguments. *) + +(* We have four kinds: simple tactics, [|]-bracketed lists, hints, and swaps *) +(* The latter two are used in forward-chaining tactics (have, suffice, wlog) *) +(* and subgoal reordering tacticals (; first & ; last), respectively. *) + +(* Force use of the tactic_expr parsing entry, to rule out tick marks. *) +let pr_ssrtacarg _ _ prt = prt tacltop +ARGUMENT EXTEND ssrtacarg TYPED AS tactic PRINTED BY pr_ssrtacarg +| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] +END +GEXTEND Gram + GLOBAL: ssrtacarg; + ssrtacarg: [[ tac = tactic_expr LEVEL "5" -> tac ]]; +END + +(* Lexically closed tactic for tacticals. *) +let pr_ssrtclarg _ _ prt tac = prt tacltop tac +ARGUMENT EXTEND ssrtclarg TYPED AS ssrtacarg + PRINTED BY pr_ssrtclarg +| [ ssrtacarg(tac) ] -> [ tac ] +END +let eval_tclarg ist tac = ssrevaltac ist tac + +let pr_ortacs prt = + let rec pr_rec = function + | [None] -> spc() ++ str "|" ++ spc() + | None :: tacs -> spc() ++ str "|" ++ pr_rec tacs + | Some tac :: tacs -> spc() ++ str "| " ++ prt tacltop tac ++ pr_rec tacs + | [] -> mt() in + function + | [None] -> spc() + | None :: tacs -> pr_rec tacs + | Some tac :: tacs -> prt tacltop tac ++ pr_rec tacs + | [] -> mt() +let pr_ssrortacs _ _ = pr_ortacs + +ARGUMENT EXTEND ssrortacs TYPED AS tactic option list PRINTED BY pr_ssrortacs +| [ ssrtacarg(tac) "|" ssrortacs(tacs) ] -> [ Some tac :: tacs ] +| [ ssrtacarg(tac) "|" ] -> [ [Some tac; None] ] +| [ ssrtacarg(tac) ] -> [ [Some tac] ] +| [ "|" ssrortacs(tacs) ] -> [ None :: tacs ] +| [ "|" ] -> [ [None; None] ] +END + +let pr_hintarg prt = function + | true, tacs -> hv 0 (str "[ " ++ pr_ortacs prt tacs ++ str " ]") + | false, [Some tac] -> prt tacltop tac + | _, _ -> mt() + +let pr_ssrhintarg _ _ = pr_hintarg + +let mk_hint tac = false, [Some tac] +let mk_orhint tacs = true, tacs +let nullhint = true, [] +let nohint = false, [] + +ARGUMENT EXTEND ssrhintarg TYPED AS bool * ssrortacs PRINTED BY pr_ssrhintarg +| [ "[" "]" ] -> [ nullhint ] +| [ "[" ssrortacs(tacs) "]" ] -> [ mk_orhint tacs ] +| [ ssrtacarg(arg) ] -> [ mk_hint arg ] +END + +ARGUMENT EXTEND ssrortacarg TYPED AS ssrhintarg PRINTED BY pr_ssrhintarg +| [ "[" ssrortacs(tacs) "]" ] -> [ mk_orhint tacs ] +END + +let hinttac ist is_by (is_or, atacs) = + let dtac = if is_by then donetac else tclIDTAC in + let mktac = function + | Some atac -> tclTHEN (ssrevaltac ist atac) dtac + | _ -> dtac in + match List.map mktac atacs with + | [] -> if is_or then dtac else tclIDTAC + | [tac] -> tac + | tacs -> tclFIRST tacs + +(** The "-"/"+"/"*" tacticals. *) + +(* These are just visual cues to flag the beginning of the script for *) +(* new subgoals, when indentation is not appropriate (typically after *) +(* tactics that generate more than two subgoals). *) + +TACTIC EXTEND ssrtclplus +| [ "YouShouldNotTypeThis" "+" ssrtclarg(arg) ] -> [ Proofview.V82.tactic (eval_tclarg ist arg) ] +END +set_pr_ssrtac "tclplus" 5 [ArgSep "+ "; ArgSsr "tclarg"] + +TACTIC EXTEND ssrtclminus +| [ "YouShouldNotTypeThis" "-" ssrtclarg(arg) ] -> [ Proofview.V82.tactic (eval_tclarg ist arg) ] +END +set_pr_ssrtac "tclminus" 5 [ArgSep "- "; ArgSsr "tclarg"] + +TACTIC EXTEND ssrtclstar +| [ "YouShouldNotTypeThis" "*" ssrtclarg(arg) ] -> [ Proofview.V82.tactic (eval_tclarg ist arg) ] +END +set_pr_ssrtac "tclstar" 5 [ArgSep "- "; ArgSsr "tclarg"] + +let gen_tclarg = in_gen (rawwit wit_ssrtclarg) + +GEXTEND Gram + GLOBAL: tactic tactic_mode; + tactic: [ + [ "+"; tac = ssrtclarg -> ssrtac_expr !@loc "tclplus" [gen_tclarg tac] + | "-"; tac = ssrtclarg -> ssrtac_expr !@loc "tclminus" [gen_tclarg tac] + | "*"; tac = ssrtclarg -> ssrtac_expr !@loc "tclstar" [gen_tclarg tac] + ] ]; + tactic_mode: [ + [ "+"; tac = G_vernac.subgoal_command -> tac None + | "-"; tac = G_vernac.subgoal_command -> tac None + | "*"; tac = G_vernac.subgoal_command -> tac None + ] ]; +END + +(** The "by" tactical. *) + +let pr_hint prt arg = + if arg = nohint then mt() else str "by " ++ pr_hintarg prt arg +let pr_ssrhint _ _ = pr_hint + +ARGUMENT EXTEND ssrhint TYPED AS ssrhintarg PRINTED BY pr_ssrhint +| [ ] -> [ nohint ] +END + +TACTIC EXTEND ssrtclby +| [ "YouShouldNotTypeThis" ssrhint(tac) ] -> [ Proofview.V82.tactic (hinttac ist true tac) ] +END +set_pr_ssrtac "tclby" 0 [ArgSsr "hint"] + +(* We can't parse "by" in ARGUMENT EXTEND because it will only be made *) +(* into a keyword in ssreflect.v; so we anticipate this in GEXTEND. *) + +GEXTEND Gram + GLOBAL: ssrhint simple_tactic; + ssrhint: [[ "by"; arg = ssrhintarg -> arg ]]; + simple_tactic: [ + [ "by"; arg = ssrhintarg -> + let garg = in_gen (rawwit wit_ssrhint) arg in + ssrtac_atom !@loc "tclby" [garg] + ] ]; +END +(* }}} *) + +(** Bound assumption argument *) + +(* The Ltac API does have a type for assumptions but it is level-dependent *) +(* and therefore impratical to use for complex arguments, so we substitute *) +(* our own to have a uniform representation. Also, we refuse to intern *) +(* idents that match global/section constants, since this would lead to *) +(* fragile Ltac scripts. *) + +type ssrhyp = SsrHyp of loc * identifier + +let hyp_id (SsrHyp (_, id)) = id +let pr_hyp (SsrHyp (_, id)) = pr_id id +let pr_ssrhyp _ _ _ = pr_hyp + +let wit_ssrhyprep = add_genarg "ssrhyprep" pr_hyp + +let hyp_err loc msg id = + Errors.user_err_loc (loc, "ssrhyp", str msg ++ pr_id id) + +let intern_hyp ist (SsrHyp (loc, id) as hyp) = + let _ = Tacintern.intern_genarg ist (in_gen (rawwit wit_var) (loc, id)) in + if not_section_id id then hyp else + hyp_err loc "Can't clear section hypothesis " id + +let interp_hyp ist gl (SsrHyp (loc, id)) = + let s, id' = interp_wit wit_var ist gl (loc, id) in + if not_section_id id' then s, SsrHyp (loc, id') else + hyp_err loc "Can't clear section hypothesis " id' + +ARGUMENT EXTEND ssrhyp TYPED AS ssrhyprep PRINTED BY pr_ssrhyp + INTERPRETED BY interp_hyp + GLOBALIZED BY intern_hyp + | [ ident(id) ] -> [ SsrHyp (loc, id) ] +END + +type ssrhyp_or_id = Hyp of ssrhyp | Id of ssrhyp + +let hoik f = function Hyp x -> f x | Id x -> f x +let hoi_id = hoik hyp_id +let pr_hoi = hoik pr_hyp +let pr_ssrhoi _ _ _ = pr_hoi + +let wit_ssrhoirep = add_genarg "ssrhoirep" pr_hoi + +let intern_ssrhoi ist = function + | Hyp h -> Hyp (intern_hyp ist h) + | Id (SsrHyp (_, id)) as hyp -> + let _ = Tacintern.intern_genarg ist (in_gen (rawwit wit_ident) id) in + hyp + +let interp_ssrhoi ist gl = function + | Hyp h -> let s, h' = interp_hyp ist gl h in s, Hyp h' + | Id (SsrHyp (loc, id)) -> + let s, id' = interp_wit wit_ident ist gl id in + s, Id (SsrHyp (loc, id')) + +ARGUMENT EXTEND ssrhoi_hyp TYPED AS ssrhoirep PRINTED BY pr_ssrhoi + INTERPRETED BY interp_ssrhoi + GLOBALIZED BY intern_ssrhoi + | [ ident(id) ] -> [ Hyp (SsrHyp(loc, id)) ] +END +ARGUMENT EXTEND ssrhoi_id TYPED AS ssrhoirep PRINTED BY pr_ssrhoi + INTERPRETED BY interp_ssrhoi + GLOBALIZED BY intern_ssrhoi + | [ ident(id) ] -> [ Id (SsrHyp(loc, id)) ] +END + +type ssrhyps = ssrhyp list + +let pr_hyps = pr_list pr_spc pr_hyp +let pr_ssrhyps _ _ _ = pr_hyps +let hyps_ids = List.map hyp_id + +let rec check_hyps_uniq ids = function + | SsrHyp (loc, id) :: _ when List.mem id ids -> + hyp_err loc "Duplicate assumption " id + | SsrHyp (_, id) :: hyps -> check_hyps_uniq (id :: ids) hyps + | [] -> () + +let check_hyp_exists hyps (SsrHyp(_, id)) = + try ignore(Context.lookup_named id hyps) + with Not_found -> errorstrm (str"No assumption is named " ++ pr_id id) + +let interp_hyps ist gl ghyps = + let hyps = List.map snd (List.map (interp_hyp ist gl) ghyps) in + check_hyps_uniq [] hyps; Tacmach.project gl, hyps + +ARGUMENT EXTEND ssrhyps TYPED AS ssrhyp list PRINTED BY pr_ssrhyps + INTERPRETED BY interp_hyps + | [ ssrhyp_list(hyps) ] -> [ check_hyps_uniq [] hyps; hyps ] +END + +(** Terms parsing. {{{ ********************************************************) + +(* Because we allow wildcards in term references, we need to stage the *) +(* interpretation of terms so that it occurs at the right time during *) +(* the execution of the tactic (e.g., so that we don't report an error *) +(* for a term that isn't actually used in the execution). *) +(* The term representation tracks whether the concrete initial term *) +(* started with an opening paren, which might avoid a conflict between *) +(* the ssrreflect term syntax and Gallina notation. *) + +(* kinds of terms *) + +type ssrtermkind = char (* print flag *) + +let input_ssrtermkind strm = match Compat.get_tok (stream_nth 0 strm) with + | Tok.KEYWORD "(" -> '(' + | Tok.KEYWORD "@" -> '@' + | _ -> ' ' + +let ssrtermkind = Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind + +let id_of_Cterm t = match id_of_cpattern t with + | Some x -> x + | None -> loc_error (loc_of_cpattern t) "Only identifiers are allowed here" +let ssrhyp_of_ssrterm = function + | k, (_, Some c) as o -> + SsrHyp (constr_loc c, id_of_Cterm (cpattern_of_term o)), String.make 1 k + | _, (_, None) -> assert false + +(* terms *) +let pr_ssrterm _ _ _ = pr_term +let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) c +let intern_term ist sigma env (_, c) = glob_constr ist env c +let interp_term ist gl (_, c) = snd (interp_open_constr ist gl c) +let force_term ist gl (_, c) = interp_constr ist gl c +let glob_ssrterm gs = function + | k, (_, Some c) -> k, Tacintern.intern_constr gs c + | ct -> ct +let subst_ssrterm s (k, c) = k, Tacsubst.subst_glob_constr_and_expr s c +let interp_ssrterm _ gl t = Tacmach.project gl, t + +ARGUMENT EXTEND ssrterm + PRINTED BY pr_ssrterm + INTERPRETED BY interp_ssrterm + GLOBALIZED BY glob_ssrterm SUBSTITUTED BY subst_ssrterm + RAW_TYPED AS cpattern RAW_PRINTED BY pr_ssrterm + GLOB_TYPED AS cpattern GLOB_PRINTED BY pr_ssrterm +| [ "YouShouldNotTypeThis" constr(c) ] -> [ mk_lterm c ] +END + +GEXTEND Gram + GLOBAL: ssrterm; + ssrterm: [[ k = ssrtermkind; c = constr -> mk_term k c ]]; +END +(* }}} *) + +(** The "in" pseudo-tactical {{{ **********************************************) + +(* We can't make "in" into a general tactical because this would create a *) +(* crippling conflict with the ltac let .. in construct. Hence, we add *) +(* explicitly an "in" suffix to all the extended tactics for which it is *) +(* relevant (including move, case, elim) and to the extended do tactical *) +(* below, which yields a general-purpose "in" of the form do [...] in ... *) + +(* This tactical needs to come before the intro tactics because the latter *) +(* must take precautions in order not to interfere with the discharged *) +(* assumptions. This is especially difficult for discharged "let"s, which *) +(* the default simpl and unfold tactics would erase blindly. *) + +(** Clear switch *) + +type ssrclear = ssrhyps + +let pr_clear_ne clr = str "{" ++ pr_hyps clr ++ str "}" +let pr_clear sep clr = if clr = [] then mt () else sep () ++ pr_clear_ne clr + +let pr_ssrclear _ _ _ = pr_clear mt + +ARGUMENT EXTEND ssrclear_ne TYPED AS ssrhyps PRINTED BY pr_ssrclear +| [ "{" ne_ssrhyp_list(clr) "}" ] -> [ check_hyps_uniq [] clr; clr ] +END + +ARGUMENT EXTEND ssrclear TYPED AS ssrclear_ne PRINTED BY pr_ssrclear +| [ ssrclear_ne(clr) ] -> [ clr ] +| [ ] -> [ [] ] +END + +let cleartac clr = check_hyps_uniq [] clr; clear (hyps_ids clr) + +(* type ssrwgen = ssrclear * ssrhyp * string *) + +let pr_wgen = function + | (clr, Some((id,k),None)) -> spc() ++ pr_clear mt clr ++ str k ++ pr_hoi id + | (clr, Some((id,k),Some p)) -> + spc() ++ pr_clear mt clr ++ str"(" ++ str k ++ pr_hoi id ++ str ":=" ++ + pr_cpattern p ++ str ")" + | (clr, None) -> spc () ++ pr_clear mt clr +let pr_ssrwgen _ _ _ = pr_wgen + +(* no globwith for char *) +ARGUMENT EXTEND ssrwgen + TYPED AS ssrclear * ((ssrhoi_hyp * string) * cpattern option) option + PRINTED BY pr_ssrwgen +| [ ssrclear_ne(clr) ] -> [ clr, None ] +| [ ssrhoi_hyp(hyp) ] -> [ [], Some((hyp, " "), None) ] +| [ "@" ssrhoi_hyp(hyp) ] -> [ [], Some((hyp, "@"), None) ] +| [ "(" ssrhoi_id(id) ":=" lcpattern(p) ")" ] -> + [ [], Some ((id," "),Some p) ] +| [ "(" ssrhoi_id(id) ")" ] -> [ [], Some ((id,"("), None) ] +| [ "(@" ssrhoi_id(id) ":=" lcpattern(p) ")" ] -> + [ [], Some ((id,"@"),Some p) ] +| [ "(" "@" ssrhoi_id(id) ":=" lcpattern(p) ")" ] -> + [ [], Some ((id,"@"),Some p) ] +END + +type ssrclseq = InGoal | InHyps + | InHypsGoal | InHypsSeqGoal | InSeqGoal | InHypsSeq | InAll | InAllHyps + +let pr_clseq = function + | InGoal | InHyps -> mt () + | InSeqGoal -> str "|- *" + | InHypsSeqGoal -> str " |- *" + | InHypsGoal -> str " *" + | InAll -> str "*" + | InHypsSeq -> str " |-" + | InAllHyps -> str "* |-" + +let wit_ssrclseq = add_genarg "ssrclseq" pr_clseq +let pr_clausehyps = pr_list pr_spc pr_wgen +let pr_ssrclausehyps _ _ _ = pr_clausehyps + +ARGUMENT EXTEND ssrclausehyps +TYPED AS ssrwgen list PRINTED BY pr_ssrclausehyps +| [ ssrwgen(hyp) "," ssrclausehyps(hyps) ] -> [ hyp :: hyps ] +| [ ssrwgen(hyp) ssrclausehyps(hyps) ] -> [ hyp :: hyps ] +| [ ssrwgen(hyp) ] -> [ [hyp] ] +END + +(* type ssrclauses = ssrahyps * ssrclseq *) + +let pr_clauses (hyps, clseq) = + if clseq = InGoal then mt () + else str "in " ++ pr_clausehyps hyps ++ pr_clseq clseq +let pr_ssrclauses _ _ _ = pr_clauses + +ARGUMENT EXTEND ssrclauses TYPED AS ssrwgen list * ssrclseq + PRINTED BY pr_ssrclauses + | [ "in" ssrclausehyps(hyps) "|-" "*" ] -> [ hyps, InHypsSeqGoal ] + | [ "in" ssrclausehyps(hyps) "|-" ] -> [ hyps, InHypsSeq ] + | [ "in" ssrclausehyps(hyps) "*" ] -> [ hyps, InHypsGoal ] + | [ "in" ssrclausehyps(hyps) ] -> [ hyps, InHyps ] + | [ "in" "|-" "*" ] -> [ [], InSeqGoal ] + | [ "in" "*" ] -> [ [], InAll ] + | [ "in" "*" "|-" ] -> [ [], InAllHyps ] + | [ ] -> [ [], InGoal ] +END + +let nohide = mkRel 0 +let hidden_goal_tag = "the_hidden_goal" + +(* Reduction that preserves the Prod/Let spine of the "in" tactical. *) + +let inc_safe n = if n = 0 then n else n + 1 +let rec safe_depth c = match kind_of_term c with +| LetIn (Name x, _, _, c') when is_discharged_id x -> safe_depth c' + 1 +| LetIn (_, _, _, c') | Prod (_, _, c') -> inc_safe (safe_depth c') +| _ -> 0 + +let red_safe r e s c0 = + let rec red_to e c n = match kind_of_term c with + | Prod (x, t, c') when n > 0 -> + let t' = r e s t in let e' = Environ.push_rel (x, None, t') e in + mkProd (x, t', red_to e' c' (n - 1)) + | LetIn (x, b, t, c') when n > 0 -> + let t' = r e s t in let e' = Environ.push_rel (x, None, t') e in + mkLetIn (x, r e s b, t', red_to e' c' (n - 1)) + | _ -> r e s c in + red_to e c0 (safe_depth c0) + +let check_wgen_uniq gens = + let clears = List.flatten (List.map fst gens) in + check_hyps_uniq [] clears; + let ids = CList.map_filter + (function (_,Some ((id,_),_)) -> Some (hoi_id id) | _ -> None) gens in + let rec check ids = function + | id :: _ when List.mem id ids -> + errorstrm (str"Duplicate generalization " ++ pr_id id) + | id :: hyps -> check (id :: ids) hyps + | [] -> () in + check [] ids + +let pf_clauseids gl gens clseq = + let keep_clears = List.map (fun x, _ -> x, None) in + if gens <> [] then (check_wgen_uniq gens; gens) else + if clseq <> InAll && clseq <> InAllHyps then keep_clears gens else + Errors.error "assumptions should be named explicitly" + +let hidden_clseq = function InHyps | InHypsSeq | InAllHyps -> true | _ -> false + +let hidetacs clseq idhide cl0 = + if not (hidden_clseq clseq) then [] else + [posetac idhide cl0; + Proofview.V82.of_tactic (convert_concl_no_check (mkVar idhide))] + +let discharge_hyp (id', (id, mode)) gl = + let cl' = subst_var id (pf_concl gl) in + match pf_get_hyp gl id, mode with + | (_, None, t), _ | (_, Some _, t), "(" -> + apply_type (mkProd (Name id', t, cl')) [mkVar id] gl + | (_, Some v, t), _ -> + Proofview.V82.of_tactic (convert_concl (mkLetIn (Name id', v, t, cl'))) gl + +let endclausestac id_map clseq gl_id cl0 gl = + let not_hyp' id = not (List.mem_assoc id id_map) in + let orig_id id = try List.assoc id id_map with _ -> id in + let dc, c = Term.decompose_prod_assum (pf_concl gl) in + let hide_goal = hidden_clseq clseq in + let c_hidden = hide_goal && c = mkVar gl_id in + let rec fits forced = function + | (id, _) :: ids, (Name id', _, _) :: dc' when id' = id -> + fits true (ids, dc') + | ids, dc' -> + forced && ids = [] && (not hide_goal || dc' = [] && c_hidden) in + let rec unmark c = match kind_of_term c with + | Var id when hidden_clseq clseq && id = gl_id -> cl0 + | Prod (Name id, t, c') when List.mem_assoc id id_map -> + mkProd (Name (orig_id id), unmark t, unmark c') + | LetIn (Name id, v, t, c') when List.mem_assoc id id_map -> + mkLetIn (Name (orig_id id), unmark v, unmark t, unmark c') + | _ -> map_constr unmark c in + let utac hyp = + Proofview.V82.of_tactic + (convert_hyp_no_check (map_named_declaration unmark hyp)) in + let utacs = List.map utac (pf_hyps gl) in + let ugtac gl' = + Proofview.V82.of_tactic + (convert_concl_no_check (unmark (pf_concl gl'))) gl' in + let ctacs = if hide_goal then [clear [gl_id]] else [] in + let mktac itacs = tclTHENLIST (itacs @ utacs @ ugtac :: ctacs) in + let itac (_, id) = Proofview.V82.of_tactic (introduction id) in + if fits false (id_map, List.rev dc) then mktac (List.map itac id_map) gl else + let all_ids = ids_of_rel_context dc @ pf_ids_of_hyps gl in + if List.for_all not_hyp' all_ids && not c_hidden then mktac [] gl else + Errors.error "tampering with discharged assumptions of \"in\" tactical" + +let is_id_constr c = match kind_of_term c with + | Lambda(_,_,c) when isRel c -> 1 = destRel c + | _ -> false + +let red_product_skip_id env sigma c = match kind_of_term c with + | App(hd,args) when Array.length args = 1 && is_id_constr hd -> args.(0) + | _ -> try Tacred.red_product env sigma c with _ -> c + +let abs_wgen keep_let ist f gen (gl,args,c) = + let sigma, env = project gl, pf_env gl in + let evar_closed t p = + if occur_existential t then + Errors.user_err_loc (loc_of_cpattern p,"ssreflect", + pr_constr_pat t ++ + str" contains holes and matches no subterm of the goal") in + match gen with + | _, Some ((x, mode), None) when mode = "@" || (mode = " " && keep_let) -> + let x = hoi_id x in + let _, bo, ty = pf_get_hyp gl x in + gl, + (if bo <> None then args else mkVar x :: args), + mkProd_or_LetIn (Name (f x),bo,ty) (subst_var x c) + | _, Some ((x, _), None) -> + let x = hoi_id x in + gl, mkVar x :: args, mkProd (Name (f x), pf_get_hyp_typ gl x, subst_var x c) + | _, Some ((x, "@"), Some p) -> + let x = hoi_id x in + let cp = interp_cpattern ist gl p None in + let (t, ucst), c = + try fill_occ_pattern ~raise_NoMatch:true env sigma c cp None 1 + with NoMatch -> redex_of_pattern env cp, c in + evar_closed t p; + let ut = red_product_skip_id env sigma t in + pf_merge_uc ucst gl, args, mkLetIn(Name (f x), ut, pf_type_of gl t, c) + | _, Some ((x, _), Some p) -> + let x = hoi_id x in + let cp = interp_cpattern ist gl p None in + let (t, ucst), c = + try fill_occ_pattern ~raise_NoMatch:true env sigma c cp None 1 + with NoMatch -> redex_of_pattern env cp, c in + evar_closed t p; + pf_merge_uc ucst gl, t :: args, mkProd(Name (f x), pf_type_of gl t, c) + | _ -> gl, args, c + +let clr_of_wgen gen clrs = match gen with + | clr, Some ((x, _), None) -> + let x = hoi_id x in + cleartac clr :: cleartac [SsrHyp(Loc.ghost,x)] :: clrs + | clr, _ -> cleartac clr :: clrs + +let tclCLAUSES ist tac (gens, clseq) gl = + if clseq = InGoal || clseq = InSeqGoal then tac gl else + let clr_gens = pf_clauseids gl gens clseq in + let clear = tclTHENLIST (List.rev(List.fold_right clr_of_wgen clr_gens [])) in + let gl_id = mk_anon_id hidden_goal_tag gl in + let cl0 = pf_concl gl in + let dtac gl = + let c = pf_concl gl in + let gl, args, c = + List.fold_right (abs_wgen true ist mk_discharged_id) gens (gl,[], c) in + apply_type c args gl in + let endtac = + let id_map = CList.map_filter (function + | _, Some ((x,_),_) -> let id = hoi_id x in Some (mk_discharged_id id, id) + | _, None -> None) gens in + endclausestac id_map clseq gl_id cl0 in + tclTHENLIST (hidetacs clseq gl_id cl0 @ [dtac; clear; tac; endtac]) gl +(* }}} *) + +(** Simpl switch *) + +type ssrsimpl = Simpl | Cut | SimplCut | Nop + +let pr_simpl = function + | Simpl -> str "/=" + | Cut -> str "//" + | SimplCut -> str "//=" + | Nop -> mt () + +let pr_ssrsimpl _ _ _ = pr_simpl + +let wit_ssrsimplrep = add_genarg "ssrsimplrep" pr_simpl + +ARGUMENT EXTEND ssrsimpl_ne TYPED AS ssrsimplrep PRINTED BY pr_ssrsimpl +| [ "/=" ] -> [ Simpl ] +| [ "//" ] -> [ Cut ] +| [ "//=" ] -> [ SimplCut ] +END + +ARGUMENT EXTEND ssrsimpl TYPED AS ssrsimplrep PRINTED BY pr_ssrsimpl +| [ ssrsimpl_ne(sim) ] -> [ sim ] +| [ ] -> [ Nop ] +END + +(* We must avoid zeta-converting any "let"s created by the "in" tactical. *) + +let safe_simpltac gl = + let cl' = red_safe Tacred.simpl (pf_env gl) (project gl) (pf_concl gl) in + Proofview.V82.of_tactic (convert_concl_no_check cl') gl + +let simpltac = function + | Simpl -> safe_simpltac + | Cut -> tclTRY donetac + | SimplCut -> tclTHEN safe_simpltac (tclTRY donetac) + | Nop -> tclIDTAC + +(** Rewriting direction *) + +let pr_dir = function L2R -> str "->" | R2L -> str "<-" +let pr_rwdir = function L2R -> mt() | R2L -> str "-" + +let rewritetac dir c = + (* Due to the new optional arg ?tac, application shouldn't be too partial *) + Proofview.V82.of_tactic begin + Equality.general_rewrite (dir = L2R) AllOccurrences true false c + end + +let wit_ssrdir = add_genarg "ssrdir" pr_dir + +let dir_org = function L2R -> 1 | R2L -> 2 + +(** Indexes *) + +(* Since SSR indexes are always positive numbers, we use the 0 value *) +(* to encode an omitted index. We reuse the in or_var type, but we *) +(* supply our own interpretation function, which checks for non *) +(* positive values, and allows the use of constr numerals, so that *) +(* e.g., "let n := eval compute in (1 + 3) in (do n!clear)" works. *) + +type ssrindex = int or_var + +let pr_index = function + | ArgVar (_, id) -> pr_id id + | ArgArg n when n > 0 -> int n + | _ -> mt () +let pr_ssrindex _ _ _ = pr_index + +let noindex = ArgArg 0 +let allocc = Some(false,[]) + +let check_index loc i = + if i > 0 then i else loc_error loc "Index not positive" +let mk_index loc = function ArgArg i -> ArgArg (check_index loc i) | iv -> iv + +let interp_index ist gl idx = + Tacmach.project gl, + match idx with + | ArgArg _ -> idx + | ArgVar (loc, id) -> + let i = + try + let v = Id.Map.find id ist.lfun in + begin match Value.to_int v with + | Some i -> i + | None -> + begin match Value.to_constr v with + | Some c -> + let rc = Detyping.detype false [] (pf_env gl) (project gl) c in + begin match Notation.uninterp_prim_token rc with + | _, Numeral bigi -> int_of_string (Bigint.to_string bigi) + | _ -> raise Not_found + end + | None -> raise Not_found + end end + with _ -> loc_error loc "Index not a number" in + ArgArg (check_index loc i) + +ARGUMENT EXTEND ssrindex TYPED AS int_or_var PRINTED BY pr_ssrindex + INTERPRETED BY interp_index +| [ int_or_var(i) ] -> [ mk_index loc i ] +END + +(** Occurrence switch *) + +(* The standard syntax of complemented occurrence lists involves a single *) +(* initial "-", e.g., {-1 3 5}. An initial *) +(* "+" may be used to indicate positive occurrences (the default). The *) +(* "+" is optional, except if the list of occurrences starts with a *) +(* variable or is empty (to avoid confusion with a clear switch). The *) +(* empty positive switch "{+}" selects no occurrences, while the empty *) +(* negative switch "{-}" selects all occurrences explicitly; this is the *) +(* default, but "{-}" prevents the implicit clear, and can be used to *) +(* force dependent elimination -- see ndefectelimtac below. *) + +type ssrocc = occ + +let pr_occ = function + | Some (true, occ) -> str "{-" ++ pr_list pr_spc int occ ++ str "}" + | Some (false, occ) -> str "{+" ++ pr_list pr_spc int occ ++ str "}" + | None -> str "{}" + +let pr_ssrocc _ _ _ = pr_occ + +ARGUMENT EXTEND ssrocc TYPED AS (bool * int list) option PRINTED BY pr_ssrocc +| [ natural(n) natural_list(occ) ] -> [ + Some (false, List.map (check_index loc) (n::occ)) ] +| [ "-" natural_list(occ) ] -> [ Some (true, occ) ] +| [ "+" natural_list(occ) ] -> [ Some (false, occ) ] +END + +let pf_mkprod gl c ?(name=constr_name c) cl = + let t = pf_type_of gl c in + if name <> Anonymous || noccurn 1 cl then mkProd (name, t, cl) else + mkProd (Name (pf_type_id gl t), t, cl) + +let pf_abs_prod name gl c cl = pf_mkprod gl c ~name (subst_term c cl) + +(** Discharge occ switch (combined occurrence / clear switch *) + +type ssrdocc = ssrclear option * ssrocc option + +let mkocc occ = None, occ +let noclr = mkocc None +let mkclr clr = Some clr, None +let nodocc = mkclr [] + +let pr_docc = function + | None, occ -> pr_occ occ + | Some clr, _ -> pr_clear mt clr + +let pr_ssrdocc _ _ _ = pr_docc + +ARGUMENT EXTEND ssrdocc TYPED AS ssrclear option * ssrocc PRINTED BY pr_ssrdocc +| [ "{" ne_ssrhyp_list(clr) "}" ] -> [ mkclr clr ] +| [ "{" ssrocc(occ) "}" ] -> [ mkocc occ ] +END + +(** View hint database and View application. {{{ ******************************) + +(* There are three databases of lemmas used to mediate the application *) +(* of reflection lemmas: one for forward chaining, one for backward *) +(* chaining, and one for secondary backward chaining. *) + +(* View hints *) + +let rec isCxHoles = function (CHole _, None) :: ch -> isCxHoles ch | _ -> false + +let pr_raw_ssrhintref prc _ _ = function + | CAppExpl (_, (None, r,x), args) when isCHoles args -> + prc (CRef (r,x)) ++ str "|" ++ int (List.length args) + | CApp (_, (_, CRef _), _) as c -> prc c + | CApp (_, (_, c), args) when isCxHoles args -> + prc c ++ str "|" ++ int (List.length args) + | c -> prc c + +let pr_rawhintref = function + | GApp (_, f, args) when isRHoles args -> + pr_glob_constr f ++ str "|" ++ int (List.length args) + | c -> pr_glob_constr c + +let pr_glob_ssrhintref _ _ _ (c, _) = pr_rawhintref c + +let pr_ssrhintref prc _ _ = prc + +let mkhintref loc c n = match c with + | CRef (r,x) -> CAppExpl (loc, (None, r, x), mkCHoles loc n) + | _ -> mkAppC (c, mkCHoles loc n) + +ARGUMENT EXTEND ssrhintref + PRINTED BY pr_ssrhintref + RAW_TYPED AS constr RAW_PRINTED BY pr_raw_ssrhintref + GLOB_TYPED AS constr GLOB_PRINTED BY pr_glob_ssrhintref + | [ constr(c) ] -> [ c ] + | [ constr(c) "|" natural(n) ] -> [ mkhintref loc c n ] +END + +(* View purpose *) + +let pr_viewpos = function + | 0 -> str " for move/" + | 1 -> str " for apply/" + | 2 -> str " for apply//" + | _ -> mt () + +let pr_ssrviewpos _ _ _ = pr_viewpos + +ARGUMENT EXTEND ssrviewpos TYPED AS int PRINTED BY pr_ssrviewpos + | [ "for" "move" "/" ] -> [ 0 ] + | [ "for" "apply" "/" ] -> [ 1 ] + | [ "for" "apply" "/" "/" ] -> [ 2 ] + | [ "for" "apply" "//" ] -> [ 2 ] + | [ ] -> [ 3 ] +END + +let pr_ssrviewposspc _ _ _ i = pr_viewpos i ++ spc () + +ARGUMENT EXTEND ssrviewposspc TYPED AS ssrviewpos PRINTED BY pr_ssrviewposspc + | [ ssrviewpos(i) ] -> [ i ] +END + +(* The table and its display command *) + +let viewtab : glob_constr list array = Array.make 3 [] + +let _ = + let init () = Array.fill viewtab 0 3 [] in + let freeze _ = Array.copy viewtab in + let unfreeze vt = Array.blit vt 0 viewtab 0 3 in + Summary.declare_summary "ssrview" + { Summary.freeze_function = freeze; + Summary.unfreeze_function = unfreeze; + Summary.init_function = init } + +let mapviewpos f n k = if n < 3 then f n else for i = 0 to k - 1 do f i done + +let print_view_hints i = + let pp_viewname = str "Hint View" ++ pr_viewpos i ++ str " " in + let pp_hints = pr_list spc pr_rawhintref viewtab.(i) in + ppnl (pp_viewname ++ hov 0 pp_hints ++ Pp.cut ()) + +VERNAC COMMAND EXTEND PrintView CLASSIFIED AS QUERY +| [ "Print" "Hint" "View" ssrviewpos(i) ] -> [ mapviewpos print_view_hints i 3 ] +END + +(* Populating the table *) + +let cache_viewhint (_, (i, lvh)) = + let mem_raw h = List.exists (Notation_ops.eq_glob_constr h) in + let add_hint h hdb = if mem_raw h hdb then hdb else h :: hdb in + viewtab.(i) <- List.fold_right add_hint lvh viewtab.(i) + +let subst_viewhint ( subst, (i, lvh as ilvh)) = + let lvh' = List.smartmap (Detyping.subst_glob_constr subst) lvh in + if lvh' == lvh then ilvh else i, lvh' + +let classify_viewhint x = Libobject.Substitute x + +let in_viewhint = + Libobject.declare_object {(Libobject.default_object "VIEW_HINTS") with + Libobject.open_function = (fun i o -> if i = 1 then cache_viewhint o); + Libobject.cache_function = cache_viewhint; + Libobject.subst_function = subst_viewhint; + Libobject.classify_function = classify_viewhint } + +let glob_view_hints lvh = + List.map (Constrintern.intern_constr (Global.env ())) lvh + +let add_view_hints lvh i = Lib.add_anonymous_leaf (in_viewhint (i, lvh)) + +VERNAC COMMAND EXTEND HintView CLASSIFIED AS QUERY + | [ "Hint" "View" ssrviewposspc(n) ne_ssrhintref_list(lvh) ] -> + [ mapviewpos (add_view_hints (glob_view_hints lvh)) n 2 ] +END + +(** Views *) + +(* Views for the "move" and "case" commands are actually open *) +(* terms, but this is handled by interp_view, which is called *) +(* by interp_casearg. We use lists, to support the *) +(* "double-view" feature of the apply command. *) + +(* type ssrview = ssrterm list *) + +let pr_view = pr_list mt (fun c -> str "/" ++ pr_term c) + +let pr_ssrview _ _ _ = pr_view + +ARGUMENT EXTEND ssrview TYPED AS ssrterm list + PRINTED BY pr_ssrview +| [ "/" constr(c) ] -> [ [mk_term ' ' c] ] +| [ "/" constr(c) ssrview(w) ] -> [ (mk_term ' ' c) :: w ] +END + +(* There are two ways of "applying" a view to term: *) +(* 1- using a view hint if the view is an instance of some *) +(* (reflection) inductive predicate. *) +(* 2- applying the view if it coerces to a function, adding *) +(* implicit arguments. *) +(* They require guessing the view hints and the number of *) +(* implicits, respectively, which we do by brute force. *) + +let view_error s gv = + errorstrm (str ("Cannot " ^ s ^ " view ") ++ pr_term gv) + +let interp_view ist si env sigma gv rid = + match intern_term ist sigma env gv with + | GApp (loc, GHole _, rargs) -> + let rv = GApp (loc, rid, rargs) in + snd (interp_open_constr ist (re_sig si sigma) (rv, None)) + | rv -> + let interp rc rargs = + interp_open_constr ist (re_sig si sigma) (mkRApp rc rargs, None) in + let rec simple_view rargs n = + if n < 0 then view_error "use" gv else + try interp rv rargs with _ -> simple_view (mkRHole :: rargs) (n - 1) in + let view_nbimps = interp_view_nbimps ist (re_sig si sigma) rv in + let view_args = [mkRApp rv (mkRHoles view_nbimps); rid] in + let rec view_with = function + | [] -> simple_view [rid] (interp_nbargs ist (re_sig si sigma) rv) + | hint :: hints -> try interp hint view_args with _ -> view_with hints in + snd (view_with (if view_nbimps < 0 then [] else viewtab.(0))) + +let top_id = mk_internal_id "top assumption" + +let with_view ist si env gl0 c name cl prune = + let c2r ist x = { ist with lfun = + Id.Map.add top_id (Value.of_constr x) ist.lfun } in + let rec loop (sigma, c') = function + | f :: view -> + let rid, ist = match kind_of_term c' with + | Var id -> mkRVar id, ist + | _ -> mkRltacVar top_id, c2r ist c' in + loop (interp_view ist si env sigma f rid) view + | [] -> + let sigma = Typeclasses.resolve_typeclasses ~fail:false env sigma in + let c' = Reductionops.nf_evar sigma c' in + let n, c', _, ucst = pf_abs_evars gl0 (sigma, c') in + let c' = if not prune then c' else pf_abs_cterm gl0 n c' in + let gl0 = pf_merge_uc ucst gl0 in + pf_abs_prod name gl0 c' (prod_applist cl [c]), c', pf_merge_uc_of sigma gl0 + in loop + +let pf_with_view ist gl (prune, view) cl c = + let env, sigma, si = pf_env gl, project gl, sig_it gl in + with_view ist si env gl c (constr_name c) cl prune (sigma, c) view +(* }}} *) + +(** Extended intro patterns {{{ ***********************************************) + +type ssrtermrep = char * glob_constr_and_expr +type ssripat = + | IpatSimpl of ssrclear * ssrsimpl + | IpatId of identifier + | IpatWild + | IpatCase of ssripats list + | IpatRw of ssrocc * ssrdir + | IpatAll + | IpatAnon + | IpatView of ssrtermrep list + | IpatNoop + | IpatNewHidden of identifier list +and ssripats = ssripat list + +let remove_loc = snd + +let rec ipat_of_intro_pattern = function + | IntroNaming (IntroIdentifier id) -> IpatId id + | IntroAction IntroWildcard -> IpatWild + | IntroAction (IntroOrAndPattern iorpat) -> + IpatCase + (List.map (List.map ipat_of_intro_pattern) + (List.map (List.map remove_loc) iorpat)) + | IntroNaming IntroAnonymous -> IpatAnon + | IntroAction (IntroRewrite b) -> IpatRw (allocc, if b then L2R else R2L) + | IntroNaming (IntroFresh id) -> IpatAnon + | IntroAction (IntroApplyOn _) -> (* to do *) Errors.error "TO DO" + | IntroAction (IntroInjection ips) -> + IpatCase [List.map ipat_of_intro_pattern (List.map remove_loc ips)] + | IntroForthcoming _ -> (* Unable to determine which kind of ipat interp_introid could return [HH] *) + assert false + +let rec pr_ipat = function + | IpatId id -> pr_id id + | IpatSimpl (clr, sim) -> pr_clear mt clr ++ pr_simpl sim + | IpatCase iorpat -> hov 1 (str "[" ++ pr_iorpat iorpat ++ str "]") + | IpatRw (occ, dir) -> pr_occ occ ++ pr_dir dir + | IpatAll -> str "*" + | IpatWild -> str "_" + | IpatAnon -> str "?" + | IpatView v -> pr_view v + | IpatNoop -> str "-" + | IpatNewHidden l -> str "[:" ++ pr_list spc pr_id l ++ str "]" +and pr_iorpat iorpat = pr_list pr_bar pr_ipats iorpat +and pr_ipats ipats = pr_list spc pr_ipat ipats + +let wit_ssripatrep = add_genarg "ssripatrep" pr_ipat + +let pr_ssripat _ _ _ = pr_ipat +let pr_ssripats _ _ _ = pr_ipats +let pr_ssriorpat _ _ _ = pr_iorpat + +let intern_ipat ist ipat = + let rec check_pat = function + | IpatSimpl (clr, _) -> ignore (List.map (intern_hyp ist) clr) + | IpatCase iorpat -> List.iter (List.iter check_pat) iorpat + | _ -> () in + check_pat ipat; ipat + +let intern_ipats ist = List.map (intern_ipat ist) + +let interp_introid ist gl id = + try IntroNaming (IntroIdentifier (hyp_id (snd (interp_hyp ist gl (SsrHyp (dummy_loc, id)))))) + with _ -> snd(snd (interp_intro_pattern ist gl (dummy_loc,IntroNaming (IntroIdentifier id)))) + +let rec add_intro_pattern_hyps (loc, ipat) hyps = match ipat with + | IntroNaming (IntroIdentifier id) -> + if not_section_id id then SsrHyp (loc, id) :: hyps else + hyp_err loc "Can't delete section hypothesis " id + | IntroAction IntroWildcard -> hyps + | IntroAction (IntroOrAndPattern iorpat) -> + List.fold_right (List.fold_right add_intro_pattern_hyps) iorpat hyps + | IntroNaming IntroAnonymous -> [] + | IntroNaming (IntroFresh _) -> [] + | IntroAction (IntroRewrite _) -> hyps + | IntroAction (IntroInjection ips) -> List.fold_right add_intro_pattern_hyps ips hyps + | IntroAction (IntroApplyOn (c,pat)) -> add_intro_pattern_hyps pat hyps + | IntroForthcoming _ -> + (* As in ipat_of_intro_pattern, was unable to determine which kind + of ipat interp_introid could return [HH] *) assert false + +let rec interp_ipat ist gl = + let ltacvar id = Id.Map.mem id ist.lfun in + let rec interp = function + | IpatId id when ltacvar id -> + ipat_of_intro_pattern (interp_introid ist gl id) + | IpatSimpl (clr, sim) -> + let add_hyps (SsrHyp (loc, id) as hyp) hyps = + if not (ltacvar id) then hyp :: hyps else + add_intro_pattern_hyps (loc, (interp_introid ist gl id)) hyps in + let clr' = List.fold_right add_hyps clr [] in + check_hyps_uniq [] clr'; IpatSimpl (clr', sim) + | IpatCase iorpat -> IpatCase (List.map (List.map interp) iorpat) + | IpatNewHidden l -> + IpatNewHidden + (List.map (function + | IntroNaming (IntroIdentifier id) -> id + | _ -> assert false) + (List.map (interp_introid ist gl) l)) + | ipat -> ipat in + interp + +let interp_ipats ist gl l = project gl, List.map (interp_ipat ist gl) l + +let pushIpatRw = function + | pats :: orpat -> (IpatRw (allocc, L2R) :: pats) :: orpat + | [] -> [] + +let pushIpatNoop = function + | pats :: orpat -> (IpatNoop :: pats) :: orpat + | [] -> [] + +ARGUMENT EXTEND ssripat TYPED AS ssripatrep list PRINTED BY pr_ssripats + INTERPRETED BY interp_ipats + GLOBALIZED BY intern_ipats + | [ "_" ] -> [ [IpatWild] ] + | [ "*" ] -> [ [IpatAll] ] + | [ ident(id) ] -> [ [IpatId id] ] + | [ "?" ] -> [ [IpatAnon] ] + | [ ssrsimpl_ne(sim) ] -> [ [IpatSimpl ([], sim)] ] + | [ ssrdocc(occ) "->" ] -> [ match occ with + | None, occ -> [IpatRw (occ, L2R)] + | Some clr, _ -> [IpatSimpl (clr, Nop); IpatRw (allocc, L2R)]] + | [ ssrdocc(occ) "<-" ] -> [ match occ with + | None, occ -> [IpatRw (occ, R2L)] + | Some clr, _ -> [IpatSimpl (clr, Nop); IpatRw (allocc, R2L)]] + | [ ssrdocc(occ) ] -> [ match occ with + | Some cl, _ -> check_hyps_uniq [] cl; [IpatSimpl (cl, Nop)] + | _ -> loc_error loc "Only identifiers are allowed here"] + | [ "->" ] -> [ [IpatRw (allocc, L2R)] ] + | [ "<-" ] -> [ [IpatRw (allocc, R2L)] ] + | [ "-" ] -> [ [IpatNoop] ] + | [ "-/" "=" ] -> [ [IpatNoop;IpatSimpl([],Simpl)] ] + | [ "-/=" ] -> [ [IpatNoop;IpatSimpl([],Simpl)] ] + | [ "-/" "/" ] -> [ [IpatNoop;IpatSimpl([],Cut)] ] + | [ "-//" ] -> [ [IpatNoop;IpatSimpl([],Cut)] ] + | [ "-/" "/=" ] -> [ [IpatNoop;IpatSimpl([],SimplCut)] ] + | [ "-//" "=" ] -> [ [IpatNoop;IpatSimpl([],SimplCut)] ] + | [ "-//=" ] -> [ [IpatNoop;IpatSimpl([],SimplCut)] ] + | [ ssrview(v) ] -> [ [IpatView v] ] + | [ "[" ":" ident_list(idl) "]" ] -> [ [IpatNewHidden idl] ] +END + +ARGUMENT EXTEND ssripats TYPED AS ssripat PRINTED BY pr_ssripats + | [ ssripat(i) ssripats(tl) ] -> [ i @ tl ] + | [ ] -> [ [] ] +END + +ARGUMENT EXTEND ssriorpat TYPED AS ssripat list PRINTED BY pr_ssriorpat +| [ ssripats(pats) "|" ssriorpat(orpat) ] -> [ pats :: orpat ] +| [ ssripats(pats) "|-" ">" ssriorpat(orpat) ] -> [ pats :: pushIpatRw orpat ] +| [ ssripats(pats) "|-" ssriorpat(orpat) ] -> [ pats :: pushIpatNoop orpat ] +| [ ssripats(pats) "|->" ssriorpat(orpat) ] -> [ pats :: pushIpatRw orpat ] +| [ ssripats(pats) "||" ssriorpat(orpat) ] -> [ pats :: [] :: orpat ] +| [ ssripats(pats) "|||" ssriorpat(orpat) ] -> [ pats :: [] :: [] :: orpat ] +| [ ssripats(pats) "||||" ssriorpat(orpat) ] -> [ [pats; []; []; []] @ orpat ] +| [ ssripats(pats) ] -> [ [pats] ] +END + +let reject_ssrhid strm = + match Compat.get_tok (stream_nth 0 strm) with + | Tok.KEYWORD "[" -> + (match Compat.get_tok (stream_nth 0 strm) with + | Tok.KEYWORD ":" -> raise Stream.Failure + | _ -> ()) + | _ -> () + +let test_nohidden = Gram.Entry.of_parser "test_ssrhid" reject_ssrhid + +ARGUMENT EXTEND ssrcpat TYPED AS ssripatrep PRINTED BY pr_ssripat + | [ "YouShouldNotTypeThis" ssriorpat(x) ] -> [ IpatCase x ] +END + +GEXTEND Gram + GLOBAL: ssrcpat; + ssrcpat: [[ test_nohidden; "["; iorpat = ssriorpat; "]" -> IpatCase iorpat ]]; +END + +GEXTEND Gram + GLOBAL: ssripat; + ssripat: [[ pat = ssrcpat -> [pat] ]]; +END + +ARGUMENT EXTEND ssripats_ne TYPED AS ssripat PRINTED BY pr_ssripats + | [ ssripat(i) ssripats(tl) ] -> [ i @ tl ] +END + +(* subsets of patterns *) +let check_ssrhpats loc w_binders ipats = + let err_loc s = Errors.user_err_loc (loc, "ssreflect", s) in + let clr, ipats = + let rec aux clr = function + | IpatSimpl (cl, Nop) :: tl -> aux (clr @ cl) tl + | IpatSimpl (cl, sim) :: tl -> clr @ cl, IpatSimpl ([], sim) :: tl + | tl -> clr, tl + in aux [] ipats in + let simpl, ipats = + match List.rev ipats with + | IpatSimpl ([],_) as s :: tl -> [s], List.rev tl + | _ -> [], ipats in + if simpl <> [] && not w_binders then + err_loc (str "No s-item allowed here: " ++ pr_ipats simpl); + let ipat, binders = + let rec loop ipat = function + | [] -> ipat, [] + | ( IpatId _| IpatAnon| IpatCase _| IpatRw _ as i) :: tl -> + if w_binders then + if simpl <> [] && tl <> [] then + err_loc(str"binders XOR s-item allowed here: "++pr_ipats(tl@simpl)) + else if not (List.for_all (function IpatId _ -> true | _ -> false) tl) + then err_loc (str "Only binders allowed here: " ++ pr_ipats tl) + else ipat @ [i], tl + else + if tl = [] then ipat @ [i], [] + else err_loc (str "No binder or s-item allowed here: " ++ pr_ipats tl) + | hd :: tl -> loop (ipat @ [hd]) tl + in loop [] ipats in + ((clr, ipat), binders), simpl + +let single loc = + function [x] -> x | _ -> loc_error loc "Only one intro pattern is allowed" + +let pr_hpats (((clr, ipat), binders), simpl) = + pr_clear mt clr ++ pr_ipats ipat ++ pr_ipats binders ++ pr_ipats simpl +let pr_ssrhpats _ _ _ = pr_hpats +let pr_ssrhpats_wtransp _ _ _ (_, x) = pr_hpats x + +ARGUMENT EXTEND ssrhpats TYPED AS ((ssrclear * ssripat) * ssripat) * ssripat +PRINTED BY pr_ssrhpats + | [ ssripats(i) ] -> [ check_ssrhpats loc true i ] +END + +ARGUMENT EXTEND ssrhpats_wtransp + TYPED AS bool * (((ssrclear * ssripat) * ssripat) * ssripat) + PRINTED BY pr_ssrhpats_wtransp + | [ ssripats(i) ] -> [ false,check_ssrhpats loc true i ] + | [ ssripats(i) "@" ssripats(j) ] -> [ true,check_ssrhpats loc true (i @ j) ] +END + +ARGUMENT EXTEND ssrhpats_nobs +TYPED AS ((ssrclear * ssripat) * ssripat) * ssripat PRINTED BY pr_ssrhpats + | [ ssripats(i) ] -> [ check_ssrhpats loc false i ] +END + +ARGUMENT EXTEND ssrrpat TYPED AS ssripatrep PRINTED BY pr_ssripat + | [ "->" ] -> [ IpatRw (allocc, L2R) ] + | [ "<-" ] -> [ IpatRw (allocc, R2L) ] +END + +type ssrintros = ssripats + +let pr_intros sep intrs = + if intrs = [] then mt() else sep () ++ str "=> " ++ pr_ipats intrs +let pr_ssrintros _ _ _ = pr_intros mt + +ARGUMENT EXTEND ssrintros_ne TYPED AS ssripat + PRINTED BY pr_ssrintros + | [ "=>" ssripats_ne(pats) ] -> [ pats ] +END + +ARGUMENT EXTEND ssrintros TYPED AS ssrintros_ne PRINTED BY pr_ssrintros + | [ ssrintros_ne(intrs) ] -> [ intrs ] + | [ ] -> [ [] ] +END + +let injecteq_id = mk_internal_id "injection equation" + +let pf_nb_prod gl = nb_prod (pf_concl gl) + +let rev_id = mk_internal_id "rev concl" + +let revtoptac n0 gl = + let n = pf_nb_prod gl - n0 in + let dc, cl = decompose_prod_n n (pf_concl gl) in + let dc' = dc @ [Name rev_id, compose_prod (List.rev dc) cl] in + let f = compose_lam dc' (mkEtaApp (mkRel (n + 1)) (-n) 1) in + refine (mkApp (f, [|Evarutil.mk_new_meta ()|])) gl + +let equality_inj l b id c gl = + let msg = ref "" in + try Proofview.V82.of_tactic (Equality.inj l b None c) gl + with + | Compat.Exc_located(_,Errors.UserError (_,s)) + | Errors.UserError (_,s) + when msg := Pp.string_of_ppcmds s; + !msg = "Not a projectable equality but a discriminable one." || + !msg = "Nothing to inject." -> + msg_warning (str !msg); + discharge_hyp (id, (id, "")) gl + +let injectidl2rtac id c gl = + tclTHEN (equality_inj None true id c) (revtoptac (pf_nb_prod gl)) gl + +let injectl2rtac c = match kind_of_term c with +| Var id -> injectidl2rtac id (mkVar id, NoBindings) +| _ -> + let id = injecteq_id in + tclTHENLIST [havetac id c; injectidl2rtac id (mkVar id, NoBindings); clear [id]] + +let is_injection_case c gl = + let (mind,_), _ = pf_reduce_to_quantified_ind gl (pf_type_of gl c) in + eq_gr (IndRef mind) (build_coq_eq ()) + +let perform_injection c gl = + let mind, t = pf_reduce_to_quantified_ind gl (pf_type_of gl c) in + let dc, eqt = decompose_prod t in + if dc = [] then injectl2rtac c gl else + if not (closed0 eqt) then + Errors.error "can't decompose a quantified equality" else + let cl = pf_concl gl in let n = List.length dc in + let c_eq = mkEtaApp c n 2 in + let cl1 = mkLambda (Anonymous, mkArrow eqt cl, mkApp (mkRel 1, [|c_eq|])) in + let id = injecteq_id in + let id_with_ebind = (mkVar id, NoBindings) in + let injtac = tclTHEN (introid id) (injectidl2rtac id id_with_ebind) in + tclTHENLAST (Proofview.V82.of_tactic (apply (compose_lam dc cl1))) injtac gl + +let simplest_newcase_ref = ref (fun t gl -> assert false) +let simplest_newcase x gl = !simplest_newcase_ref x gl + +let ssrscasetac c gl = + if is_injection_case c gl then perform_injection c gl + else simplest_newcase c gl + +let intro_all gl = + let dc, _ = Term.decompose_prod_assum (pf_concl gl) in + tclTHENLIST (List.map anontac (List.rev dc)) gl + +let rec intro_anon gl = + try anontac (List.hd (fst (Term.decompose_prod_n_assum 1 (pf_concl gl)))) gl + with err0 -> try tclTHEN red_in_concl intro_anon gl with _ -> raise err0 + (* with _ -> Errors.error "No product even after reduction" *) + +let with_top tac = + tclTHENLIST [introid top_id; tac (mkVar top_id); clear [top_id]] + +let rec mapLR f = function [] -> [] | x :: s -> let y = f x in y :: mapLR f s + +let wild_ids = ref [] + +let new_wild_id () = + let i = 1 + List.length !wild_ids in + let id = mk_wildcard_id i in + wild_ids := id :: !wild_ids; + id + +let clear_wilds wilds gl = + clear (List.filter (fun id -> List.mem id wilds) (pf_ids_of_hyps gl)) gl + +let clear_with_wilds wilds clr0 gl = + let extend_clr clr (id, _, _ as nd) = + if List.mem id clr || not (List.mem id wilds) then clr else + let vars = global_vars_set_of_decl (pf_env gl) nd in + let occurs id' = Idset.mem id' vars in + if List.exists occurs clr then id :: clr else clr in + clear (Context.fold_named_context_reverse extend_clr ~init:clr0 (pf_hyps gl)) gl + +let tclTHENS_nonstrict tac tacl taclname gl = + let tacres = tac gl in + let n_gls = List.length (sig_it tacres) in + let n_tac = List.length tacl in + if n_gls = n_tac then tclTHENS (fun _ -> tacres) tacl gl else + if n_gls = 0 then tacres else + let pr_only n1 n2 = if n1 < n2 then str "only " else mt () in + let pr_nb n1 n2 name = + pr_only n1 n2 ++ int n1 ++ str (" " ^ String.plural n1 name) in + errorstrm (pr_nb n_tac n_gls taclname ++ spc () + ++ str "for " ++ pr_nb n_gls n_tac "subgoal") + +(* Forward reference to extended rewrite *) +let ipat_rewritetac = ref (fun _ -> rewritetac) + +let rec is_name_in_ipats name = function + | IpatSimpl(clr,_) :: tl -> + List.exists (function SsrHyp(_,id) -> id = name) clr + || is_name_in_ipats name tl + | IpatId id :: tl -> id = name || is_name_in_ipats name tl + | IpatCase l :: tl -> is_name_in_ipats name (List.flatten l @ tl) + | _ :: tl -> is_name_in_ipats name tl + | [] -> false + +let move_top_with_view = ref (fun _ -> assert false) + +let rec nat_of_n n = + if n = 0 then mkConstruct path_of_O + else mkApp (mkConstruct path_of_S, [|nat_of_n (n-1)|]) + +let ssr_abstract_id = Summary.ref "~name:SSR:abstractid" 0 + +let mk_abstract_id () = incr ssr_abstract_id; nat_of_n !ssr_abstract_id + +let ssrmkabs id gl = + let env, concl = pf_env gl, pf_concl gl in + let step sigma = + let sigma, abstract_proof, abstract_ty = + let sigma, (ty, _) = + Evarutil.new_type_evar env sigma Evd.univ_flexible_alg in + let sigma, ablock = mkSsrConst "abstract_lock" env sigma in + let sigma, lock = Evarutil.new_evar env sigma ablock in + let sigma, abstract = mkSsrConst "abstract" env sigma in + let abstract_ty = mkApp(abstract, [|ty;mk_abstract_id ();lock|]) in + let sigma, m = Evarutil.new_evar env sigma abstract_ty in + sigma, m, abstract_ty in + let sigma, kont = + let rd = Name id, None, abstract_ty in + Evarutil.new_evar (Environ.push_rel rd env) sigma concl in + pp(lazy(pr_constr concl)); + let term = mkApp (mkLambda(Name id,abstract_ty,kont) ,[|abstract_proof|]) in + let sigma, _ = Typing.e_type_of env sigma term in + sigma, term in + Proofview.V82.of_tactic + (Proofview.tclTHEN + (Tactics.New.refine step) + (Proofview.tclFOCUS 1 3 Proofview.shelve)) gl + +let ssrmkabstac ids = + List.fold_right (fun id tac -> tclTHENFIRST (ssrmkabs id) tac) ids tclIDTAC + +(* introstac: for "move" and "clear", tclEQINTROS: for "case" and "elim" *) +(* This block hides the spaghetti-code needed to implement the only two *) +(* tactics that should be used to process intro patters. *) +(* The difficulty is that we don't want to always rename, but we can *) +(* compute needeed renamings only at runtime, so we theread a tree like *) +(* imperativestructure so that outer renamigs are inherited by inner *) +(* ipts and that the cler performed at the end of ipatstac clears hyps *) +(* eventually renamed at runtime. *) +(* TODO: hide wild_ids in this block too *) +let introstac, tclEQINTROS = + let rec map_acc_k f k = function + | [] -> (* tricky: we save wilds now, we get to_cler (aka k) later *) + let clear_ww = clear_with_wilds !wild_ids in + [fun gl -> clear_ww (hyps_ids (List.flatten (List.map (!) k))) gl] + | x :: xs -> let k, x = f k xs x in x :: map_acc_k f k xs in + let rename force to_clr rest clr gl = + let hyps = pf_hyps gl in + pp(lazy(str"rename " ++ pr_clear spc clr)); + let () = if not force then List.iter (check_hyp_exists hyps) clr in + if List.exists (fun x -> force || is_name_in_ipats (hyp_id x) rest) clr then + let ren_clr, ren = + List.split (List.map (fun x -> let x = hyp_id x in + let x' = mk_anon_id (string_of_id x) gl in + SsrHyp (dummy_loc, x'), (x, x')) clr) in + let () = to_clr := ren_clr in + Proofview.V82.of_tactic (rename_hyp ren) gl + else + let () = to_clr := clr in + tclIDTAC gl in + let rec ipattac ?ist k rest = function + | IpatWild -> k, introid (new_wild_id ()) + | IpatCase iorpat -> k, tclIORPAT ?ist k (with_top ssrscasetac) iorpat + | IpatRw (occ, dir) -> k, with_top (!ipat_rewritetac occ dir) + | IpatId id -> k, introid id + | IpatNewHidden idl -> k, ssrmkabstac idl + | IpatSimpl (clr, sim) -> + let to_clr = ref [] in + to_clr :: k, tclTHEN (rename false to_clr rest clr) (simpltac sim) + | IpatAll -> k, intro_all + | IpatAnon -> k, intro_anon + | IpatNoop -> k, tclIDTAC + | IpatView v -> match ist with + | None -> anomaly "ipattac with no ist but view" + | Some ist -> match rest with + | (IpatCase _ | IpatRw _)::_ -> + let to_clr = ref [] in let top_id = ref top_id in + to_clr :: k, + tclTHEN + (!move_top_with_view false top_id (false,v) ist) + (fun gl -> + rename true to_clr rest [SsrHyp (dummy_loc, !top_id)]gl) + | _ -> k, !move_top_with_view true (ref top_id) (true,v) ist + and tclIORPAT ?ist k tac = function + | [[]] -> tac + | orp -> + tclTHENS_nonstrict tac (mapLR (ipatstac ?ist k) orp) "intro pattern" + and ipatstac ?ist k ipats = + tclTHENLIST (map_acc_k (ipattac ?ist) k ipats) in + let introstac ?ist ipats = + wild_ids := []; + let tac = ipatstac ?ist [] ipats in + tclTHENLIST [tac; clear_wilds !wild_ids] in + let tclEQINTROS ?ist tac eqtac ipats = + wild_ids := []; + let rec split_itacs to_clr tac' = function + | (IpatSimpl _ as spat) :: ipats' -> + let to_clr, tac = ipattac ?ist to_clr ipats' spat in + split_itacs to_clr (tclTHEN tac' tac) ipats' + | IpatCase iorpat :: ipats' -> + to_clr, tclIORPAT ?ist to_clr tac' iorpat, ipats' + | ipats' -> to_clr, tac', ipats' in + let to_clr, tac1, ipats' = split_itacs [] tac ipats in + let tac2 = ipatstac ?ist to_clr ipats' in + tclTHENLIST [tac1; eqtac; tac2; clear_wilds !wild_ids] in + introstac, tclEQINTROS +;; + +let rec eqmoveipats eqpat = function + | (IpatSimpl _ as ipat) :: ipats -> ipat :: eqmoveipats eqpat ipats + | (IpatAll :: _ | []) as ipats -> IpatAnon :: eqpat :: ipats + | ipat :: ipats -> ipat :: eqpat :: ipats + +(* General case *) +let tclINTROS ist tac ipats = + tclEQINTROS ~ist (tac ist) tclIDTAC ipats + +(** The "=>" tactical *) + +let ssrintros_sep = + let atom_sep = function + | TacSplit (_, [NoBindings]) -> mt + (* | TacExtend (_, "ssrapply", []) -> mt *) + | _ -> spc in + function + | TacId [] -> mt + | TacArg (_, Tacexp _) -> mt + | TacArg (_, Reference _) -> mt + | TacAtom (_, atom) -> atom_sep atom + | _ -> spc + +let pr_ssrintrosarg _ _ prt (tac, ipats) = + prt tacltop tac ++ pr_intros (ssrintros_sep tac) ipats + +ARGUMENT EXTEND ssrintrosarg TYPED AS tactic * ssrintros + PRINTED BY pr_ssrintrosarg +| [ "YouShouldNotTypeThis" ssrtacarg(arg) ssrintros_ne(ipats) ] -> [ arg, ipats ] +END + +TACTIC EXTEND ssrtclintros +| [ "YouShouldNotTypeThis" ssrintrosarg(arg) ] -> + [ let tac, intros = arg in + Proofview.V82.tactic (tclINTROS ist (fun ist -> ssrevaltac ist tac) intros) ] +END +set_pr_ssrtac "tclintros" 0 [ArgSsr "introsarg"] + +let tclintros_expr loc tac ipats = + let args = [in_gen (rawwit wit_ssrintrosarg) (tac, ipats)] in + ssrtac_expr loc "tclintros" args + +GEXTEND Gram + GLOBAL: tactic_expr; + tactic_expr: LEVEL "1" [ RIGHTA + [ tac = tactic_expr; intros = ssrintros_ne -> tclintros_expr !@loc tac intros + ] ]; +END +(* }}} *) + +(** Multipliers {{{ ***********************************************************) + +(* modality *) + +type ssrmmod = May | Must | Once + +let pr_mmod = function May -> str "?" | Must -> str "!" | Once -> mt () + +let wit_ssrmmod = add_genarg "ssrmmod" pr_mmod +let ssrmmod = Pcoq.create_generic_entry "ssrmmod" (Genarg.rawwit wit_ssrmmod) +GEXTEND Gram + GLOBAL: ssrmmod; + ssrmmod: [[ "!" -> Must | LEFTQMARK -> May | "?" -> May]]; +END + +(* tactical *) + +let tclID tac = tac + +let tclDOTRY n tac = + if n <= 0 then tclIDTAC else + let rec loop i gl = + if i = n then tclTRY tac gl else + tclTRY (tclTHEN tac (loop (i + 1))) gl in + loop 1 + +let tclDO n tac = + let prefix i = str"At iteration " ++ int i ++ str": " in + let tac_err_at i gl = + try tac gl + with + | Errors.UserError (l, s) as e -> + let _, info = Errors.push e in + let e' = Errors.UserError (l, prefix i ++ s) in + Util.iraise (e', info) + | Compat.Exc_located(loc, Errors.UserError (l, s)) -> + raise (Compat.Exc_located(loc, Errors.UserError (l, prefix i ++ s))) in + let rec loop i gl = + if i = n then tac_err_at i gl else + (tclTHEN (tac_err_at i) (loop (i + 1))) gl in + loop 1 + +let tclMULT = function + | 0, May -> tclREPEAT + | 1, May -> tclTRY + | n, May -> tclDOTRY n + | 0, Must -> tclAT_LEAST_ONCE + | n, Must when n > 1 -> tclDO n + | _ -> tclID + +(** The "do" tactical. ********************************************************) + +(* +type ssrdoarg = ((ssrindex * ssrmmod) * ssrhint) * ssrclauses +*) + +let pr_ssrdoarg prc _ prt (((n, m), tac), clauses) = + pr_index n ++ pr_mmod m ++ pr_hintarg prt tac ++ pr_clauses clauses + +ARGUMENT EXTEND ssrdoarg + TYPED AS ((ssrindex * ssrmmod) * ssrhintarg) * ssrclauses + PRINTED BY pr_ssrdoarg +| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] +END + +let ssrdotac ist (((n, m), tac), clauses) = + let mul = get_index n, m in + tclCLAUSES ist (tclMULT mul (hinttac ist false tac)) clauses + +TACTIC EXTEND ssrtcldo +| [ "YouShouldNotTypeThis" "do" ssrdoarg(arg) ] -> [ Proofview.V82.tactic (ssrdotac ist arg) ] +END +set_pr_ssrtac "tcldo" 3 [ArgSep "do "; ArgSsr "doarg"] + +let ssrdotac_expr loc n m tac clauses = + let arg = ((n, m), tac), clauses in + ssrtac_expr loc "tcldo" [in_gen (rawwit wit_ssrdoarg) arg] + +GEXTEND Gram + GLOBAL: tactic_expr; + ssrdotac: [ + [ tac = tactic_expr LEVEL "3" -> mk_hint tac + | tacs = ssrortacarg -> tacs + ] ]; + tactic_expr: LEVEL "3" [ RIGHTA + [ IDENT "do"; m = ssrmmod; tac = ssrdotac; clauses = ssrclauses -> + ssrdotac_expr !@loc noindex m tac clauses + | IDENT "do"; tac = ssrortacarg; clauses = ssrclauses -> + ssrdotac_expr !@loc noindex Once tac clauses + | IDENT "do"; n = int_or_var; m = ssrmmod; + tac = ssrdotac; clauses = ssrclauses -> + ssrdotac_expr !@loc (mk_index !@loc n) m tac clauses + ] ]; +END +(* }}} *) + +(** The "first" and "last" tacticals. {{{ *************************************) + +(* type ssrseqarg = ssrindex * (ssrtacarg * ssrtac option) *) + +let pr_seqtacarg prt = function + | (is_first, []), _ -> str (if is_first then "first" else "last") + | tac, Some dtac -> + hv 0 (pr_hintarg prt tac ++ spc() ++ str "|| " ++ prt tacltop dtac) + | tac, _ -> pr_hintarg prt tac + +let pr_ssrseqarg _ _ prt = function + | ArgArg 0, tac -> pr_seqtacarg prt tac + | i, tac -> pr_index i ++ str " " ++ pr_seqtacarg prt tac + +(* We must parse the index separately to resolve the conflict with *) +(* an unindexed tactic. *) +ARGUMENT EXTEND ssrseqarg TYPED AS ssrindex * (ssrhintarg * tactic option) + PRINTED BY pr_ssrseqarg +| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] +END + +let sq_brace_tacnames = + ["first"; "solve"; "do"; "rewrite"; "have"; "suffices"; "wlog"] + (* "by" is a keyword *) +let accept_ssrseqvar strm = + match Compat.get_tok (stream_nth 0 strm) with + | Tok.IDENT id when not (List.mem id sq_brace_tacnames) -> + accept_before_syms_or_ids ["["] ["first";"last"] strm + | _ -> raise Stream.Failure + +let test_ssrseqvar = Gram.Entry.of_parser "test_ssrseqvar" accept_ssrseqvar + +let swaptacarg (loc, b) = (b, []), Some (TacId []) + +let check_seqtacarg dir arg = match snd arg, dir with + | ((true, []), Some (TacAtom (loc, _))), L2R -> + loc_error loc "expected \"last\"" + | ((false, []), Some (TacAtom (loc, _))), R2L -> + loc_error loc "expected \"first\"" + | _, _ -> arg + +let ssrorelse = Gram.Entry.create "ssrorelse" +GEXTEND Gram + GLOBAL: ssrorelse ssrseqarg; + ssrseqidx: [ + [ test_ssrseqvar; id = Prim.ident -> ArgVar (!@loc, id) + | n = Prim.natural -> ArgArg (check_index !@loc n) + ] ]; + ssrswap: [[ IDENT "first" -> !@loc, true | IDENT "last" -> !@loc, false ]]; + ssrorelse: [[ "||"; tac = tactic_expr LEVEL "2" -> tac ]]; + ssrseqarg: [ + [ arg = ssrswap -> noindex, swaptacarg arg + | i = ssrseqidx; tac = ssrortacarg; def = OPT ssrorelse -> i, (tac, def) + | i = ssrseqidx; arg = ssrswap -> i, swaptacarg arg + | tac = tactic_expr LEVEL "3" -> noindex, (mk_hint tac, None) + ] ]; +END + +let tclPERM perm tac gls = + let subgls = tac gls in + let sigma, subgll = Refiner.unpackage subgls in + let subgll' = perm subgll in + Refiner.repackage sigma subgll' +(* +let tclPERM perm tac gls = + let mkpft n g r = + {Proof_type.open_subgoals = n; Proof_type.goal = g; Proof_type.ref = r} in + let mkleaf g = mkpft 0 g None in + let mkprpft n g pr a = mkpft n g (Some (Proof_type.Prim pr, a)) in + let mkrpft n g c = mkprpft n g (Proof_type.Refine c) in + let mkipft n g = + let mki pft (id, _, _ as d) = + let g' = {g with evar_concl = mkNamedProd_or_LetIn d g.evar_concl} in + mkprpft n g' (Proof_type.Intro id) [pft] in + List.fold_left mki in + let gl = Refiner.sig_it gls in + let mkhyp subgl = + let rec chop_section = function + | (x, _, _ as d) :: e when not_section_id x -> d :: chop_section e + | _ -> [] in + let lhyps = Environ.named_context_of_val subgl.evar_hyps in + mk_perm_id (), subgl, chop_section lhyps in + let mkpfvar (hyp, subgl, lhyps) = + let mkarg args (lhyp, body, _) = + if body = None then mkVar lhyp :: args else args in + mkrpft 0 subgl (applist (mkVar hyp, List.fold_left mkarg [] lhyps)) [] in + let mkpfleaf (_, subgl, lhyps) = mkipft 1 gl (mkleaf subgl) lhyps in + let mkmeta _ = Evarutil.mk_new_meta () in + let mkhypdecl (hyp, subgl, lhyps) = + hyp, None, it_mkNamedProd_or_LetIn subgl.evar_concl lhyps in + let subgls, v as res0 = tac gls in + let sigma, subgll = Refiner.unpackage subgls in + let n = List.length subgll in if n = 0 then res0 else + let hyps = List.map mkhyp subgll in + let hyp_decls = List.map mkhypdecl (List.rev (perm hyps)) in + let c = applist (mkmeta (), List.map mkmeta subgll) in + let pft0 = mkipft 0 gl (v (List.map mkpfvar hyps)) hyp_decls in + let pft1 = mkrpft n gl c (pft0 :: List.map mkpfleaf (perm hyps)) in + let subgll', v' = Refiner.frontier pft1 in + Refiner.repackage sigma subgll', v' +*) + +let tclREV tac gl = tclPERM List.rev tac gl + +let rot_hyps dir i hyps = + let n = List.length hyps in + if i = 0 then List.rev hyps else + if i > n then Errors.error "Not enough subgoals" else + let rec rot i l_hyps = function + | hyp :: hyps' when i > 0 -> rot (i - 1) (hyp :: l_hyps) hyps' + | hyps' -> hyps' @ (List.rev l_hyps) in + rot (match dir with L2R -> i | R2L -> n - i) [] hyps + +let tclSEQAT ist atac1 dir (ivar, ((_, atacs2), atac3)) = + let i = get_index ivar in + let evtac = ssrevaltac ist in + let tac1 = evtac atac1 in + if atacs2 = [] && atac3 <> None then tclPERM (rot_hyps dir i) tac1 else + let evotac = function Some atac -> evtac atac | _ -> tclIDTAC in + let tac3 = evotac atac3 in + let rec mk_pad n = if n > 0 then tac3 :: mk_pad (n - 1) else [] in + match dir, mk_pad (i - 1), List.map evotac atacs2 with + | L2R, [], [tac2] when atac3 = None -> tclTHENFIRST tac1 tac2 + | L2R, [], [tac2] when atac3 = None -> tclTHENLAST tac1 tac2 + | L2R, pad, tacs2 -> tclTHENSFIRSTn tac1 (Array.of_list (pad @ tacs2)) tac3 + | R2L, pad, tacs2 -> tclTHENSLASTn tac1 tac3 (Array.of_list (tacs2 @ pad)) + +(* We can't actually parse the direction separately because this *) +(* would introduce conflicts with the basic ltac syntax. *) +let pr_ssrseqdir _ _ _ = function + | L2R -> str ";" ++ spc () ++ str "first " + | R2L -> str ";" ++ spc () ++ str "last " + +ARGUMENT EXTEND ssrseqdir TYPED AS ssrdir PRINTED BY pr_ssrseqdir +| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] +END + +TACTIC EXTEND ssrtclseq +| [ "YouShouldNotTypeThis" ssrtclarg(tac) ssrseqdir(dir) ssrseqarg(arg) ] -> + [ Proofview.V82.tactic (tclSEQAT ist tac dir arg) ] +END +set_pr_ssrtac "tclseq" 5 [ArgSsr "tclarg"; ArgSsr "seqdir"; ArgSsr "seqarg"] + +let tclseq_expr loc tac dir arg = + let arg1 = in_gen (rawwit wit_ssrtclarg) tac in + let arg2 = in_gen (rawwit wit_ssrseqdir) dir in + let arg3 = in_gen (rawwit wit_ssrseqarg) (check_seqtacarg dir arg) in + ssrtac_expr loc "tclseq" [arg1; arg2; arg3] + +GEXTEND Gram + GLOBAL: tactic_expr; + ssr_first: [ + [ tac = ssr_first; ipats = ssrintros_ne -> tclintros_expr !@loc tac ipats + | "["; tacl = LIST0 tactic_expr SEP "|"; "]" -> TacFirst tacl + ] ]; + ssr_first_else: [ + [ tac1 = ssr_first; tac2 = ssrorelse -> TacOrelse (tac1, tac2) + | tac = ssr_first -> tac ]]; + tactic_expr: LEVEL "4" [ LEFTA + [ tac1 = tactic_expr; ";"; IDENT "first"; tac2 = ssr_first_else -> + TacThen (tac1, tac2) + | tac = tactic_expr; ";"; IDENT "first"; arg = ssrseqarg -> + tclseq_expr !@loc tac L2R arg + | tac = tactic_expr; ";"; IDENT "last"; arg = ssrseqarg -> + tclseq_expr !@loc tac R2L arg + ] ]; +END +(* }}} *) + +(** 5. Bookkeeping tactics (clear, move, case, elim) *) + +(* post-interpretation of terms *) +let all_ok _ _ = true + +let pf_abs_ssrterm ?(resolve_typeclasses=false) ist gl t = + let sigma, ct as t = interp_term ist gl t in + let sigma, _ as t = + let env = pf_env gl in + if not resolve_typeclasses then t + else + let sigma = Typeclasses.resolve_typeclasses ~fail:false env sigma in + sigma, Evarutil.nf_evar sigma ct in + let n, c, abstracted_away, ucst = pf_abs_evars gl t in + List.fold_left Evd.remove sigma abstracted_away, pf_abs_cterm gl n c, ucst + +let pf_interp_ty ?(resolve_typeclasses=false) ist gl ty = + let n_binders = ref 0 in + let ty = match ty with + | a, (t, None) -> + let rec force_type = function + | GProd (l, x, k, s, t) -> incr n_binders; GProd (l, x, k, s, force_type t) + | GLetIn (l, x, v, t) -> incr n_binders; GLetIn (l, x, v, force_type t) + | ty -> mkRCast ty mkRType in + a, (force_type t, None) + | _, (_, Some ty) -> + let rec force_type = function + | CProdN (l, abs, t) -> + n_binders := !n_binders + List.length (List.flatten (List.map pi1 abs)); + CProdN (l, abs, force_type t) + | CLetIn (l, n, v, t) -> incr n_binders; CLetIn (l, n, v, force_type t) + | ty -> mkCCast dummy_loc ty (mkCType dummy_loc) in + mk_term ' ' (force_type ty) in + let strip_cast (sigma, t) = + let rec aux t = match kind_of_type t with + | CastType (t, ty) when !n_binders = 0 && isSort ty -> t + | ProdType(n,s,t) -> decr n_binders; mkProd (n, s, aux t) + | LetInType(n,v,ty,t) -> decr n_binders; mkLetIn (n, v, ty, aux t) + | _ -> anomaly "pf_interp_ty: ssr Type cast deleted by typecheck" in + sigma, aux t in + let sigma, cty as ty = strip_cast (interp_term ist gl ty) in + let ty = + let env = pf_env gl in + if not resolve_typeclasses then ty + else + let sigma = Typeclasses.resolve_typeclasses ~fail:false env sigma in + sigma, Evarutil.nf_evar sigma cty in + let n, c, _, ucst = pf_abs_evars gl ty in + let lam_c = pf_abs_cterm gl n c in + let ctx, c = decompose_lam_n n lam_c in + n, compose_prod ctx c, lam_c, ucst +;; + +let whd_app f args = Reductionops.whd_betaiota Evd.empty (mkApp (f, args)) + +let pr_cargs a = + str "[" ++ pr_list pr_spc pr_constr (Array.to_list a) ++ str "]" + +let pp_term gl t = + let t = Reductionops.nf_evar (project gl) t in pr_constr t +let pp_concat hd ?(sep=str", ") = function [] -> hd | x :: xs -> + hd ++ List.fold_left (fun acc x -> acc ++ sep ++ x) x xs + +let fake_pmatcher_end () = + mkProp, L2R, (Evd.empty, Evd.empty_evar_universe_context, mkProp) + +(* TASSI: given (c : ty), generates (c ??? : ty[???/...]) with m evars *) +exception NotEnoughProducts +let prof_saturate_whd = mk_profiler "saturate.whd";; +let saturate ?(beta=false) ?(bi_types=false) env sigma c ?(ty=Retyping.get_type_of env sigma c) m += + let rec loop ty args sigma n = + if n = 0 then + let args = List.rev args in + (if beta then Reductionops.whd_beta sigma else fun x -> x) + (mkApp (c, Array.of_list (List.map snd args))), ty, args, sigma + else match kind_of_type ty with + | ProdType (_, src, tgt) -> + let sigma, x = + Evarutil.new_evar env (create_evar_defs sigma) + (if bi_types then Reductionops.nf_betaiota sigma src else src) in + loop (subst1 x tgt) ((m - n,x) :: args) sigma (n-1) + | CastType (t, _) -> loop t args sigma n + | LetInType (_, v, _, t) -> loop (subst1 v t) args sigma n + | SortType _ -> assert false + | AtomicType _ -> + let ty = + prof_saturate_whd.profile + (Reductionops.whd_betadeltaiota env sigma) ty in + match kind_of_type ty with + | ProdType _ -> loop ty args sigma n + | _ -> anomaly "saturate did not find enough products" + in + loop ty [] sigma m + +let pf_saturate ?beta ?bi_types gl c ?ty m = + let env, sigma, si = pf_env gl, project gl, sig_it gl in + let t, ty, args, sigma = saturate ?beta ?bi_types env sigma c ?ty m in + t, ty, args, re_sig si sigma + +(** Rewrite redex switch *) + +(** Generalization (discharge) item *) + +(* An item is a switch + term pair. *) + +(* type ssrgen = ssrdocc * ssrterm *) + +let pr_gen (docc, dt) = pr_docc docc ++ pr_cpattern dt + +let pr_ssrgen _ _ _ = pr_gen + +ARGUMENT EXTEND ssrgen TYPED AS ssrdocc * cpattern PRINTED BY pr_ssrgen +| [ ssrdocc(docc) cpattern(dt) ] -> [ docc, dt ] +| [ cpattern(dt) ] -> [ nodocc, dt ] +END + +let has_occ ((_, occ), _) = occ <> None +let hyp_of_var v = SsrHyp (dummy_loc, destVar v) + +let interp_clr = function +| Some clr, (k, c) + when (k = ' ' || k = '@') && is_pf_var c -> hyp_of_var c :: clr +| Some clr, _ -> clr +| None, _ -> [] + +(* XXX the k of the redex should percolate out *) +let pf_interp_gen_aux ist gl to_ind ((oclr, occ), t) = + let pat = interp_cpattern ist gl t None in (* UGLY API *) + let cl, env, sigma = pf_concl gl, pf_env gl, project gl in + let (c, ucst), cl = + try fill_occ_pattern ~raise_NoMatch:true env sigma cl pat occ 1 + with NoMatch -> redex_of_pattern env pat, cl in + let clr = interp_clr (oclr, (tag_of_cpattern t, c)) in + if not(occur_existential c) then + if tag_of_cpattern t = '@' then + if not (isVar c) then + errorstrm (str "@ can be used with variables only") + else match pf_get_hyp gl (destVar c) with + | _, None, _ -> errorstrm (str "@ can be used with let-ins only") + | name, Some bo, ty -> true, pat, mkLetIn (Name name,bo,ty,cl),c,clr,ucst + else false, pat, pf_mkprod gl c cl, c, clr,ucst + else if to_ind && occ = None then + let nv, p, _, ucst' = pf_abs_evars gl (fst pat, c) in + let ucst = Evd.union_evar_universe_context ucst ucst' in + if nv = 0 then anomaly "occur_existential but no evars" else + false, pat, mkProd (constr_name c, pf_type_of gl p, pf_concl gl), p, clr,ucst + else loc_error (loc_of_cpattern t) "generalized term didn't match" + +let genclrtac cl cs clr = + let tclmyORELSE tac1 tac2 gl = + try tac1 gl + with e when Errors.noncritical e -> tac2 e gl in + (* apply_type may give a type error, but the useful message is + * the one of clear. You type "move: x" and you get + * "x is used in hyp H" instead of + * "The term H has type T x but is expected to have type T x0". *) + tclTHEN + (tclmyORELSE + (apply_type cl cs) + (fun type_err gl -> + tclTHEN + (tclTHEN (Proofview.V82.of_tactic (elim_type (build_coq_False ()))) (cleartac clr)) + (fun gl -> raise type_err) + gl)) + (cleartac clr) + +let gentac ist gen gl = +(* pp(lazy(str"sigma@gentac=" ++ pr_evar_map None (project gl))); *) + let conv, _, cl, c, clr, ucst = pf_interp_gen_aux ist gl false gen in + pp(lazy(str"c@gentac=" ++ pr_constr c)); + let gl = pf_merge_uc ucst gl in + if conv + then tclTHEN (Proofview.V82.of_tactic (convert_concl cl)) (cleartac clr) gl + else genclrtac cl [c] clr gl + +let pf_interp_gen ist gl to_ind gen = + let _, _, a, b, c, ucst = pf_interp_gen_aux ist gl to_ind gen in + a, b ,c, pf_merge_uc ucst gl + +(** Generalization (discharge) sequence *) + +(* A discharge sequence is represented as a list of up to two *) +(* lists of d-items, plus an ident list set (the possibly empty *) +(* final clear switch). The main list is empty iff the command *) +(* is defective, and has length two if there is a sequence of *) +(* dependent terms (and in that case it is the first of the two *) +(* lists). Thus, the first of the two lists is never empty. *) + +(* type ssrgens = ssrgen list *) +(* type ssrdgens = ssrgens list * ssrclear *) + +let gens_sep = function [], [] -> mt | _ -> spc + +let pr_dgens pr_gen (gensl, clr) = + let prgens s gens = str s ++ pr_list spc pr_gen gens in + let prdeps deps = prgens ": " deps ++ spc () ++ str "/" in + match gensl with + | [deps; []] -> prdeps deps ++ pr_clear pr_spc clr + | [deps; gens] -> prdeps deps ++ prgens " " gens ++ pr_clear spc clr + | [gens] -> prgens ": " gens ++ pr_clear spc clr + | _ -> pr_clear pr_spc clr + +let pr_ssrdgens _ _ _ = pr_dgens pr_gen + +let cons_gen gen = function + | gens :: gensl, clr -> (gen :: gens) :: gensl, clr + | _ -> anomaly "missing gen list" + +let cons_dep (gensl, clr) = + if List.length gensl = 1 then ([] :: gensl, clr) else + Errors.error "multiple dependents switches '/'" + +ARGUMENT EXTEND ssrdgens_tl TYPED AS ssrgen list list * ssrclear + PRINTED BY pr_ssrdgens +| [ "{" ne_ssrhyp_list(clr) "}" cpattern(dt) ssrdgens_tl(dgens) ] -> + [ cons_gen (mkclr clr, dt) dgens ] +| [ "{" ne_ssrhyp_list(clr) "}" ] -> + [ [[]], clr ] +| [ "{" ssrocc(occ) "}" cpattern(dt) ssrdgens_tl(dgens) ] -> + [ cons_gen (mkocc occ, dt) dgens ] +| [ "/" ssrdgens_tl(dgens) ] -> + [ cons_dep dgens ] +| [ cpattern(dt) ssrdgens_tl(dgens) ] -> + [ cons_gen (nodocc, dt) dgens ] +| [ ] -> + [ [[]], [] ] +END + +ARGUMENT EXTEND ssrdgens TYPED AS ssrdgens_tl PRINTED BY pr_ssrdgens +| [ ":" ssrgen(gen) ssrdgens_tl(dgens) ] -> [ cons_gen gen dgens ] +END + +let genstac (gens, clr) ist = + tclTHENLIST (cleartac clr :: List.rev_map (gentac ist) gens) + +(* Common code to handle generalization lists along with the defective case *) + +let with_defective maintac deps clr ist gl = + let top_id = + match kind_of_type (pf_concl gl) with + | ProdType (Name id, _, _) + when has_discharged_tag (string_of_id id) -> id + | _ -> top_id in + let top_gen = mkclr clr, cpattern_of_id top_id in + tclTHEN (introid top_id) (maintac deps top_gen ist) gl + +let with_dgens (gensl, clr) maintac ist = match gensl with + | [deps; []] -> with_defective maintac deps clr ist + | [deps; gen :: gens] -> + tclTHEN (genstac (gens, clr) ist) (maintac deps gen ist) + | [gen :: gens] -> tclTHEN (genstac (gens, clr) ist) (maintac [] gen ist) + | _ -> with_defective maintac [] clr ist + +let first_goal gls = + let gl = gls.Evd.it and sig_0 = gls.Evd.sigma in + if List.is_empty gl then Errors.error "first_goal"; + { Evd.it = List.hd gl; Evd.sigma = sig_0; } + +let with_deps deps0 maintac cl0 cs0 clr0 ist gl0 = + let rec loop gl cl cs clr args clrs = function + | [] -> + let n = List.length args in + maintac (if n > 0 then applist (to_lambda n cl, args) else cl) clrs ist gl0 + | dep :: deps -> + let gl' = first_goal (genclrtac cl cs clr gl) in + let cl', c', clr',gl' = pf_interp_gen ist gl' false dep in + loop gl' cl' [c'] clr' (c' :: args) (clr' :: clrs) deps in + loop gl0 cl0 cs0 clr0 cs0 [clr0] (List.rev deps0) + +(** Equations *) + +(* argument *) + +type ssreqid = ssripat option + +let pr_eqid = function Some pat -> str " " ++ pr_ipat pat | None -> mt () +let pr_ssreqid _ _ _ = pr_eqid + +(* We must use primitive parsing here to avoid conflicts with the *) +(* basic move, case, and elim tactics. *) +ARGUMENT EXTEND ssreqid TYPED AS ssripatrep option PRINTED BY pr_ssreqid +| [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] +END + +let accept_ssreqid strm = + match Compat.get_tok (Util.stream_nth 0 strm) with + | Tok.IDENT _ -> accept_before_syms [":"] strm + | Tok.KEYWORD ":" -> () + | Tok.KEYWORD pat when List.mem pat ["_"; "?"; "->"; "<-"] -> + accept_before_syms [":"] strm + | _ -> raise Stream.Failure + +let test_ssreqid = Gram.Entry.of_parser "test_ssreqid" accept_ssreqid + +GEXTEND Gram + GLOBAL: ssreqid; + ssreqpat: [ + [ id = Prim.ident -> IpatId id + | "_" -> IpatWild + | "?" -> IpatAnon + | occ = ssrdocc; "->" -> (match occ with + | None, occ -> IpatRw (occ, L2R) + | _ -> loc_error !@loc "Only occurrences are allowed here") + | occ = ssrdocc; "<-" -> (match occ with + | None, occ -> IpatRw (occ, R2L) + | _ -> loc_error !@loc "Only occurrences are allowed here") + | "->" -> IpatRw (allocc, L2R) + | "<-" -> IpatRw (allocc, R2L) + ]]; + ssreqid: [ + [ test_ssreqid; pat = ssreqpat -> Some pat + | test_ssreqid -> None + ]]; +END + +(* creation *) + +let mkEq dir cl c t n gl = + let eqargs = [|t; c; c|] in eqargs.(dir_org dir) <- mkRel n; + let eq, gl = pf_fresh_global (build_coq_eq()) gl in + let refl, gl = mkRefl t c gl in + mkArrow (mkApp (eq, eqargs)) (lift 1 cl), refl, gl + +let pushmoveeqtac cl c gl = + let x, t, cl1 = destProd cl in + let cl2, eqc, gl = mkEq R2L cl1 c t 1 gl in + apply_type (mkProd (x, t, cl2)) [c; eqc] gl + +let pushcaseeqtac cl gl = + let cl1, args = destApplication cl in + let n = Array.length args in + let dc, cl2 = decompose_lam_n n cl1 in + let _, t = List.nth dc (n - 1) in + let cl3, eqc, gl = mkEq R2L cl2 args.(0) t n gl in + let prot, gl = mkProt (pf_type_of gl cl) cl3 gl in + let cl4 = mkApp (compose_lam dc prot, args) in + let gl, _ = pf_e_type_of gl cl4 in + tclTHEN (apply_type cl4 [eqc]) + (Proofview.V82.of_tactic (convert_concl cl4)) gl + +let pushelimeqtac gl = + let _, args = destApplication (pf_concl gl) in + let x, t, _ = destLambda args.(1) in + let cl1 = mkApp (args.(1), Array.sub args 2 (Array.length args - 2)) in + let cl2, eqc, gl = mkEq L2R cl1 args.(2) t 1 gl in + tclTHEN (apply_type (mkProd (x, t, cl2)) [args.(2); eqc]) + (Proofview.V82.of_tactic intro) gl + +(** Bookkeeping (discharge-intro) argument *) + +(* Since all bookkeeping ssr commands have the same discharge-intro *) +(* argument format we use a single grammar entry point to parse them. *) +(* the entry point parses only non-empty arguments to avoid conflicts *) +(* with the basic Coq tactics. *) + +(* type ssrarg = ssrview * (ssreqid * (ssrdgens * ssripats)) *) + +let pr_ssrarg _ _ _ (view, (eqid, (dgens, ipats))) = + let pri = pr_intros (gens_sep dgens) in + pr_view view ++ pr_eqid eqid ++ pr_dgens pr_gen dgens ++ pri ipats + +ARGUMENT EXTEND ssrarg TYPED AS ssrview * (ssreqid * (ssrdgens * ssrintros)) + PRINTED BY pr_ssrarg +| [ ssrview(view) ssreqid(eqid) ssrdgens(dgens) ssrintros(ipats) ] -> + [ view, (eqid, (dgens, ipats)) ] +| [ ssrview(view) ssrclear(clr) ssrintros(ipats) ] -> + [ view, (None, (([], clr), ipats)) ] +| [ ssreqid(eqid) ssrdgens(dgens) ssrintros(ipats) ] -> + [ [], (eqid, (dgens, ipats)) ] +| [ ssrclear_ne(clr) ssrintros(ipats) ] -> + [ [], (None, (([], clr), ipats)) ] +| [ ssrintros_ne(ipats) ] -> + [ [], (None, (([], []), ipats)) ] +END + +(** The "clear" tactic *) + +(* We just add a numeric version that clears the n top assumptions. *) + +let poptac ist n = introstac ~ist (List.init n (fun _ -> IpatWild)) + +TACTIC EXTEND ssrclear + | [ "clear" natural(n) ] -> [ Proofview.V82.tactic (poptac ist n) ] +END + +(** The "move" tactic *) + +let rec improper_intros = function + | IpatSimpl _ :: ipats -> improper_intros ipats + | (IpatId _ | IpatAnon | IpatCase _ | IpatAll) :: _ -> false + | _ -> true + +let check_movearg = function + | view, (eqid, _) when view <> [] && eqid <> None -> + Errors.error "incompatible view and equation in move tactic" + | view, (_, (([gen :: _], _), _)) when view <> [] && has_occ gen -> + Errors.error "incompatible view and occurrence switch in move tactic" + | _, (_, ((dgens, _), _)) when List.length dgens > 1 -> + Errors.error "dependents switch `/' in move tactic" + | _, (eqid, (_, ipats)) when eqid <> None && improper_intros ipats -> + Errors.error "no proper intro pattern for equation in move tactic" + | arg -> arg + +ARGUMENT EXTEND ssrmovearg TYPED AS ssrarg PRINTED BY pr_ssrarg +| [ ssrarg(arg) ] -> [ check_movearg arg ] +END + +let viewmovetac_aux clear name_ref (_, vl as v) _ gen ist gl = + let cl, c, clr, gl = pf_interp_gen ist gl false gen in + let cl, c, gl = if vl = [] then cl, c, gl else pf_with_view ist gl v cl c in + let clr = if clear then clr else [] in + name_ref := (match id_of_cpattern (snd gen) with Some id -> id | _ -> top_id); + genclrtac cl [c] clr gl + +let () = move_top_with_view := + (fun c r v -> with_defective (viewmovetac_aux c r v) [] []) + +let viewmovetac v deps gen ist gl = + viewmovetac_aux true (ref top_id) v deps gen ist gl + +let eqmovetac _ gen ist gl = + let cl, c, _, gl = pf_interp_gen ist gl false gen in pushmoveeqtac cl c gl + +let movehnftac gl = match kind_of_term (pf_concl gl) with + | Prod _ | LetIn _ -> tclIDTAC gl + | _ -> hnf_in_concl gl + +let ssrmovetac ist = function + | _::_ as view, (_, (dgens, ipats)) -> + let dgentac = with_dgens dgens (viewmovetac (true, view)) ist in + tclTHEN dgentac (introstac ~ist ipats) + | _, (Some pat, (dgens, ipats)) -> + let dgentac = with_dgens dgens eqmovetac ist in + tclTHEN dgentac (introstac ~ist (eqmoveipats pat ipats)) + | _, (_, (([gens], clr), ipats)) -> + let gentac = genstac (gens, clr) ist in + tclTHEN gentac (introstac ~ist ipats) + | _, (_, ((_, clr), ipats)) -> + tclTHENLIST [movehnftac; cleartac clr; introstac ~ist ipats] + +TACTIC EXTEND ssrmove +| [ "move" ssrmovearg(arg) ssrrpat(pat) ] -> + [ Proofview.V82.tactic (tclTHEN (ssrmovetac ist arg) (introstac ~ist [pat])) ] +| [ "move" ssrmovearg(arg) ssrclauses(clauses) ] -> + [ Proofview.V82.tactic (tclCLAUSES ist (ssrmovetac ist arg) clauses) ] +| [ "move" ssrrpat(pat) ] -> [ Proofview.V82.tactic (introstac ~ist [pat]) ] +| [ "move" ] -> [ Proofview.V82.tactic (movehnftac) ] +END + +(* TASSI: given the type of an elimination principle, it finds the higher order + * argument (index), it computes it's arity and the arity of the eliminator and + * checks if the eliminator is recursive or not *) +let analyze_eliminator elimty env sigma = + let rec loop ctx t = match kind_of_type t with + | AtomicType (hd, args) when isRel hd -> + ctx, destRel hd, not (noccurn 1 t), Array.length args + | CastType (t, _) -> loop ctx t + | ProdType (x, ty, t) -> loop ((x,None,ty) :: ctx) t + | LetInType (x,b,ty,t) -> loop ((x,Some b,ty) :: ctx) (subst1 b t) + | _ -> + let env' = Environ.push_rel_context ctx env in + let t' = Reductionops.whd_betadeltaiota env' sigma t in + if not (Term.eq_constr t t') then loop ctx t' else + errorstrm (str"The eliminator has the wrong shape."++spc()++ + str"A (applied) bound variable was expected as the conclusion of "++ + str"the eliminator's"++Pp.cut()++str"type:"++spc()++pr_constr elimty) in + let ctx, pred_id, elim_is_dep, n_pred_args = loop [] elimty in + let n_elim_args = rel_context_nhyps ctx in + let is_rec_elim = + let count_occurn n term = + let count = ref 0 in + let rec occur_rec n c = match kind_of_term c with + | Rel m -> if m = n then incr count + | _ -> iter_constr_with_binders succ occur_rec n c + in + occur_rec n term; !count in + let occurr2 n t = count_occurn n t > 1 in + not (List.for_all_i + (fun i (_,rd) -> pred_id <= i || not (occurr2 (pred_id - i) rd)) + 1 (assums_of_rel_context ctx)) + in + n_elim_args - pred_id, n_elim_args, is_rec_elim, elim_is_dep, n_pred_args + +(* TASSI: This version of unprotects inlines the unfold tactic definition, + * since we don't want to wipe out let-ins, and it seems there is no flag + * to change that behaviour in the standard unfold code *) +let unprotecttac gl = + let c, gl = pf_mkSsrConst "protect_term" gl in + let prot, _ = destConst c in + onClause (fun idopt -> + let hyploc = Option.map (fun id -> id, InHyp) idopt in + reduct_option + (Reductionops.clos_norm_flags + (Closure.RedFlags.mkflags + [Closure.RedFlags.fBETA; + Closure.RedFlags.fCONST prot; + Closure.RedFlags.fIOTA]), DEFAULTcast) hyploc) + allHypsAndConcl gl + +let dependent_apply_error = + try Errors.error "Could not fill dependent hole in \"apply\"" with err -> err + +(* TASSI: Sometimes Coq's apply fails. According to my experience it may be + * related to goals that are products and with beta redexes. In that case it + * guesses the wrong number of implicit arguments for your lemma. What follows + * is just like apply, but with a user-provided number n of implicits. + * + * Refine.refine function that handles type classes and evars but fails to + * handle "dependently typed higher order evars". + * + * Refiner.refiner that does not handle metas with a non ground type but works + * with dependently typed higher order metas. *) +let applyn ~with_evars ?beta ?(with_shelve=false) n t gl = + if with_evars then + let refine gl = + let t, ty, args, gl = pf_saturate ?beta ~bi_types:true gl t n in +(* pp(lazy(str"sigma@saturate=" ++ pr_evar_map None (project gl))); *) + let gl = pf_unify_HO gl ty (pf_concl gl) in + let gs = CList.map_filter (fun (_, e) -> + if isEvar (pf_nf_evar gl e) then Some e else None) + args in + pf_partial_solution gl t gs + in + Proofview.(V82.of_tactic + (tclTHEN (V82.tactic refine) + (if with_shelve then shelve_unifiable else tclUNIT ()))) gl + else + let t, gl = if n = 0 then t, gl else + let sigma, si = project gl, sig_it gl in + let rec loop sigma bo args = function (* saturate with metas *) + | 0 -> mkApp (t, Array.of_list (List.rev args)), re_sig si sigma + | n -> match kind_of_term bo with + | Lambda (_, ty, bo) -> + if not (closed0 ty) then raise dependent_apply_error; + let m = Evarutil.new_meta () in + loop (meta_declare m ty sigma) bo ((mkMeta m)::args) (n-1) + | _ -> assert false + in loop sigma t [] n in + pp(lazy(str"Refiner.refiner " ++ pr_constr t)); + Refiner.refiner (Proof_type.Refine t) gl + +let refine_with ?(first_goes_last=false) ?beta ?(with_evars=true) oc gl = + let rec mkRels = function 1 -> [] | n -> mkRel n :: mkRels (n-1) in + let uct = Evd.evar_universe_context (fst oc) in + let n, oc = pf_abs_evars_pirrel gl oc in + let gl = pf_merge_uc uct gl in + let oc = if not first_goes_last || n <= 1 then oc else + let l, c = decompose_lam oc in + if not (List.for_all_i (fun i (_,t) -> closedn ~-i t) (1-n) l) then oc else + compose_lam (let xs,y = List.chop (n-1) l in y @ xs) + (mkApp (compose_lam l c, Array.of_list (mkRel 1 :: mkRels n))) + in + pp(lazy(str"after: " ++ pr_constr oc)); + try applyn ~with_evars ~with_shelve:true ?beta n oc gl with _ -> raise dependent_apply_error + +let pf_fresh_inductive_instance ind gl = + let sigma, env, it = project gl, pf_env gl, sig_it gl in + let sigma, indu = Evd.fresh_inductive_instance env sigma ind in + indu, re_sig it sigma + +(** The "case" and "elim" tactic *) + +(* A case without explicit dependent terms but with both a view and an *) +(* occurrence switch and/or an equation is treated as dependent, with the *) +(* viewed term as the dependent term (the occurrence switch would be *) +(* meaningless otherwise). When both a view and explicit dependents are *) +(* present, it is forbidden to put a (meaningless) occurrence switch on *) +(* the viewed term. *) + +(* This is both elim and case (defaulting to the former). If ~elim is omitted + * the standard eliminator is chosen. The code is made of 4 parts: + * 1. find the eliminator if not given as ~elim and analyze it + * 2. build the patterns to be matched against the conclusion, looking at + * (occ, c), deps and the pattern inferred from the type of the eliminator + * 3. build the new predicate matching the patterns, and the tactic to + * generalize the equality in case eqid is not None + * 4. build the tactic handle intructions and clears as required in ipats and + * by eqid *) +let ssrelim ?(is_case=false) ?ist deps what ?elim eqid ipats gl = + (* some sanity checks *) + let oc, orig_clr, occ, c_gen, gl = match what with + | `EConstr(_,_,t) when isEvar t -> + anomaly "elim called on a constr evar" + | `EGen _ when ist = None -> + anomaly "no ist and non simple elimination" + | `EGen (_, g) when elim = None && is_wildcard g -> + errorstrm(str"Indeterminate pattern and no eliminator") + | `EGen ((Some clr,occ), g) when is_wildcard g -> + None, clr, occ, None, gl + | `EGen ((None, occ), g) when is_wildcard g -> None,[],occ,None,gl + | `EGen ((_, occ), p as gen) -> + let _, c, clr,gl = pf_interp_gen (Option.get ist) gl true gen in + Some c, clr, occ, Some p,gl + | `EConstr (clr, occ, c) -> Some c, clr, occ, None,gl in + let orig_gl, concl, env = gl, pf_concl gl, pf_env gl in + pp(lazy(str(if is_case then "==CASE==" else "==ELIM=="))); + (* Utils of local interest only *) + let iD s ?t gl = let t = match t with None -> pf_concl gl | Some x -> x in + pp(lazy(str s ++ pr_constr t)); tclIDTAC gl in + let eq, gl = pf_fresh_global (build_coq_eq ()) gl in + let protectC, gl = pf_mkSsrConst "protect_term" gl in + let fire_subst gl t = Reductionops.nf_evar (project gl) t in + let fire_sigma sigma t = Reductionops.nf_evar sigma t in + let is_undef_pat = function + | sigma, T t -> + (match kind_of_term (fire_sigma sigma t) with Evar _ -> true | _ -> false) + | _ -> false in + let match_pat env p occ h cl = + let sigma0 = project orig_gl in + pp(lazy(str"matching: " ++ pp_pattern p)); + let (c,ucst), cl = + fill_occ_pattern ~raise_NoMatch:true env sigma0 cl p occ h in + pp(lazy(str" got: " ++ pr_constr c)); + c, cl, ucst in + let mkTpat gl t = (* takes a term, refreshes it and makes a T pattern *) + let n, t, _, ucst = pf_abs_evars orig_gl (project gl, fire_subst gl t) in + let t, _, _, sigma = saturate ~beta:true env (project gl) t n in + Evd.merge_universe_context sigma ucst, T t in + let unif_redex gl (sigma, r as p) t = (* t is a hint for the redex of p *) + let n, t, _, ucst = pf_abs_evars orig_gl (project gl, fire_subst gl t) in + let t, _, _, sigma = saturate ~beta:true env sigma t n in + let sigma = Evd.merge_universe_context sigma ucst in + match r with + | X_In_T (e, p) -> sigma, E_As_X_In_T (t, e, p) + | _ -> + try unify_HO env sigma t (fst (redex_of_pattern env p)), r + with e when Errors.noncritical e -> p in + (* finds the eliminator applies it to evars and c saturated as needed *) + (* obtaining "elim ??? (c ???)". pred is the higher order evar *) + let cty, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl = + match elim with + | Some elim -> + let gl, elimty = pf_e_type_of gl elim in + let pred_id, n_elim_args, is_rec, elim_is_dep, n_pred_args = + analyze_eliminator elimty env (project gl) in + let elim, elimty, elim_args, gl = + pf_saturate ~beta:is_case gl elim ~ty:elimty n_elim_args in + let pred = List.assoc pred_id elim_args in + let elimty = Reductionops.whd_betadeltaiota env (project gl) elimty in + None, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl + | None -> + let c = Option.get oc in let c_ty = pf_type_of gl c in + let ((kn, i) as ind, _ as indu), unfolded_c_ty = + pf_reduce_to_quantified_ind gl c_ty in + let sort = elimination_sort_of_goal gl in + let gl, elim = + if not is_case then + let t, gl = pf_fresh_global (Indrec.lookup_eliminator ind sort) gl in + gl, t + else + pf_eapply (fun env sigma -> + Indrec.build_case_analysis_scheme env sigma indu true) gl sort in + let elimty = pf_type_of gl elim in + let pred_id,n_elim_args,is_rec,elim_is_dep,n_pred_args = + analyze_eliminator elimty env (project gl) in + let rctx = fst (decompose_prod_assum unfolded_c_ty) in + let n_c_args = rel_context_length rctx in + let c, c_ty, t_args, gl = pf_saturate gl c ~ty:c_ty n_c_args in + let elim, elimty, elim_args, gl = + pf_saturate ~beta:is_case gl elim ~ty:elimty n_elim_args in + let pred = List.assoc pred_id elim_args in + let pc = match n_c_args, c_gen with + | 0, Some p -> interp_cpattern (Option.get ist) orig_gl p None + | _ -> mkTpat gl c in + let cty = Some (c, c_ty, pc) in + let elimty = Reductionops.whd_betadeltaiota env (project gl) elimty in + cty, elim, elimty, elim_args, n_elim_args, elim_is_dep, is_rec, pred, gl + in + pp(lazy(str"elim= "++ pr_constr_pat elim)); + pp(lazy(str"elimty= "++ pr_constr elimty)); + let inf_deps_r = match kind_of_type elimty with + | AtomicType (_, args) -> List.rev (Array.to_list args) + | _ -> assert false in + let saturate_until gl c c_ty f = + let rec loop n = try + let c, c_ty, _, gl = pf_saturate gl c ~ty:c_ty n in + let gl' = f c c_ty gl in + Some (c, c_ty, gl, gl') + with NotEnoughProducts -> None | _ -> loop (n+1) in loop 0 in + let elim_is_dep, gl = match cty with + | None -> true, gl + | Some (c, c_ty, _) -> + let res = + if elim_is_dep then None else + let arg = List.assoc (n_elim_args - 1) elim_args in + let arg_ty = pf_type_of gl arg in + match saturate_until gl c c_ty (fun c c_ty gl -> + pf_unify_HO (pf_unify_HO gl c_ty arg_ty) arg c) with + | Some (c, _, _, gl) -> Some (false, gl) + | None -> None in + match res with + | Some x -> x + | None -> + let inf_arg = List.hd inf_deps_r in + let inf_arg_ty = pf_type_of gl inf_arg in + match saturate_until gl c c_ty (fun _ c_ty gl -> + pf_unify_HO gl c_ty inf_arg_ty) with + | Some (c, _, _,gl) -> true, gl + | None -> + errorstrm (str"Unable to apply the eliminator to the term"++ + spc()++pr_constr c++spc()++str"or to unify it's type with"++ + pr_constr inf_arg_ty) in + pp(lazy(str"elim_is_dep= " ++ bool elim_is_dep)); + let predty = pf_type_of gl pred in + (* Patterns for the inductive types indexes to be bound in pred are computed + * looking at the ones provided by the user and the inferred ones looking at + * the type of the elimination principle *) + let pp_pat (_,p,_,_) = pp_pattern p in + let pp_inf_pat gl (_,_,t,_) = pr_constr_pat (fire_subst gl t) in + let patterns, clr, gl = + let rec loop patterns clr i = function + | [],[] -> patterns, clr, gl + | ((oclr, occ), t):: deps, inf_t :: inf_deps -> + let ist = match ist with Some x -> x | None -> assert false in + let p = interp_cpattern ist orig_gl t None in + let clr_t = + interp_clr (oclr,(tag_of_cpattern t,fst (redex_of_pattern env p)))in + (* if we are the index for the equation we do not clear *) + let clr_t = if deps = [] && eqid <> None then [] else clr_t in + let p = if is_undef_pat p then mkTpat gl inf_t else p in + loop (patterns @ [i, p, inf_t, occ]) + (clr_t @ clr) (i+1) (deps, inf_deps) + | [], c :: inf_deps -> + pp(lazy(str"adding inferred pattern " ++ pr_constr_pat c)); + loop (patterns @ [i, mkTpat gl c, c, allocc]) + clr (i+1) ([], inf_deps) + | _::_, [] -> errorstrm (str "Too many dependent abstractions") in + let deps, head_p, inf_deps_r = match what, elim_is_dep, cty with + | `EConstr _, _, None -> anomaly "Simple welim with no term" + | _, false, _ -> deps, [], inf_deps_r + | `EGen gen, true, None -> deps @ [gen], [], inf_deps_r + | _, true, Some (c, _, pc) -> + let occ = if occ = None then allocc else occ in + let inf_p, inf_deps_r = List.hd inf_deps_r, List.tl inf_deps_r in + deps, [1, pc, inf_p, occ], inf_deps_r in + let patterns, clr, gl = + loop [] orig_clr (List.length head_p+1) (List.rev deps, inf_deps_r) in + head_p @ patterns, Util.List.uniquize clr, gl + in + pp(lazy(pp_concat (str"patterns=") (List.map pp_pat patterns))); + pp(lazy(pp_concat (str"inf. patterns=") (List.map (pp_inf_pat gl) patterns))); + (* Predicate generation, and (if necessary) tactic to generalize the + * equation asked by the user *) + let elim_pred, gen_eq_tac, clr, gl = + let error gl t inf_t = errorstrm (str"The given pattern matches the term"++ + spc()++pp_term gl t++spc()++str"while the inferred pattern"++ + spc()++pr_constr_pat (fire_subst gl inf_t)++spc()++ str"doesn't") in + let match_or_postpone (cl, gl, post) (h, p, inf_t, occ) = + let p = unif_redex gl p inf_t in + if is_undef_pat p then + let () = pp(lazy(str"postponing " ++ pp_pattern p)) in + cl, gl, post @ [h, p, inf_t, occ] + else try + let c, cl, ucst = match_pat env p occ h cl in + let gl = pf_merge_uc ucst gl in + let gl = try pf_unify_HO gl inf_t c with _ -> error gl c inf_t in + cl, gl, post + with + | NoMatch | NoProgress -> + let e, ucst = redex_of_pattern env p in + let gl = pf_merge_uc ucst gl in + let n, e, _, _ucst = pf_abs_evars gl (fst p, e) in + let e, _, _, gl = pf_saturate ~beta:true gl e n in + let gl = try pf_unify_HO gl inf_t e with _ -> error gl e inf_t in + cl, gl, post + in + let rec match_all concl gl patterns = + let concl, gl, postponed = + List.fold_left match_or_postpone (concl, gl, []) patterns in + if postponed = [] then concl, gl + else if List.length postponed = List.length patterns then + errorstrm (str "Some patterns are undefined even after all"++spc()++ + str"the defined ones matched") + else match_all concl gl postponed in + let concl, gl = match_all concl gl patterns in + let pred_rctx, _ = decompose_prod_assum (fire_subst gl predty) in + let concl, gen_eq_tac, clr, gl = match eqid with + | Some (IpatId _) when not is_rec -> + let k = List.length deps in + let c = fire_subst gl (List.assoc (n_elim_args - k - 1) elim_args) in + let t = pf_type_of gl c in + let gen_eq_tac, gl = + let refl = mkApp (eq, [|t; c; c|]) in + let new_concl = mkArrow refl (lift 1 (pf_concl orig_gl)) in + let new_concl = fire_subst gl new_concl in + let erefl, gl = mkRefl t c gl in + let erefl = fire_subst gl erefl in + apply_type new_concl [erefl], gl in + let rel = k + if elim_is_dep then 1 else 0 in + let src, gl = mkProt mkProp (mkApp (eq,[|t; c; mkRel rel|])) gl in + let concl = mkArrow src (lift 1 concl) in + let clr = if deps <> [] then clr else [] in + concl, gen_eq_tac, clr, gl + | _ -> concl, tclIDTAC, clr, gl in + let mk_lam t r = mkLambda_or_LetIn r t in + let concl = List.fold_left mk_lam concl pred_rctx in + let gl, concl = + if eqid <> None && is_rec then + let concl, gl = mkProt (pf_type_of gl concl) concl gl in + let gl, _ = pf_e_type_of gl concl in + gl, concl + else gl, concl in + concl, gen_eq_tac, clr, gl in + let gl, pty = pf_e_type_of gl elim_pred in + pp(lazy(str"elim_pred=" ++ pp_term gl elim_pred)); + pp(lazy(str"elim_pred_ty=" ++ pp_term gl pty)); + let gl = pf_unify_HO gl pred elim_pred in + let elim = fire_subst gl elim in + let gl, _ = pf_e_type_of gl elim in + (* check that the patterns do not contain non instantiated dependent metas *) + let () = + let evars_of_term = Evarutil.undefined_evars_of_term (project gl) in + let patterns = List.map (fun (_,_,t,_) -> fire_subst gl t) patterns in + let patterns_ev = List.map evars_of_term patterns in + let ev = List.fold_left Intset.union Intset.empty patterns_ev in + let ty_ev = Intset.fold (fun i e -> + let ex = i in + let i_ty = Evd.evar_concl (Evd.find (project gl) ex) in + Intset.union e (evars_of_term i_ty)) + ev Intset.empty in + let inter = Intset.inter ev ty_ev in + if not (Intset.is_empty inter) then begin + let i = Intset.choose inter in + let pat = List.find (fun t -> Intset.mem i (evars_of_term t)) patterns in + errorstrm(str"Pattern"++spc()++pr_constr_pat pat++spc()++ + str"was not completely instantiated and one of its variables"++spc()++ + str"occurs in the type of another non instantieted pattern variable"); + end + in + (* the elim tactic, with the eliminator and the predicated we computed *) + let elim = project gl, elim in + let elim_tac gl = + tclTHENLIST [refine_with ~with_evars:false elim; cleartac clr] gl in + (* handling of following intro patterns and equation introduction if any *) + let elim_intro_tac gl = + let intro_eq = + match eqid with + | Some (IpatId ipat) when not is_rec -> + let rec intro_eq gl = match kind_of_type (pf_concl gl) with + | ProdType (_, src, tgt) -> + (match kind_of_type src with + | AtomicType (hd, _) when Term.eq_constr hd protectC -> + tclTHENLIST [unprotecttac; introid ipat] gl + | _ -> tclTHENLIST [ iD "IA"; introstac [IpatAnon]; intro_eq] gl) + |_ -> errorstrm (str "Too many names in intro pattern") in + intro_eq + | Some (IpatId ipat) -> + let name gl = mk_anon_id "K" gl in + let intro_lhs gl = + let elim_name = match orig_clr, what with + | [SsrHyp(_, x)], _ -> x + | _, `EConstr(_,_,t) when isVar t -> destVar t + | _ -> name gl in + if is_name_in_ipats elim_name ipats then introid (name gl) gl + else introid elim_name gl + in + let rec gen_eq_tac gl = + let concl = pf_concl gl in + let ctx, last = decompose_prod_assum concl in + let args = match kind_of_type last with + | AtomicType (hd, args) -> assert(Term.eq_constr hd protectC); args + | _ -> assert false in + let case = args.(Array.length args-1) in + if not(closed0 case) then tclTHEN (introstac [IpatAnon]) gen_eq_tac gl + else + let case_ty = pf_type_of gl case in + let refl = mkApp (eq, [|lift 1 case_ty; mkRel 1; lift 1 case|]) in + let new_concl = fire_subst gl + (mkProd (Name (name gl), case_ty, mkArrow refl (lift 2 concl))) in + let erefl, gl = mkRefl case_ty case gl in + let erefl = fire_subst gl erefl in + apply_type new_concl [case;erefl] gl in + tclTHENLIST [gen_eq_tac; intro_lhs; introid ipat] + | _ -> tclIDTAC in + let unprot = if eqid <> None && is_rec then unprotecttac else tclIDTAC in + tclEQINTROS ?ist elim_tac (tclTHENLIST [intro_eq; unprot]) ipats gl + in + tclTHENLIST [gen_eq_tac; elim_intro_tac] orig_gl +;; + +let simplest_newelim x= ssrelim ~is_case:false [] (`EConstr ([],None,x)) None [] +let simplest_newcase x= ssrelim ~is_case:true [] (`EConstr ([],None,x)) None [] +let _ = simplest_newcase_ref := simplest_newcase + +let check_casearg = function +| view, (_, (([_; gen :: _], _), _)) when view <> [] && has_occ gen -> + Errors.error "incompatible view and occurrence switch in dependent case tactic" +| arg -> arg + +ARGUMENT EXTEND ssrcasearg TYPED AS ssrarg PRINTED BY pr_ssrarg +| [ ssrarg(arg) ] -> [ check_casearg arg ] +END + +let ssrcasetac ist (view, (eqid, (dgens, ipats))) = + let ndefectcasetac view eqid ipats deps ((_, occ), _ as gen) ist gl = + let simple = (eqid = None && deps = [] && occ = None) in + let cl, c, clr, gl = pf_interp_gen ist gl true gen in + let _, vc, gl = + if view = [] then c, c, gl else pf_with_view ist gl (false, view) cl c in + if simple && is_injection_case vc gl then + tclTHENLIST [perform_injection vc; cleartac clr; introstac ~ist ipats] gl + else + (* macro for "case/v E: x" ---> "case E: x / (v x)" *) + let deps, clr, occ = + if view <> [] && eqid <> None && deps = [] then [gen], [], None + else deps, clr, occ in + ssrelim ~is_case:true ~ist deps (`EConstr (clr,occ, vc)) eqid ipats gl + in + with_dgens dgens (ndefectcasetac view eqid ipats) ist + +TACTIC EXTEND ssrcase +| [ "case" ssrcasearg(arg) ssrclauses(clauses) ] -> + [ Proofview.V82.tactic (tclCLAUSES ist (ssrcasetac ist arg) clauses) ] +| [ "case" ] -> [ Proofview.V82.tactic (with_top ssrscasetac) ] +END + +(** The "elim" tactic *) + +(* Elim views are elimination lemmas, so the eliminated term is not addded *) +(* to the dependent terms as for "case", unless it actually occurs in the *) +(* goal, the "all occurrences" {+} switch is used, or the equation switch *) +(* is used and there are no dependents. *) + +let ssrelimtac ist (view, (eqid, (dgens, ipats))) = + let ndefectelimtac view eqid ipats deps gen ist gl = + let elim = match view with [v] -> Some (snd(force_term ist gl v)) | _ -> None in + ssrelim ~ist deps (`EGen gen) ?elim eqid ipats gl + in + with_dgens dgens (ndefectelimtac view eqid ipats) ist + +TACTIC EXTEND ssrelim +| [ "elim" ssrarg(arg) ssrclauses(clauses) ] -> + [ Proofview.V82.tactic (tclCLAUSES ist (ssrelimtac ist arg) clauses) ] +| [ "elim" ] -> [ Proofview.V82.tactic (with_top simplest_newelim) ] +END + +(** 6. Backward chaining tactics: apply, exact, congr. *) + +(** The "apply" tactic *) + +let pr_agen (docc, dt) = pr_docc docc ++ pr_term dt +let pr_ssragen _ _ _ = pr_agen +let pr_ssragens _ _ _ = pr_dgens pr_agen + +ARGUMENT EXTEND ssragen TYPED AS ssrdocc * ssrterm PRINTED BY pr_ssragen +| [ "{" ne_ssrhyp_list(clr) "}" ssrterm(dt) ] -> [ mkclr clr, dt ] +| [ ssrterm(dt) ] -> [ nodocc, dt ] +END + +ARGUMENT EXTEND ssragens TYPED AS ssragen list list * ssrclear +PRINTED BY pr_ssragens +| [ "{" ne_ssrhyp_list(clr) "}" ssrterm(dt) ssragens(agens) ] -> + [ cons_gen (mkclr clr, dt) agens ] +| [ "{" ne_ssrhyp_list(clr) "}" ] -> [ [[]], clr] +| [ ssrterm(dt) ssragens(agens) ] -> + [ cons_gen (nodocc, dt) agens ] +| [ ] -> [ [[]], [] ] +END + +let mk_applyarg views agens intros = views, (None, (agens, intros)) + +let pr_ssraarg _ _ _ (view, (eqid, (dgens, ipats))) = + let pri = pr_intros (gens_sep dgens) in + pr_view view ++ pr_eqid eqid ++ pr_dgens pr_agen dgens ++ pri ipats + +ARGUMENT EXTEND ssrapplyarg +TYPED AS ssrview * (ssreqid * (ssragens * ssrintros)) +PRINTED BY pr_ssraarg +| [ ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] -> + [ mk_applyarg [] (cons_gen gen dgens) intros ] +| [ ssrclear_ne(clr) ssrintros(intros) ] -> + [ mk_applyarg [] ([], clr) intros ] +| [ ssrintros_ne(intros) ] -> + [ mk_applyarg [] ([], []) intros ] +| [ ssrview(view) ":" ssragen(gen) ssragens(dgens) ssrintros(intros) ] -> + [ mk_applyarg view (cons_gen gen dgens) intros ] +| [ ssrview(view) ssrclear(clr) ssrintros(intros) ] -> + [ mk_applyarg view ([], clr) intros ] +END + +let interp_agen ist gl ((goclr, _), (k, gc)) (clr, rcs) = +(* pp(lazy(str"sigma@interp_agen=" ++ pr_evar_map None (project gl))); *) + let rc = glob_constr ist (pf_env gl) gc in + let rcs' = rc :: rcs in + match goclr with + | None -> clr, rcs' + | Some ghyps -> + let clr' = snd (interp_hyps ist gl ghyps) @ clr in + if k <> ' ' then clr', rcs' else + match rc with + | GVar (loc, id) when not_section_id id -> SsrHyp (loc, id) :: clr', rcs' + | GRef (loc, VarRef id, _) when not_section_id id -> + SsrHyp (loc, id) :: clr', rcs' + | _ -> clr', rcs' + +let interp_agens ist gl gagens = + match List.fold_right (interp_agen ist gl) gagens ([], []) with + | clr, rlemma :: args -> + let n = interp_nbargs ist gl rlemma - List.length args in + let rec loop i = + if i > n then + errorstrm (str "Cannot apply lemma " ++ pf_pr_glob_constr gl rlemma) + else + try interp_refine ist gl (mkRApp rlemma (mkRHoles i @ args)) + with _ -> loop (i + 1) in + clr, loop 0 + | _ -> assert false + +let apply_rconstr ?ist t gl = +(* pp(lazy(str"sigma@apply_rconstr=" ++ pr_evar_map None (project gl))); *) + let n = match ist, t with + | None, (GVar (_, id) | GRef (_, VarRef id,_)) -> pf_nbargs gl (mkVar id) + | Some ist, _ -> interp_nbargs ist gl t + | _ -> anomaly "apply_rconstr without ist and not RVar" in + let mkRlemma i = mkRApp t (mkRHoles i) in + let cl = pf_concl gl in + let rec loop i = + if i > n then + errorstrm (str"Cannot apply lemma "++pf_pr_glob_constr gl t) + else try pf_match gl (mkRlemma i) (OfType cl) with _ -> loop (i + 1) in + refine_with (loop 0) gl + +let mkRAppView ist gl rv gv = + let nb_view_imps = interp_view_nbimps ist gl rv in + mkRApp rv (mkRHoles (abs nb_view_imps)) + +let prof_apply_interp_with = mk_profiler "ssrapplytac.interp_with";; + +let refine_interp_apply_view i ist gl gv = + let pair i = List.map (fun x -> i, x) in + let rv = pf_intern_term ist gl gv in + let v = mkRAppView ist gl rv gv in + let interp_with (i, hint) = + interp_refine ist gl (mkRApp hint (v :: mkRHoles i)) in + let interp_with x = prof_apply_interp_with.profile interp_with x in + let rec loop = function + | [] -> (try apply_rconstr ~ist rv gl with _ -> view_error "apply" gv) + | h :: hs -> (try refine_with (snd (interp_with h)) gl with _ -> loop hs) in + loop (pair i viewtab.(i) @ if i = 2 then pair 1 viewtab.(1) else []) + +let apply_top_tac gl = + tclTHENLIST [introid top_id; apply_rconstr (mkRVar top_id); clear [top_id]] gl + +let inner_ssrapplytac gviews ggenl gclr ist gl = + let _, clr = interp_hyps ist gl gclr in + let vtac gv i gl' = refine_interp_apply_view i ist gl' gv in + let ggenl, tclGENTAC = + if gviews <> [] && ggenl <> [] then + let ggenl= List.map (fun (x,g) -> x, cpattern_of_term g) (List.hd ggenl) in + [], tclTHEN (genstac (ggenl,[]) ist) + else ggenl, tclTHEN tclIDTAC in + tclGENTAC (fun gl -> + match gviews, ggenl with + | v :: tl, [] -> + let dbl = if List.length tl = 1 then 2 else 1 in + tclTHEN + (List.fold_left (fun acc v -> tclTHENLAST acc (vtac v dbl)) (vtac v 1) tl) + (cleartac clr) gl + | [], [agens] -> + let clr', (sigma, lemma) = interp_agens ist gl agens in + let gl = pf_merge_uc_of sigma gl in + tclTHENLIST [cleartac clr; refine_with ~beta:true lemma; cleartac clr'] gl + | _, _ -> tclTHEN apply_top_tac (cleartac clr) gl) gl + +let ssrapplytac ist (views, (_, ((gens, clr), intros))) = + tclINTROS ist (inner_ssrapplytac views gens clr) intros + +TACTIC EXTEND ssrapply +| [ "apply" ssrapplyarg(arg) ] -> [ Proofview.V82.tactic (ssrapplytac ist arg) ] +| [ "apply" ] -> [ Proofview.V82.tactic apply_top_tac ] +END + +(** The "exact" tactic *) + +let mk_exactarg views dgens = mk_applyarg views dgens [] + +ARGUMENT EXTEND ssrexactarg TYPED AS ssrapplyarg PRINTED BY pr_ssraarg +| [ ":" ssragen(gen) ssragens(dgens) ] -> + [ mk_exactarg [] (cons_gen gen dgens) ] +| [ ssrview(view) ssrclear(clr) ] -> + [ mk_exactarg view ([], clr) ] +| [ ssrclear_ne(clr) ] -> + [ mk_exactarg [] ([], clr) ] +END + +let vmexacttac pf gl = exact_no_check (mkCast (pf, VMcast, pf_concl gl)) gl + +TACTIC EXTEND ssrexact +| [ "exact" ssrexactarg(arg) ] -> [ Proofview.V82.tactic (tclBY (ssrapplytac ist arg)) ] +| [ "exact" ] -> [ Proofview.V82.tactic (tclORELSE donetac (tclBY apply_top_tac)) ] +| [ "exact" "<:" lconstr(pf) ] -> [ Proofview.V82.tactic (vmexacttac pf) ] +END + +(** The "congr" tactic *) + +(* type ssrcongrarg = open_constr * (int * constr) *) + +let pr_ssrcongrarg _ _ _ ((n, f), dgens) = + (if n <= 0 then mt () else str " " ++ int n) ++ + str " " ++ pr_term f ++ pr_dgens pr_gen dgens + +ARGUMENT EXTEND ssrcongrarg TYPED AS (int * ssrterm) * ssrdgens + PRINTED BY pr_ssrcongrarg +| [ natural(n) constr(c) ssrdgens(dgens) ] -> [ (n, mk_term ' ' c), dgens ] +| [ natural(n) constr(c) ] -> [ (n, mk_term ' ' c),([[]],[]) ] +| [ constr(c) ssrdgens(dgens) ] -> [ (0, mk_term ' ' c), dgens ] +| [ constr(c) ] -> [ (0, mk_term ' ' c), ([[]],[]) ] +END + +let rec mkRnat n = + if n <= 0 then GRef (dummy_loc, glob_O, None) else + mkRApp (GRef (dummy_loc, glob_S, None)) [mkRnat (n - 1)] + +let interp_congrarg_at ist gl n rf ty m = + pp(lazy(str"===interp_congrarg_at===")); + let congrn, _ = mkSsrRRef "nary_congruence" in + let args1 = mkRnat n :: mkRHoles n @ [ty] in + let args2 = mkRHoles (3 * n) in + let rec loop i = + if i + n > m then None else + try + let rt = mkRApp congrn (args1 @ mkRApp rf (mkRHoles i) :: args2) in + pp(lazy(str"rt=" ++ pr_glob_constr rt)); + Some (interp_refine ist gl rt) + with _ -> loop (i + 1) in + loop 0 + +let pattern_id = mk_internal_id "pattern value" + +let congrtac ((n, t), ty) ist gl = + pp(lazy(str"===congr===")); + pp(lazy(str"concl=" ++ pr_constr (pf_concl gl))); + let sigma, _ as it = interp_term ist gl t in + let gl = pf_merge_uc_of sigma gl in + let _, f, _, _ucst = pf_abs_evars gl it in + let ist' = {ist with lfun = + Id.Map.add pattern_id (Value.of_constr f) Id.Map.empty } in + let rf = mkRltacVar pattern_id in + let m = pf_nbargs gl f in + let _, cf = if n > 0 then + match interp_congrarg_at ist' gl n rf ty m with + | Some cf -> cf + | None -> errorstrm (str "No " ++ int n ++ str "-congruence with " + ++ pr_term t) + else let rec loop i = + if i > m then errorstrm (str "No congruence with " ++ pr_term t) + else match interp_congrarg_at ist' gl i rf ty m with + | Some cf -> cf + | None -> loop (i + 1) in + loop 1 in + tclTHEN (refine_with cf) (tclTRY (Proofview.V82.of_tactic reflexivity)) gl + +let newssrcongrtac arg ist gl = + pp(lazy(str"===newcongr===")); + pp(lazy(str"concl=" ++ pr_constr (pf_concl gl))); + (* utils *) + let fs gl t = Reductionops.nf_evar (project gl) t in + let tclMATCH_GOAL (c, gl_c) proj t_ok t_fail gl = + match try Some (pf_unify_HO gl_c (pf_concl gl) c) with _ -> None with + | Some gl_c -> + tclTHEN (Proofview.V82.of_tactic (convert_concl (fs gl_c c))) + (t_ok (proj gl_c)) gl + | None -> t_fail () gl in + let mk_evar gl ty = + let env, sigma, si = pf_env gl, project gl, sig_it gl in + let sigma, x = Evarutil.new_evar env (create_evar_defs sigma) ty in + x, re_sig si sigma in + let arr, gl = pf_mkSsrConst "ssr_congr_arrow" gl in + let ssr_congr lr = mkApp (arr, lr) in + (* here thw two cases: simple equality or arrow *) + let equality, _, eq_args, gl' = + let eq, gl = pf_fresh_global (build_coq_eq ()) gl in + pf_saturate gl eq 3 in + tclMATCH_GOAL (equality, gl') (fun gl' -> fs gl' (List.assoc 0 eq_args)) + (fun ty -> congrtac (arg, Detyping.detype false [] (pf_env gl) (project gl) ty) ist) + (fun () -> + let lhs, gl' = mk_evar gl mkProp in let rhs, gl' = mk_evar gl' mkProp in + let arrow = mkArrow lhs (lift 1 rhs) in + tclMATCH_GOAL (arrow, gl') (fun gl' -> [|fs gl' lhs;fs gl' rhs|]) + (fun lr -> tclTHEN (Proofview.V82.of_tactic (apply (ssr_congr lr))) (congrtac (arg, mkRType) ist)) + (fun _ _ -> errorstrm (str"Conclusion is not an equality nor an arrow"))) + gl +;; + +TACTIC EXTEND ssrcongr +| [ "congr" ssrcongrarg(arg) ] -> +[ let arg, dgens = arg in + Proofview.V82.tactic begin + match dgens with + | [gens], clr -> tclTHEN (genstac (gens,clr) ist) (newssrcongrtac arg ist) + | _ -> errorstrm (str"Dependent family abstractions not allowed in congr") + end] +END + +(** 7. Rewriting tactics (rewrite, unlock) *) + +(** Coq rewrite compatibility flag *) + +let ssr_strict_match = ref false + +let _ = + Goptions.declare_bool_option + { Goptions.optsync = true; + Goptions.optname = "strict redex matching"; + Goptions.optkey = ["Match"; "Strict"]; + Goptions.optread = (fun () -> !ssr_strict_match); + Goptions.optdepr = false; + Goptions.optwrite = (fun b -> ssr_strict_match := b) } + +(** Rewrite multiplier *) + +type ssrmult = int * ssrmmod + +let notimes = 0 +let nomult = 1, Once + +let pr_mult (n, m) = + if n > 0 && m <> Once then int n ++ pr_mmod m else pr_mmod m + +let pr_ssrmult _ _ _ = pr_mult + +ARGUMENT EXTEND ssrmult_ne TYPED AS int * ssrmmod PRINTED BY pr_ssrmult + | [ natural(n) ssrmmod(m) ] -> [ check_index loc n, m ] + | [ ssrmmod(m) ] -> [ notimes, m ] +END + +ARGUMENT EXTEND ssrmult TYPED AS ssrmult_ne PRINTED BY pr_ssrmult + | [ ssrmult_ne(m) ] -> [ m ] + | [ ] -> [ nomult ] +END + +(** Rewrite clear/occ switches *) + +let pr_rwocc = function + | None, None -> mt () + | None, occ -> pr_occ occ + | Some clr, _ -> pr_clear_ne clr + +let pr_ssrrwocc _ _ _ = pr_rwocc + +ARGUMENT EXTEND ssrrwocc TYPED AS ssrdocc PRINTED BY pr_ssrrwocc +| [ "{" ssrhyp_list(clr) "}" ] -> [ mkclr clr ] +| [ "{" ssrocc(occ) "}" ] -> [ mkocc occ ] +| [ ] -> [ noclr ] +END + +(** Rewrite rules *) + +type ssrwkind = RWred of ssrsimpl | RWdef | RWeq +(* type ssrrule = ssrwkind * ssrterm *) + +let pr_rwkind = function + | RWred s -> pr_simpl s + | RWdef -> str "/" + | RWeq -> mt () + +let wit_ssrrwkind = add_genarg "ssrrwkind" pr_rwkind + +let pr_rule = function + | RWred s, _ -> pr_simpl s + | RWdef, r-> str "/" ++ pr_term r + | RWeq, r -> pr_term r + +let pr_ssrrule _ _ _ = pr_rule + +let noruleterm loc = mk_term ' ' (mkCProp loc) + +ARGUMENT EXTEND ssrrule_ne TYPED AS ssrrwkind * ssrterm PRINTED BY pr_ssrrule + | [ ssrsimpl_ne(s) ] -> [ RWred s, noruleterm loc ] + | [ "/" ssrterm(t) ] -> [ RWdef, t ] + | [ ssrterm(t) ] -> [ RWeq, t ] +END + +ARGUMENT EXTEND ssrrule TYPED AS ssrrule_ne PRINTED BY pr_ssrrule + | [ ssrrule_ne(r) ] -> [ r ] + | [ ] -> [ RWred Nop, noruleterm loc ] +END + +(** Rewrite arguments *) + +(* type ssrrwarg = (ssrdir * ssrmult) * ((ssrdocc * ssrpattern) * ssrrule) *) + +let pr_option f = function None -> mt() | Some x -> f x +let pr_pattern_squarep= pr_option (fun r -> str "[" ++ pr_rpattern r ++ str "]") +let pr_ssrpattern_squarep _ _ _ = pr_pattern_squarep +let pr_rwarg ((d, m), ((docc, rx), r)) = + pr_rwdir d ++ pr_mult m ++ pr_rwocc docc ++ pr_pattern_squarep rx ++ pr_rule r + +let pr_ssrrwarg _ _ _ = pr_rwarg + +let mk_rwarg (d, (n, _ as m)) ((clr, occ as docc), rx) (rt, _ as r) = + if rt <> RWeq then begin + if rt = RWred Nop && not (m = nomult && occ = None && rx = None) + && (clr = None || clr = Some []) then + anomaly "Improper rewrite clear switch"; + if d = R2L && rt <> RWdef then + Errors.error "Right-to-left switch on simplification"; + if n <> 1 && rt = RWred Cut then + Errors.error "Bad or useless multiplier"; + if occ <> None && rx = None && rt <> RWdef then + Errors.error "Missing redex for simplification occurrence" + end; (d, m), ((docc, rx), r) + +let norwmult = L2R, nomult +let norwocc = noclr, None + +(* +let pattern_ident = Prim.pattern_ident in +GEXTEND Gram +GLOBAL: pattern_ident; +pattern_ident: +[[c = pattern_ident -> (CRef (Ident (loc,c)), NoBindings)]]; +END +*) + +ARGUMENT EXTEND ssrpattern_squarep +TYPED AS rpattern option PRINTED BY pr_ssrpattern_squarep + | [ "[" rpattern(rdx) "]" ] -> [ Some rdx ] + | [ ] -> [ None ] +END + +ARGUMENT EXTEND ssrpattern_ne_squarep +TYPED AS rpattern option PRINTED BY pr_ssrpattern_squarep + | [ "[" rpattern(rdx) "]" ] -> [ Some rdx ] +END + + +ARGUMENT EXTEND ssrrwarg + TYPED AS (ssrdir * ssrmult) * ((ssrdocc * rpattern option) * ssrrule) + PRINTED BY pr_ssrrwarg + | [ "-" ssrmult(m) ssrrwocc(docc) ssrpattern_squarep(rx) ssrrule_ne(r) ] -> + [ mk_rwarg (R2L, m) (docc, rx) r ] + | [ "-/" ssrterm(t) ] -> (* just in case '-/' should become a token *) + [ mk_rwarg (R2L, nomult) norwocc (RWdef, t) ] + | [ ssrmult_ne(m) ssrrwocc(docc) ssrpattern_squarep(rx) ssrrule_ne(r) ] -> + [ mk_rwarg (L2R, m) (docc, rx) r ] + | [ "{" ne_ssrhyp_list(clr) "}" ssrpattern_ne_squarep(rx) ssrrule_ne(r) ] -> + [ mk_rwarg norwmult (mkclr clr, rx) r ] + | [ "{" ne_ssrhyp_list(clr) "}" ssrrule(r) ] -> + [ mk_rwarg norwmult (mkclr clr, None) r ] + | [ "{" ssrocc(occ) "}" ssrpattern_squarep(rx) ssrrule_ne(r) ] -> + [ mk_rwarg norwmult (mkocc occ, rx) r ] + | [ "{" "}" ssrpattern_squarep(rx) ssrrule_ne(r) ] -> + [ mk_rwarg norwmult (nodocc, rx) r ] + | [ ssrpattern_ne_squarep(rx) ssrrule_ne(r) ] -> + [ mk_rwarg norwmult (noclr, rx) r ] + | [ ssrrule_ne(r) ] -> + [ mk_rwarg norwmult norwocc r ] +END + +let simplintac occ rdx sim gl = + let simptac gl = + let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in + let simp env c _ _ = red_safe Tacred.simpl env sigma0 c in + Proofview.V82.of_tactic + (convert_concl_no_check (eval_pattern env0 sigma0 concl0 rdx occ simp)) + gl in + match sim with + | Simpl -> simptac gl + | SimplCut -> tclTHEN simptac (tclTRY donetac) gl + | _ -> simpltac sim gl + +let rec get_evalref c = match kind_of_term c with + | Var id -> EvalVarRef id + | Const (k,_) -> EvalConstRef k + | App (c', _) -> get_evalref c' + | Cast (c', _, _) -> get_evalref c' + | _ -> errorstrm (str "The term " ++ pr_constr c ++ str " is not unfoldable") + +(* Strip a pattern generated by a prenex implicit to its constant. *) +let strip_unfold_term ((sigma, t) as p) kt = match kind_of_term t with + | App (f, a) when kt = ' ' && Array.for_all isEvar a && isConst f -> + (sigma, f), true + | Const _ | Var _ -> p, true + | _ -> p, false + +let unfoldintac occ rdx t (kt,_) gl = + let fs sigma x = Reductionops.nf_evar sigma x in + let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in + let (sigma, t), const = strip_unfold_term t kt in + let body env t c = + Tacred.unfoldn [OnlyOccurrences [1], get_evalref t] env sigma0 c in + let easy = occ = None && rdx = None in + let red_flags = if easy then Closure.betaiotazeta else Closure.betaiota in + let beta env = Reductionops.clos_norm_flags red_flags env sigma0 in + let unfold, conclude = match rdx with + | Some (_, (In_T _ | In_X_In_T _)) | None -> + let ise = create_evar_defs sigma in + let ise, u = mk_tpattern env0 sigma0 (ise,t) all_ok L2R t in + let find_T, end_T = + mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[u]) in + (fun env c _ h -> + try find_T env c h (fun env c _ _ -> body env t c) + with NoMatch when easy -> c + | NoMatch | NoProgress -> errorstrm (str"No occurrence of " + ++ pr_constr_pat t ++ spc() ++ str "in " ++ pr_constr c)), + (fun () -> try end_T () with + | NoMatch when easy -> fake_pmatcher_end () + | NoMatch -> anomaly "unfoldintac") + | _ -> + (fun env (c as orig_c) _ h -> + if const then + let rec aux c = + match kind_of_term c with + | Const _ when Term.eq_constr c t -> body env t t + | App (f,a) when Term.eq_constr f t -> mkApp (body env f f,a) + | _ -> let c = Reductionops.whd_betaiotazeta sigma0 c in + match kind_of_term c with + | Const _ when Term.eq_constr c t -> body env t t + | App (f,a) when Term.eq_constr f t -> mkApp (body env f f,a) + | Const f -> aux (body env c c) + | App (f, a) -> aux (mkApp (body env f f, a)) + | _ -> errorstrm (str "The term "++pr_constr orig_c++ + str" contains no " ++ pr_constr t ++ str" even after unfolding") + in aux c + else + try body env t (fs (unify_HO env sigma c t) t) + with _ -> errorstrm (str "The term " ++ + pr_constr c ++spc()++ str "does not unify with " ++ pr_constr_pat t)), + fake_pmatcher_end in + let concl = + try beta env0 (eval_pattern env0 sigma0 concl0 rdx occ unfold) + with Option.IsNone -> errorstrm (str"Failed to unfold " ++ pr_constr_pat t) in + let _ = conclude () in + Proofview.V82.of_tactic (convert_concl concl) gl +;; + +let foldtac occ rdx ft gl = + let fs sigma x = Reductionops.nf_evar sigma x in + let sigma0, concl0, env0 = project gl, pf_concl gl, pf_env gl in + let sigma, t = ft in + let fold, conclude = match rdx with + | Some (_, (In_T _ | In_X_In_T _)) | None -> + let ise = create_evar_defs sigma in + let ut = red_product_skip_id env0 sigma t in + let ise, ut = mk_tpattern env0 sigma0 (ise,t) all_ok L2R ut in + let find_T, end_T = + mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[ut]) in + (fun env c _ h -> try find_T env c h (fun env t _ _ -> t) with NoMatch ->c), + (fun () -> try end_T () with NoMatch -> fake_pmatcher_end ()) + | _ -> + (fun env c _ h -> try let sigma = unify_HO env sigma c t in fs sigma t + with _ -> errorstrm (str "fold pattern " ++ pr_constr_pat t ++ spc () + ++ str "does not match redex " ++ pr_constr_pat c)), + fake_pmatcher_end in + let concl = eval_pattern env0 sigma0 concl0 rdx occ fold in + let _ = conclude () in + Proofview.V82.of_tactic (convert_concl concl) gl +;; + +let converse_dir = function L2R -> R2L | R2L -> L2R + +let rw_progress rhs lhs ise = not (Term.eq_constr lhs (Evarutil.nf_evar ise rhs)) + +(* Coq has a more general form of "equation" (any type with a single *) +(* constructor with no arguments with_rect_r elimination lemmas). *) +(* However there is no clear way of determining the LHS and RHS of *) +(* such a generic Leibnitz equation -- short of inspecting the type *) +(* of the elimination lemmas. *) + +let rec strip_prod_assum c = match kind_of_term c with + | Prod (_, _, c') -> strip_prod_assum c' + | LetIn (_, v, _, c') -> strip_prod_assum (subst1 v c) + | Cast (c', _, _) -> strip_prod_assum c' + | _ -> c + +let rule_id = mk_internal_id "rewrite rule" + +exception PRtype_error +exception PRindetermined_rhs of constr + +let pirrel_rewrite pred rdx rdx_ty new_rdx dir (sigma, c) c_ty gl = +(* pp(lazy(str"sigma@pirrel_rewrite=" ++ pr_evar_map None sigma)); *) + let env = pf_env gl in + let beta = Reductionops.clos_norm_flags Closure.beta env sigma in + let sigma, p = + let sigma = create_evar_defs sigma in + Evarutil.new_evar env sigma (beta (subst1 new_rdx pred)) in + let pred = mkNamedLambda pattern_id rdx_ty pred in + let elim, gl = + let ((kn, i) as ind, _), unfolded_c_ty = pf_reduce_to_quantified_ind gl c_ty in + let sort = elimination_sort_of_goal gl in + let elim, gl = pf_fresh_global (Indrec.lookup_eliminator ind sort) gl in + if dir = R2L then elim, gl else (* taken from Coq's rewrite *) + let elim, _ = destConst elim in + let mp,dp,l = repr_con (constant_of_kn (canonical_con elim)) in + let l' = label_of_id (Nameops.add_suffix (id_of_label l) "_r") in + let c1' = Global.constant_of_delta_kn (canonical_con (make_con mp dp l')) in + mkConst c1', gl in + let proof = mkApp (elim, [| rdx_ty; new_rdx; pred; p; rdx; c |]) in + (* We check the proof is well typed *) + let sigma, proof_ty = + try Typing.e_type_of env sigma proof with _ -> raise PRtype_error in + pp(lazy(str"pirrel_rewrite proof term of type: " ++ pr_constr proof_ty)); + try refine_with + ~first_goes_last:(not !ssroldreworder) ~with_evars:false (sigma, proof) gl + with _ -> + (* we generate a msg like: "Unable to find an instance for the variable" *) + let c = Reductionops.nf_evar sigma c in + let hd_ty, miss = match kind_of_term c with + | App (hd, args) -> + let hd_ty = Retyping.get_type_of env sigma hd in + let names = let rec aux t = function 0 -> [] | n -> + let t = Reductionops.whd_betadeltaiota env sigma t in + match kind_of_type t with + | ProdType (name, _, t) -> name :: aux t (n-1) + | _ -> assert false in aux hd_ty (Array.length args) in + hd_ty, Util.List.map_filter (fun (t, name) -> + let evs = Intset.elements (Evarutil.undefined_evars_of_term sigma t) in + let open_evs = List.filter (fun k -> + InProp <> Retyping.get_sort_family_of + env sigma (Evd.evar_concl (Evd.find sigma k))) + evs in + if open_evs <> [] then Some name else None) + (List.combine (Array.to_list args) names) + | _ -> anomaly "rewrite rule not an application" in + errorstrm (Himsg.explain_refiner_error (Logic.UnresolvedBindings miss)++ + (Pp.fnl()++str"Rule's type:" ++ spc() ++ pr_constr hd_ty)) +;; + +let is_const_ref c r = isConst c && eq_gr (ConstRef (fst(destConst c))) r +let is_construct_ref c r = + isConstruct c && eq_gr (ConstructRef (fst(destConstruct c))) r +let is_ind_ref c r = isInd c && eq_gr (IndRef (fst(destInd c))) r + +let rwcltac cl rdx dir sr gl = + let n, r_n,_, ucst = pf_abs_evars gl sr in + let r_n' = pf_abs_cterm gl n r_n in + let r' = subst_var pattern_id r_n' in + let gl = pf_merge_uc ucst gl in + let rdxt = Retyping.get_type_of (pf_env gl) (fst sr) rdx in +(* pp(lazy(str"sigma@rwcltac=" ++ pr_evar_map None (fst sr))); *) + pp(lazy(str"r@rwcltac=" ++ pr_constr (snd sr))); + let cvtac, rwtac, gl = + if closed0 r' then + let env, sigma, c, c_eq = pf_env gl, fst sr, snd sr, build_coq_eq () in + let sigma, c_ty = Typing.e_type_of env sigma c in + pp(lazy(str"c_ty@rwcltac=" ++ pr_constr c_ty)); + match kind_of_type (Reductionops.whd_betadeltaiota env sigma c_ty) with + | AtomicType(e, a) when is_ind_ref e c_eq -> + let new_rdx = if dir = L2R then a.(2) else a.(1) in + pirrel_rewrite cl rdx rdxt new_rdx dir (sigma,c) c_ty, tclIDTAC, gl + | _ -> + let cl' = mkApp (mkNamedLambda pattern_id rdxt cl, [|rdx|]) in + let sigma, _ = Typing.e_type_of env sigma cl' in + let gl = pf_merge_uc_of sigma gl in + Proofview.V82.of_tactic (convert_concl cl'), rewritetac dir r', gl + else + let dc, r2 = decompose_lam_n n r' in + let r3, _, r3t = + try destCast r2 with _ -> + errorstrm (str "no cast from " ++ pr_constr_pat (snd sr) + ++ str " to " ++ pr_constr r2) in + let cl' = mkNamedProd rule_id (compose_prod dc r3t) (lift 1 cl) in + let cl'' = mkNamedProd pattern_id rdxt cl' in + let itacs = [introid pattern_id; introid rule_id] in + let cltac = clear [pattern_id; rule_id] in + let rwtacs = [rewritetac dir (mkVar rule_id); cltac] in + apply_type cl'' [rdx; compose_lam dc r3], tclTHENLIST (itacs @ rwtacs), gl + in + let cvtac' _ = + try cvtac gl with + | PRtype_error -> + if occur_existential (pf_concl gl) + then errorstrm (str "Rewriting impacts evars") + else errorstrm (str "Dependent type error in rewrite of " + ++ pf_pr_constr gl (project gl) (mkNamedLambda pattern_id rdxt cl)) + | Errors.UserError _ as e -> raise e + | e -> anomaly ("cvtac's exception: " ^ Printexc.to_string e); + in + tclTHEN cvtac' rwtac gl + +let prof_rwcltac = mk_profiler "rwrxtac.rwcltac";; +let rwcltac cl rdx dir sr gl = + prof_rwcltac.profile (rwcltac cl rdx dir sr) gl +;; + + +let lz_coq_prod = + let prod = lazy (build_prod ()) in fun () -> Lazy.force prod + +let lz_setoid_relation = + let sdir = ["Classes"; "RelationClasses"] in + let last_srel = ref (Environ.empty_env, None) in + fun env -> match !last_srel with + | env', srel when env' == env -> srel + | _ -> + let srel = + try Some (coq_constant "Class_setoid" sdir "RewriteRelation") + with _ -> None in + last_srel := (env, srel); srel + +let ssr_is_setoid env = + match lz_setoid_relation env with + | None -> fun _ _ _ -> false + | Some srel -> + fun sigma r args -> + Rewrite.is_applied_rewrite_relation env + sigma [] (mkApp (r, args)) <> None + +let prof_rwxrtac_find_rule = mk_profiler "rwrxtac.find_rule";; + +let closed0_check cl p gl = + if closed0 cl then + errorstrm (str"No occurrence of redex "++pf_pr_constr gl (project gl) p) + +let rwprocess_rule dir rule gl = + let env = pf_env gl in + let coq_prod = lz_coq_prod () in + let is_setoid = ssr_is_setoid env in + let r_sigma, rules = + let rec loop d sigma r t0 rs red = + let t = + if red = 1 then Tacred.hnf_constr env sigma t0 + else Reductionops.whd_betaiotazeta sigma t0 in + pp(lazy(str"rewrule="++pr_constr t)); + match kind_of_term t with + | Prod (_, xt, at) -> + let ise, x = Evarutil.new_evar env (create_evar_defs sigma) xt in + loop d ise (mkApp (r, [|x|])) (subst1 x at) rs 0 + | App (pr, a) when is_ind_ref pr coq_prod.Coqlib.typ -> + let sr sigma = match kind_of_term (Tacred.hnf_constr env sigma r) with + | App (c, ra) when is_construct_ref c coq_prod.Coqlib.intro -> + fun i -> ra.(i + 1), sigma + | _ -> let ra = Array.append a [|r|] in + function 1 -> + let sigma, pi1 = Evd.fresh_global env sigma coq_prod.Coqlib.proj1 in + mkApp (pi1, ra), sigma + | _ -> + let sigma, pi2 = Evd.fresh_global env sigma coq_prod.Coqlib.proj2 in + mkApp (pi2, ra), sigma in + if Term.eq_constr a.(0) (build_coq_True ()) then + let s, sigma = sr sigma 2 in + loop (converse_dir d) sigma s a.(1) rs 0 + else + let s, sigma = sr sigma 2 in + let sigma, rs2 = loop d sigma s a.(1) rs 0 in + let s, sigma = sr sigma 1 in + loop d sigma s a.(0) rs2 0 + | App (r_eq, a) when Hipattern.match_with_equality_type t != None -> + let indu = destInd r_eq and rhs = Array.last a in + let np = Inductiveops.inductive_nparamdecls (fst indu) in + let ind_ct = Inductiveops.type_of_constructors env indu in + let lhs0 = last_arg (strip_prod_assum ind_ct.(0)) in + let rdesc = match kind_of_term lhs0 with + | Rel i -> + let lhs = a.(np - i) in + let lhs, rhs = if d = L2R then lhs, rhs else rhs, lhs in +(* msgnl (str "RW: " ++ pr_rwdir d ++ str " " ++ pr_constr_pat r ++ str " : " + ++ pr_constr_pat lhs ++ str " ~> " ++ pr_constr_pat rhs); *) + d, r, lhs, rhs +(* + let l_i, r_i = if d = L2R then i, 1 - ndep else 1 - ndep, i in + let lhs = a.(np - l_i) and rhs = a.(np - r_i) in + let a' = Array.copy a in let _ = a'.(np - l_i) <- mkVar pattern_id in + let r' = mkCast (r, DEFAULTcast, mkApp (r_eq, a')) in + (d, r', lhs, rhs) +*) + | _ -> + let lhs = substl (array_list_of_tl (Array.sub a 0 np)) lhs0 in + let lhs, rhs = if d = R2L then lhs, rhs else rhs, lhs in + let d' = if Array.length a = 1 then d else converse_dir d in + d', r, lhs, rhs in + sigma, rdesc :: rs + | App (s_eq, a) when is_setoid sigma s_eq a -> + let np = Array.length a and i = 3 - dir_org d in + let lhs = a.(np - i) and rhs = a.(np + i - 3) in + let a' = Array.copy a in let _ = a'.(np - i) <- mkVar pattern_id in + let r' = mkCast (r, DEFAULTcast, mkApp (s_eq, a')) in + sigma, (d, r', lhs, rhs) :: rs + | _ -> + if red = 0 then loop d sigma r t rs 1 + else errorstrm (str "not a rewritable relation: " ++ pr_constr_pat t + ++ spc() ++ str "in rule " ++ pr_constr_pat (snd rule)) + in + let sigma, r = rule in + let t = Retyping.get_type_of env sigma r in + loop dir sigma r t [] 0 + in + r_sigma, rules + +let rwrxtac occ rdx_pat dir rule gl = + let env = pf_env gl in + let r_sigma, rules = rwprocess_rule dir rule gl in + let find_rule rdx = + let rec rwtac = function + | [] -> + errorstrm (str "pattern " ++ pr_constr_pat rdx ++ + str " does not match " ++ pr_dir_side dir ++ + str " of " ++ pr_constr_pat (snd rule)) + | (d, r, lhs, rhs) :: rs -> + try + let ise = unify_HO env (create_evar_defs r_sigma) lhs rdx in + if not (rw_progress rhs rdx ise) then raise NoMatch else + d, (ise, Evd.evar_universe_context ise, Reductionops.nf_evar ise r) + with _ -> rwtac rs in + rwtac rules in + let find_rule rdx = prof_rwxrtac_find_rule.profile find_rule rdx in + let sigma0, env0, concl0 = project gl, pf_env gl, pf_concl gl in + let find_R, conclude = match rdx_pat with + | Some (_, (In_T _ | In_X_In_T _)) | None -> + let upats_origin = dir, snd rule in + let rpat env sigma0 (sigma, pats) (d, r, lhs, rhs) = + let sigma, pat = + mk_tpattern env sigma0 (sigma,r) (rw_progress rhs) d lhs in + sigma, pats @ [pat] in + let rpats = List.fold_left (rpat env0 sigma0) (r_sigma,[]) rules in + let find_R, end_R = mk_tpattern_matcher sigma0 occ ~upats_origin rpats in + (fun e c _ i -> find_R ~k:(fun _ _ _ h -> mkRel h) e c i), + fun cl -> let rdx,d,r = end_R () in closed0_check cl rdx gl; (d,r),rdx + | Some(_, (T e | X_In_T (_,e) | E_As_X_In_T (e,_,_) | E_In_X_In_T (e,_,_))) -> + let r = ref None in + (fun env c _ h -> do_once r (fun () -> find_rule c, c); mkRel h), + (fun concl -> closed0_check concl e gl; assert_done r) in + let concl = eval_pattern env0 sigma0 concl0 rdx_pat occ find_R in + let (d, r), rdx = conclude concl in + let r = Evd.merge_universe_context (pi1 r) (pi2 r), pi3 r in + rwcltac concl rdx d r gl +;; + +let prof_rwxrtac = mk_profiler "rwrxtac";; +let rwrxtac occ rdx_pat dir rule gl = + prof_rwxrtac.profile (rwrxtac occ rdx_pat dir rule) gl +;; + +let ssrinstancesofrule ist dir arg gl = + let sigma0, env0, concl0 = project gl, pf_env gl, pf_concl gl in + let rule = interp_term ist gl arg in + let r_sigma, rules = rwprocess_rule dir rule gl in + let find, conclude = + let upats_origin = dir, snd rule in + let rpat env sigma0 (sigma, pats) (d, r, lhs, rhs) = + let sigma, pat = + mk_tpattern env sigma0 (sigma,r) (rw_progress rhs) d lhs in + sigma, pats @ [pat] in + let rpats = List.fold_left (rpat env0 sigma0) (r_sigma,[]) rules in + mk_tpattern_matcher ~all_instances:true ~raise_NoMatch:true sigma0 None ~upats_origin rpats in + let print env p c _ = ppnl (hov 1 (str"instance:" ++ spc() ++ pr_constr p ++ spc() ++ str "matches:" ++ spc() ++ pr_constr c)); c in + ppnl (str"BEGIN INSTANCES"); + try + while true do + ignore(find env0 concl0 1 ~k:print) + done; raise NoMatch + with NoMatch -> ppnl (str"END INSTANCES"); tclIDTAC gl + +TACTIC EXTEND ssrinstofruleL2R +| [ "ssrinstancesofruleL2R" ssrterm(arg) ] -> [ Proofview.V82.tactic (ssrinstancesofrule ist L2R arg) ] +END +TACTIC EXTEND ssrinstofruleR2L +| [ "ssrinstancesofruleR2L" ssrterm(arg) ] -> [ Proofview.V82.tactic (ssrinstancesofrule ist R2L arg) ] +END + +(* Resolve forward reference *) +let _ = + ipat_rewritetac := fun occ dir c gl -> rwrxtac occ None dir (project gl, c) gl + +let rwargtac ist ((dir, mult), (((oclr, occ), grx), (kind, gt))) gl = + let fail = ref false in + let interp_rpattern ist gl gc = + try interp_rpattern ist gl gc + with _ when snd mult = May -> fail := true; project gl, T mkProp in + let interp gc gl = + try interp_term ist gl gc + with _ when snd mult = May -> fail := true; (project gl, mkProp) in + let rwtac gl = + let rx = Option.map (interp_rpattern ist gl) grx in + let t = interp gt gl in + (match kind with + | RWred sim -> simplintac occ rx sim + | RWdef -> if dir = R2L then foldtac occ rx t else unfoldintac occ rx t gt + | RWeq -> rwrxtac occ rx dir t) gl in + let ctac = cleartac (interp_clr (oclr, (fst gt, snd (interp gt gl)))) in + if !fail then ctac gl else tclTHEN (tclMULT mult rwtac) ctac gl + +(** Rewrite argument sequence *) + +(* type ssrrwargs = ssrrwarg list *) + +let pr_ssrrwargs _ _ _ rwargs = pr_list spc pr_rwarg rwargs + +ARGUMENT EXTEND ssrrwargs TYPED AS ssrrwarg list PRINTED BY pr_ssrrwargs + | [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] +END + +let ssr_rw_syntax = Summary.ref ~name:"SSR:rewrite" true + +let _ = + Goptions.declare_bool_option + { Goptions.optsync = true; + Goptions.optname = "ssreflect rewrite"; + Goptions.optkey = ["SsrRewrite"]; + Goptions.optread = (fun _ -> !ssr_rw_syntax); + Goptions.optdepr = false; + Goptions.optwrite = (fun b -> ssr_rw_syntax := b) } + +let test_ssr_rw_syntax = + let test strm = + if not !ssr_rw_syntax then raise Stream.Failure else + if is_ssr_loaded () then () else + match Compat.get_tok (Util.stream_nth 0 strm) with + | Tok.KEYWORD key when List.mem key.[0] ['{'; '['; '/'] -> () + | _ -> raise Stream.Failure in + Gram.Entry.of_parser "test_ssr_rw_syntax" test + +GEXTEND Gram + GLOBAL: ssrrwargs; + ssrrwargs: [[ test_ssr_rw_syntax; a = LIST1 ssrrwarg -> a ]]; +END + +(** The "rewrite" tactic *) + +let ssrrewritetac ist rwargs = + tclTHENLIST (List.map (rwargtac ist) rwargs) + +TACTIC EXTEND ssrrewrite + | [ "rewrite" ssrrwargs(args) ssrclauses(clauses) ] -> + [ Proofview.V82.tactic (tclCLAUSES ist (ssrrewritetac ist args) clauses) ] +END + +(** The "unlock" tactic *) + +let pr_unlockarg (occ, t) = pr_occ occ ++ pr_term t +let pr_ssrunlockarg _ _ _ = pr_unlockarg + +ARGUMENT EXTEND ssrunlockarg TYPED AS ssrocc * ssrterm + PRINTED BY pr_ssrunlockarg + | [ "{" ssrocc(occ) "}" ssrterm(t) ] -> [ occ, t ] + | [ ssrterm(t) ] -> [ None, t ] +END + +let pr_ssrunlockargs _ _ _ args = pr_list spc pr_unlockarg args + +ARGUMENT EXTEND ssrunlockargs TYPED AS ssrunlockarg list + PRINTED BY pr_ssrunlockargs + | [ ssrunlockarg_list(args) ] -> [ args ] +END + +let unfoldtac occ ko t kt gl = + let cl, c = pf_fill_occ_term gl occ (fst (strip_unfold_term t kt)) in + let cl' = subst1 (pf_unfoldn [OnlyOccurrences [1], get_evalref c] gl c) cl in + let f = if ko = None then Closure.betaiotazeta else Closure.betaiota in + Proofview.V82.of_tactic + (convert_concl (pf_reduce (Reductionops.clos_norm_flags f) gl cl')) gl + +let unlocktac ist args gl = + let utac (occ, gt) gl = + unfoldtac occ occ (interp_term ist gl gt) (fst gt) gl in + let locked, gl = pf_mkSsrConst "locked" gl in + let key, gl = pf_mkSsrConst "master_key" gl in + let ktacs = [ + (fun gl -> unfoldtac None None (project gl,locked) '(' gl); + simplest_newcase key ] in + tclTHENLIST (List.map utac args @ ktacs) gl + +TACTIC EXTEND ssrunlock + | [ "unlock" ssrunlockargs(args) ssrclauses(clauses) ] -> +[ Proofview.V82.tactic (tclCLAUSES ist (unlocktac ist args) clauses) ] +END + +(** 8. Forward chaining tactics (pose, set, have, suffice, wlog) *) + +(** Defined identifier *) + +type ssrfwdid = identifier + +let pr_ssrfwdid _ _ _ id = pr_spc () ++ pr_id id + +(* We use a primitive parser for the head identifier of forward *) +(* tactis to avoid syntactic conflicts with basic Coq tactics. *) +ARGUMENT EXTEND ssrfwdid TYPED AS ident PRINTED BY pr_ssrfwdid + | [ "YouShouldNotTypeThis" ] -> [ anomaly "Grammar placeholder match" ] +END + +let accept_ssrfwdid strm = + match Compat.get_tok (stream_nth 0 strm) with + | Tok.IDENT id -> accept_before_syms_or_any_id [":"; ":="; "("] strm + | _ -> raise Stream.Failure + + +let test_ssrfwdid = Gram.Entry.of_parser "test_ssrfwdid" accept_ssrfwdid + +GEXTEND Gram + GLOBAL: ssrfwdid; + ssrfwdid: [[ test_ssrfwdid; id = Prim.ident -> id ]]; + END + + + +(** Definition value formatting *) + +(* We use an intermediate structure to correctly render the binder list *) +(* abbreviations. We use a list of hints to extract the binders and *) +(* base term from a term, for the two first levels of representation of *) +(* of constr terms. *) + +type 'term ssrbind = + | Bvar of name + | Bdecl of name list * 'term + | Bdef of name * 'term option * 'term + | Bstruct of name + | Bcast of 'term + +let pr_binder prl = function + | Bvar x -> + pr_name x + | Bdecl (xs, t) -> + str "(" ++ pr_list pr_spc pr_name xs ++ str " : " ++ prl t ++ str ")" + | Bdef (x, None, v) -> + str "(" ++ pr_name x ++ str " := " ++ prl v ++ str ")" + | Bdef (x, Some t, v) -> + str "(" ++ pr_name x ++ str " : " ++ prl t ++ + str " := " ++ prl v ++ str ")" + | Bstruct x -> + str "{struct " ++ pr_name x ++ str "}" + | Bcast t -> + str ": " ++ prl t + +type 'term ssrbindval = 'term ssrbind list * 'term + +type ssrbindfmt = + | BFvar + | BFdecl of int (* #xs *) + | BFcast (* final cast *) + | BFdef of bool (* has cast? *) + | BFrec of bool * bool (* has struct? * has cast? *) + +let rec mkBstruct i = function + | Bvar x :: b -> + if i = 0 then [Bstruct x] else mkBstruct (i - 1) b + | Bdecl (xs, _) :: b -> + let i' = i - List.length xs in + if i' < 0 then [Bstruct (List.nth xs i)] else mkBstruct i' b + | _ :: b -> mkBstruct i b + | [] -> [] + +let rec format_local_binders h0 bl0 = match h0, bl0 with + | BFvar :: h, LocalRawAssum ([_, x], _, _) :: bl -> + Bvar x :: format_local_binders h bl + | BFdecl _ :: h, LocalRawAssum (lxs, _, t) :: bl -> + Bdecl (List.map snd lxs, t) :: format_local_binders h bl + | BFdef false :: h, LocalRawDef ((_, x), v) :: bl -> + Bdef (x, None, v) :: format_local_binders h bl + | BFdef true :: h, + LocalRawDef ((_, x), CCast (_, v, CastConv t)) :: bl -> + Bdef (x, Some t, v) :: format_local_binders h bl + | _ -> [] + +let rec format_constr_expr h0 c0 = match h0, c0 with + | BFvar :: h, CLambdaN (_, [[_, x], _, _], c) -> + let bs, c' = format_constr_expr h c in + Bvar x :: bs, c' + | BFdecl _:: h, CLambdaN (_, [lxs, _, t], c) -> + let bs, c' = format_constr_expr h c in + Bdecl (List.map snd lxs, t) :: bs, c' + | BFdef false :: h, CLetIn(_, (_, x), v, c) -> + let bs, c' = format_constr_expr h c in + Bdef (x, None, v) :: bs, c' + | BFdef true :: h, CLetIn(_, (_, x), CCast (_, v, CastConv t), c) -> + let bs, c' = format_constr_expr h c in + Bdef (x, Some t, v) :: bs, c' + | [BFcast], CCast (_, c, CastConv t) -> + [Bcast t], c + | BFrec (has_str, has_cast) :: h, + CFix (_, _, [_, (Some locn, CStructRec), bl, t, c]) -> + let bs = format_local_binders h bl in + let bstr = if has_str then [Bstruct (Name (snd locn))] else [] in + bs @ bstr @ (if has_cast then [Bcast t] else []), c + | BFrec (_, has_cast) :: h, CCoFix (_, _, [_, bl, t, c]) -> + format_local_binders h bl @ (if has_cast then [Bcast t] else []), c + | _, c -> + [], c + +let rec format_glob_decl h0 d0 = match h0, d0 with + | BFvar :: h, (x, _, None, _) :: d -> + Bvar x :: format_glob_decl h d + | BFdecl 1 :: h, (x, _, None, t) :: d -> + Bdecl ([x], t) :: format_glob_decl h d + | BFdecl n :: h, (x, _, None, t) :: d when n > 1 -> + begin match format_glob_decl (BFdecl (n - 1) :: h) d with + | Bdecl (xs, _) :: bs -> Bdecl (x :: xs, t) :: bs + | bs -> Bdecl ([x], t) :: bs + end + | BFdef false :: h, (x, _, Some v, _) :: d -> + Bdef (x, None, v) :: format_glob_decl h d + | BFdef true:: h, (x, _, Some (GCast (_, v, CastConv t)), _) :: d -> + Bdef (x, Some t, v) :: format_glob_decl h d + | _, (x, _, None, t) :: d -> + Bdecl ([x], t) :: format_glob_decl [] d + | _, (x, _, Some v, _) :: d -> + Bdef (x, None, v) :: format_glob_decl [] d + | _, [] -> [] + +let rec format_glob_constr h0 c0 = match h0, c0 with + | BFvar :: h, GLambda (_, x, _, _, c) -> + let bs, c' = format_glob_constr h c in + Bvar x :: bs, c' + | BFdecl 1 :: h, GLambda (_, x, _, t, c) -> + let bs, c' = format_glob_constr h c in + Bdecl ([x], t) :: bs, c' + | BFdecl n :: h, GLambda (_, x, _, t, c) when n > 1 -> + begin match format_glob_constr (BFdecl (n - 1) :: h) c with + | Bdecl (xs, _) :: bs, c' -> Bdecl (x :: xs, t) :: bs, c' + | _ -> [Bdecl ([x], t)], c + end + | BFdef false :: h, GLetIn(_, x, v, c) -> + let bs, c' = format_glob_constr h c in + Bdef (x, None, v) :: bs, c' + | BFdef true :: h, GLetIn(_, x, GCast (_, v, CastConv t), c) -> + let bs, c' = format_glob_constr h c in + Bdef (x, Some t, v) :: bs, c' + | [BFcast], GCast (_, c, CastConv t) -> + [Bcast t], c + | BFrec (has_str, has_cast) :: h, GRec (_, f, _, bl, t, c) + when Array.length c = 1 -> + let bs = format_glob_decl h bl.(0) in + let bstr = match has_str, f with + | true, GFix ([|Some i, GStructRec|], _) -> mkBstruct i bs + | _ -> [] in + bs @ bstr @ (if has_cast then [Bcast t.(0)] else []), c.(0) + | _, c -> + [], c + +(** Forward chaining argument *) + +(* There are three kinds of forward definitions: *) +(* - Hint: type only, cast to Type, may have proof hint. *) +(* - Have: type option + value, no space before type *) +(* - Pose: binders + value, space before binders. *) + +type ssrfwdkind = FwdHint of string * bool | FwdHave | FwdPose + +type ssrfwdfmt = ssrfwdkind * ssrbindfmt list + +let pr_fwdkind = function + | FwdHint (s,_) -> str (s ^ " ") | _ -> str " :=" ++ spc () +let pr_fwdfmt (fk, _ : ssrfwdfmt) = pr_fwdkind fk + +let wit_ssrfwdfmt = add_genarg "ssrfwdfmt" pr_fwdfmt + +(* type ssrfwd = ssrfwdfmt * ssrterm *) + +let mkFwdVal fk c = ((fk, []), mk_term ' ' c) +let mkssrFwdVal fk c = ((fk, []), (c,None)) + +let mkFwdCast fk loc t c = ((fk, [BFcast]), mk_term ' ' (CCast (loc, c, dC t))) +let mkssrFwdCast fk loc t c = ((fk, [BFcast]), (c, Some t)) + +let mkFwdHint s t = + let loc = constr_loc t in + mkFwdCast (FwdHint (s,false)) loc t (mkCHole loc) +let mkFwdHintNoTC s t = + let loc = constr_loc t in + mkFwdCast (FwdHint (s,true)) loc t (mkCHole loc) + +let pr_gen_fwd prval prc prlc fk (bs, c) = + let prc s = str s ++ spc () ++ prval prc prlc c in + match fk, bs with + | FwdHint (s,_), [Bcast t] -> str s ++ spc () ++ prlc t + | FwdHint (s,_), _ -> prc (s ^ "(* typeof *)") + | FwdHave, [Bcast t] -> str ":" ++ spc () ++ prlc t ++ prc " :=" + | _, [] -> prc " :=" + | _, _ -> spc () ++ pr_list spc (pr_binder prlc) bs ++ prc " :=" + +let pr_fwd_guarded prval prval' = function +| (fk, h), (_, (_, Some c)) -> + pr_gen_fwd prval pr_constr_expr prl_constr_expr fk (format_constr_expr h c) +| (fk, h), (_, (c, None)) -> + pr_gen_fwd prval' pr_glob_constr prl_glob_constr fk (format_glob_constr h c) + +let pr_unguarded prc prlc = prlc + +let pr_fwd = pr_fwd_guarded pr_unguarded pr_unguarded +let pr_ssrfwd _ _ _ = pr_fwd + +ARGUMENT EXTEND ssrfwd TYPED AS (ssrfwdfmt * ssrterm) PRINTED BY pr_ssrfwd + | [ ":=" lconstr(c) ] -> [ mkFwdVal FwdPose c ] + | [ ":" lconstr (t) ":=" lconstr(c) ] -> [ mkFwdCast FwdPose loc t c ] +END + +(** Independent parsing for binders *) + +(* The pose, pose fix, and pose cofix tactics use these internally to *) +(* parse argument fragments. *) + +let pr_ssrbvar prc _ _ v = prc v + +ARGUMENT EXTEND ssrbvar TYPED AS constr PRINTED BY pr_ssrbvar +| [ ident(id) ] -> [ mkCVar loc id ] +| [ "_" ] -> [ mkCHole loc ] +END + +let bvar_lname = function + | CRef (Ident (loc, id), _) -> loc, Name id + | c -> constr_loc c, Anonymous + +let pr_ssrbinder prc _ _ (_, c) = prc c + +ARGUMENT EXTEND ssrbinder TYPED AS ssrfwdfmt * constr PRINTED BY pr_ssrbinder + | [ ssrbvar(bv) ] -> + [ let xloc, _ as x = bvar_lname bv in + (FwdPose, [BFvar]), + CLambdaN (loc,[[x],Default Explicit,mkCHole xloc],mkCHole loc) ] + | [ "(" ssrbvar(bv) ")" ] -> + [ let xloc, _ as x = bvar_lname bv in + (FwdPose, [BFvar]), + CLambdaN (loc,[[x],Default Explicit,mkCHole xloc],mkCHole loc) ] + | [ "(" ssrbvar(bv) ":" lconstr(t) ")" ] -> + [ let x = bvar_lname bv in + (FwdPose, [BFdecl 1]), + CLambdaN (loc, [[x], Default Explicit, t], mkCHole loc) ] + | [ "(" ssrbvar(bv) ne_ssrbvar_list(bvs) ":" lconstr(t) ")" ] -> + [ let xs = List.map bvar_lname (bv :: bvs) in + let n = List.length xs in + (FwdPose, [BFdecl n]), + CLambdaN (loc, [xs, Default Explicit, t], mkCHole loc) ] + | [ "(" ssrbvar(id) ":" lconstr(t) ":=" lconstr(v) ")" ] -> + [ let loc' = Loc.join_loc (constr_loc t) (constr_loc v) in + let v' = CCast (loc', v, dC t) in + (FwdPose,[BFdef true]), CLetIn (loc,bvar_lname id, v',mkCHole loc) ] + | [ "(" ssrbvar(id) ":=" lconstr(v) ")" ] -> + [ (FwdPose,[BFdef false]), CLetIn (loc,bvar_lname id, v,mkCHole loc) ] +END + +GEXTEND Gram + GLOBAL: ssrbinder; + ssrbinder: [ + [ ["of" | "&"]; c = operconstr LEVEL "99" -> + let loc = !@loc in + (FwdPose, [BFvar]), + CLambdaN (loc,[[loc,Anonymous],Default Explicit,c],mkCHole loc) ] + ]; +END + +let rec binders_fmts = function + | ((_, h), _) :: bs -> h @ binders_fmts bs + | _ -> [] + +let push_binders c2 bs = + let loc2 = constr_loc c2 in let mkloc loc1 = Loc.join_loc loc1 loc2 in + let rec loop ty c = function + | (_, CLambdaN (loc1, b, _)) :: bs when ty -> + CProdN (mkloc loc1, b, loop ty c bs) + | (_, CLambdaN (loc1, b, _)) :: bs -> + CLambdaN (mkloc loc1, b, loop ty c bs) + | (_, CLetIn (loc1, x, v, _)) :: bs -> + CLetIn (mkloc loc1, x, v, loop ty c bs) + | [] -> c + | _ -> anomaly "binder not a lambda nor a let in" in + match c2 with + | CCast (x, ct, CastConv cty) -> + (CCast (x, loop false ct bs, CastConv (loop true cty bs))) + | ct -> loop false ct bs + +let rec fix_binders = function + | (_, CLambdaN (_, [xs, _, t], _)) :: bs -> + LocalRawAssum (xs, Default Explicit, t) :: fix_binders bs + | (_, CLetIn (_, x, v, _)) :: bs -> + LocalRawDef (x, v) :: fix_binders bs + | _ -> [] + +let pr_ssrstruct _ _ _ = function + | Some id -> str "{struct " ++ pr_id id ++ str "}" + | None -> mt () + +ARGUMENT EXTEND ssrstruct TYPED AS ident option PRINTED BY pr_ssrstruct +| [ "{" "struct" ident(id) "}" ] -> [ Some id ] +| [ ] -> [ None ] +END + +(** The "pose" tactic *) + +(* The plain pose form. *) + +let bind_fwd bs = function + | (fk, h), (ck, (rc, Some c)) -> + (fk,binders_fmts bs @ h), (ck,(rc,Some (push_binders c bs))) + | fwd -> fwd + +ARGUMENT EXTEND ssrposefwd TYPED AS ssrfwd PRINTED BY pr_ssrfwd + | [ ssrbinder_list(bs) ssrfwd(fwd) ] -> [ bind_fwd bs fwd ] +END + +(* The pose fix form. *) + +let pr_ssrfixfwd _ _ _ (id, fwd) = str " fix " ++ pr_id id ++ pr_fwd fwd + +let bvar_locid = function + | CRef (Ident (loc, id), _) -> loc, id + | _ -> Errors.error "Missing identifier after \"(co)fix\"" + + +ARGUMENT EXTEND ssrfixfwd TYPED AS ident * ssrfwd PRINTED BY pr_ssrfixfwd + | [ "fix" ssrbvar(bv) ssrbinder_list(bs) ssrstruct(sid) ssrfwd(fwd) ] -> + [ let (_, id) as lid = bvar_locid bv in + let (fk, h), (ck, (rc, oc)) = fwd in + let c = Option.get oc in + let has_cast, t', c' = match format_constr_expr h c with + | [Bcast t'], c' -> true, t', c' + | _ -> false, mkCHole (constr_loc c), c in + let lb = fix_binders bs in + let has_struct, i = + let rec loop = function + (l', Name id') :: _ when Option.equal Id.equal sid (Some id') -> true, (l', id') + | [l', Name id'] when sid = None -> false, (l', id') + | _ :: bn -> loop bn + | [] -> Errors.error "Bad structural argument" in + loop (names_of_local_assums lb) in + let h' = BFrec (has_struct, has_cast) :: binders_fmts bs in + let fix = CFix (loc, lid, [lid, (Some i, CStructRec), lb, t', c']) in + id, ((fk, h'), (ck, (rc, Some fix))) ] +END + + +(* The pose cofix form. *) + +let pr_ssrcofixfwd _ _ _ (id, fwd) = str " cofix " ++ pr_id id ++ pr_fwd fwd + +ARGUMENT EXTEND ssrcofixfwd TYPED AS ssrfixfwd PRINTED BY pr_ssrcofixfwd + | [ "cofix" ssrbvar(bv) ssrbinder_list(bs) ssrfwd(fwd) ] -> + [ let _, id as lid = bvar_locid bv in + let (fk, h), (ck, (rc, oc)) = fwd in + let c = Option.get oc in + let has_cast, t', c' = match format_constr_expr h c with + | [Bcast t'], c' -> true, t', c' + | _ -> false, mkCHole (constr_loc c), c in + let h' = BFrec (false, has_cast) :: binders_fmts bs in + let cofix = CCoFix (loc, lid, [lid, fix_binders bs, t', c']) in + id, ((fk, h'), (ck, (rc, Some cofix))) + ] +END + +let ssrposetac ist (id, (_, t)) gl = + let sigma, t, ucst = pf_abs_ssrterm ist gl t in + posetac id t (pf_merge_uc ucst gl) + + +let prof_ssrposetac = mk_profiler "ssrposetac";; +let ssrposetac arg gl = prof_ssrposetac.profile (ssrposetac arg) gl;; + +TACTIC EXTEND ssrpose +| [ "pose" ssrfixfwd(ffwd) ] -> [ Proofview.V82.tactic (ssrposetac ist ffwd) ] +| [ "pose" ssrcofixfwd(ffwd) ] -> [ Proofview.V82.tactic (ssrposetac ist ffwd) ] +| [ "pose" ssrfwdid(id) ssrposefwd(fwd) ] -> [ Proofview.V82.tactic (ssrposetac ist (id, fwd)) ] +END + +(** The "set" tactic *) + +(* type ssrsetfwd = ssrfwd * ssrdocc *) + +let guard_setrhs s i = s.[i] = '{' + +let pr_setrhs occ prc prlc c = + if occ = nodocc then pr_guarded guard_setrhs prlc c else pr_docc occ ++ prc c + +let pr_fwd_guarded prval prval' = function +| (fk, h), (_, (_, Some c)) -> + pr_gen_fwd prval pr_constr_expr prl_constr_expr fk (format_constr_expr h c) +| (fk, h), (_, (c, None)) -> + pr_gen_fwd prval' pr_glob_constr prl_glob_constr fk (format_glob_constr h c) + +(* This does not print the type, it should be fixed... *) +let pr_ssrsetfwd _ _ _ (((fk,_),(t,_)), docc) = + pr_gen_fwd (fun _ _ -> pr_cpattern) + (fun _ -> mt()) (fun _ -> mt()) fk ([Bcast ()],t) + +ARGUMENT EXTEND ssrsetfwd +TYPED AS (ssrfwdfmt * (lcpattern * ssrterm option)) * ssrdocc +PRINTED BY pr_ssrsetfwd +| [ ":" lconstr(t) ":=" "{" ssrocc(occ) "}" cpattern(c) ] -> + [ mkssrFwdCast FwdPose loc (mk_lterm t) c, mkocc occ ] +| [ ":" lconstr(t) ":=" lcpattern(c) ] -> + [ mkssrFwdCast FwdPose loc (mk_lterm t) c, nodocc ] +| [ ":=" "{" ssrocc(occ) "}" cpattern(c) ] -> + [ mkssrFwdVal FwdPose c, mkocc occ ] +| [ ":=" lcpattern(c) ] -> [ mkssrFwdVal FwdPose c, nodocc ] +END + +let ssrsettac ist id ((_, (pat, pty)), (_, occ)) gl = + let pat = interp_cpattern ist gl pat (Option.map snd pty) in + let cl, sigma, env = pf_concl gl, project gl, pf_env gl in + let (c, ucst), cl = + try fill_occ_pattern ~raise_NoMatch:true env sigma cl pat occ 1 + with NoMatch -> redex_of_pattern ~resolve_typeclasses:true env pat, cl in + if occur_existential c then errorstrm(str"The pattern"++spc()++ + pr_constr_pat c++spc()++str"did not match and has holes."++spc()++ + str"Did you mean pose?") else + let c, cty = match kind_of_term c with + | Cast(t, DEFAULTcast, ty) -> t, ty + | _ -> c, pf_type_of gl c in + let cl' = mkLetIn (Name id, c, cty, cl) in + let gl = pf_merge_uc ucst gl in + tclTHEN (Proofview.V82.of_tactic (convert_concl cl')) (introid id) gl + +TACTIC EXTEND ssrset +| [ "set" ssrfwdid(id) ssrsetfwd(fwd) ssrclauses(clauses) ] -> + [ Proofview.V82.tactic (tclCLAUSES ist (ssrsettac ist id fwd) clauses) ] +END + +(** The "have" tactic *) + +(* type ssrhavefwd = ssrfwd * ssrhint *) + +let pr_ssrhavefwd _ _ prt (fwd, hint) = pr_fwd fwd ++ pr_hint prt hint + +ARGUMENT EXTEND ssrhavefwd TYPED AS ssrfwd * ssrhint PRINTED BY pr_ssrhavefwd +| [ ":" lconstr(t) ssrhint(hint) ] -> [ mkFwdHint ":" t, hint ] +| [ ":" lconstr(t) ":=" lconstr(c) ] -> [ mkFwdCast FwdHave loc t c, nohint ] +| [ ":" lconstr(t) ":=" ] -> [ mkFwdHintNoTC ":" t, nohint ] +| [ ":=" lconstr(c) ] -> [ mkFwdVal FwdHave c, nohint ] +END + +let intro_id_to_binder = List.map (function + | IpatId id -> + let xloc, _ as x = bvar_lname (mkCVar dummy_loc id) in + (FwdPose, [BFvar]), + CLambdaN (dummy_loc, [[x], Default Explicit, mkCHole xloc], + mkCHole dummy_loc) + | _ -> anomaly "non-id accepted as binder") + +let binder_to_intro_id = List.map (function + | (FwdPose, [BFvar]), CLambdaN (_,[ids,_,_],_) + | (FwdPose, [BFdecl _]), CLambdaN (_,[ids,_,_],_) -> + List.map (function (_, Name id) -> IpatId id | _ -> IpatAnon) ids + | (FwdPose, [BFdef _]), CLetIn (_,(_,Name id),_,_) -> [IpatId id] + | (FwdPose, [BFdef _]), CLetIn (_,(_,Anonymous),_,_) -> [IpatAnon] + | _ -> anomaly "ssrbinder is not a binder") + +let pr_ssrhavefwdwbinders _ _ prt (tr,((hpats, (fwd, hint)))) = + pr_hpats hpats ++ pr_fwd fwd ++ pr_hint prt hint + +ARGUMENT EXTEND ssrhavefwdwbinders + TYPED AS bool * (ssrhpats * (ssrfwd * ssrhint)) + PRINTED BY pr_ssrhavefwdwbinders +| [ ssrhpats_wtransp(trpats) ssrbinder_list(bs) ssrhavefwd(fwd) ] -> + [ let tr, pats = trpats in + let ((clr, pats), binders), simpl = pats in + let allbs = intro_id_to_binder binders @ bs in + let allbinders = binders @ List.flatten (binder_to_intro_id bs) in + let hint = bind_fwd allbs (fst fwd), snd fwd in + tr, ((((clr, pats), allbinders), simpl), hint) ] +END + +(* Tactic. *) + +let is_Evar_or_CastedMeta x = + isEvar_or_Meta x || + (isCast x && isEvar_or_Meta (pi1 (destCast x))) + +let occur_existential_or_casted_meta c = + let rec occrec c = match kind_of_term c with + | Evar _ -> raise Not_found + | Cast (m,_,_) when isMeta m -> raise Not_found + | _ -> iter_constr occrec c + in try occrec c; false with Not_found -> true + +let examine_abstract id gl = + let tid = pf_type_of gl id in + let abstract, gl = pf_mkSsrConst "abstract" gl in + if not (isApp tid) || not (Term.eq_constr (fst(destApp tid)) abstract) then + errorstrm(strbrk"not an abstract constant: "++pr_constr id); + let _, args_id = destApp tid in + if Array.length args_id <> 3 then + errorstrm(strbrk"not a proper abstract constant: "++pr_constr id); + if not (is_Evar_or_CastedMeta args_id.(2)) then + errorstrm(strbrk"abstract constant "++pr_constr id++str" already used"); + tid, args_id + +let pf_find_abstract_proof check_lock gl abstract_n = + let fire gl t = Reductionops.nf_evar (project gl) t in + let abstract, gl = pf_mkSsrConst "abstract" gl in + let l = Evd.fold_undefined (fun e ei l -> + match kind_of_term ei.Evd.evar_concl with + | App(hd, [|ty; n; lock|]) + when (not check_lock || + (occur_existential_or_casted_meta (fire gl ty) && + is_Evar_or_CastedMeta (fire gl lock))) && + Term.eq_constr hd abstract && Term.eq_constr n abstract_n -> e::l + | _ -> l) (project gl) [] in + match l with + | [e] -> e + | _ -> errorstrm(strbrk"abstract constant "++pr_constr abstract_n++ + strbrk" not found in the evar map exactly once. "++ + strbrk"Did you tamper with it?") + +let unfold cl = + let module R = Reductionops in let module F = Closure.RedFlags in + reduct_in_concl (R.clos_norm_flags (F.mkflags + (List.map (fun c -> F.fCONST (fst (destConst c))) cl @ [F.fBETA; F.fIOTA]))) + +let havegentac ist t gl = + let sigma, c, ucst = pf_abs_ssrterm ist gl t in + let gl = pf_merge_uc ucst gl in + apply_type (mkArrow (pf_type_of gl c) (pf_concl gl)) [c] gl + +let havetac ist + (transp,((((clr, pats), binders), simpl), (((fk, _), t), hint))) + suff namefst gl += + let concl = pf_concl gl in + let skols, pats = + List.partition (function IpatNewHidden _ -> true | _ -> false) pats in + let itac_mkabs = introstac ~ist skols in + let itac_c = introstac ~ist (IpatSimpl(clr,Nop) :: pats) in + let itac, id, clr = introstac ~ist pats, tclIDTAC, cleartac clr in + let binderstac n = + let rec aux = function 0 -> [] | n -> IpatAnon :: aux (n-1) in + tclTHEN (if binders <> [] then introstac ~ist (aux n) else tclIDTAC) + (introstac ~ist binders) in + let simpltac = introstac ~ist simpl in + let fixtc = + not !ssrhaveNOtcresolution && + match fk with FwdHint(_,true) -> false | _ -> true in + let hint = hinttac ist true hint in + let cuttac t gl = + if transp then + let have_let, gl = pf_mkSsrConst "ssr_have_let" gl in + let step = mkApp (have_let, [|concl;t|]) in + let gl, _ = pf_e_type_of gl step in + applyn ~with_evars:true ~with_shelve:false 2 step gl + else basecuttac "ssr_have" t gl in + (* Introduce now abstract constants, so that everything sees them *) + let abstract_key, gl = pf_mkSsrConst "abstract_key" gl in + let unlock_abs (idty,args_id) gl = + let gl, _ = pf_e_type_of gl idty in + pf_unify_HO gl args_id.(2) abstract_key in + tclTHENFIRST itac_mkabs (fun gl -> + let mkt t = mk_term ' ' t in + let mkl t = (' ', (t, None)) in + let interp gl rtc t = pf_abs_ssrterm ~resolve_typeclasses:rtc ist gl t in + let interp_ty gl rtc t = + let a,b,_,u = pf_interp_ty ~resolve_typeclasses:rtc ist gl t in a,b,u in + let ct, cty, hole, loc = match t with + | _, (_, Some (CCast (loc, ct, CastConv cty))) -> + mkt ct, mkt cty, mkt (mkCHole dummy_loc), loc + | _, (_, Some ct) -> + mkt ct, mkt (mkCHole dummy_loc), mkt (mkCHole dummy_loc), dummy_loc + | _, (GCast (loc, ct, CastConv cty), None) -> + mkl ct, mkl cty, mkl mkRHole, loc + | _, (t, None) -> mkl t, mkl mkRHole, mkl mkRHole, dummy_loc in + let gl, cut, sol, itac1, itac2 = + match fk, namefst, suff with + | FwdHave, true, true -> + errorstrm (str"Suff have does not accept a proof term") + | FwdHave, false, true -> + let cty = combineCG cty hole (mkCArrow loc) mkRArrow in + let _, t, uc = interp gl false (combineCG ct cty (mkCCast loc) mkRCast) in + let gl = pf_merge_uc uc gl in + let ty = pf_type_of gl t in + let ctx, _ = decompose_prod_n 1 ty in + let assert_is_conv gl = + try Proofview.V82.of_tactic (convert_concl (compose_prod ctx concl)) gl + with _ -> errorstrm (str "Given proof term is not of type " ++ + pr_constr (mkArrow (mkVar (id_of_string "_")) concl)) in + gl, ty, tclTHEN assert_is_conv (Proofview.V82.of_tactic (apply t)), id, itac_c + | FwdHave, false, false -> + let skols = List.flatten (List.map (function + | IpatNewHidden ids -> ids + | _ -> assert false) skols) in + let skols_args = + List.map (fun id -> examine_abstract (mkVar id) gl) skols in + let gl = List.fold_right unlock_abs skols_args gl in + let sigma, t, uc = + interp gl false (combineCG ct cty (mkCCast loc) mkRCast) in + let gl = re_sig (sig_it gl) (Evd.merge_universe_context sigma uc) in + let gs = + List.map (fun (_,a) -> + pf_find_abstract_proof false gl a.(1)) skols_args in + let tacopen_skols gl = + let stuff, g = Refiner.unpackage gl in + Refiner.repackage stuff (gs @ [g]) in + let gl, ty = pf_e_type_of gl t in + gl, ty, Proofview.V82.of_tactic (apply t), id, + tclTHEN (tclTHEN itac_c simpltac) + (tclTHEN tacopen_skols (fun gl -> + let abstract, gl = pf_mkSsrConst "abstract" gl in + unfold [abstract; abstract_key] gl)) + | _,true,true -> + let _, ty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in + gl, mkArrow ty concl, hint, itac, clr + | _,false,true -> + let _, ty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in + gl, mkArrow ty concl, hint, id, itac_c + | _, false, false -> + let n, cty, uc = interp_ty gl fixtc cty in let gl = pf_merge_uc uc gl in + gl, cty, tclTHEN (binderstac n) hint, id, tclTHEN itac_c simpltac + | _, true, false -> assert false in + tclTHENS (cuttac cut) [ tclTHEN sol itac1; itac2 ] gl) + gl +;; + +(* to extend the abstract value one needs: + Utility lemma to partially instantiate an abstract constant type. + Lemma use_abstract T n l (x : abstract T n l) : T. + Proof. by case: l x. Qed. +*) +let ssrabstract ist gens (*last*) gl = + let main _ (_,cid) ist gl = +(* + let proj1, proj2, prod = + let pdata = build_prod () in + pdata.Coqlib.proj1, pdata.Coqlib.proj2, pdata.Coqlib.typ in +*) + let concl, env = pf_concl gl, pf_env gl in + let fire gl t = Reductionops.nf_evar (project gl) t in + let abstract, gl = pf_mkSsrConst "abstract" gl in + let abstract_key, gl = pf_mkSsrConst "abstract_key" gl in + let id = mkVar (Option.get (id_of_cpattern cid)) in + let idty, args_id = examine_abstract id gl in + let abstract_n = args_id.(1) in + let abstract_proof = pf_find_abstract_proof true gl abstract_n in + let gl, proof = + let pf_unify_HO gl a b = + try pf_unify_HO gl a b + with _ -> errorstrm(strbrk"The abstract variable "++pr_constr id++ + strbrk" cannot abstract this goal. Did you generalize it?") in + let rec find_hole p t = + match kind_of_term t with + | Evar _ (*when last*) -> pf_unify_HO gl concl t, p + | Meta _ (*when last*) -> pf_unify_HO gl concl t, p + | Cast(m,_,_) when isEvar_or_Meta m (*when last*) -> pf_unify_HO gl concl t, p +(* + | Evar _ -> + let sigma, it = project gl, sig_it gl in + let sigma, ty = Evarutil.new_type_evar sigma env in + let gl = re_sig it sigma in + let p = mkApp (proj2,[|ty;concl;p|]) in + let concl = mkApp(prod,[|ty; concl|]) in + pf_unify_HO gl concl t, p + | App(hd, [|left; right|]) when Term.eq_constr hd prod -> + find_hole (mkApp (proj1,[|left;right;p|])) left +*) + | _ -> errorstrm(strbrk"abstract constant "++pr_constr abstract_n++ + strbrk" has an unexpected shape. Did you tamper with it?") + in + find_hole + ((*if last then*) id + (*else mkApp(mkSsrConst "use_abstract",Array.append args_id [|id|])*)) + (fire gl args_id.(0)) in + let gl = (*if last then*) pf_unify_HO gl abstract_key args_id.(2) (*else gl*) in + let gl, _ = pf_e_type_of gl idty in + let proof = fire gl proof in +(* if last then *) + let tacopen gl = + let stuff, g = Refiner.unpackage gl in + Refiner.repackage stuff [ g; abstract_proof ] in + tclTHENS tacopen [tclSOLVE [Proofview.V82.of_tactic (apply proof)];unfold[abstract;abstract_key]] gl +(* else apply proof gl *) + in + let introback ist (gens, _) = + introstac ~ist + (List.map (fun (_,cp) -> match id_of_cpattern cp with + | None -> IpatAnon + | Some id -> IpatId id) + (List.tl (List.hd gens))) in + tclTHEN (with_dgens gens main ist) (introback ist gens) gl + +(* The standard TACTIC EXTEND does not work for abstract *) +GEXTEND Gram + GLOBAL: tactic_expr; + tactic_expr: LEVEL "3" + [ RIGHTA [ IDENT "abstract"; gens = ssrdgens -> + Tacexpr.TacML (!@loc, ssrtac_name "abstract", + [Genarg.in_gen (Genarg.rawwit wit_ssrdgens) gens]) ]]; +END +TACTIC EXTEND ssrabstract +| [ "abstract" ssrdgens(gens) ] -> [ + if List.length (fst gens) <> 1 then + errorstrm (str"dependents switches '/' not allowed here"); + Proofview.V82.tactic (ssrabstract ist gens) ] +END + +let prof_havetac = mk_profiler "havetac";; +let havetac arg a b gl = prof_havetac.profile (havetac arg a b) gl;; + +TACTIC EXTEND ssrhave +| [ "have" ssrhavefwdwbinders(fwd) ] -> + [ Proofview.V82.tactic (havetac ist fwd false false) ] +END + +TACTIC EXTEND ssrhavesuff +| [ "have" "suff" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> + [ Proofview.V82.tactic (havetac ist (false,(pats,fwd)) true false) ] +END + +TACTIC EXTEND ssrhavesuffices +| [ "have" "suffices" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> + [ Proofview.V82.tactic (havetac ist (false,(pats,fwd)) true false) ] +END + +TACTIC EXTEND ssrsuffhave +| [ "suff" "have" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> + [ Proofview.V82.tactic (havetac ist (false,(pats,fwd)) true true) ] +END + +TACTIC EXTEND ssrsufficeshave +| [ "suffices" "have" ssrhpats_nobs(pats) ssrhavefwd(fwd) ] -> + [ Proofview.V82.tactic (havetac ist (false,(pats,fwd)) true true) ] +END + +(** The "suffice" tactic *) + +let pr_ssrsufffwdwbinders _ _ prt (hpats, (fwd, hint)) = + pr_hpats hpats ++ pr_fwd fwd ++ pr_hint prt hint + +ARGUMENT EXTEND ssrsufffwd + TYPED AS ssrhpats * (ssrfwd * ssrhint) PRINTED BY pr_ssrsufffwdwbinders +| [ ssrhpats(pats) ssrbinder_list(bs) ":" lconstr(t) ssrhint(hint) ] -> + [ let ((clr, pats), binders), simpl = pats in + let allbs = intro_id_to_binder binders @ bs in + let allbinders = binders @ List.flatten (binder_to_intro_id bs) in + let fwd = mkFwdHint ":" t in + (((clr, pats), allbinders), simpl), (bind_fwd allbs fwd, hint) ] +END + +let sufftac ist ((((clr, pats),binders),simpl), ((_, c), hint)) = + let htac = tclTHEN (introstac ~ist pats) (hinttac ist true hint) in + let c = match c with + | (a, (b, Some (CCast (_, _, CastConv cty)))) -> a, (b, Some cty) + | (a, (GCast (_, _, CastConv cty), None)) -> a, (cty, None) + | _ -> anomaly "suff: ssr cast hole deleted by typecheck" in + let ctac gl = + let _,ty,_,uc = pf_interp_ty ist gl c in let gl = pf_merge_uc uc gl in + basecuttac "ssr_suff" ty gl in + tclTHENS ctac [htac; tclTHEN (cleartac clr) (introstac ~ist (binders@simpl))] + +TACTIC EXTEND ssrsuff +| [ "suff" ssrsufffwd(fwd) ] -> [ Proofview.V82.tactic (sufftac ist fwd) ] +END + +TACTIC EXTEND ssrsuffices +| [ "suffices" ssrsufffwd(fwd) ] -> [ Proofview.V82.tactic (sufftac ist fwd) ] +END + +(** The "wlog" (Without Loss Of Generality) tactic *) + +(* type ssrwlogfwd = ssrwgen list * ssrfwd *) + +let pr_ssrwlogfwd _ _ _ (gens, t) = + str ":" ++ pr_list mt pr_wgen gens ++ spc() ++ pr_fwd t + +ARGUMENT EXTEND ssrwlogfwd TYPED AS ssrwgen list * ssrfwd + PRINTED BY pr_ssrwlogfwd +| [ ":" ssrwgen_list(gens) "/" lconstr(t) ] -> [ gens, mkFwdHint "/" t] +END + +let destProd_or_LetIn c = + match kind_of_term c with + | Prod (n,ty,c) -> (n,None,ty), c + | LetIn (n,bo,ty,c) -> (n,Some bo,ty), c + | _ -> raise DestKO + +let wlogtac ist (((clr0, pats),_),_) (gens, ((_, ct))) hint suff ghave gl = + let mkabs gen = abs_wgen false ist (fun x -> x) gen in + let mkclr gen clrs = clr_of_wgen gen clrs in + let mkpats = function + | _, Some ((x, _), _) -> fun pats -> IpatId (hoi_id x) :: pats + | _ -> fun x -> x in + let ct = match ct with + | (a, (b, Some (CCast (_, _, CastConv cty)))) -> a, (b, Some cty) + | (a, (GCast (_, _, CastConv cty), None)) -> a, (cty, None) + | _ -> anomaly "wlog: ssr cast hole deleted by typecheck" in + let cut_implies_goal = not (suff || ghave <> `NoGen) in + let c, args, ct, gl = + let gens = List.filter (function _, Some _ -> true | _ -> false) gens in + let concl = pf_concl gl in + let c = mkProp in + let c = if cut_implies_goal then mkArrow c concl else c in + let gl, args, c = List.fold_right mkabs gens (gl,[],c) in + let env, _ = + List.fold_left (fun (env, c) _ -> + let rd, c = destProd_or_LetIn c in + Environ.push_rel rd env, c) (pf_env gl, c) gens in + let sigma, ev = Evarutil.new_evar env (project gl) Term.mkProp in + let k, _ = Term.destEvar ev in + let fake_gl = {Evd.it = k; Evd.sigma = sigma} in + let _, ct, _, uc = pf_interp_ty ist fake_gl ct in + let rec var2rel c g s = match kind_of_term c, g with + | Prod(Anonymous,_,c), [] -> mkProd(Anonymous, Vars.subst_vars s ct, c) + | Sort _, [] -> Vars.subst_vars s ct + | LetIn(Name id as n,b,ty,c), _::g -> mkLetIn (n,b,ty,var2rel c g (id::s)) + | Prod(Name id as n,ty,c), _::g -> mkProd (n,ty,var2rel c g (id::s)) + | _ -> Errors.anomaly(str"SSR: wlog: var2rel: " ++ pr_constr c) in + let c = var2rel c gens [] in + let rec pired c = function + | [] -> c + | t::ts as args -> match kind_of_term c with + | Prod(_,_,c) -> pired (subst1 t c) ts + | LetIn(id,b,ty,c) -> mkLetIn (id,b,ty,pired c args) + | _ -> Errors.anomaly(str"SSR: wlog: pired: " ++ pr_constr c) in + c, args, pired c args, pf_merge_uc uc gl in + let tacipat pats = introstac ~ist pats in + let tacigens = + tclTHEN + (tclTHENLIST(List.rev(List.fold_right mkclr gens [cleartac clr0]))) + (introstac ~ist (List.fold_right mkpats gens [])) in + let hinttac = hinttac ist true hint in + let cut_kind, fst_goal_tac, snd_goal_tac = + match suff, ghave with + | true, `NoGen -> "ssr_wlog", tclTHEN hinttac (tacipat pats), tacigens + | false, `NoGen -> "ssr_wlog", hinttac, tclTHEN tacigens (tacipat pats) + | true, `Gen _ -> assert false + | false, `Gen id -> + if gens = [] then errorstrm(str"gen have requires some generalizations"); + let clear0 = cleartac clr0 in + let id, name_general_hyp, cleanup, pats = match id, pats with + | None, (IpatId id as ip)::pats -> Some id, tacipat [ip], clear0, pats + | None, _ -> None, tclIDTAC, clear0, pats + | Some (Some id),_ -> Some id, introid id, clear0, pats + | Some _,_ -> + let id = mk_anon_id "tmp" gl in + Some id, introid id, tclTHEN clear0 (clear [id]), pats in + let tac_specialize = match id with + | None -> tclIDTAC + | Some id -> + if pats = [] then tclIDTAC else + let args = Array.of_list args in + pp(lazy(str"specialized="++pr_constr (mkApp (mkVar id,args)))); + pp(lazy(str"specialized_ty="++pr_constr ct)); + tclTHENS (basecuttac "ssr_have" ct) + [Proofview.V82.of_tactic (apply (mkApp (mkVar id,args))); tclIDTAC] in + "ssr_have", + (if hint = nohint then tacigens else hinttac), + tclTHENLIST [name_general_hyp; tac_specialize; tacipat pats; cleanup] + in + tclTHENS (basecuttac cut_kind c) [fst_goal_tac; snd_goal_tac] gl + +TACTIC EXTEND ssrwlog +| [ "wlog" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> + [ Proofview.V82.tactic (wlogtac ist pats fwd hint false `NoGen) ] +END + +TACTIC EXTEND ssrwlogs +| [ "wlog" "suff" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> + [ Proofview.V82.tactic (wlogtac ist pats fwd hint true `NoGen) ] +END + +TACTIC EXTEND ssrwlogss +| [ "wlog" "suffices" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]-> + [ Proofview.V82.tactic (wlogtac ist pats fwd hint true `NoGen) ] +END + +TACTIC EXTEND ssrwithoutloss +| [ "without" "loss" ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> + [ Proofview.V82.tactic (wlogtac ist pats fwd hint false `NoGen) ] +END + +TACTIC EXTEND ssrwithoutlosss +| [ "without" "loss" "suff" + ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> + [ Proofview.V82.tactic (wlogtac ist pats fwd hint true `NoGen) ] +END + +TACTIC EXTEND ssrwithoutlossss +| [ "without" "loss" "suffices" + ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ]-> + [ Proofview.V82.tactic (wlogtac ist pats fwd hint true `NoGen) ] +END + +(* Generally have *) +let pr_idcomma _ _ _ = function + | None -> mt() + | Some None -> str"_, " + | Some (Some id) -> pr_id id ++ str", " + +ARGUMENT EXTEND ssr_idcomma TYPED AS ident option option PRINTED BY pr_idcomma + | [ ] -> [ None ] +END + +let accept_idcomma strm = + match Compat.get_tok (stream_nth 0 strm) with + | Tok.IDENT _ | Tok.KEYWORD "_" -> accept_before_syms [","] strm + | _ -> raise Stream.Failure + +let test_idcomma = Gram.Entry.of_parser "test_idcomma" accept_idcomma + +GEXTEND Gram + GLOBAL: ssr_idcomma; + ssr_idcomma: [ [ test_idcomma; + ip = [ id = IDENT -> Some (id_of_string id) | "_" -> None ]; "," -> + Some ip + ] ]; +END + +let augment_preclr clr1 (((clr0, x),y),z) = (((clr1 @ clr0, x),y),z) + +TACTIC EXTEND ssrgenhave +| [ "gen" "have" ssrclear(clr) + ssr_idcomma(id) ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> + [ let pats = augment_preclr clr pats in + Proofview.V82.tactic (wlogtac ist pats fwd hint false (`Gen id)) ] +END + +TACTIC EXTEND ssrgenhave2 +| [ "generally" "have" ssrclear(clr) + ssr_idcomma(id) ssrhpats_nobs(pats) ssrwlogfwd(fwd) ssrhint(hint) ] -> + [ let pats = augment_preclr clr pats in + Proofview.V82.tactic (wlogtac ist pats fwd hint false (`Gen id)) ] +END + +(** Canonical Structure alias *) + +let def_body : Vernacexpr.definition_expr Gram.Entry.e = Obj.magic + (Grammar.Entry.find (Obj.magic gallina_ext) "vernac:def_body") in + +GEXTEND Gram + GLOBAL: gallina_ext; + + gallina_ext: + (* Canonical structure *) + [[ IDENT "Canonical"; qid = Constr.global -> + Vernacexpr.VernacCanonical (AN qid) + | IDENT "Canonical"; ntn = Prim.by_notation -> + Vernacexpr.VernacCanonical (ByNotation ntn) + | IDENT "Canonical"; qid = Constr.global; + d = def_body -> + let s = coerce_reference_to_id qid in + Vernacexpr.VernacDefinition + ((Some Decl_kinds.Global,Decl_kinds.CanonicalStructure), + (dummy_loc,s),(d )) + ]]; +END + +(** 9. Keyword compatibility fixes. *) + +(* Coq v8.1 notation uses "by" and "of" quasi-keywords, i.e., reserved *) +(* identifiers used as keywords. This is incompatible with ssreflect.v *) +(* which makes "by" and "of" true keywords, because of technicalities *) +(* in the internal lexer-parser API of Coq. We patch this here by *) +(* adding new parsing rules that recognize the new keywords. *) +(* To make matters worse, the Coq grammar for tactics fails to *) +(* export the non-terminals we need to patch. Fortunately, the CamlP5 *) +(* API provides a backdoor access (with loads of Obj.magic trickery). *) + +(* Coq v8.3 defines "by" as a keyword, some hacks are not needed any *) +(* longer and thus comment out. Such comments are marked with v8.3 *) + +let tac_ent = List.fold_left Grammar.Entry.find (Obj.magic simple_tactic) in +let hypident_ent = + tac_ent ["clause_dft_all"; "in_clause"; "hypident_occ"; "hypident"] in +let id_or_meta : Obj.t Gram.Entry.e = Obj.magic + (Grammar.Entry.find hypident_ent "id_or_meta") in +let hypident : (Obj.t * hyp_location_flag) Gram.Entry.e = + Obj.magic hypident_ent in +GEXTEND Gram + GLOBAL: hypident; +hypident: [ + [ "("; IDENT "type"; "of"; id = id_or_meta; ")" -> id, InHypTypeOnly + | "("; IDENT "value"; "of"; id = id_or_meta; ")" -> id, InHypValueOnly + ] ]; +END + +GEXTEND Gram + GLOBAL: hloc; +hloc: [ + [ "in"; "("; "Type"; "of"; id = ident; ")" -> + HypLocation ((dummy_loc,id), InHypTypeOnly) + | "in"; "("; IDENT "Value"; "of"; id = ident; ")" -> + HypLocation ((dummy_loc,id), InHypValueOnly) + ] ]; +END + +let constr_eval + : (Constrexpr.constr_expr,Obj.t,Obj.t) Genredexpr.may_eval Gram.Entry.e + = Obj.magic (Grammar.Entry.find (Obj.magic constr_may_eval) "constr_eval") + +GEXTEND Gram + GLOBAL: constr_eval; + constr_eval: [ + [ IDENT "type"; "of"; c = Constr.constr -> Genredexpr.ConstrTypeOf c ] + ]; +END + +(* We wipe out all the keywords generated by the grammar rules we defined. *) +(* The user is supposed to Require Import ssreflect or Require ssreflect *) +(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) +(* consequence the extended ssreflect grammar. *) +let () = Lexer.unfreeze frozen_lexer ;; + +(* vim: set filetype=ocaml foldmethod=marker: *) diff --git a/mathcomp/ssreflect/plugin/v8.5beta2/ssreflect.mllib b/mathcomp/ssreflect/plugin/v8.5beta2/ssreflect.mllib new file mode 100644 index 0000000..006b70f --- /dev/null +++ b/mathcomp/ssreflect/plugin/v8.5beta2/ssreflect.mllib @@ -0,0 +1,2 @@ +Ssrmatching +Ssreflect diff --git a/mathcomp/ssreflect/plugin/v8.5beta2/ssrmatching.ml4 b/mathcomp/ssreflect/plugin/v8.5beta2/ssrmatching.ml4 new file mode 100644 index 0000000..2fd0fe6 --- /dev/null +++ b/mathcomp/ssreflect/plugin/v8.5beta2/ssrmatching.ml4 @@ -0,0 +1,1290 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) + +(* Defining grammar rules with "xx" in it automatically declares keywords too, + * we thus save the lexer to restore it at the end of the file *) +let frozen_lexer = Lexer.freeze () ;; + +(*i camlp4use: "pa_extend.cmo" i*) +(*i camlp4deps: "grammar/grammar.cma" i*) + +open Names +open Pp +open Pcoq +open Genarg +open Constrarg +open Term +open Vars +open Topconstr +open Libnames +open Tactics +open Tacticals +open Termops +open Namegen +open Recordops +open Tacmach +open Coqlib +open Glob_term +open Util +open Evd +open Extend +open Goptions +open Tacexpr +open Tacinterp +open Pretyping +open Constr +open Tactic +open Extraargs +open Ppconstr +open Printer + +open Globnames +open Misctypes +open Decl_kinds +open Evar_kinds +open Constrexpr +open Constrexpr_ops +open Notation_term +open Notation_ops +open Locus +open Locusops + +DECLARE PLUGIN "ssreflect" + +type loc = Loc.t +let dummy_loc = Loc.ghost +let errorstrm = Errors.errorlabstrm "ssreflect" +let loc_error loc msg = Errors.user_err_loc (loc, msg, str msg) + +(* 0 cost pp function. Active only if env variable SSRDEBUG is set *) +(* or if SsrDebug is Set *) +let pp_ref = ref (fun _ -> ()) +let ssr_pp s = pperrnl (str"SSR: "++Lazy.force s) +let _ = try ignore(Sys.getenv "SSRDEBUG"); pp_ref := ssr_pp with Not_found -> () +let debug b = + if b then pp_ref := ssr_pp else pp_ref := fun _ -> () +let _ = + Goptions.declare_bool_option + { Goptions.optsync = false; + Goptions.optname = "ssrmatching debugging"; + Goptions.optkey = ["SsrMatchingDebug"]; + Goptions.optdepr = false; + Goptions.optread = (fun _ -> !pp_ref == ssr_pp); + Goptions.optwrite = debug } +let pp s = !pp_ref s + +(** Utils {{{ *****************************************************************) +let env_size env = List.length (Environ.named_context env) +let safeDestApp c = + match kind_of_term c with App (f, a) -> f, a | _ -> c, [| |] +let get_index = function ArgArg i -> i | _ -> + Errors.anomaly (str"Uninterpreted index") +(* Toplevel constr must be globalized twice ! *) +let glob_constr ist genv = function + | _, Some ce -> + let vars = Id.Map.fold (fun x _ accu -> Id.Set.add x accu) ist.lfun Id.Set.empty in + let ltacvars = { Constrintern.empty_ltac_sign with Constrintern.ltac_vars = vars } in + Constrintern.intern_gen WithoutTypeConstraint ~ltacvars:ltacvars genv ce + | rc, None -> rc + +(* Term printing utilities functions for deciding bracketing. *) +let pr_paren prx x = hov 1 (str "(" ++ prx x ++ str ")") +(* String lexing utilities *) +let skip_wschars s = + let rec loop i = match s.[i] with '\n'..' ' -> loop (i + 1) | _ -> i in loop +(* We also guard characters that might interfere with the ssreflect *) +(* tactic syntax. *) +let guard_term ch1 s i = match s.[i] with + | '(' -> false + | '{' | '/' | '=' -> true + | _ -> ch1 = '(' +(* The call 'guard s i' should return true if the contents of s *) +(* starting at i need bracketing to avoid ambiguities. *) +let pr_guarded guard prc c = + msg_with Format.str_formatter (prc c); + let s = Format.flush_str_formatter () ^ "$" in + if guard s (skip_wschars s 0) then pr_paren prc c else prc c +(* More sensible names for constr printers *) +let pr_constr = pr_constr +let prl_glob_constr c = pr_lglob_constr_env (Global.env ()) c +let pr_glob_constr c = pr_glob_constr_env (Global.env ()) c +let prl_constr_expr = pr_lconstr_expr +let pr_constr_expr = pr_constr_expr +let prl_glob_constr_and_expr = function + | _, Some c -> prl_constr_expr c + | c, None -> prl_glob_constr c +let pr_glob_constr_and_expr = function + | _, Some c -> pr_constr_expr c + | c, None -> pr_glob_constr c +let pr_term (k, c) = pr_guarded (guard_term k) pr_glob_constr_and_expr c +let prl_term (k, c) = pr_guarded (guard_term k) prl_glob_constr_and_expr c + +(** Adding a new uninterpreted generic argument type *) +let add_genarg tag pr = + let wit = Genarg.make0 None tag in + let glob ist x = (ist, x) in + let subst _ x = x in + let interp ist gl x = (gl.Evd.sigma, x) in + let gen_pr _ _ _ = pr in + let () = Genintern.register_intern0 wit glob in + let () = Genintern.register_subst0 wit subst in + let () = Geninterp.register_interp0 wit interp in + Pptactic.declare_extra_genarg_pprule wit gen_pr gen_pr gen_pr; + wit + +(** Constructors for cast type *) +let dC t = CastConv t +(** Constructors for constr_expr *) +let isCVar = function CRef (Ident _, _) -> true | _ -> false +let destCVar = function CRef (Ident (_, id), _) -> id | _ -> + Errors.anomaly (str"not a CRef") +let mkCHole loc = CHole (loc, None, IntroAnonymous, None) +let mkCLambda loc name ty t = + CLambdaN (loc, [[loc, name], Default Explicit, ty], t) +let mkCLetIn loc name bo t = + CLetIn (loc, (loc, name), bo, t) +let mkCCast loc t ty = CCast (loc,t, dC ty) +(** Constructors for rawconstr *) +let mkRHole = GHole (dummy_loc, InternalHole, IntroAnonymous, None) +let mkRApp f args = if args = [] then f else GApp (dummy_loc, f, args) +let mkRCast rc rt = GCast (dummy_loc, rc, dC rt) +let mkRLambda n s t = GLambda (dummy_loc, n, Explicit, s, t) + +(* ssrterm conbinators *) +let combineCG t1 t2 f g = match t1, t2 with + | (x, (t1, None)), (_, (t2, None)) -> x, (g t1 t2, None) + | (x, (_, Some t1)), (_, (_, Some t2)) -> x, (mkRHole, Some (f t1 t2)) + | _, (_, (_, None)) -> Errors.anomaly (str"have: mixed C-G constr") + | _ -> Errors.anomaly (str"have: mixed G-C constr") +let loc_ofCG = function + | (_, (s, None)) -> Glob_ops.loc_of_glob_constr s + | (_, (_, Some s)) -> Constrexpr_ops.constr_loc s + +let mk_term k c = k, (mkRHole, Some c) +let mk_lterm = mk_term ' ' + +(* }}} *) + +(** Profiling {{{ *************************************************************) +type profiler = { + profile : 'a 'b. ('a -> 'b) -> 'a -> 'b; + reset : unit -> unit; + print : unit -> unit } +let profile_now = ref false +let something_profiled = ref false +let profilers = ref [] +let add_profiler f = profilers := f :: !profilers;; +let profile b = + profile_now := b; + if b then List.iter (fun f -> f.reset ()) !profilers; + if not b then List.iter (fun f -> f.print ()) !profilers +;; +let _ = + Goptions.declare_bool_option + { Goptions.optsync = false; + Goptions.optname = "ssrmatching profiling"; + Goptions.optkey = ["SsrMatchingProfiling"]; + Goptions.optread = (fun _ -> !profile_now); + Goptions.optdepr = false; + Goptions.optwrite = profile } +let () = + let prof_total = + let init = ref 0.0 in { + profile = (fun f x -> assert false); + reset = (fun () -> init := Unix.gettimeofday ()); + print = (fun () -> if !something_profiled then + prerr_endline + (Printf.sprintf "!! %-39s %10d %9.4f %9.4f %9.4f" + "total" 0 (Unix.gettimeofday() -. !init) 0.0 0.0)) } in + let prof_legenda = { + profile = (fun f x -> assert false); + reset = (fun () -> ()); + print = (fun () -> if !something_profiled then begin + prerr_endline + (Printf.sprintf "!! %39s ---------- --------- --------- ---------" + (String.make 39 '-')); + prerr_endline + (Printf.sprintf "!! %-39s %10s %9s %9s %9s" + "function" "#calls" "total" "max" "average") end) } in + add_profiler prof_legenda; + add_profiler prof_total +;; + +let mk_profiler s = + let total, calls, max = ref 0.0, ref 0, ref 0.0 in + let reset () = total := 0.0; calls := 0; max := 0.0 in + let profile f x = + if not !profile_now then f x else + let before = Unix.gettimeofday () in + try + incr calls; + let res = f x in + let after = Unix.gettimeofday () in + let delta = after -. before in + total := !total +. delta; + if delta > !max then max := delta; + res + with exc -> + let after = Unix.gettimeofday () in + let delta = after -. before in + total := !total +. delta; + if delta > !max then max := delta; + raise exc in + let print () = + if !calls <> 0 then begin + something_profiled := true; + prerr_endline + (Printf.sprintf "!! %-39s %10d %9.4f %9.4f %9.4f" + s !calls !total !max (!total /. (float_of_int !calls))) end in + let prof = { profile = profile; reset = reset; print = print } in + add_profiler prof; + prof +;; +(* }}} *) + +exception NoProgress + +(** Unification procedures. *) + +(* To enforce the rigidity of the rooted match we always split *) +(* top applications, so the unification procedures operate on *) +(* arrays of patterns and terms. *) +(* We perform three kinds of unification: *) +(* EQ: exact conversion check *) +(* FO: first-order unification of evars, without conversion *) +(* HO: higher-order unification with conversion *) +(* The subterm unification strategy is to find the first FO *) +(* match, if possible, and the first HO match otherwise, then *) +(* compute all the occurrences that are EQ matches for the *) +(* relevant subterm. *) +(* Additional twists: *) +(* - If FO/HO fails then we attempt to fill evars using *) +(* typeclasses before raising an outright error. We also *) +(* fill typeclasses even after a successful match, since *) +(* beta-reduction and canonical instances may leave *) +(* undefined evars. *) +(* - We do postchecks to rule out matches that are not *) +(* closed or that assign to a global evar; these can be *) +(* disabled for rewrite or dependent family matches. *) +(* - We do a full FO scan before turning to HO, as the FO *) +(* comparison can be much faster than the HO one. *) + +let unif_EQ env sigma p c = + let evars = existential_opt_value sigma in + try let _ = Reduction.conv env p ~evars c in true with _ -> false + +let unif_EQ_args env sigma pa a = + let n = Array.length pa in + let rec loop i = (i = n) || unif_EQ env sigma pa.(i) a.(i) && loop (i + 1) in + loop 0 + +let prof_unif_eq_args = mk_profiler "unif_EQ_args";; +let unif_EQ_args env sigma pa a = + prof_unif_eq_args.profile (unif_EQ_args env sigma pa) a +;; + +let unif_HO env ise p c = Evarconv.the_conv_x env p c ise + +let unif_HOtype env ise p c = Evarconv.the_conv_x_leq env p c ise + +let unif_HO_args env ise0 pa i ca = + let n = Array.length pa in + let rec loop ise j = + if j = n then ise else loop (unif_HO env ise pa.(j) ca.(i + j)) (j + 1) in + loop ise0 0 + +(* FO unification should boil down to calling w_unify with no_delta, but *) +(* alas things are not so simple: w_unify does partial type-checking, *) +(* which breaks down when the no-delta flag is on (as the Coq type system *) +(* requires full convertibility. The workaround here is to convert all *) +(* evars into metas, since 8.2 does not TC metas. This means some lossage *) +(* for HO evars, though hopefully Miller patterns can pick up some of *) +(* those cases, and HO matching will mop up the rest. *) +let flags_FO = + let flags = + { (Unification.default_no_delta_unify_flags ()).Unification.core_unify_flags + with + Unification.modulo_conv_on_closed_terms = None; + Unification.modulo_eta = true; + Unification.modulo_betaiota = true; + Unification.modulo_delta_types = full_transparent_state} + in + { Unification.core_unify_flags = flags; + Unification.merge_unify_flags = flags; + Unification.subterm_unify_flags = flags; + Unification.allow_K_in_toplevel_higher_order_unification = false; + Unification.resolve_evars = + (Unification.default_no_delta_unify_flags ()).Unification.resolve_evars + } +let unif_FO env ise p c = + Unification.w_unify env ise Reduction.CONV ~flags:flags_FO p c + +(* Perform evar substitution in main term and prune substitution. *) +let nf_open_term sigma0 ise c = + let s = ise and s' = ref sigma0 in + let rec nf c' = match kind_of_term c' with + | Evar ex -> + begin try nf (existential_value s ex) with _ -> + let k, a = ex in let a' = Array.map nf a in + if not (Evd.mem !s' k) then + s' := Evd.add !s' k (Evarutil.nf_evar_info s (Evd.find s k)); + mkEvar (k, a') + end + | _ -> map_constr nf c' in + let copy_def k evi () = + if evar_body evi != Evd.Evar_empty then () else + match Evd.evar_body (Evd.find s k) with + | Evar_defined c' -> s' := Evd.define k (nf c') !s' + | _ -> () in + let c' = nf c in let _ = Evd.fold copy_def sigma0 () in + !s', Evd.evar_universe_context s, c' + +let unif_end env sigma0 ise0 pt ok = + let ise = Evarconv.consider_remaining_unif_problems env ise0 in + let s, uc, t = nf_open_term sigma0 ise pt in + let ise1 = create_evar_defs s in + let ise2 = Typeclasses.resolve_typeclasses ~fail:true env ise1 in + if not (ok ise) then raise NoProgress else + if ise2 == ise1 then (s, uc, t) + else + let s, uc', t = nf_open_term sigma0 ise2 t in + s, Evd.union_evar_universe_context uc uc', t + +let pf_unif_HO gl sigma pt p c = + let env = pf_env gl in + let ise = unif_HO env (create_evar_defs sigma) p c in + unif_end env (project gl) ise pt (fun _ -> true) + +let unify_HO env sigma0 t1 t2 = + let sigma = unif_HO env sigma0 t1 t2 in + let sigma, uc, _ = unif_end env sigma0 sigma t2 (fun _ -> true) in + Evd.merge_universe_context sigma uc + +let pf_unify_HO gl t1 t2 = + let env, sigma0, si = pf_env gl, project gl, sig_it gl in + let sigma = unify_HO env sigma0 t1 t2 in + re_sig si sigma + +(* This is what the definition of iter_constr should be... *) +let iter_constr_LR f c = match kind_of_term c with + | Evar (k, a) -> Array.iter f a + | Cast (cc, _, t) -> f cc; f t + | Prod (_, t, b) | Lambda (_, t, b) -> f t; f b + | LetIn (_, v, t, b) -> f v; f t; f b + | App (cf, a) -> f cf; Array.iter f a + | Case (_, p, v, b) -> f v; f p; Array.iter f b + | Fix (_, (_, t, b)) | CoFix (_, (_, t, b)) -> + for i = 0 to Array.length t - 1 do f t.(i); f b.(i) done + | _ -> () + +(* The comparison used to determine which subterms matches is KEYED *) +(* CONVERSION. This looks for convertible terms that either have the same *) +(* same head constant as pat if pat is an application (after beta-iota), *) +(* or start with the same constr constructor (esp. for LetIn); this is *) +(* disregarded if the head term is let x := ... in x, and casts are always *) +(* ignored and removed). *) +(* Record projections get special treatment: in addition to the projection *) +(* constant itself, ssreflect also recognizes head constants of canonical *) +(* projections. *) + +exception NoMatch +type ssrdir = L2R | R2L +let pr_dir_side = function L2R -> str "LHS" | R2L -> str "RHS" +let inv_dir = function L2R -> R2L | R2L -> L2R + + +type pattern_class = + | KpatFixed + | KpatConst + | KpatEvar of existential_key + | KpatLet + | KpatLam + | KpatRigid + | KpatFlex + | KpatProj of constant + +type tpattern = { + up_k : pattern_class; + up_FO : constr; + up_f : constr; + up_a : constr array; + up_t : constr; (* equation proof term or matched term *) + up_dir : ssrdir; (* direction of the rule *) + up_ok : constr -> evar_map -> bool; (* progess test for rewrite *) + } + +let all_ok _ _ = true + +let proj_nparams c = + try 1 + Recordops.find_projection_nparams (ConstRef c) with _ -> 0 + +let isFixed c = match kind_of_term c with + | Var _ | Ind _ | Construct _ | Const _ -> true + | _ -> false + +let isRigid c = match kind_of_term c with + | Prod _ | Sort _ | Lambda _ | Case _ | Fix _ | CoFix _ -> true + | _ -> false + +exception UndefPat + +let hole_var = mkVar (id_of_string "_") +let pr_constr_pat c0 = + let rec wipe_evar c = + if isEvar c then hole_var else map_constr wipe_evar c in + pr_constr (wipe_evar c0) + +(* Turn (new) evars into metas *) +let evars_for_FO ~hack env sigma0 (ise0:evar_map) c0 = + let ise = ref ise0 in + let sigma = ref ise0 in + let nenv = env_size env + if hack then 1 else 0 in + let rec put c = match kind_of_term c with + | Evar (k, a as ex) -> + begin try put (existential_value !sigma ex) + with NotInstantiatedEvar -> + if Evd.mem sigma0 k then map_constr put c else + let evi = Evd.find !sigma k in + let dc = List.firstn (max 0 (Array.length a - nenv)) (evar_filtered_context evi) in + let abs_dc (d, c) = function + | x, Some b, t -> d, mkNamedLetIn x (put b) (put t) c + | x, None, t -> mkVar x :: d, mkNamedProd x (put t) c in + let a, t = + Context.fold_named_context_reverse abs_dc ~init:([], (put evi.evar_concl)) dc in + let m = Evarutil.new_meta () in + ise := meta_declare m t !ise; + sigma := Evd.define k (applist (mkMeta m, a)) !sigma; + put (existential_value !sigma ex) + end + | _ -> map_constr put c in + let c1 = put c0 in !ise, c1 + +(* Compile a match pattern from a term; t is the term to fill. *) +(* p_origin can be passed to obtain a better error message *) +let mk_tpattern ?p_origin ?(hack=false) env sigma0 (ise, t) ok dir p = + let k, f, a = + let f, a = Reductionops.whd_betaiota_stack ise p in + match kind_of_term f with + | Const (p,_) -> + let np = proj_nparams p in + if np = 0 || np > List.length a then KpatConst, f, a else + let a1, a2 = List.chop np a in KpatProj p, applist(f, a1), a2 + | Var _ | Ind _ | Construct _ -> KpatFixed, f, a + | Evar (k, _) -> + if Evd.mem sigma0 k then KpatEvar k, f, a else + if a <> [] then KpatFlex, f, a else + (match p_origin with None -> Errors.error "indeterminate pattern" + | Some (dir, rule) -> + errorstrm (str "indeterminate " ++ pr_dir_side dir + ++ str " in " ++ pr_constr_pat rule)) + | LetIn (_, v, _, b) -> + if b <> mkRel 1 then KpatLet, f, a else KpatFlex, v, a + | Lambda _ -> KpatLam, f, a + | _ -> KpatRigid, f, a in + let aa = Array.of_list a in + let ise', p' = evars_for_FO ~hack env sigma0 ise (mkApp (f, aa)) in + ise', + { up_k = k; up_FO = p'; up_f = f; + up_a = aa; up_ok = ok; up_dir = dir; up_t = t} + +(* Specialize a pattern after a successful match: assign a precise head *) +(* kind and arity for Proj and Flex patterns. *) +let ungen_upat lhs (sigma, uc, t) u = + let f, a = safeDestApp lhs in + let k = match kind_of_term f with + | Var _ | Ind _ | Construct _ -> KpatFixed + | Const _ -> KpatConst + | Evar (k, _) -> if is_defined sigma k then raise NoMatch else KpatEvar k + | LetIn _ -> KpatLet + | Lambda _ -> KpatLam + | _ -> KpatRigid in + sigma, uc, {u with up_k = k; up_FO = lhs; up_f = f; up_a = a; up_t = t} + +let nb_cs_proj_args pc f u = + let na k = + List.length (snd (lookup_canonical_conversion (ConstRef pc, k))).o_TCOMPS in + try match kind_of_term f with + | Prod _ -> na Prod_cs + | Sort s -> na (Sort_cs (family_of_sort s)) + | Const (c',_) when Constant.equal c' pc -> Array.length (snd (destApp u.up_f)) + | Var _ | Ind _ | Construct _ | Const _ -> na (Const_cs (global_of_constr f)) + | _ -> -1 + with Not_found -> -1 + +let isEvar_k k f = + match kind_of_term f with Evar (k', _) -> k = k' | _ -> false + +let nb_args c = + match kind_of_term c with App (_, a) -> Array.length a | _ -> 0 + +let mkSubArg i a = if i = Array.length a then a else Array.sub a 0 i +let mkSubApp f i a = if i = 0 then f else mkApp (f, mkSubArg i a) + +let splay_app ise = + let rec loop c a = match kind_of_term c with + | App (f, a') -> loop f (Array.append a' a) + | Cast (c', _, _) -> loop c' a + | Evar ex -> + (try loop (existential_value ise ex) a with _ -> c, a) + | _ -> c, a in + fun c -> match kind_of_term c with + | App (f, a) -> loop f a + | Cast _ | Evar _ -> loop c [| |] + | _ -> c, [| |] + +let filter_upat i0 f n u fpats = + let na = Array.length u.up_a in + if n < na then fpats else + let np = match u.up_k with + | KpatConst when Term.eq_constr u.up_f f -> na + | KpatFixed when Term.eq_constr u.up_f f -> na + | KpatEvar k when isEvar_k k f -> na + | KpatLet when isLetIn f -> na + | KpatLam when isLambda f -> na + | KpatRigid when isRigid f -> na + | KpatFlex -> na + | KpatProj pc -> + let np = na + nb_cs_proj_args pc f u in if n < np then -1 else np + | _ -> -1 in + if np < na then fpats else + let () = if !i0 < np then i0 := n in (u, np) :: fpats + +let filter_upat_FO i0 f n u fpats = + let np = nb_args u.up_FO in + if n < np then fpats else + let ok = match u.up_k with + | KpatConst -> Term.eq_constr u.up_f f + | KpatFixed -> Term.eq_constr u.up_f f + | KpatEvar k -> isEvar_k k f + | KpatLet -> isLetIn f + | KpatLam -> isLambda f + | KpatRigid -> isRigid f + | KpatProj pc -> Term.eq_constr f (mkConst pc) + | KpatFlex -> i0 := n; true in + if ok then begin if !i0 < np then i0 := np; (u, np) :: fpats end else fpats + +exception FoundUnif of (evar_map * evar_universe_context * tpattern) +(* Note: we don't update env as we descend into the term, as the primitive *) +(* unification procedure always rejects subterms with bound variables. *) + +(* We are forced to duplicate code between the FO/HO matching because we *) +(* have to work around several kludges in unify.ml: *) +(* - w_unify drops into second-order unification when the pattern is an *) +(* application whose head is a meta. *) +(* - w_unify tries to unify types without subsumption when the pattern *) +(* head is an evar or meta (e.g., it fails on ?1 = nat when ?1 : Type). *) +(* - w_unify expands let-in (zeta conversion) eagerly, whereas we want to *) +(* match a head let rigidly. *) +let match_upats_FO upats env sigma0 ise = + let rec loop c = + let f, a = splay_app ise c in let i0 = ref (-1) in + let fpats = + List.fold_right (filter_upat_FO i0 f (Array.length a)) upats [] in + while !i0 >= 0 do + let i = !i0 in i0 := -1; + let c' = mkSubApp f i a in + let one_match (u, np) = + let skip = + if i <= np then i < np else + if u.up_k == KpatFlex then begin i0 := i - 1; false end else + begin if !i0 < np then i0 := np; true end in + if skip || not (closed0 c') then () else try + let _ = match u.up_k with + | KpatFlex -> + let kludge v = mkLambda (Anonymous, mkProp, v) in + unif_FO env ise (kludge u.up_FO) (kludge c') + | KpatLet -> + let kludge vla = + let vl, a = safeDestApp vla in + let x, v, t, b = destLetIn vl in + mkApp (mkLambda (x, t, b), Array.cons v a) in + unif_FO env ise (kludge u.up_FO) (kludge c') + | _ -> unif_FO env ise u.up_FO c' in + let ise' = (* Unify again using HO to assign evars *) + let p = mkApp (u.up_f, u.up_a) in + try unif_HO env ise p c' with _ -> raise NoMatch in + let lhs = mkSubApp f i a in + let pt' = unif_end env sigma0 ise' u.up_t (u.up_ok lhs) in + raise (FoundUnif (ungen_upat lhs pt' u)) + with FoundUnif _ as sigma_u -> raise sigma_u + | Not_found -> Errors.anomaly (str"incomplete ise in match_upats_FO") + | _ -> () in + List.iter one_match fpats + done; + iter_constr_LR loop f; Array.iter loop a in + fun c -> try loop c with Invalid_argument _ -> Errors.anomaly (str"IN FO") + +let prof_FO = mk_profiler "match_upats_FO";; +let match_upats_FO upats env sigma0 ise c = + prof_FO.profile (match_upats_FO upats env sigma0) ise c +;; + + +let match_upats_HO ~on_instance upats env sigma0 ise c = + let it_did_match = ref false in + let rec aux upats env sigma0 ise c = + let f, a = splay_app ise c in let i0 = ref (-1) in + let fpats = List.fold_right (filter_upat i0 f (Array.length a)) upats [] in + while !i0 >= 0 do + let i = !i0 in i0 := -1; + let one_match (u, np) = + let skip = + if i <= np then i < np else + if u.up_k == KpatFlex then begin i0 := i - 1; false end else + begin if !i0 < np then i0 := np; true end in + if skip then () else try + let ise' = match u.up_k with + | KpatFixed | KpatConst -> ise + | KpatEvar _ -> + let _, pka = destEvar u.up_f and _, ka = destEvar f in + unif_HO_args env ise pka 0 ka + | KpatLet -> + let x, v, t, b = destLetIn f in + let _, pv, _, pb = destLetIn u.up_f in + let ise' = unif_HO env ise pv v in + unif_HO (Environ.push_rel (x, None, t) env) ise' pb b + | KpatFlex | KpatProj _ -> + unif_HO env ise u.up_f (mkSubApp f (i - Array.length u.up_a) a) + | _ -> unif_HO env ise u.up_f f in + let ise'' = unif_HO_args env ise' u.up_a (i - Array.length u.up_a) a in + let lhs = mkSubApp f i a in + let pt' = unif_end env sigma0 ise'' u.up_t (u.up_ok lhs) in + on_instance (ungen_upat lhs pt' u) + with FoundUnif _ as sigma_u -> raise sigma_u + | NoProgress -> it_did_match := true + | _ -> () in + List.iter one_match fpats + done; + iter_constr_LR (aux upats env sigma0 ise) f; + Array.iter (aux upats env sigma0 ise) a + in + aux upats env sigma0 ise c; + if !it_did_match then raise NoProgress + +let prof_HO = mk_profiler "match_upats_HO";; +let match_upats_HO ~on_instance upats env sigma0 ise c = + prof_HO.profile (match_upats_HO ~on_instance upats env sigma0) ise c +;; + + +let fixed_upat = function +| {up_k = KpatFlex | KpatEvar _ | KpatProj _} -> false +| {up_t = t} -> not (occur_existential t) + +let do_once r f = match !r with Some _ -> () | None -> r := Some (f ()) + +let assert_done r = + match !r with Some x -> x | None -> Errors.anomaly (str"do_once never called") + +let assert_done_multires r = + match !r with + | None -> Errors.anomaly (str"do_once never called") + | Some (n, xs) -> + r := Some (n+1,xs); + try List.nth xs n with Failure _ -> raise NoMatch + +type subst = Environ.env -> Term.constr -> Term.constr -> int -> Term.constr +type find_P = + Environ.env -> Term.constr -> int -> + k:subst -> + Term.constr +type conclude = unit -> + Term.constr * ssrdir * (Evd.evar_map * Evd.evar_universe_context * Term.constr) + +(* upats_origin makes a better error message only *) +let mk_tpattern_matcher ?(all_instances=false) + ?(raise_NoMatch=false) ?upats_origin sigma0 occ (ise, upats) += + let nocc = ref 0 and skip_occ = ref false in + let use_occ, occ_list = match occ with + | Some (true, ol) -> ol = [], ol + | Some (false, ol) -> ol <> [], ol + | None -> false, [] in + let max_occ = List.fold_right max occ_list 0 in + let subst_occ = + let occ_set = Array.make max_occ (not use_occ) in + let _ = List.iter (fun i -> occ_set.(i - 1) <- use_occ) occ_list in + let _ = if max_occ = 0 then skip_occ := use_occ in + fun () -> incr nocc; + if !nocc = max_occ then skip_occ := use_occ; + if !nocc <= max_occ then occ_set.(!nocc - 1) else not use_occ in + let upat_that_matched = ref None in + let match_EQ env sigma u = + match u.up_k with + | KpatLet -> + let x, pv, t, pb = destLetIn u.up_f in + let env' = Environ.push_rel (x, None, t) env in + let match_let f = match kind_of_term f with + | LetIn (_, v, _, b) -> unif_EQ env sigma pv v && unif_EQ env' sigma pb b + | _ -> false in match_let + | KpatFixed -> Term.eq_constr u.up_f + | KpatConst -> Term.eq_constr u.up_f + | KpatLam -> fun c -> + (match kind_of_term c with + | Lambda _ -> unif_EQ env sigma u.up_f c + | _ -> false) + | _ -> unif_EQ env sigma u.up_f in +let p2t p = mkApp(p.up_f,p.up_a) in +let source () = match upats_origin, upats with + | None, [p] -> + (if fixed_upat p then str"term " else str"partial term ") ++ + pr_constr_pat (p2t p) ++ spc() + | Some (dir,rule), [p] -> str"The " ++ pr_dir_side dir ++ str" of " ++ + pr_constr_pat rule ++ fnl() ++ ws 4 ++ pr_constr_pat (p2t p) ++ fnl() + | Some (dir,rule), _ -> str"The " ++ pr_dir_side dir ++ str" of " ++ + pr_constr_pat rule ++ spc() + | _, [] | None, _::_::_ -> + Errors.anomaly (str"mk_tpattern_matcher with no upats_origin") in +let on_instance, instances = + let instances = ref [] in + (fun x -> + if all_instances then instances := !instances @ [x] + else raise (FoundUnif x)), + (fun () -> !instances) in +let rec uniquize = function + | [] -> [] + | (sigma,_,{ up_f = f; up_a = a; up_t = t } as x) :: xs -> + let t = Reductionops.nf_evar sigma t in + let f = Reductionops.nf_evar sigma f in + let a = Array.map (Reductionops.nf_evar sigma) a in + let neq (sigma1,_,{ up_f = f1; up_a = a1; up_t = t1 }) = + let t1 = Reductionops.nf_evar sigma1 t1 in + let f1 = Reductionops.nf_evar sigma1 f1 in + let a1 = Array.map (Reductionops.nf_evar sigma1) a1 in + not (Term.eq_constr t t1 && + Term.eq_constr f f1 && CArray.for_all2 Term.eq_constr a a1) in + x :: uniquize (List.filter neq xs) in + +((fun env c h ~k -> + do_once upat_that_matched (fun () -> + try + if not all_instances then match_upats_FO upats env sigma0 ise c; + match_upats_HO ~on_instance upats env sigma0 ise c; + raise NoMatch + with FoundUnif sigma_u -> 0,[sigma_u] + | (NoMatch|NoProgress) when all_instances && instances () <> [] -> + 0, uniquize (instances ()) + | NoMatch when (not raise_NoMatch) -> + errorstrm (source () ++ str "does not match any subterm of the goal") + | NoProgress when (not raise_NoMatch) -> + let dir = match upats_origin with Some (d,_) -> d | _ -> + Errors.anomaly (str"mk_tpattern_matcher with no upats_origin") in + errorstrm (str"all matches of "++source()++ + str"are equal to the " ++ pr_dir_side (inv_dir dir)) + | NoProgress -> raise NoMatch); + let sigma, _, ({up_f = pf; up_a = pa} as u) = + if all_instances then assert_done_multires upat_that_matched + else List.hd (snd(assert_done upat_that_matched)) in + pp(lazy(str"sigma@tmatch=" ++ pr_evar_map None sigma)); + if !skip_occ then ((*ignore(k env u.up_t 0);*) c) else + let match_EQ = match_EQ env sigma u in + let pn = Array.length pa in + let rec subst_loop (env,h as acc) c' = + if !skip_occ then c' else + let f, a = splay_app sigma c' in + if Array.length a >= pn && match_EQ f && unif_EQ_args env sigma pa a then + let a1, a2 = Array.chop (Array.length pa) a in + let fa1 = mkApp (f, a1) in + let f' = if subst_occ () then k env u.up_t fa1 h else fa1 in + mkApp (f', Array.map_left (subst_loop acc) a2) + else + (* TASSI: clear letin values to avoid unfolding *) + let inc_h (n,_,ty) (env,h') = Environ.push_rel (n,None,ty) env, h' + 1 in + let f' = map_constr_with_binders_left_to_right inc_h subst_loop acc f in + mkApp (f', Array.map_left (subst_loop acc) a) in + subst_loop (env,h) c) : find_P), +((fun () -> + let sigma, uc, ({up_f = pf; up_a = pa} as u) = + match !upat_that_matched with + | Some (_,x) -> List.hd x | None when raise_NoMatch -> raise NoMatch + | None -> Errors.anomaly (str"companion function never called") in + let p' = mkApp (pf, pa) in + if max_occ <= !nocc then p', u.up_dir, (sigma, uc, u.up_t) + else errorstrm (str"Only " ++ int !nocc ++ str" < " ++ int max_occ ++ + str(String.plural !nocc " occurence") ++ match upats_origin with + | None -> str" of" ++ spc() ++ pr_constr_pat p' + | Some (dir,rule) -> str" of the " ++ pr_dir_side dir ++ fnl() ++ + ws 4 ++ pr_constr_pat p' ++ fnl () ++ + str"of " ++ pr_constr_pat rule)) : conclude) + +type ('ident, 'term) ssrpattern = + | T of 'term + | In_T of 'term + | X_In_T of 'ident * 'term + | In_X_In_T of 'ident * 'term + | E_In_X_In_T of 'term * 'ident * 'term + | E_As_X_In_T of 'term * 'ident * 'term + +let pr_pattern = function + | T t -> prl_term t + | In_T t -> str "in " ++ prl_term t + | X_In_T (x,t) -> prl_term x ++ str " in " ++ prl_term t + | In_X_In_T (x,t) -> str "in " ++ prl_term x ++ str " in " ++ prl_term t + | E_In_X_In_T (e,x,t) -> + prl_term e ++ str " in " ++ prl_term x ++ str " in " ++ prl_term t + | E_As_X_In_T (e,x,t) -> + prl_term e ++ str " as " ++ prl_term x ++ str " in " ++ prl_term t + +let pr_pattern_w_ids = function + | T t -> prl_term t + | In_T t -> str "in " ++ prl_term t + | X_In_T (x,t) -> pr_id x ++ str " in " ++ prl_term t + | In_X_In_T (x,t) -> str "in " ++ pr_id x ++ str " in " ++ prl_term t + | E_In_X_In_T (e,x,t) -> + prl_term e ++ str " in " ++ pr_id x ++ str " in " ++ prl_term t + | E_As_X_In_T (e,x,t) -> + prl_term e ++ str " as " ++ pr_id x ++ str " in " ++ prl_term t + +let pr_pattern_aux pr_constr = function + | T t -> pr_constr t + | In_T t -> str "in " ++ pr_constr t + | X_In_T (x,t) -> pr_constr x ++ str " in " ++ pr_constr t + | In_X_In_T (x,t) -> str "in " ++ pr_constr x ++ str " in " ++ pr_constr t + | E_In_X_In_T (e,x,t) -> + pr_constr e ++ str " in " ++ pr_constr x ++ str " in " ++ pr_constr t + | E_As_X_In_T (e,x,t) -> + pr_constr e ++ str " as " ++ pr_constr x ++ str " in " ++ pr_constr t +let pp_pattern (sigma, p) = + pr_pattern_aux (fun t -> pr_constr (pi3 (nf_open_term sigma sigma t))) p +let pr_cpattern = pr_term +let pr_rpattern _ _ _ = pr_pattern + +let pr_option f = function None -> mt() | Some x -> f x +let pr_ssrpattern _ _ _ = pr_option pr_pattern +let pr_pattern_squarep = pr_option (fun r -> str "[" ++ pr_pattern r ++ str "]") +let pr_ssrpattern_squarep _ _ _ = pr_pattern_squarep +let pr_pattern_roundp = pr_option (fun r -> str "(" ++ pr_pattern r ++ str ")") +let pr_ssrpattern_roundp _ _ _ = pr_pattern_roundp + +let wit_rpatternty = add_genarg "rpatternty" pr_pattern + +ARGUMENT EXTEND rpattern TYPED AS rpatternty PRINTED BY pr_rpattern + | [ lconstr(c) ] -> [ T (mk_lterm c) ] + | [ "in" lconstr(c) ] -> [ In_T (mk_lterm c) ] + | [ lconstr(x) "in" lconstr(c) ] -> + [ X_In_T (mk_lterm x, mk_lterm c) ] + | [ "in" lconstr(x) "in" lconstr(c) ] -> + [ In_X_In_T (mk_lterm x, mk_lterm c) ] + | [ lconstr(e) "in" lconstr(x) "in" lconstr(c) ] -> + [ E_In_X_In_T (mk_lterm e, mk_lterm x, mk_lterm c) ] + | [ lconstr(e) "as" lconstr(x) "in" lconstr(c) ] -> + [ E_As_X_In_T (mk_lterm e, mk_lterm x, mk_lterm c) ] +END + +type cpattern = char * glob_constr_and_expr +let tag_of_cpattern = fst +let loc_of_cpattern = loc_ofCG +let cpattern_of_term t = t +type occ = (bool * int list) option + +type rpattern = (cpattern, cpattern) ssrpattern +let pr_rpattern = pr_pattern + +type pattern = Evd.evar_map * (Term.constr, Term.constr) ssrpattern + + +let id_of_cpattern = function + | _,(_,Some (CRef (Ident (_, x), _))) -> Some x + | _,(_,Some (CAppExpl (_, (_, Ident (_, x), _), []))) -> Some x + | _,(GRef (_, VarRef x, _) ,None) -> Some x + | _ -> None +let id_of_Cterm t = match id_of_cpattern t with + | Some x -> x + | None -> loc_error (loc_of_cpattern t) "Only identifiers are allowed here" + +let interp_wit wit ist gl x = + let globarg = in_gen (glbwit wit) x in + let sigma, arg = interp_genarg ist (pf_env gl) (project gl) (pf_concl gl) gl.Evd.it globarg in + sigma, out_gen (topwit wit) arg +let interp_constr = interp_wit wit_constr +let interp_open_constr ist gl gc = + interp_wit wit_open_constr ist gl ((), gc) +let pf_intern_term ist gl (_, c) = glob_constr ist (pf_env gl) c +let interp_term ist gl (_, c) = snd (interp_open_constr ist gl c) +let glob_ssrterm gs = function + | k, (_, Some c) -> k, Tacintern.intern_constr gs c + | ct -> ct +let subst_ssrterm s (k, c) = k, Tacsubst.subst_glob_constr_and_expr s c +let pr_ssrterm _ _ _ = pr_term +let input_ssrtermkind strm = match Compat.get_tok (stream_nth 0 strm) with + | Tok.KEYWORD "(" -> '(' + | Tok.KEYWORD "@" -> '@' + | _ -> ' ' +let ssrtermkind = Gram.Entry.of_parser "ssrtermkind" input_ssrtermkind + +(* This piece of code asserts the following notations are reserved *) +(* Reserved Notation "( a 'in' b )" (at level 0). *) +(* Reserved Notation "( a 'as' b )" (at level 0). *) +(* Reserved Notation "( a 'in' b 'in' c )" (at level 0). *) +(* Reserved Notation "( a 'as' b 'in' c )" (at level 0). *) +let glob_cpattern gs p = + pp(lazy(str"globbing pattern: " ++ pr_term p)); + let glob x = snd (glob_ssrterm gs (mk_lterm x)) in + let encode k s l = + let name = Name (id_of_string ("_ssrpat_" ^ s)) in + k, (mkRCast mkRHole (mkRLambda name mkRHole (mkRApp mkRHole l)), None) in + let bind_in t1 t2 = + let d = dummy_loc in let n = Name (destCVar t1) in + fst (glob (mkCCast d (mkCHole d) (mkCLambda d n (mkCHole d) t2))) in + let check_var t2 = if not (isCVar t2) then + loc_error (constr_loc t2) "Only identifiers are allowed here" in + match p with + | _, (_, None) as x -> x + | k, (v, Some t) as orig -> + if k = 'x' then glob_ssrterm gs ('(', (v, Some t)) else + match t with + | CNotation(_, "( _ in _ )", ([t1; t2], [], [])) -> + (try match glob t1, glob t2 with + | (r1, None), (r2, None) -> encode k "In" [r1;r2] + | (r1, Some _), (r2, Some _) when isCVar t1 -> + encode k "In" [r1; r2; bind_in t1 t2] + | (r1, Some _), (r2, Some _) -> encode k "In" [r1; r2] + | _ -> Errors.anomaly (str"where are we?") + with _ when isCVar t1 -> encode k "In" [bind_in t1 t2]) + | CNotation(_, "( _ in _ in _ )", ([t1; t2; t3], [], [])) -> + check_var t2; encode k "In" [fst (glob t1); bind_in t2 t3] + | CNotation(_, "( _ as _ )", ([t1; t2], [], [])) -> + encode k "As" [fst (glob t1); fst (glob t2)] + | CNotation(_, "( _ as _ in _ )", ([t1; t2; t3], [], [])) -> + check_var t2; encode k "As" [fst (glob t1); bind_in t2 t3] + | _ -> glob_ssrterm gs orig +;; + +let interp_ssrterm _ gl t = Tacmach.project gl, t + +ARGUMENT EXTEND cpattern + PRINTED BY pr_ssrterm + INTERPRETED BY interp_ssrterm + GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm + RAW_TYPED AS cpattern RAW_PRINTED BY pr_ssrterm + GLOB_TYPED AS cpattern GLOB_PRINTED BY pr_ssrterm +| [ "Qed" constr(c) ] -> [ mk_lterm c ] +END + +let (!@) = Compat.to_coqloc + +GEXTEND Gram + GLOBAL: cpattern; + cpattern: [[ k = ssrtermkind; c = constr -> + let pattern = mk_term k c in + if loc_ofCG pattern <> !@loc && k = '(' then mk_term 'x' c else pattern ]]; +END + +ARGUMENT EXTEND lcpattern + PRINTED BY pr_ssrterm + INTERPRETED BY interp_ssrterm + GLOBALIZED BY glob_cpattern SUBSTITUTED BY subst_ssrterm + RAW_TYPED AS cpattern RAW_PRINTED BY pr_ssrterm + GLOB_TYPED AS cpattern GLOB_PRINTED BY pr_ssrterm +| [ "Qed" lconstr(c) ] -> [ mk_lterm c ] +END + +GEXTEND Gram + GLOBAL: lcpattern; + lcpattern: [[ k = ssrtermkind; c = lconstr -> + let pattern = mk_term k c in + if loc_ofCG pattern <> !@loc && k = '(' then mk_term 'x' c else pattern ]]; +END + +let interp_pattern ist gl red redty = + pp(lazy(str"interpreting: " ++ pr_pattern red)); + let xInT x y = X_In_T(x,y) and inXInT x y = In_X_In_T(x,y) in + let inT x = In_T x and eInXInT e x t = E_In_X_In_T(e,x,t) in + let eAsXInT e x t = E_As_X_In_T(e,x,t) in + let mkG ?(k=' ') x = k,(x,None) in + let decode t f g = + try match (pf_intern_term ist gl t) with + | GCast(_,GHole _,CastConv(GLambda(_,Name x,_,_,c))) -> f x (' ',(c,None)) + | it -> g t with _ -> g t in + let decodeG t f g = decode (mkG t) f g in + let bad_enc id _ = Errors.anomaly (str"bad encoding for pattern "++str id) in + let cleanup_XinE h x rp sigma = + let h_k = match kind_of_term h with Evar (k,_) -> k | _ -> assert false in + let to_clean, update = (* handle rename if x is already used *) + let ctx = pf_hyps gl in + let len = Context.named_context_length ctx in + let name = ref None in + try ignore(Context.lookup_named x ctx); (name, fun k -> + if !name = None then + let nctx = Evd.evar_context (Evd.find sigma k) in + let nlen = Context.named_context_length nctx in + if nlen > len then begin + name := Some (pi1 (List.nth nctx (nlen - len - 1))) + end) + with Not_found -> ref (Some x), fun _ -> () in + let sigma0 = project gl in + let new_evars = + let rec aux acc t = match kind_of_term t with + | Evar (k,_) -> + if k = h_k || List.mem k acc || Evd.mem sigma0 k then acc else + (update k; k::acc) + | _ -> fold_constr aux acc t in + aux [] (Evarutil.nf_evar sigma rp) in + let sigma = + List.fold_left (fun sigma e -> + if Evd.is_defined sigma e then sigma else (* clear may be recursiv *) + let name = Option.get !to_clean in + pp(lazy(pr_id name)); + try snd(Logic.prim_refiner (Proof_type.Thin [name]) sigma e) + with Evarutil.ClearDependencyError _ -> sigma) + sigma new_evars in + sigma in + let red = match red with + | T(k,(GCast (_,GHole _,(CastConv(GLambda (_,Name id,_,_,t)))),None)) + when let id = string_of_id id in let len = String.length id in + (len > 8 && String.sub id 0 8 = "_ssrpat_") -> + let id = string_of_id id in let len = String.length id in + (match String.sub id 8 (len - 8), t with + | "In", GApp(_, _, [t]) -> decodeG t xInT (fun x -> T x) + | "In", GApp(_, _, [e; t]) -> decodeG t (eInXInT (mkG e)) (bad_enc id) + | "In", GApp(_, _, [e; t; e_in_t]) -> + decodeG t (eInXInT (mkG e)) + (fun _ -> decodeG e_in_t xInT (fun _ -> assert false)) + | "As", GApp(_, _, [e; t]) -> decodeG t (eAsXInT (mkG e)) (bad_enc id) + | _ -> bad_enc id ()) + | T t -> decode t xInT (fun x -> T x) + | In_T t -> decode t inXInT inT + | X_In_T (e,t) -> decode t (eInXInT e) (fun x -> xInT (id_of_Cterm e) x) + | In_X_In_T (e,t) -> inXInT (id_of_Cterm e) t + | E_In_X_In_T (e,x,rp) -> eInXInT e (id_of_Cterm x) rp + | E_As_X_In_T (e,x,rp) -> eAsXInT e (id_of_Cterm x) rp in + pp(lazy(str"decoded as: " ++ pr_pattern_w_ids red)); + let red = match redty with None -> red | Some ty -> let ty = ' ', ty in + match red with + | T t -> T (combineCG t ty (mkCCast (loc_ofCG t)) mkRCast) + | X_In_T (x,t) -> + let ty = pf_intern_term ist gl ty in + E_As_X_In_T (mkG (mkRCast mkRHole ty), x, t) + | E_In_X_In_T (e,x,t) -> + let ty = mkG (pf_intern_term ist gl ty) in + E_In_X_In_T (combineCG e ty (mkCCast (loc_ofCG t)) mkRCast, x, t) + | E_As_X_In_T (e,x,t) -> + let ty = mkG (pf_intern_term ist gl ty) in + E_As_X_In_T (combineCG e ty (mkCCast (loc_ofCG t)) mkRCast, x, t) + | red -> red in + pp(lazy(str"typed as: " ++ pr_pattern_w_ids red)); + let mkXLetIn loc x (a,(g,c)) = match c with + | Some b -> a,(g,Some (mkCLetIn loc x (mkCHole loc) b)) + | None -> a,(GLetIn (loc,x,(GHole (loc, BinderType x, IntroAnonymous, None)), g), None) in + match red with + | T t -> let sigma, t = interp_term ist gl t in sigma, T t + | In_T t -> let sigma, t = interp_term ist gl t in sigma, In_T t + | X_In_T (x, rp) | In_X_In_T (x, rp) -> + let mk x p = match red with X_In_T _ -> X_In_T(x,p) | _ -> In_X_In_T(x,p) in + let rp = mkXLetIn dummy_loc (Name x) rp in + let sigma, rp = interp_term ist gl rp in + let _, h, _, rp = destLetIn rp in + let sigma = cleanup_XinE h x rp sigma in + let rp = subst1 h (Evarutil.nf_evar sigma rp) in + sigma, mk h rp + | E_In_X_In_T(e, x, rp) | E_As_X_In_T (e, x, rp) -> + let mk e x p = + match red with E_In_X_In_T _ ->E_In_X_In_T(e,x,p)|_->E_As_X_In_T(e,x,p) in + let rp = mkXLetIn dummy_loc (Name x) rp in + let sigma, rp = interp_term ist gl rp in + let _, h, _, rp = destLetIn rp in + let sigma = cleanup_XinE h x rp sigma in + let rp = subst1 h (Evarutil.nf_evar sigma rp) in + let sigma, e = interp_term ist (re_sig (sig_it gl) sigma) e in + sigma, mk e h rp +;; +let interp_cpattern ist gl red redty = interp_pattern ist gl (T red) redty;; +let interp_rpattern ist gl red = interp_pattern ist gl red None;; + +(* The full occurrence set *) +let noindex = Some(false,[]) + +(* calls do_subst on every sub-term identified by (pattern,occ) *) +let eval_pattern ?raise_NoMatch env0 sigma0 concl0 pattern occ do_subst = + let fs sigma x = Reductionops.nf_evar sigma x in + let pop_evar sigma e p = + let { Evd.evar_body = e_body } as e_def = Evd.find sigma e in + let e_body = match e_body with Evar_defined c -> c + | _ -> errorstrm (str "Matching the pattern " ++ pr_constr p ++ + str " did not instantiate ?" ++ int (Evar.repr e) ++ spc () ++ + str "Does the variable bound by the \"in\" construct occur "++ + str "in the pattern?") in + let sigma = + Evd.add (Evd.remove sigma e) e {e_def with Evd.evar_body = Evar_empty} in + sigma, e_body in + let ex_value hole = + match kind_of_term hole with Evar (e,_) -> e | _ -> assert false in + let mk_upat_for ?hack env sigma0 (sigma, t) ?(p=t) ok = + let sigma,pat= mk_tpattern ?hack env sigma0 (sigma,p) ok L2R (fs sigma t) in + sigma, [pat] in + match pattern with + | None -> do_subst env0 concl0 concl0 1 + | Some (sigma, (T rp | In_T rp)) -> + let rp = fs sigma rp in + let ise = create_evar_defs sigma in + let occ = match pattern with Some (_, T _) -> occ | _ -> noindex in + let rp = mk_upat_for env0 sigma0 (ise, rp) all_ok in + let find_T, end_T = mk_tpattern_matcher ?raise_NoMatch sigma0 occ rp in + let concl = find_T env0 concl0 1 do_subst in + let _ = end_T () in + concl + | Some (sigma, (X_In_T (hole, p) | In_X_In_T (hole, p))) -> + let p = fs sigma p in + let occ = match pattern with Some (_, X_In_T _) -> occ | _ -> noindex in + let ex = ex_value hole in + let rp = mk_upat_for ~hack:true env0 sigma0 (sigma, p) all_ok in + let find_T, end_T = mk_tpattern_matcher sigma0 noindex rp in + (* we start from sigma, so hole is considered a rigid head *) + let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in + let find_X, end_X = mk_tpattern_matcher ?raise_NoMatch sigma occ holep in + let concl = find_T env0 concl0 1 (fun env c _ h -> + let p_sigma = unify_HO env (create_evar_defs sigma) c p in + let sigma, e_body = pop_evar p_sigma ex p in + fs p_sigma (find_X env (fs sigma p) h + (fun env _ -> do_subst env e_body))) in + let _ = end_X () in let _ = end_T () in + concl + | Some (sigma, E_In_X_In_T (e, hole, p)) -> + let p, e = fs sigma p, fs sigma e in + let ex = ex_value hole in + let rp = mk_upat_for ~hack:true env0 sigma0 (sigma, p) all_ok in + let find_T, end_T = mk_tpattern_matcher sigma0 noindex rp in + let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in + let find_X, end_X = mk_tpattern_matcher sigma noindex holep in + let re = mk_upat_for env0 sigma0 (sigma, e) all_ok in + let find_E, end_E = mk_tpattern_matcher ?raise_NoMatch sigma0 occ re in + let concl = find_T env0 concl0 1 (fun env c _ h -> + let p_sigma = unify_HO env (create_evar_defs sigma) c p in + let sigma, e_body = pop_evar p_sigma ex p in + fs p_sigma (find_X env (fs sigma p) h (fun env c _ h -> + find_E env e_body h do_subst))) in + let _ = end_E () in let _ = end_X () in let _ = end_T () in + concl + | Some (sigma, E_As_X_In_T (e, hole, p)) -> + let p, e = fs sigma p, fs sigma e in + let ex = ex_value hole in + let rp = + let e_sigma = unify_HO env0 sigma hole e in + e_sigma, fs e_sigma p in + let rp = mk_upat_for ~hack:true env0 sigma0 rp all_ok in + let find_TE, end_TE = mk_tpattern_matcher sigma0 noindex rp in + let holep = mk_upat_for env0 sigma (sigma, hole) all_ok in + let find_X, end_X = mk_tpattern_matcher sigma occ holep in + let concl = find_TE env0 concl0 1 (fun env c _ h -> + let p_sigma = unify_HO env (create_evar_defs sigma) c p in + let sigma, e_body = pop_evar p_sigma ex p in + fs p_sigma (find_X env (fs sigma p) h (fun env c _ h -> + let e_sigma = unify_HO env sigma e_body e in + let e_body = fs e_sigma e in + do_subst env e_body e_body h))) in + let _ = end_X () in let _ = end_TE () in + concl +;; + +let redex_of_pattern ?(resolve_typeclasses=false) env (sigma, p) = + let e = match p with + | In_T _ | In_X_In_T _ -> Errors.anomaly (str"pattern without redex") + | T e | X_In_T (e, _) | E_As_X_In_T (e, _, _) | E_In_X_In_T (e, _, _) -> e in + let sigma = + if not resolve_typeclasses then sigma + else Typeclasses.resolve_typeclasses ~fail:false env sigma in + Reductionops.nf_evar sigma e, Evd.evar_universe_context sigma + +let fill_occ_pattern ?raise_NoMatch env sigma cl pat occ h = + let find_R, conclude = let r = ref None in + (fun env c _ h' -> do_once r (fun () -> c, Evd.empty_evar_universe_context); + mkRel (h'+h-1)), + (fun _ -> if !r = None then redex_of_pattern env pat else assert_done r) in + let cl = eval_pattern ?raise_NoMatch env sigma cl (Some pat) occ find_R in + let e = conclude cl in + e, cl +;; + +(* clenup interface for external use *) +let mk_tpattern ?p_origin env sigma0 sigma_t f dir c = + mk_tpattern ?p_origin env sigma0 sigma_t f dir c +;; + +let pf_fill_occ env concl occ sigma0 p (sigma, t) ok h = + let ise = create_evar_defs sigma in + let ise, u = mk_tpattern env sigma0 (ise,t) ok L2R p in + let find_U, end_U = + mk_tpattern_matcher ~raise_NoMatch:true sigma0 occ (ise,[u]) in + let concl = find_U env concl h (fun _ _ _ -> mkRel) in + let rdx, _, (sigma, uc, p) = end_U () in + sigma, uc, p, concl, rdx + +let fill_occ_term env cl occ sigma0 (sigma, t) = + try + let sigma',uc,t',cl,_= pf_fill_occ env cl occ sigma0 t (sigma, t) all_ok 1 in + if sigma' != sigma0 then Errors.error "matching impacts evars" + else cl, (Evd.merge_universe_context sigma' uc, t') + with NoMatch -> try + let sigma', uc, t' = + unif_end env sigma0 (create_evar_defs sigma) t (fun _ -> true) in + if sigma' != sigma0 then raise NoMatch + else cl, (Evd.merge_universe_context sigma' uc, t') + with _ -> + errorstrm (str "partial term " ++ pr_constr_pat t + ++ str " does not match any subterm of the goal") + +let pf_fill_occ_term gl occ t = + let sigma0 = project gl and env = pf_env gl and concl = pf_concl gl in + let cl,(_,t) = fill_occ_term env concl occ sigma0 t in + cl, t + +let cpattern_of_id id = ' ', (GRef (dummy_loc, VarRef id, None), None) + +let is_wildcard = function + | _,(_,Some (CHole _)|GHole _,None) -> true + | _ -> false + +(* "ssrpattern" *) +let pr_ssrpatternarg _ _ _ cpat = pr_rpattern cpat + +ARGUMENT EXTEND ssrpatternarg + TYPED AS rpattern + PRINTED BY pr_ssrpatternarg +| [ "[" rpattern(pat) "]" ] -> [ pat ] +END + +let pf_merge_uc uc gl = + re_sig (sig_it gl) (Evd.merge_universe_context (project gl) uc) + +let ssrpatterntac ist arg gl = + let pat = interp_rpattern ist gl arg in + let sigma0 = project gl in + let concl0 = pf_concl gl in + let (t, uc), concl_x = + fill_occ_pattern (Global.env()) sigma0 concl0 pat noindex 1 in + let tty = pf_type_of gl t in + let concl = mkLetIn (Name (id_of_string "toto"), t, tty, concl_x) in + Proofview.V82.of_tactic (convert_concl concl DEFAULTcast) gl + +TACTIC EXTEND ssrat +| [ "ssrpattern" ssrpatternarg(arg) ] -> [ Proofview.V82.tactic (ssrpatterntac ist arg) ] +END + +let ssrinstancesof ist arg gl = + let ok rhs lhs ise = true in +(* not (Term.eq_constr lhs (Evarutil.nf_evar ise rhs)) in *) + let env, sigma, concl = pf_env gl, project gl, pf_concl gl in + let sigma0, cpat = interp_cpattern ist gl arg None in + let pat = match cpat with T x -> x | _ -> errorstrm (str"Not supported") in + let etpat, tpat = mk_tpattern env sigma (sigma0,pat) (ok pat) L2R pat in + let find, conclude = + mk_tpattern_matcher ~all_instances:true ~raise_NoMatch:true + sigma None (etpat,[tpat]) in + let print env p c _ = ppnl (hov 1 (str"instance:" ++ spc() ++ pr_constr p ++ spc() ++ str "matches:" ++ spc() ++ pr_constr c)); c in + ppnl (str"BEGIN INSTANCES"); + try + while true do + ignore(find env concl 1 ~k:print) + done; raise NoMatch + with NoMatch -> ppnl (str"END INSTANCES"); tclIDTAC gl + +TACTIC EXTEND ssrinstoftpat +| [ "ssrinstancesoftpat" cpattern(arg) ] -> [ Proofview.V82.tactic (ssrinstancesof ist arg) ] +END + +(* We wipe out all the keywords generated by the grammar rules we defined. *) +(* The user is supposed to Require Import ssreflect or Require ssreflect *) +(* and Import ssreflect.SsrSyntax to obtain these keywords and as a *) +(* consequence the extended ssreflect grammar. *) +let () = Lexer.unfreeze frozen_lexer ;; + +(* vim: set filetype=ocaml foldmethod=marker: *) diff --git a/mathcomp/ssreflect/plugin/v8.5beta2/ssrmatching.mli b/mathcomp/ssreflect/plugin/v8.5beta2/ssrmatching.mli new file mode 100644 index 0000000..e8b4d81 --- /dev/null +++ b/mathcomp/ssreflect/plugin/v8.5beta2/ssrmatching.mli @@ -0,0 +1,239 @@ +(* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) + +open Genarg +open Tacexpr +open Environ +open Evd +open Proof_type +open Term + +(** ******** Small Scale Reflection pattern matching facilities ************* *) + +(** Pattern parsing *) + +(** The type of context patterns, the patterns of the [set] tactic and + [:] tactical. These are patterns that identify a precise subterm. *) +type cpattern +val pr_cpattern : cpattern -> Pp.std_ppcmds + +(** CS cpattern: (f _), (X in t), (t in X in t), (t as X in t) *) +val cpattern : cpattern Pcoq.Gram.entry +val wit_cpattern : cpattern uniform_genarg_type + +(** OS cpattern: f _, (X in t), (t in X in t), (t as X in t) *) +val lcpattern : cpattern Pcoq.Gram.entry +val wit_lcpattern : cpattern uniform_genarg_type + +(** The type of rewrite patterns, the patterns of the [rewrite] tactic. + These patterns also include patterns that identify all the subterms + of a context (i.e. "in" prefix) *) +type rpattern +val pr_rpattern : rpattern -> Pp.std_ppcmds + +(** OS rpattern: f _, in t, X in t, in X in t, t in X in t, t as X in t *) +val rpattern : rpattern Pcoq.Gram.entry +val wit_rpattern : rpattern uniform_genarg_type + +(** Pattern interpretation and matching *) + +exception NoMatch +exception NoProgress + +(** AST for [rpattern] (and consequently [cpattern]) *) +type ('ident, 'term) ssrpattern = + | T of 'term + | In_T of 'term + | X_In_T of 'ident * 'term + | In_X_In_T of 'ident * 'term + | E_In_X_In_T of 'term * 'ident * 'term + | E_As_X_In_T of 'term * 'ident * 'term + +type pattern = evar_map * (constr, constr) ssrpattern +val pp_pattern : pattern -> Pp.std_ppcmds + +(** Extracts the redex and applies to it the substitution part of the pattern. + @raise Anomaly if called on [In_T] or [In_X_In_T] *) +val redex_of_pattern : + ?resolve_typeclasses:bool -> env -> pattern -> + constr Evd.in_evar_universe_context + +(** [interp_rpattern ise gl rpat] "internalizes" and "interprets" [rpat] + in the current [Ltac] interpretation signature [ise] and tactic input [gl]*) +val interp_rpattern : + Tacinterp.interp_sign -> goal Tacmach.sigma -> + rpattern -> + pattern + +(** [interp_cpattern ise gl cpat ty] "internalizes" and "interprets" [cpat] + in the current [Ltac] interpretation signature [ise] and tactic input [gl]. + [ty] is an optional type for the redex of [cpat] *) +val interp_cpattern : + Tacinterp.interp_sign -> goal Tacmach.sigma -> + cpattern -> glob_constr_and_expr option -> + pattern + +(** The set of occurrences to be matched. The boolean is set to true + * to signal the complement of this set (i.e. {-1 3}) *) +type occ = (bool * int list) option + +(** [subst e p t i]. [i] is the number of binders + traversed so far, [p] the term from the pattern, [t] the matched one *) +type subst = env -> constr -> constr -> int -> constr + +(** [eval_pattern b env sigma t pat occ subst] maps [t] calling [subst] on every + [occ] occurrence of [pat]. The [int] argument is the number of + binders traversed. If [pat] is [None] then then subst is called on [t]. + [t] must live in [env] and [sigma], [pat] must have been interpreted in + (an extension of) [sigma]. + @raise NoMatch if [pat] has no occurrence and [b] is [true] (default [false]) + @return [t] where all [occ] occurrences of [pat] have been mapped using + [subst] *) +val eval_pattern : + ?raise_NoMatch:bool -> + env -> evar_map -> constr -> + pattern option -> occ -> subst -> + constr + +(** [fill_occ_pattern b env sigma t pat occ h] is a simplified version of + [eval_pattern]. + It replaces all [occ] occurrences of [pat] in [t] with Rel [h]. + [t] must live in [env] and [sigma], [pat] must have been interpreted in + (an extension of) [sigma]. + @raise NoMatch if [pat] has no occurrence and [b] is [true] (default [false]) + @return the instance of the redex of [pat] that was matched and [t] + transformed as described above. *) +val fill_occ_pattern : + ?raise_NoMatch:bool -> + env -> evar_map -> constr -> + pattern -> occ -> int -> + constr Evd.in_evar_universe_context * constr + +(** *************************** Low level APIs ****************************** *) + +(* The primitive matching facility. It matches of a term with holes, like + the T pattern above, and calls a continuation on its occurrences. *) + +type ssrdir = L2R | R2L +val pr_dir_side : ssrdir -> Pp.std_ppcmds + +(** a pattern for a term with wildcards *) +type tpattern + +(** [mk_tpattern env sigma0 sigma_p ok p_origin dir t] compiles a term [t] + living in [env] [sigma] (an extension of [sigma0]) intro a [tpattern]. + The [tpattern] can hold a (proof) term [p] and a diction [dir]. The [ok] + callback is used to filter occurrences. + @return the compiled [tpattern] and its [evar_map] + @raise UserEerror is the pattern is a wildcard *) +val mk_tpattern : + ?p_origin:ssrdir * constr -> + env -> evar_map -> + evar_map * constr -> + (constr -> evar_map -> bool) -> + ssrdir -> constr -> + evar_map * tpattern + +(** [findP env t i k] is a stateful function that finds the next occurrence + of a tpattern and calls the callback [k] to map the subterm matched. + The [int] argument passed to [k] is the number of binders traversed so far + plus the initial value [i]. + @return [t] where the subterms identified by the selected occurrences of + the patter have been mapped using [k] + @raise NoMatch if the raise_NoMatch flag given to [mk_tpattern_matcher] is + [true] and if the pattern did not match + @raise UserEerror if the raise_NoMatch flag given to [mk_tpattern_matcher] is + [false] and if the pattern did not match *) +type find_P = + env -> constr -> int -> k:subst -> constr + +(** [conclude ()] asserts that all mentioned ocurrences have been visited. + @return the instance of the pattern, the evarmap after the pattern + instantiation, the proof term and the ssrdit stored in the tpattern + @raise UserEerror if too many occurrences were specified *) +type conclude = + unit -> constr * ssrdir * (evar_map * Evd.evar_universe_context * constr) + +(** [mk_tpattern_matcher b o sigma0 occ sigma_tplist] creates a pair + a function [find_P] and [conclude] with the behaviour explained above. + The flag [b] (default [false]) changes the error reporting behaviour + of [find_P] if none of the [tpattern] matches. The argument [o] can + be passed to tune the [UserError] eventually raised (useful if the + pattern is coming from the LHS/RHS of an equation) *) +val mk_tpattern_matcher : + ?all_instances:bool -> + ?raise_NoMatch:bool -> + ?upats_origin:ssrdir * constr -> + evar_map -> occ -> evar_map * tpattern list -> + find_P * conclude + +(** Example of [mk_tpattern_matcher] to implement + [rewrite \{occ\}\[in t\]rules]. + It first matches "in t" (called [pat]), then in all matched subterms + it matches the LHS of the rules using [find_R]. + [concl0] is the initial goal, [concl] will be the goal where some terms + are replaced by a De Bruijn index. The [rw_progress] extra check + selects only occurrences that are not rewritten to themselves (e.g. + an occurrence "x + x" rewritten with the commutativity law of addition + is skipped) {[ + let find_R, conclude = match pat with + | Some (_, In_T _) -> + let aux (sigma, pats) (d, r, lhs, rhs) = + let sigma, pat = + mk_tpattern env0 sigma0 (sigma, r) (rw_progress rhs) d lhs in + sigma, pats @ [pat] in + let rpats = List.fold_left aux (r_sigma, []) rules in + let find_R, end_R = mk_tpattern_matcher sigma0 occ rpats in + find_R ~k:(fun _ _ h -> mkRel h), + fun cl -> let rdx, d, r = end_R () in (d,r),rdx + | _ -> ... in + let concl = eval_pattern env0 sigma0 concl0 pat occ find_R in + let (d, r), rdx = conclude concl in ]} *) + +(* convenience shortcut: [pf_fill_occ_term gl occ (sigma,t)] returns + * the conclusion of [gl] where [occ] occurrences of [t] have been replaced + * by [Rel 1] and the instance of [t] *) +val pf_fill_occ_term : + goal Tacmach.sigma -> occ -> evar_map * constr -> constr * constr + +(* It may be handy to inject a simple term into the first form of cpattern *) +val cpattern_of_term : char * glob_constr_and_expr -> cpattern + +(** Helpers to make stateful closures. Example: a [find_P] function may be + called many times, but the pattern instantiation phase is performed only the + first time. The corresponding [conclude] has to return the instantiated + pattern redex. Since it is up to [find_P] to raise [NoMatch] if the pattern + has no instance, [conclude] considers it an anomaly if the pattern did + not match *) + +(** [do_once r f] calls [f] and updates the ref only once *) +val do_once : 'a option ref -> (unit -> 'a) -> unit +(** [assert_done r] return the content of r. @raise Anomaly is r is [None] *) +val assert_done : 'a option ref -> 'a + +(** Very low level APIs. + these are calls to evarconv's [the_conv_x] followed by + [consider_remaining_unif_problems] and [resolve_typeclasses]. + In case of failure they raise [NoMatch] *) + +val unify_HO : env -> evar_map -> constr -> constr -> evar_map +val pf_unify_HO : goal Tacmach.sigma -> constr -> constr -> goal Tacmach.sigma + +(** Some more low level functions needed to implement the full SSR language + on top of the former APIs *) +val tag_of_cpattern : cpattern -> char +val loc_of_cpattern : cpattern -> Loc.t +val id_of_cpattern : cpattern -> Names.variable option +val is_wildcard : cpattern -> bool +val cpattern_of_id : Names.variable -> cpattern +val cpattern_of_id : Names.variable -> cpattern +val pr_constr_pat : constr -> Pp.std_ppcmds +val pf_merge_uc : Evd.evar_universe_context -> goal Evd.sigma -> goal Evd.sigma + +(* One can also "Set SsrMatchingDebug" from a .v *) +val debug : bool -> unit + +(* One should delimit a snippet with "Set SsrMatchingProfiling" and + * "Unset SsrMatchingProfiling" to get timings *) +val profile : bool -> unit + +(* eof *) diff --git a/mathcomp/ssreflect/seq.v b/mathcomp/ssreflect/seq.v index 8522017..15a93e8 100644 --- a/mathcomp/ssreflect/seq.v +++ b/mathcomp/ssreflect/seq.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat. (******************************************************************************) (* The seq type is the ssreflect type for sequences; it is an alias for the *) @@ -16,14 +18,14 @@ Require Import ssreflect ssrfun ssrbool eqtype ssrnat. (* by a notation. *) (* Here is the list of seq operations: *) (* ** Constructors: *) -(* seq T == the type of sequences with items of type T *) -(* bitseq == seq bool *) -(* [::], nil, Nil T == the empty sequence (of type T) *) -(* x :: s, cons x s, Cons T x s == the sequence x followed by s (of type T) *) -(* [:: x] == the singleton sequence *) -(* [:: x_0; ...; x_n] == the explicit sequence of the x_i *) -(* [:: x_0, ..., x_n & s] == the sequence of the x_i, followed by s *) -(* rcons s x == the sequence s, followed by x *) +(* seq T == the type of sequences of items of type T. *) +(* bitseq == seq bool. *) +(* [::], nil, Nil T == the empty sequence (of type T). *) +(* x :: s, cons x s, Cons T x s == the sequence x followed by s (of type T). *) +(* [:: x] == the singleton sequence. *) +(* [:: x_0; ...; x_n] == the explicit sequence of the x_i. *) +(* [:: x_0, ..., x_n & s] == the sequence of the x_i, followed by s. *) +(* rcons s x == the sequence s, followed by x. *) (* All of the above, except rcons, can be used in patterns. We define a view *) (* lastP and an induction principle last_ind that can be used to decompose *) (* or traverse a sequence in a right to left order. The view lemma lastP has *) @@ -31,21 +33,21 @@ Require Import ssreflect ssrfun ssrbool eqtype ssrnat. (* will generate two subgoals in which p has been replaced by [::] and by *) (* rcons p' x, respectively. *) (* ** Factories: *) -(* nseq n x == a sequence of n x's *) -(* ncons n x s == a sequence of n x's, followed by s *) -(* seqn n x_0 ... x_n-1 == the sequence of the x_i (can be partially applied) *) -(* iota m n == the sequence m, m + 1, ..., m + n - 1 *) -(* mkseq f n == the sequence f 0, f 1, ..., f (n - 1) *) +(* nseq n x == a sequence of n x's. *) +(* ncons n x s == a sequence of n x's, followed by s. *) +(* seqn n x_0 ... x_n-1 == the sequence of the x_i; can be partially applied. *) +(* iota m n == the sequence m, m + 1, ..., m + n - 1. *) +(* mkseq f n == the sequence f 0, f 1, ..., f (n - 1). *) (* ** Sequential access: *) -(* head x0 s == the head (zero'th item) of s if s is non-empty, else x0 *) -(* ohead s == None if s is empty, else Some x where x is the head of s *) -(* behead s == s minus its head, i.e., s' if s = x :: s', else [::] *) -(* last x s == the last element of x :: s (which is non-empty) *) -(* belast x s == x :: s minus its last item *) +(* head x0 s == the head (zero'th item) of s if s is non-empty, else x0. *) +(* ohead s == None if s is empty, else Some x when the head of s is x. *) +(* behead s == s minus its head, i.e., s' if s = x :: s', else [::]. *) +(* last x s == the last element of x :: s (which is non-empty). *) +(* belast x s == x :: s minus its last item. *) (* ** Dimensions: *) -(* size s == the number of items (length) in s *) +(* size s == the number of items (length) in s. *) (* shape ss == the sequence of sizes of the items of the sequence of *) -(* sequences ss *) +(* sequences ss. *) (* ** Random access: *) (* nth x0 s i == the item i of s (numbered from 0), or x0 if s does *) (* not have at least i+1 items (i.e., size x <= i) *) @@ -57,32 +59,32 @@ Require Import ssreflect ssrfun ssrbool eqtype ssrnat. (* incr_nth s i == the nat sequence s with item i incremented (s is *) (* first padded with 0's to size i+1, if needed). *) (* ** Predicates: *) -(* nilp s == s is [::] *) -(* := (size s == 0) *) -(* x \in s == x appears in s (this requires an eqType for T) *) +(* nilp s == s is [::]. *) +(* := (size s == 0). *) +(* x \in s == x appears in s (this requires an eqType for T). *) (* index x s == the first index at which x appears in s, or size s if *) -(* x \notin s *) +(* x \notin s. *) (* has p s == the (applicative, boolean) predicate p holds for some *) -(* item in s *) -(* all p s == p holds for all items in s *) +(* item in s. *) +(* all p s == p holds for all items in s. *) (* find p s == the index of the first item in s for which p holds, *) -(* or size s if no such item is found *) -(* count p s == the number of items of s for which p holds *) -(* count_mem x s == the number of times x occurs in s := count (pred1 x) s *) +(* or size s if no such item is found. *) +(* count p s == the number of items of s for which p holds. *) +(* count_mem x s == the number of times x occurs in s := count (pred1 x) s. *) (* constant s == all items in s are identical (trivial if s = [::]) *) -(* uniq s == all the items in s are pairwise different *) +(* uniq s == all the items in s are pairwise different. *) (* subseq s1 s2 == s1 is a subsequence of s2, i.e., s1 = mask m s2 for *) (* some m : bitseq (see below). *) (* perm_eq s1 s2 == s2 is a permutation of s1, i.e., s1 and s2 have the *) (* items (with the same repetitions), but possibly in a *) (* different order. *) -(* perm_eql s1 s2 <-> s1 and s2 behave identically on the left of perm_eq *) -(* perm_eqr s1 s2 <-> s1 and s2 behave identically on the rightt of perm_eq *) +(* perm_eql s1 s2 <-> s1 and s2 behave identically on the left of perm_eq. *) +(* perm_eqr s1 s2 <-> s1 and s2 behave identically on the right of perm_eq. *) (* These left/right transitive versions of perm_eq make it easier to chain *) (* a sequence of equivalences. *) (* ** Filtering: *) (* filter p s == the subsequence of s consisting of all the items *) -(* for which the (boolean) predicate p holds *) +(* for which the (boolean) predicate p holds. *) (* subfilter s : seq sT == when sT has a subType p structure, the sequence *) (* of items of type sT corresponding to items of s *) (* for which p holds. *) @@ -325,7 +327,7 @@ Fixpoint set_nth s n y {struct n} := Lemma nth0 s : nth s 0 = head s. Proof. by []. Qed. Lemma nth_default s n : size s <= n -> nth s n = x0. -Proof. by elim: s n => [|x s IHs] [|n]; try exact (IHs n). Qed. +Proof. by elim: s n => [|x s IHs] []. Qed. Lemma nth_nil n : nth [::] n = x0. Proof. by case: n. Qed. @@ -341,26 +343,26 @@ Proof. by case: s n => [|x s] [|n]. Qed. Lemma nth_cat s1 s2 n : nth (s1 ++ s2) n = if n < size s1 then nth s1 n else nth s2 (n - size s1). -Proof. by elim: s1 n => [|x s1 IHs] [|n]; try exact (IHs n). Qed. +Proof. by elim: s1 n => [|x s1 IHs] []. Qed. Lemma nth_rcons s x n : nth (rcons s x) n = if n < size s then nth s n else if n == size s then x else x0. -Proof. by elim: s n => [|y s IHs] [|n] //=; rewrite ?nth_nil ?IHs. Qed. +Proof. by elim: s n => [|y s IHs] [] //=; apply: nth_nil. Qed. Lemma nth_ncons m x s n : nth (ncons m x s) n = if n < m then x else nth s (n - m). -Proof. by elim: m n => [|m IHm] [|n] //=; exact: IHm. Qed. +Proof. by elim: m n => [|m IHm] []. Qed. Lemma nth_nseq m x n : nth (nseq m x) n = (if n < m then x else x0). -Proof. by elim: m n => [|m IHm] [|n] //=; exact: IHm. Qed. +Proof. by elim: m n => [|m IHm] []. Qed. Lemma eq_from_nth s1 s2 : size s1 = size s2 -> (forall i, i < size s1 -> nth s1 i = nth s2 i) -> s1 = s2. Proof. elim: s1 s2 => [|x1 s1 IHs1] [|x2 s2] //= [eq_sz] eq_s12. -rewrite [x1](eq_s12 0) // (IHs1 s2) // => i; exact: (eq_s12 i.+1). +by rewrite [x1](eq_s12 0) // (IHs1 s2) // => i; apply: (eq_s12 i.+1). Qed. Lemma size_set_nth s n y : size (set_nth s n y) = maxn n.+1 (size s). @@ -416,7 +418,7 @@ Lemma has_count s : has s = (0 < count s). Proof. by elim: s => //= x s ->; case (a x). Qed. Lemma count_size s : count s <= size s. -Proof. by elim: s => //= x s; case: (a x); last exact: leqW. Qed. +Proof. by elim: s => //= x s; case: (a x); last apply: leqW. Qed. Lemma all_count s : all s = (count s == size s). Proof. @@ -434,7 +436,7 @@ by elim: s => //= x s IHs /andP[-> Hs]; rewrite IHs. Qed. Lemma filter_id s : filter (filter s) = filter s. -Proof. by apply/all_filterP; exact: filter_all. Qed. +Proof. by apply/all_filterP; apply: filter_all. Qed. Lemma has_find s : has s = (find s < size s). Proof. by elim: s => //= x s IHs; case (a x); rewrite ?leqnn. Qed. @@ -529,7 +531,7 @@ Lemma sub_find s : find a2 s <= find a1 s. Proof. by elim: s => //= x s IHs; case: ifP => // /(contraFF (@s12 x))->. Qed. Lemma sub_has s : has a1 s -> has a2 s. -Proof. by rewrite !has_find; exact: leq_ltn_trans (sub_find s). Qed. +Proof. by rewrite !has_find; apply: leq_ltn_trans (sub_find s). Qed. Lemma sub_count s : count a1 s <= count a2 s. Proof. @@ -618,7 +620,7 @@ Lemma drop0 s : drop 0 s = s. Proof. by case: s. Qed. Lemma drop1 : drop 1 =1 behead. Proof. by case=> [|x [|y s]]. Qed. Lemma drop_oversize n s : size s <= n -> drop n s = [::]. -Proof. by elim: n s => [|n IHn] [|x s]; try exact (IHn s). Qed. +Proof. by elim: s n => [|x s IHs] []. Qed. Lemma drop_size s : drop (size s) s = [::]. Proof. by rewrite drop_oversize // leqnn. Qed. @@ -628,18 +630,18 @@ Lemma drop_cons x s : Proof. by []. Qed. Lemma size_drop s : size (drop n0 s) = size s - n0. -Proof. by elim: s n0 => [|x s IHs] [|n]; try exact (IHs n). Qed. +Proof. by elim: s n0 => [|x s IHs] []. Qed. Lemma drop_cat s1 s2 : drop n0 (s1 ++ s2) = if n0 < size s1 then drop n0 s1 ++ s2 else drop (n0 - size s1) s2. -Proof. by elim: s1 n0 => [|x s1 IHs] [|n]; try exact (IHs n). Qed. +Proof. by elim: s1 n0 => [|x s1 IHs] []. Qed. Lemma drop_size_cat n s1 s2 : size s1 = n -> drop n (s1 ++ s2) = s2. Proof. by move <-; elim: s1 => //=; rewrite drop0. Qed. Lemma nconsK n x : cancel (ncons n x) (drop n). -Proof. by elim: n => //; case. Qed. +Proof. by elim: n => // -[]. Qed. Fixpoint take n s {struct s} := match s, n with @@ -650,7 +652,7 @@ Fixpoint take n s {struct s} := Lemma take0 s : take 0 s = [::]. Proof. by case: s. Qed. Lemma take_oversize n s : size s <= n -> take n s = s. -Proof. by elim: n s => [|n IHn] [|x s] Hsn; try (congr cons; apply: IHn). Qed. +Proof. by elim: s n => [|x s IHs] [|n] //= /IHs->. Qed. Lemma take_size s : take (size s) s = s. Proof. by rewrite take_oversize // leqnn. Qed. @@ -661,14 +663,14 @@ Proof. by []. Qed. Lemma drop_rcons s : n0 <= size s -> forall x, drop n0 (rcons s x) = rcons (drop n0 s) x. -Proof. by elim: s n0 => [|y s IHs] [|n]; try exact (IHs n). Qed. +Proof. by elim: s n0 => [|y s IHs] []. Qed. Lemma cat_take_drop s : take n0 s ++ drop n0 s = s. -Proof. by elim: s n0 => [|x s IHs] [|n]; try exact (congr1 _ (IHs n)). Qed. +Proof. by elim: s n0 => [|x s IHs] [|n] //=; rewrite IHs. Qed. Lemma size_takel s : n0 <= size s -> size (take n0 s) = n0. Proof. -by move/subnKC; rewrite -{2}(cat_take_drop s) size_cat size_drop => /addIn. +by move/subKn; rewrite -size_drop -[in size s](cat_take_drop s) size_cat addnK. Qed. Lemma size_take s : size (take n0 s) = if n0 < size s then n0 else size s. @@ -950,7 +952,7 @@ Lemma mem_last x s : last x s \in x :: s. Proof. by rewrite lastI mem_rcons mem_head. Qed. Lemma mem_behead s : {subset behead s <= s}. -Proof. by case: s => // y s x; exact: predU1r. Qed. +Proof. by case: s => // y s x; apply: predU1r. Qed. Lemma mem_belast s y : {subset belast y s <= y :: s}. Proof. by move=> x ys'x; rewrite lastI mem_rcons mem_behead. Qed. @@ -982,23 +984,23 @@ Lemma hasPn s : reflect (forall x, x \in s -> ~~ a x) (~~ has a s). Proof. apply: (iffP idP) => not_a_s => [x s_x|]. by apply: contra not_a_s => a_x; apply/hasP; exists x. -by apply/hasP=> [[x s_x]]; apply/negP; exact: not_a_s. +by apply/hasP=> [[x s_x]]; apply/negP; apply: not_a_s. Qed. Lemma allP s : reflect (forall x, x \in s -> a x) (all a s). Proof. elim: s => [|x s IHs]; first by left. rewrite /= andbC; case: IHs => IHs /=. - apply: (iffP idP) => [Hx y|]; last by apply; exact: mem_head. + apply: (iffP idP) => [Hx y|]; last by apply; apply: mem_head. by case/predU1P=> [->|Hy]; auto. -by right=> H; case IHs => y Hy; apply H; exact: mem_behead. +by right=> H; case IHs => y Hy; apply H; apply: mem_behead. Qed. Lemma allPn s : reflect (exists2 x, x \in s & ~~ a x) (~~ all a s). Proof. elim: s => [|x s IHs]; first by right=> [[x Hx _]]. rewrite /= andbC negb_and; case: IHs => IHs /=. - by left; case: IHs => y Hy Hay; exists y; first exact: mem_behead. + by left; case: IHs => y Hy Hay; exists y; first apply: mem_behead. apply: (iffP idP) => [|[y]]; first by exists x; rewrite ?mem_head. by case/predU1P=> [-> // | s_y not_a_y]; case: IHs; exists y. Qed. @@ -1019,7 +1021,7 @@ Variables a1 a2 : pred T. Lemma eq_in_filter s : {in s, a1 =1 a2} -> filter a1 s = filter a2 s. Proof. elim: s => //= x s IHs eq_a. -rewrite eq_a ?mem_head ?IHs // => y s_y; apply: eq_a; exact: mem_behead. +by rewrite eq_a ?mem_head ?IHs // => y s_y; apply: eq_a; apply: mem_behead. Qed. Lemma eq_in_find s : {in s, a1 =1 a2} -> find a1 s = find a2 s. @@ -1154,7 +1156,7 @@ Fixpoint undup s := if s is x :: s' then if x \in s' then undup s' else x :: undup s' else [::]. Lemma size_undup s : size (undup s) <= size s. -Proof. by elim: s => //= x s IHs; case: (x \in s) => //=; exact: ltnW. Qed. +Proof. by elim: s => //= x s IHs; case: (x \in s) => //=; apply: ltnW. Qed. Lemma mem_undup s : undup s =i s. Proof. @@ -1178,7 +1180,7 @@ Qed. Lemma filter_undup p s : filter p (undup s) = undup (filter p s). Proof. elim: s => //= x s IHs; rewrite (fun_if undup) fun_if /= mem_filter /=. -by rewrite (fun_if (filter p)) /= IHs; case: ifP => -> //=; exact: if_same. +by rewrite (fun_if (filter p)) /= IHs; case: ifP => -> //=; apply: if_same. Qed. Lemma undup_nil s : undup s = [::] -> s = [::]. @@ -1228,7 +1230,7 @@ Lemma mem_rot s : rot n0 s =i s. Proof. by move=> x; rewrite -{2}(cat_take_drop n0 s) !mem_cat /= orbC. Qed. Lemma eqseq_rot s1 s2 : (rot n0 s1 == rot n0 s2) = (s1 == s2). -Proof. by apply: inj_eq; exact: rot_inj. Qed. +Proof. by apply: inj_eq; apply: rot_inj. Qed. CoInductive rot_to_spec s x := RotToSpec i s' of rot i s = x :: s'. @@ -1316,7 +1318,7 @@ Lemma size_incr_nth v i : size (incr_nth v i) = if i < size v then size v else i.+1. Proof. elim: v i => [|n v IHv] [|i] //=; first by rewrite size_ncons /= addn1. -rewrite IHv; exact: fun_if. +by rewrite IHv; apply: fun_if. Qed. (* Equality up to permutation *) @@ -1385,7 +1387,7 @@ Lemma perm_cons x s1 s2 : perm_eq (x :: s1) (x :: s2) = perm_eq s1 s2. Proof. exact: (perm_cat2l [::x]). Qed. Lemma perm_cat2r s1 s2 s3 : perm_eq (s2 ++ s1) (s3 ++ s1) = perm_eq s2 s3. -Proof. by do 2!rewrite perm_eq_sym perm_catC; exact: perm_cat2l. Qed. +Proof. by do 2!rewrite perm_eq_sym perm_catC; apply: perm_cat2l. Qed. Lemma perm_catAC s1 s2 s3 : perm_eql ((s1 ++ s2) ++ s3) ((s1 ++ s3) ++ s2). Proof. by apply/perm_eqlP; rewrite -!catA perm_cat2l perm_catC. Qed. @@ -1437,7 +1439,7 @@ rewrite -(rot_uniq i) -(size_rot i) def_s2 /= in le_s21 *. have ss13 y (s1y : y \in s1): y \in s3. by have:= ss12 y s1y; rewrite -(mem_rot i) def_s2 inE (negPf (memPn _ y s1y)). rewrite IHs // andbT; apply: contraL _ le_s21 => s3x; rewrite -leqNgt. -by apply/(uniq_leq_size Us1x)/allP; rewrite /= s3x; exact/allP. +by apply/(uniq_leq_size Us1x)/allP; rewrite /= s3x; apply/allP. Qed. Lemma uniq_size_uniq s1 s2 : @@ -1541,7 +1543,7 @@ Proof. by rewrite rot_uniq. Qed. Lemma rotrK : cancel (@rotr T n0) (rot n0). Proof. move=> s; have [lt_n0s | ge_n0s] := ltnP n0 (size s). - by rewrite -{1}(subKn (ltnW lt_n0s)) -{1}[size s]size_rotr; exact: rotK. + by rewrite -{1}(subKn (ltnW lt_n0s)) -{1}[size s]size_rotr; apply: rotK. by rewrite -{2}(rot_oversize ge_n0s) /rotr (eqnP ge_n0s) rot0. Qed. @@ -1676,7 +1678,7 @@ Lemma mask_uniq s : uniq s -> forall m, uniq (mask m s). Proof. elim: s => [|x s IHs] Uxs [|b m] //=. case: b Uxs => //= /andP[s'x Us]; rewrite {}IHs // andbT. -by apply: contra s'x; exact: mem_mask. +by apply: contra s'x; apply: mem_mask. Qed. Lemma mem_mask_rot m s : @@ -1756,13 +1758,13 @@ by exists (m1 ++ m2); rewrite ?size_cat ?mask_cat ?sz_m1 ?sz_m2. Qed. Lemma prefix_subseq s1 s2 : subseq s1 (s1 ++ s2). -Proof. by rewrite -{1}[s1]cats0 cat_subseq ?sub0seq. Qed. +Proof. by rewrite -[s1 in subseq s1]cats0 cat_subseq ?sub0seq. Qed. Lemma suffix_subseq s1 s2 : subseq s2 (s1 ++ s2). -Proof. by rewrite -{1}[s2]cat0s cat_subseq ?sub0seq. Qed. +Proof. exact: cat_subseq (sub0seq s1) _. Qed. Lemma mem_subseq s1 s2 : subseq s1 s2 -> {subset s1 <= s2}. -Proof. by case/subseqP=> m _ -> x; exact: mem_mask. Qed. +Proof. by case/subseqP=> m _ -> x; apply: mem_mask. Qed. Lemma sub1seq x s : subseq [:: x] s = (x \in s). Proof. @@ -1782,13 +1784,13 @@ by move/(@all_pred1P _ true)->; rewrite sz_m mask_true. Qed. Lemma subseq_cons s x : subseq s (x :: s). -Proof. by exact: (@cat_subseq [::] s [:: x]). Qed. +Proof. exact: suffix_subseq [:: x] s. Qed. Lemma subseq_rcons s x : subseq s (rcons s x). -Proof. by rewrite -{1}[s]cats0 -cats1 cat_subseq. Qed. +Proof. by rewrite -cats1 prefix_subseq. Qed. Lemma subseq_uniq s1 s2 : subseq s1 s2 -> uniq s2 -> uniq s1. -Proof. by case/subseqP=> m _ -> Us2; exact: mask_uniq. Qed. +Proof. by case/subseqP=> m _ -> Us2; apply: mask_uniq. Qed. End Subseq. @@ -1821,11 +1823,11 @@ Proof. by move/perm_to_rem/perm_eq_size->. Qed. Lemma rem_subseq s : subseq (rem s) s. Proof. elim: s => //= y s IHs; rewrite eq_sym. -by case: ifP => _; [exact: subseq_cons | rewrite eqxx]. +by case: ifP => _; [apply: subseq_cons | rewrite eqxx]. Qed. Lemma rem_uniq s : uniq s -> uniq (rem s). -Proof. by apply: subseq_uniq; exact: rem_subseq. Qed. +Proof. by apply: subseq_uniq; apply: rem_subseq. Qed. Lemma mem_rem s : {subset rem s <= s}. Proof. exact: mem_subseq (rem_subseq s). Qed. @@ -1864,7 +1866,7 @@ Lemma behead_map s : behead (map s) = map (behead s). Proof. by case: s. Qed. Lemma nth_map n s : n < size s -> nth x2 (map s) n = f (nth x1 s n). -Proof. by elim: s n => [|x s IHs] [|n]; try exact (IHs n). Qed. +Proof. by elim: s n => [|x s IHs] []. Qed. Lemma map_rcons s x : map (rcons s x) = rcons (map s) (f x). Proof. by rewrite -!cats1 map_cat. Qed. @@ -1955,7 +1957,7 @@ move=> uniq_s2; apply: (iffP idP) => [ss12 | ->]; last exact: filter_subseq. apply/eqP; rewrite -size_subseq_leqif ?subseq_filter ?(introT allP) //. apply/eqP/esym/perm_eq_size. rewrite uniq_perm_eq ?filter_uniq ?(subseq_uniq ss12) // => x. -by rewrite mem_filter; apply: andb_idr; exact: (mem_subseq ss12). +by rewrite mem_filter; apply: andb_idr; apply: (mem_subseq ss12). Qed. Lemma perm_to_subseq s1 s2 : @@ -1980,7 +1982,7 @@ Implicit Type s : seq T1. Lemma map_f s x : x \in s -> f x \in map f s. Proof. elim: s => [|y s IHs] //=. -by case/predU1P=> [->|Hx]; [exact: predU1l | apply: predU1r; auto]. +by case/predU1P=> [->|Hx]; [apply: predU1l | apply: predU1r; auto]. Qed. Lemma mapP s y : reflect (exists2 x, x \in s & y = f x) (y \in map f s). @@ -1989,7 +1991,7 @@ elim: s => [|x s IHs]; first by right; case. rewrite /= in_cons eq_sym; case Hxy: (f x == y). by left; exists x; [rewrite mem_head | rewrite (eqP Hxy)]. apply: (iffP IHs) => [[x' Hx' ->]|[x' Hx' Dy]]. - by exists x'; first exact: predU1r. + by exists x'; first apply: predU1r. by move: Dy Hxy => ->; case/predU1P: Hx' => [->|]; [rewrite eqxx | exists x']. Qed. @@ -2004,7 +2006,7 @@ Proof. elim: s => //= x s IHs //= injf; congr (~~ _ && _). apply/mapP/idP=> [[y sy /injf] | ]; last by exists x. by rewrite mem_head mem_behead // => ->. -apply: IHs => y z sy sz; apply: injf => //; exact: predU1r. +by apply: IHs => y z sy sz; apply: injf => //; apply: predU1r. Qed. Lemma map_subseq s1 s2 : subseq s1 s2 -> subseq (map f s1) (map f s2). @@ -2018,7 +2020,7 @@ Lemma nth_index_map s x0 x : Proof. elim: s => //= y s IHs inj_f s_x; rewrite (inj_in_eq inj_f) ?mem_head //. move: s_x; rewrite inE eq_sym; case: eqP => [-> | _] //=; apply: IHs. -by apply: sub_in2 inj_f => z; exact: predU1r. +by apply: sub_in2 inj_f => z; apply: predU1r. Qed. Lemma perm_map s t : perm_eq s t -> perm_eq (map f s) (map f t). @@ -2033,7 +2035,7 @@ Lemma index_map s x : index (f x) (map f s) = index x s. Proof. by rewrite /index; elim: s => //= y s IHs; rewrite (inj_eq Hf) IHs. Qed. Lemma map_inj_uniq s : uniq (map f s) = uniq s. -Proof. apply: map_inj_in_uniq; exact: in2W. Qed. +Proof. by apply: map_inj_in_uniq; apply: in2W. Qed. End EqMap. @@ -2045,7 +2047,7 @@ Lemma map_of_seq (T1 : eqType) T2 (s : seq T1) (fs : seq T2) (y0 : T2) : Proof. exists (fun x => nth y0 fs (index x s)) => uAs eq_sz. apply/esym/(@eq_from_nth _ y0); rewrite ?size_map eq_sz // => i ltis. -have x0 : T1 by [case: (s) ltis]; by rewrite (nth_map x0) // index_uniq. +by have x0 : T1 by [case: (s) ltis]; rewrite (nth_map x0) // index_uniq. Qed. Section MapComp. @@ -2124,7 +2126,7 @@ Qed. Lemma pmap_uniq s : uniq s -> uniq (pmap f s). Proof. -by move/(filter_uniq [eta f]); rewrite -(pmap_filter fK); exact: map_uniq. +by move/(filter_uniq [eta f]); rewrite -(pmap_filter fK); apply: map_uniq. Qed. End EqPmap. @@ -2145,7 +2147,7 @@ Variables (T : eqType) (p : pred T) (sT : subType p). Let insT : T -> option sT := insub. Lemma mem_pmap_sub s u : (u \in pmap insT s) = (val u \in s). -Proof. apply: (can2_mem_pmap (insubK _)); exact: valK. Qed. +Proof. exact/(can2_mem_pmap (insubK _))/valK. Qed. Lemma pmap_sub_uniq s : uniq s -> uniq (pmap insT s). Proof. exact: (pmap_uniq (insubK _)). Qed. @@ -2176,7 +2178,7 @@ Lemma mem_iota m n i : (i \in iota m n) = (m <= i) && (i < m + n). Proof. elim: n m => [|n IHn] /= m; first by rewrite addn0 ltnNge andbN. rewrite -addSnnS leq_eqVlt in_cons eq_sym. -by case: eqP => [->|_]; [rewrite leq_addr | exact: IHn]. +by case: eqP => [->|_]; [rewrite leq_addr | apply: IHn]. Qed. Lemma iota_uniq m n : uniq (iota m n). @@ -2194,7 +2196,7 @@ Lemma size_mkseq f n : size (mkseq f n) = n. Proof. by rewrite size_map size_iota. Qed. Lemma eq_mkseq f g : f =1 g -> mkseq f =1 mkseq g. -Proof. by move=> Efg n; exact: eq_map Efg _. Qed. +Proof. by move=> Efg n; apply: eq_map Efg _. Qed. Lemma nth_mkseq f n i : i < n -> nth x0 (mkseq f n) i = f i. Proof. by move=> Hi; rewrite (nth_map 0) ?nth_iota ?size_iota. Qed. @@ -2538,7 +2540,7 @@ Lemma allpairs_uniq R (f : S -> T -> R) s t : {in [seq (x, y) | x <- s, y <- t] &, injective (prod_curry f)} -> uniq (allpairs f s t). Proof. -move=> Us Ut inj_f; have: all (mem s) s by exact/allP. +move=> Us Ut inj_f; have: all (mem s) s by apply/allP. elim: {-2}s Us => //= x s1 IHs /andP[s1'x Us1] /andP[sx1 ss1]. rewrite cat_uniq {}IHs // andbT map_inj_in_uniq ?Ut // => [|y1 y2 *]. apply/hasPn=> _ /allpairsP[z [s1z tz ->]]; apply/mapP=> [[y ty Dy]]. diff --git a/mathcomp/ssreflect/ssrbool.v b/mathcomp/ssreflect/ssrbool.v index aeaa266..4b24c32 100644 --- a/mathcomp/ssreflect/ssrbool.v +++ b/mathcomp/ssreflect/ssrbool.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun. (******************************************************************************) (* A theory of boolean predicates and operators. A large part of this file is *) @@ -37,15 +39,27 @@ Require Import ssreflect ssrfun. (* boolP my_formula generates two subgoals with *) (* assumtions my_formula and ~~ myformula. As *) (* with altP, my_formula must be an application. *) -(* unless C P <-> hP : P may be assumed when proving P. *) -(* := (P -> C) -> C (Pierce's law). *) -(* This is slightly weaker but easier to use than *) -(* P \/ C when P C : Prop. *) -(* classically P <-> hP : P can be assumed when proving is_true b *) +(* \unless C, P <-> we can assume property P when a something that *) +(* holds under condition C (such as C itself). *) +(* := forall G : Prop, (C -> G) -> (P -> G) -> G. *) +(* This is just C \/ P or rather its impredicative *) +(* encoding, whose usage better fits the above *) +(* description: given a lemma UCP whose conclusion *) +(* is \unless C, P we can assume P by writing: *) +(* wlog hP: / P by apply/UCP; (prove C -> goal). *) +(* or even apply: UCP id _ => hP if the goal is C. *) +(* classically P <-> we can assume P when proving is_true b. *) (* := forall b : bool, (P -> b) -> b. *) (* This is equivalent to ~ (~ P) when P : Prop. *) +(* implies P Q == wrapper coinductive type that coerces to P -> Q *) +(* and can be used as a P -> Q view unambigously. *) +(* Useful to avoid spurious insertion of <-> views *) +(* when Q is a conjunction of foralls, as in Lemma *) +(* all_and2 below; conversely, avoids confusion in *) +(* apply views for impredicative properties, such *) +(* as \unless C, P. Also supports contrapositives. *) (* a && b == the boolean conjunction of a and b. *) -(* a || b == then boolean disjunction of a and b. *) +(* a || b == the boolean disjunction of a and b. *) (* a ==> b == the boolean implication of b by a. *) (* ~~ a == the boolean negation of a. *) (* a (+) b == the boolean exclusive or (or sum) of a and b. *) @@ -507,10 +521,10 @@ Lemma elimTFn : b = c -> if c then ~ P else P. Proof. by move <-; apply: (elimNTF Hb); case b. Qed. Lemma equivPifn : (Q -> P) -> (P -> Q) -> if b then ~ Q else Q. -Proof. rewrite -if_neg; exact: equivPif. Qed. +Proof. by rewrite -if_neg; apply: equivPif. Qed. Lemma xorPifn : Q \/ P -> ~ (Q /\ P) -> if b then Q else ~ Q. -Proof. rewrite -if_neg; exact: xorPif. Qed. +Proof. by rewrite -if_neg; apply: xorPif. Qed. End ReflectNegCore. @@ -541,7 +555,7 @@ Lemma iffP : (P -> Q) -> (Q -> P) -> reflect Q b. Proof. by case: Pb; constructor; auto. Qed. Lemma equivP : (P <-> Q) -> reflect Q b. -Proof. by case; exact: iffP. Qed. +Proof. by case; apply: iffP. Qed. Lemma sumboolP (decQ : decidable Q) : reflect Q decQ. Proof. by case: decQ; constructor. Qed. @@ -550,16 +564,16 @@ Lemma appP : reflect Q b -> P -> Q. Proof. by move=> Qb; move/introT; case: Qb. Qed. Lemma sameP : reflect P c -> b = c. -Proof. case; [exact: introT | exact: introF]. Qed. +Proof. by case; [apply: introT | apply: introF]. Qed. Lemma decPcases : if b then P else ~ P. Proof. by case Pb. Qed. Definition decP : decidable P. by case: b decPcases; [left | right]. Defined. -Lemma rwP : P <-> b. Proof. by split; [exact: introT | exact: elimT]. Qed. +Lemma rwP : P <-> b. Proof. by split; [apply: introT | apply: elimT]. Qed. Lemma rwP2 : reflect Q b -> (P <-> Q). -Proof. by move=> Qb; split=> ?; [exact: appP | apply: elimT; case: Qb]. Qed. +Proof. by move=> Qb; split=> ?; [apply: appP | apply: elimT; case: Qb]. Qed. (* Predicate family to reflect excluded middle in bool. *) CoInductive alt_spec : bool -> Type := @@ -580,50 +594,75 @@ Hint View for apply// equivPif|3 xorPif|3 equivPifn|3 xorPifn|3. (* Allow the direct application of a reflection lemma to a boolean assertion. *) Coercion elimT : reflect >-> Funclass. -(* Pierce's law, a weak form of classical reasoning. *) -Definition unless condition property := (property -> condition) -> condition. +CoInductive implies P Q := Implies of P -> Q. +Lemma impliesP P Q : implies P Q -> P -> Q. Proof. by case. Qed. +Lemma impliesPn (P Q : Prop) : implies P Q -> ~ Q -> ~ P. +Proof. by case=> iP ? /iP. Qed. +Coercion impliesP : implies >-> Funclass. +Hint View for move/ impliesPn|2 impliesP|2. +Hint View for apply/ impliesPn|2 impliesP|2. -Lemma bind_unless C P {Q} : unless C P -> unless (unless C Q) P. -Proof. by move=> haveP suffPQ suffQ; apply: haveP => /suffPQ; exact. Qed. +(* Impredicative or, which can emulate a classical not-implies. *) +Definition unless condition property : Prop := + forall goal : Prop, (condition -> goal) -> (property -> goal) -> goal. -Lemma unless_contra b C : (~~ b -> C) -> unless C b. -Proof. by case: b => [_ haveC | haveC _]; exact: haveC. Qed. +Notation "\unless C , P" := (unless C P) + (at level 200, C at level 100, + format "'[' \unless C , '/ ' P ']'") : type_scope. + +Lemma unlessL C P : implies C (\unless C, P). +Proof. by split=> hC G /(_ hC). Qed. + +Lemma unlessR C P : implies P (\unless C, P). +Proof. by split=> hP G _ /(_ hP). Qed. + +Lemma unless_sym C P : implies (\unless C, P) (\unless P, C). +Proof. by split; apply; [apply/unlessR | apply/unlessL]. Qed. + +Lemma unlessP (C P : Prop) : (\unless C, P) <-> C \/ P. +Proof. by split=> [|[/unlessL | /unlessR]]; apply; [left | right]. Qed. + +Lemma bind_unless C P {Q} : implies (\unless C, P) (\unless (\unless C, Q), P). +Proof. by split; apply=> [hC|hP]; [apply/unlessL/unlessL | apply/unlessR]. Qed. + +Lemma unless_contra b C : implies (~~ b -> C) (\unless C, b). +Proof. by split; case: b => [_ | hC]; [apply/unlessR | apply/unlessL/hC]. Qed. (* Classical reasoning becomes directly accessible for any bool subgoal. *) (* Note that we cannot use "unless" here for lack of universe polymorphism. *) Definition classically P : Prop := forall b : bool, (P -> b) -> b. -Lemma classicP : forall P : Prop, classically P <-> ~ ~ P. +Lemma classicP (P : Prop) : classically P <-> ~ ~ P. Proof. -move=> P; split=> [cP nP | nnP [] // nP]; last by case nnP; move/nP. +split=> [cP nP | nnP [] // nP]; last by case nnP; move/nP. by have: P -> false; [move/nP | move/cP]. Qed. -Lemma classic_bind : forall P Q, - (P -> classically Q) -> (classically P -> classically Q). -Proof. by move=> P Q IH IH_P b IH_Q; apply: IH_P; move/IH; exact. Qed. +Lemma classicW P : P -> classically P. Proof. by move=> hP _ ->. Qed. -Lemma classic_EM : forall P, classically (decidable P). +Lemma classic_bind P Q : (P -> classically Q) -> classically P -> classically Q. +Proof. by move=> iPQ cP b /iPQ-/cP. Qed. + +Lemma classic_EM P : classically (decidable P). Proof. -by move=> P [] // IH; apply IH; right => ?; apply: notF (IH _); left. +by case=> // undecP; apply/undecP; right=> notP; apply/notF/undecP; left. Qed. -Lemma classic_imply : forall P Q, (P -> classically Q) -> classically (P -> Q). +Lemma classic_pick T P : classically ({x : T | P x} + (forall x, ~ P x)). Proof. -move=> P Q IH [] // notPQ; apply notPQ; move/IH=> hQ; case: notF. -by apply: hQ => hQ; case: notF; exact: notPQ. +case=> // undecP; apply/undecP; right=> x Px. +by apply/notF/undecP; left; exists x. Qed. -Lemma classic_pick : forall T P, - classically ({x : T | P x} + (forall x, ~ P x)). +Lemma classic_imply P Q : (P -> classically Q) -> classically (P -> Q). Proof. -move=> T P [] // IH; apply IH; right=> x Px; case: notF. -by apply: IH; left; exists x. +move=> iPQ []// notPQ; apply/notPQ=> /iPQ-cQ. +by case: notF; apply: cQ => hQ; apply: notPQ. Qed. (* List notations for wider connectives; the Prop connectives have a fixed *) (* width so as to avoid iterated destruction (we go up to width 5 for /\, and *) -(* width 4 for or. The bool connectives have arbitrary widths, but denote *) +(* width 4 for or). The bool connectives have arbitrary widths, but denote *) (* expressions that associate to the RIGHT. This is consistent with the right *) (* associativity of list expressions and thus more convenient in most proofs. *) @@ -665,23 +704,28 @@ Section AllAnd. Variables (T : Type) (P1 P2 P3 P4 P5 : T -> Prop). Local Notation a P := (forall x, P x). -Lemma all_and2 (hP : forall x, [/\ P1 x & P2 x]) : [/\ a P1 & a P2]. -Proof. by split=> x; case: (hP x). Qed. +Lemma all_and2 : implies (forall x, [/\ P1 x & P2 x]) [/\ a P1 & a P2]. +Proof. by split=> haveP; split=> x; case: (haveP x). Qed. -Lemma all_and3 (hP : forall x, [/\ P1 x, P2 x & P3 x]) : - [/\ a P1, a P2 & a P3]. -Proof. by split=> x; case: (hP x). Qed. +Lemma all_and3 : implies (forall x, [/\ P1 x, P2 x & P3 x]) + [/\ a P1, a P2 & a P3]. +Proof. by split=> haveP; split=> x; case: (haveP x). Qed. -Lemma all_and4 (hP : forall x, [/\ P1 x, P2 x, P3 x & P4 x]) : - [/\ a P1, a P2, a P3 & a P4]. -Proof. by split=> x; case: (hP x). Qed. +Lemma all_and4 : implies (forall x, [/\ P1 x, P2 x, P3 x & P4 x]) + [/\ a P1, a P2, a P3 & a P4]. +Proof. by split=> haveP; split=> x; case: (haveP x). Qed. -Lemma all_and5 (hP : forall x, [/\ P1 x, P2 x, P3 x, P4 x & P5 x]) : - [/\ a P1, a P2, a P3, a P4 & a P5]. -Proof. by split=> x; case: (hP x). Qed. +Lemma all_and5 : implies (forall x, [/\ P1 x, P2 x, P3 x, P4 x & P5 x]) + [/\ a P1, a P2, a P3, a P4 & a P5]. +Proof. by split=> haveP; split=> x; case: (haveP x). Qed. End AllAnd. +Implicit Arguments all_and2 [[T] [P1] [P2]]. +Implicit Arguments all_and3 [[T] [P1] [P2] [P3]]. +Implicit Arguments all_and4 [[T] [P1] [P2] [P3] [P4]]. +Implicit Arguments all_and5 [[T] [P1] [P2] [P3] [P4] [P5]]. + Lemma pair_andP P Q : P /\ Q <-> P * Q. Proof. by split; case. Qed. Section ReflectConnectives. @@ -1423,7 +1467,7 @@ Definition antisymmetric := forall x y, R x y && R y x -> x = y. Definition pre_symmetric := forall x y, R x y -> R y x. Lemma symmetric_from_pre : pre_symmetric -> symmetric. -Proof. move=> symR x y; apply/idP/idP; exact: symR. Qed. +Proof. by move=> symR x y; apply/idP/idP; apply: symR. Qed. Definition reflexive := forall x, R x x. Definition irreflexive := forall x, R x x = false. @@ -1457,7 +1501,7 @@ Qed. End RelationProperties. Lemma rev_trans T (R : rel T) : transitive R -> transitive (fun x y => R y x). -Proof. by move=> trR x y z Ryx Rzy; exact: trR Rzy Ryx. Qed. +Proof. by move=> trR x y z Ryx Rzy; apply: trR Rzy Ryx. Qed. (* Property localization *) @@ -1608,14 +1652,14 @@ Lemma in3T : {in T1 & T2 & T3, {all3 P3}} -> {all3 P3}. Proof. by move=> ? ?; auto. Qed. Lemma sub_in1 (Ph : ph {all1 P1}) : prop_in1 d1' Ph -> prop_in1 d1 Ph. -Proof. move=> allP x /sub1; exact: allP. Qed. +Proof. by move=> allP x /sub1; apply: allP. Qed. Lemma sub_in11 (Ph : ph {all2 P2}) : prop_in11 d1' d2' Ph -> prop_in11 d1 d2 Ph. -Proof. move=> allP x1 x2 /sub1 d1x1 /sub2; exact: allP. Qed. +Proof. by move=> allP x1 x2 /sub1 d1x1 /sub2; apply: allP. Qed. Lemma sub_in111 (Ph : ph {all3 P3}) : prop_in111 d1' d2' d3' Ph -> prop_in111 d1 d2 d3 Ph. -Proof. by move=> allP x1 x2 x3 /sub1 d1x1 /sub2 d2x2 /sub3; exact: allP. Qed. +Proof. by move=> allP x1 x2 x3 /sub1 d1x1 /sub2 d2x2 /sub3; apply: allP. Qed. Let allQ1 f'' := {all1 Q1 f''}. Let allQ1l f'' h' := {all1 Q1l f'' h'}. @@ -1637,15 +1681,15 @@ Proof. by move=> ? ?; auto. Qed. Lemma subon1 (Phf : ph (allQ1 f)) (Ph : ph (allQ1 f)) : prop_on1 d2' Phf Ph -> prop_on1 d2 Phf Ph. -Proof. by move=> allQ x /sub2; exact: allQ. Qed. +Proof. by move=> allQ x /sub2; apply: allQ. Qed. Lemma subon1l (Phf : ph (allQ1l f)) (Ph : ph (allQ1l f h)) : prop_on1 d2' Phf Ph -> prop_on1 d2 Phf Ph. -Proof. by move=> allQ x /sub2; exact: allQ. Qed. +Proof. by move=> allQ x /sub2; apply: allQ. Qed. Lemma subon2 (Phf : ph (allQ2 f)) (Ph : ph (allQ2 f)) : prop_on2 d2' Phf Ph -> prop_on2 d2 Phf Ph. -Proof. by move=> allQ x y /sub2=> d2fx /sub2; exact: allQ. Qed. +Proof. by move=> allQ x y /sub2=> d2fx /sub2; apply: allQ. Qed. Lemma can_in_inj : {in D1, cancel f g} -> {in D1 &, injective f}. Proof. by move=> fK x y /fK{2}<- /fK{2}<- ->. Qed. @@ -1680,40 +1724,40 @@ Proof. by case=> g' fK g'K; exists g' => * ? *; auto. Qed. Lemma sub_in_bij (D1' : pred T1) : {subset D1 <= D1'} -> {in D1', bijective f} -> {in D1, bijective f}. Proof. -by move=> subD [g' fK g'K]; exists g' => x; move/subD; [exact: fK | exact: g'K]. +by move=> subD [g' fK g'K]; exists g' => x; move/subD; [apply: fK | apply: g'K]. Qed. Lemma subon_bij (D2' : pred T2) : {subset D2 <= D2'} -> {on D2', bijective f} -> {on D2, bijective f}. Proof. -by move=> subD [g' fK g'K]; exists g' => x; move/subD; [exact: fK | exact: g'K]. +by move=> subD [g' fK g'K]; exists g' => x; move/subD; [apply: fK | apply: g'K]. Qed. End LocalGlobal. Lemma sub_in2 T d d' (P : T -> T -> Prop) : sub_mem d d' -> forall Ph : ph {all2 P}, prop_in2 d' Ph -> prop_in2 d Ph. -Proof. by move=> /= sub_dd'; exact: sub_in11. Qed. +Proof. by move=> /= sub_dd'; apply: sub_in11. Qed. Lemma sub_in3 T d d' (P : T -> T -> T -> Prop) : sub_mem d d' -> forall Ph : ph {all3 P}, prop_in3 d' Ph -> prop_in3 d Ph. -Proof. by move=> /= sub_dd'; exact: sub_in111. Qed. +Proof. by move=> /= sub_dd'; apply: sub_in111. Qed. Lemma sub_in12 T1 T d1 d1' d d' (P : T1 -> T -> T -> Prop) : sub_mem d1 d1' -> sub_mem d d' -> forall Ph : ph {all3 P}, prop_in12 d1' d' Ph -> prop_in12 d1 d Ph. -Proof. by move=> /= sub1 sub; exact: sub_in111. Qed. +Proof. by move=> /= sub1 sub; apply: sub_in111. Qed. Lemma sub_in21 T T3 d d' d3 d3' (P : T -> T -> T3 -> Prop) : sub_mem d d' -> sub_mem d3 d3' -> forall Ph : ph {all3 P}, prop_in21 d' d3' Ph -> prop_in21 d d3 Ph. -Proof. by move=> /= sub sub3; exact: sub_in111. Qed. +Proof. by move=> /= sub sub3; apply: sub_in111. Qed. Lemma equivalence_relP_in T (R : rel T) (A : pred T) : {in A & &, equivalence_rel R} <-> {in A, reflexive R} /\ {in A &, forall x y, R x y -> {in A, R x =1 R y}}. Proof. -split=> [eqiR | [Rxx trR] x y z *]; last by split=> [|/trR-> //]; exact: Rxx. +split=> [eqiR | [Rxx trR] x y z *]; last by split=> [|/trR-> //]; apply: Rxx. by split=> [x Ax|x y Ax Ay Rxy z Az]; [rewrite (eqiR x x) | rewrite (eqiR x y)]. Qed. diff --git a/mathcomp/ssreflect/ssreflect.v b/mathcomp/ssreflect/ssreflect.v index 815a0fb..705c4ec 100644 --- a/mathcomp/ssreflect/ssreflect.v +++ b/mathcomp/ssreflect/ssreflect.v @@ -1,6 +1,6 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) Require Import Bool. (* For bool_scope delimiter 'bool'. *) -Require Import ssrmatching. +Require Import mathcomp.ssreflect.ssrmatching. Declare ML Module "ssreflect". Set SsrAstVersion. diff --git a/mathcomp/ssreflect/ssrfun.v b/mathcomp/ssreflect/ssrfun.v index f83724e..6b82548 100644 --- a/mathcomp/ssreflect/ssrfun.v +++ b/mathcomp/ssreflect/ssrfun.v @@ -1,5 +1,6 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect. +Require Import mathcomp.ssreflect.ssreflect. + (******************************************************************************) (* This file contains the basic definitions and notations for working with *) @@ -719,7 +720,7 @@ Lemma pcan_inj g : pcancel g -> injective. Proof. by move=> fK x y /(congr1 g); rewrite !fK => [[]]. Qed. Lemma can_inj g : cancel g -> injective. -Proof. by move/can_pcan; exact: pcan_inj. Qed. +Proof. by move/can_pcan; apply: pcan_inj. Qed. Lemma canLR g x y : cancel g -> x = f y -> g x = y. Proof. by move=> fK ->. Qed. @@ -746,10 +747,10 @@ Lemma inj_id : injective (@id A). Proof. by []. Qed. Lemma inj_can_sym f' : cancel f f' -> injective f' -> cancel f' f. -Proof. move=> fK injf' x; exact: injf'. Qed. +Proof. by move=> fK injf' x; apply: injf'. Qed. Lemma inj_comp : injective f -> injective h -> injective (f \o h). -Proof. move=> injf injh x y /injf; exact: injh. Qed. +Proof. by move=> injf injh x y /injf; apply: injh. Qed. Lemma can_comp f' h' : cancel f f' -> cancel h h' -> cancel (f \o h) (h' \o f'). Proof. by move=> fK hK x; rewrite /= fK hK. Qed. @@ -759,7 +760,7 @@ Lemma pcan_pcomp f' h' : Proof. by move=> fK hK x; rewrite /pcomp fK /= hK. Qed. Lemma eq_inj : injective f -> f =1 g -> injective g. -Proof. by move=> injf eqfg x y; rewrite -2!eqfg; exact: injf. Qed. +Proof. by move=> injf eqfg x y; rewrite -2!eqfg; apply: injf. Qed. Lemma eq_can f' g' : cancel f f' -> f =1 g -> f' =1 g' -> cancel g g'. Proof. by move=> fK eqfg eqfg' x; rewrite -eqfg -eqfg'. Qed. diff --git a/mathcomp/ssreflect/ssrnat.v b/mathcomp/ssreflect/ssrnat.v index bd54e57..50dcd7f 100644 --- a/mathcomp/ssreflect/ssrnat.v +++ b/mathcomp/ssreflect/ssrnat.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype. Require Import BinNat. Require BinPos Ndec. Require Export Ring. @@ -278,7 +280,7 @@ Lemma subSnn n : n.+1 - n = 1. Proof. exact (addnK n 1). Qed. Lemma subnDA m n p : n - (m + p) = (n - m) - p. -Proof. by elim: m n => [|m IHm] [|n]; try exact (IHm n). Qed. +Proof. by elim: m n => [|m IHm] []. Qed. Lemma subnAC : right_commutative subn. Proof. by move=> m n p; rewrite -!subnDA addnC. Qed. @@ -328,7 +330,7 @@ Lemma prednK n : 0 < n -> n.-1.+1 = n. Proof. exact: ltn_predK. Qed. Lemma leqNgt m n : (m <= n) = ~~ (n < m). -Proof. by elim: m n => [|m IHm] [|n] //; exact: IHm n. Qed. +Proof. by elim: m n => [|m IHm] []. Qed. Lemma ltnNge m n : (m < n) = ~~ (n <= m). Proof. by rewrite leqNgt. Qed. @@ -344,7 +346,7 @@ Lemma neq0_lt0n n : (n == 0) = false -> 0 < n. Proof. by case: n. Qed. Hint Resolve lt0n_neq0 neq0_lt0n. Lemma eqn_leq m n : (m == n) = (m <= n <= m). -Proof. elim: m n => [|m IHm] [|n] //; exact: IHm n. Qed. +Proof. by elim: m n => [|m IHm] []. Qed. Lemma anti_leq : antisymmetric leq. Proof. by move=> m n; rewrite -eqn_leq => /eqP. Qed. @@ -359,29 +361,29 @@ Lemma ltn_eqF m n : m < n -> m == n = false. Proof. by move/gtn_eqF; rewrite eq_sym. Qed. Lemma leq_eqVlt m n : (m <= n) = (m == n) || (m < n). -Proof. elim: m n => [|m IHm] [|n] //; exact: IHm n. Qed. +Proof. by elim: m n => [|m IHm] []. Qed. Lemma ltn_neqAle m n : (m < n) = (m != n) && (m <= n). Proof. by rewrite ltnNge leq_eqVlt negb_or -leqNgt eq_sym. Qed. Lemma leq_trans n m p : m <= n -> n <= p -> m <= p. -Proof. by elim: n m p => [|i IHn] [|m] [|p] //; exact: IHn m p. Qed. +Proof. by elim: n m p => [|i IHn] [|m] [|p] //; apply: IHn m p. Qed. Lemma leq_ltn_trans n m p : m <= n -> n < p -> m < p. -Proof. move=> Hmn; exact: leq_trans. Qed. +Proof. by move=> Hmn; apply: leq_trans. Qed. Lemma ltnW m n : m < n -> m <= n. Proof. exact: leq_trans. Qed. Hint Resolve ltnW. Lemma leqW m n : m <= n -> m <= n.+1. -Proof. by move=> le_mn; exact: ltnW. Qed. +Proof. by move=> le_mn; apply: ltnW. Qed. Lemma ltn_trans n m p : m < n -> n < p -> m < p. -Proof. by move=> lt_mn /ltnW; exact: leq_trans. Qed. +Proof. by move=> lt_mn /ltnW; apply: leq_trans. Qed. Lemma leq_total m n : (m <= n) || (m >= n). -Proof. by rewrite -implyNb -ltnNge; apply/implyP; exact: ltnW. Qed. +Proof. by rewrite -implyNb -ltnNge; apply/implyP; apply: ltnW. Qed. (* Link to the legacy comparison predicates. *) @@ -403,7 +405,7 @@ case: n1 / le_mn1 def_n1 => [|n1 le_mn1] def_n1 [|n2 le_mn2] def_n2. - by move/leP: (le_mn2); rewrite -{1}def_n2 ltnn. - by move/leP: (le_mn1); rewrite {1}def_n2 ltnn. case: def_n2 (def_n2) => ->{n2} def_n2 in le_mn2 *. -by rewrite [def_n2]eq_axiomK /=; congr le_S; exact: IHn. +by rewrite [def_n2]eq_axiomK /=; congr le_S; apply: IHn. Qed. Lemma ltP m n : reflect (m < n)%coq_nat (m < n). @@ -446,7 +448,7 @@ CoInductive compare_nat m n : bool -> bool -> bool -> Set := Lemma ltngtP m n : compare_nat m n (m < n) (n < m) (m == n). Proof. rewrite ltn_neqAle eqn_leq; case: ltnP; first by constructor. -by rewrite leq_eqVlt orbC; case: leqP; constructor; first exact/eqnP. +by rewrite leq_eqVlt orbC; case: leqP; constructor; first apply/eqnP. Qed. (* Monotonicity lemmas *) @@ -455,10 +457,10 @@ Lemma leq_add2l p m n : (p + m <= p + n) = (m <= n). Proof. by elim: p. Qed. Lemma ltn_add2l p m n : (p + m < p + n) = (m < n). -Proof. by rewrite -addnS; exact: leq_add2l. Qed. +Proof. by rewrite -addnS; apply: leq_add2l. Qed. Lemma leq_add2r p m n : (m + p <= n + p) = (m <= n). -Proof. by rewrite -!(addnC p); exact: leq_add2l. Qed. +Proof. by rewrite -!(addnC p); apply: leq_add2l. Qed. Lemma ltn_add2r p m n : (m + p < n + p) = (m < n). Proof. exact: leq_add2r p m.+1 n. Qed. @@ -475,16 +477,16 @@ Lemma leq_addl m n : n <= m + n. Proof. by rewrite addnC leq_addr. Qed. Lemma ltn_addr m n p : m < n -> m < n + p. -Proof. by move/leq_trans=> -> //; exact: leq_addr. Qed. +Proof. by move/leq_trans=> -> //; apply: leq_addr. Qed. Lemma ltn_addl m n p : m < n -> m < p + n. -Proof. by move/leq_trans=> -> //; exact: leq_addl. Qed. +Proof. by move/leq_trans=> -> //; apply: leq_addl. Qed. Lemma addn_gt0 m n : (0 < m + n) = (0 < m) || (0 < n). Proof. by rewrite !lt0n -negb_and addn_eq0. Qed. Lemma subn_gt0 m n : (0 < n - m) = (m < n). -Proof. by elim: m n => [|m IHm] [|n] //; exact: IHm n. Qed. +Proof. by elim: m n => [|m IHm] [|n] //; apply: IHm n. Qed. Lemma subn_eq0 m n : (m - n == 0) = (m <= n). Proof. by []. Qed. @@ -499,7 +501,7 @@ Lemma subnKC m n : m <= n -> m + (n - m) = n. Proof. by elim: m n => [|m IHm] [|n] // /(IHm n) {2}<-. Qed. Lemma subnK m n : m <= n -> (n - m) + m = n. -Proof. by rewrite addnC; exact: subnKC. Qed. +Proof. by rewrite addnC; apply: subnKC. Qed. Lemma addnBA m n p : p <= n -> m + (n - p) = m + n - p. Proof. by move=> le_pn; rewrite -{2}(subnK le_pn) addnA addnK. Qed. @@ -531,10 +533,10 @@ Lemma leq_sub m1 m2 n1 n2 : m1 <= m2 -> n2 <= n1 -> m1 - n1 <= m2 - n2. Proof. by move/(leq_sub2r n1)=> le_m12 /(leq_sub2l m2); apply: leq_trans. Qed. Lemma ltn_sub2r p m n : p < n -> m < n -> m - p < n - p. -Proof. by move/subnSK <-; exact: (@leq_sub2r p.+1). Qed. +Proof. by move/subnSK <-; apply: (@leq_sub2r p.+1). Qed. Lemma ltn_sub2l p m n : m < p -> m < n -> p - n < p - m. -Proof. by move/subnSK <-; exact: leq_sub2l. Qed. +Proof. by move/subnSK <-; apply: leq_sub2l. Qed. Lemma ltn_subRL m n p : (n < p - m) = (m + n < p). Proof. by rewrite !ltnNge leq_subLR. Qed. @@ -613,7 +615,7 @@ Lemma minnC : commutative minn. Proof. by move=> m n; rewrite /minn; case ltngtP. Qed. Lemma addn_min_max m n : minn m n + maxn m n = m + n. -Proof. by rewrite /minn /maxn; case: ltngtP => // [_|->] //; exact: addnC. Qed. +Proof. by rewrite /minn /maxn; case: ltngtP => // [_|->] //; apply: addnC. Qed. Lemma minnE m n : minn m n = m - (m - n). Proof. by rewrite -(subnDl n) -maxnE -addn_min_max addnK minnC. Qed. @@ -717,7 +719,7 @@ Proof. have: forall n, P n -> n >= 0 by []. have: acc_nat 0. case exP => n; rewrite -(addn0 n); elim: n 0 => [|n IHn] j; first by left. - rewrite addSnnS; right; exact: IHn. + by rewrite addSnnS; right; apply: IHn. move: 0; fix 2 => m IHm m_lb; case Pm: (P m); first by exists m. apply: find_ex_minn m.+1 _ _ => [|n Pn]; first by case: IHm; rewrite ?Pm. by rewrite ltn_neqAle m_lb //; case: eqP Pm => // -> /idP[]. @@ -891,7 +893,7 @@ Lemma muln_gt0 m n : (0 < m * n) = (0 < m) && (0 < n). Proof. by case: m n => // m [|n] //=; rewrite muln0. Qed. Lemma leq_pmull m n : n > 0 -> m <= n * m. -Proof. by move/prednK <-; exact: leq_addr. Qed. +Proof. by move/prednK <-; apply: leq_addr. Qed. Lemma leq_pmulr m n : n > 0 -> m <= m * n. Proof. by move/leq_pmull; rewrite mulnC. Qed. @@ -1021,7 +1023,7 @@ Proof. by rewrite !eqn0Ngt expn_gt0 negb_or -lt0n. Qed. Lemma ltn_expl m n : 1 < m -> n < m ^ n. Proof. move=> m_gt1; elim: n => //= n; rewrite -(leq_pmul2l (ltnW m_gt1)) expnS. -by apply: leq_trans; exact: ltn_Pmull. +by apply: leq_trans; apply: ltn_Pmull. Qed. Lemma leq_exp2l m n1 n2 : 1 < m -> (m ^ n1 <= m ^ n2) = (n1 <= n2). @@ -1154,7 +1156,7 @@ Lemma doubleD m n : (m + n).*2 = m.*2 + n.*2. Proof. by rewrite -!addnn -!addnA (addnCA n). Qed. Lemma doubleB m n : (m - n).*2 = m.*2 - n.*2. -Proof. elim: m n => [|m IHm] [|n] //; exact: IHm n. Qed. +Proof. by elim: m n => [|m IHm] []. Qed. Lemma leq_double m n : (m.*2 <= n.*2) = (m <= n). Proof. by rewrite /leq -doubleB; case (m - n). Qed. @@ -1295,7 +1297,7 @@ Coercion leq_of_leqif m n C (H : m <= n ?= iff C) := H.1 : m <= n. Lemma leqifP m n C : reflect (m <= n ?= iff C) (if C then m == n else m < n). Proof. rewrite ltn_neqAle; apply: (iffP idP) => [|lte]; last by rewrite !lte; case C. -by case C => [/eqP-> | /andP[/negPf]]; split=> //; exact: eqxx. +by case C => [/eqP-> | /andP[/negPf]]; split=> //; apply: eqxx. Qed. Lemma leqif_refl m C : reflect (m <= m ?= iff C) C. @@ -1338,22 +1340,15 @@ Lemma leqif_mul m1 n1 C1 m2 n2 C2 : m1 <= n1 ?= iff C1 -> m2 <= n2 ?= iff C2 -> m1 * m2 <= n1 * n2 ?= iff (n1 * n2 == 0) || (C1 && C2). Proof. -move=> le1 le2; case: posnP => [n12_0 | ]. - rewrite n12_0; move/eqP: n12_0 {le1 le2}le1.1 le2.1; rewrite muln_eq0. - by case/orP=> /eqP->; case: m1 m2 => [|m1] [|m2] // _ _; - rewrite ?muln0; exact/leqif_refl. -rewrite muln_gt0 => /andP[n1_gt0 n2_gt0]. -have [m2_0 | m2_gt0] := posnP m2. - apply/leqifP; rewrite -le2 andbC eq_sym eqn_leq leqNgt m2_0 muln0. - by rewrite muln_gt0 n1_gt0 n2_gt0. -have mono_n1 := leq_pmul2l n1_gt0; have mono_m2 := leq_pmul2r m2_gt0. -rewrite -(mono_leqif mono_m2) in le1; rewrite -(mono_leqif mono_n1) in le2. -exact: leqif_trans le1 le2. +case: n1 => [|n1] le1; first by case: m1 le1 => [|m1] [_ <-] //. +case: n2 m2 => [|n2] [|m2] /=; try by case=> // _ <-; rewrite !muln0 ?andbF. +have /leq_pmul2l-/mono_leqif<-: 0 < n1.+1 by []. +by apply: leqif_trans; have /leq_pmul2r-/mono_leqif->: 0 < m2.+1. Qed. Lemma nat_Cauchy m n : 2 * (m * n) <= m ^ 2 + n ^ 2 ?= iff (m == n). Proof. -wlog le_nm: m n / n <= m. +without loss le_nm: m n / n <= m. by case: (leqP m n); auto; rewrite eq_sym addnC (mulnC m); auto. apply/leqifP; case: ifP => [/eqP-> | ne_mn]; first by rewrite mulnn addnn mul2n. by rewrite -subn_gt0 -sqrn_sub // sqrn_gt0 subn_gt0 ltn_neqAle eq_sym ne_mn. @@ -1500,7 +1495,7 @@ elim: p q => [p IHp|p IHp|] [q|q|] //=; rewrite !natTrecE //; Qed. Lemma nat_of_add_bin b1 b2 : (b1 + b2)%num = b1 + b2 :> nat. -Proof. case: b1 b2 => [|p] [|q] //=; exact: nat_of_addn_gt0. Qed. +Proof. by case: b1 b2 => [|p] [|q] //=; apply: nat_of_addn_gt0. Qed. Lemma nat_of_mul_bin b1 b2 : (b1 * b2)%num = b1 * b2 :> nat. Proof. @@ -1510,8 +1505,7 @@ Qed. Lemma nat_of_exp_bin n (b : N) : n ^ b = pow_N 1 muln n b. Proof. -case: b => [|p] /=; first exact: expn0. -by elim: p => //= p <-; rewrite natTrecE mulnn -expnM muln2 ?expnS. +by case: b; last (elim=> //= p <-; rewrite natTrecE mulnn -expnM muln2 ?expnS). Qed. End NumberInterpretation. @@ -1548,7 +1542,7 @@ by move: nat_of_add_bin nat_of_mul_bin; split=> //= m n; move/eqP->. Qed. Lemma nat_power_theory : power_theory 1 muln (@eq _) nat_of_bin expn. -Proof. split; exact: nat_of_exp_bin. Qed. +Proof. by split; apply: nat_of_exp_bin. Qed. (* Interface to the ring tactic machinery. *) diff --git a/mathcomp/ssrtest/Make b/mathcomp/ssrtest/Make index adcc2d2..ab4c666 100644 --- a/mathcomp/ssrtest/Make +++ b/mathcomp/ssrtest/Make @@ -1,22 +1,20 @@ --R ../theories Ssreflect --I ../src/ absevarprop.v -binders.v binders_of.v +binders.v caseview.v congr.v deferclear.v dependent_type_err.v -elim.v elim2.v elim_pattern.v +elim.v first_n.v -gen_pattern.v gen_have.v +gen_pattern.v havesuff.v -have_view_idiom.v have_TC.v have_transp.v +have_view_idiom.v if_isnt.v indetLHS.v intro_beta.v @@ -33,12 +31,14 @@ set_lamda.v set_pattern.v ssrsyntax1.v ssrsyntax2.v -testmx.v tc.v +testmx.v typeof.v unkeyed.v view_case.v wlogletin.v -wlong_intro.v wlog_suff.v +wlong_intro.v +-R ../theories Ssreflect +-I ../src/ diff --git a/mathcomp/ssrtest/absevarprop.v b/mathcomp/ssrtest/absevarprop.v index 513e53f..4fef29e 100644 --- a/mathcomp/ssrtest/absevarprop.v +++ b/mathcomp/ssrtest/absevarprop.v @@ -1,5 +1,8 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq. +From mathcomp Require Import fintype. Lemma test15: forall (y : nat) (x : 'I_2), y < 1 -> val x = y -> Some x = insub y. diff --git a/mathcomp/ssrtest/binders.v b/mathcomp/ssrtest/binders.v index 6a63167..11a8d26 100644 --- a/mathcomp/ssrtest/binders.v +++ b/mathcomp/ssrtest/binders.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool eqtype ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool eqtype ssrnat. Lemma test (x : bool) : True. have H1 x := x. diff --git a/mathcomp/ssrtest/binders_of.v b/mathcomp/ssrtest/binders_of.v index 70a822e..e8366f6 100644 --- a/mathcomp/ssrtest/binders_of.v +++ b/mathcomp/ssrtest/binders_of.v @@ -1,5 +1,7 @@ -Require Import ssreflect seq. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import seq. Lemma test1 : True. have f of seq nat & nat : nat. @@ -7,6 +9,4 @@ have f of seq nat & nat : nat. have g of nat := 3. have h of nat : nat := 3. have _ : f [::] 3 = g 3 + h 4. - by admit. -by admit. -Qed.
\ No newline at end of file +Admitted. diff --git a/mathcomp/ssrtest/caseview.v b/mathcomp/ssrtest/caseview.v index 108cf46..7b0d4ab 100644 --- a/mathcomp/ssrtest/caseview.v +++ b/mathcomp/ssrtest/caseview.v @@ -1,4 +1,5 @@ -Require Import ssreflect. +Require Import mathcomp.ssreflect.ssreflect. + Lemma test (A B : Prop) : A /\ B -> True. Proof. by case=> _ /id _. Qed.
\ No newline at end of file diff --git a/mathcomp/ssrtest/congr.v b/mathcomp/ssrtest/congr.v index edd52fe..0314772 100644 --- a/mathcomp/ssrtest/congr.v +++ b/mathcomp/ssrtest/congr.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool eqtype ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool eqtype ssrnat. Lemma test1 : forall a b : nat, a == b -> a == 0 -> b == 0. Proof. move=> a b Eab Eac; congr (_ == 0) : Eac; exact: eqP Eab. Qed. diff --git a/mathcomp/ssrtest/deferclear.v b/mathcomp/ssrtest/deferclear.v index a13a20e..ed57fca 100644 --- a/mathcomp/ssrtest/deferclear.v +++ b/mathcomp/ssrtest/deferclear.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect. +Require Import mathcomp.ssreflect.ssreflect. + +From mathcomp Require Import ssrbool eqtype fintype ssrnat. Variable T : Type. diff --git a/mathcomp/ssrtest/dependent_type_err.v b/mathcomp/ssrtest/dependent_type_err.v index cd9570b..b2835c7 100644 --- a/mathcomp/ssrtest/dependent_type_err.v +++ b/mathcomp/ssrtest/dependent_type_err.v @@ -1,4 +1,6 @@ -Require Import ssreflect ssrfun ssrbool eqtype ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat. Lemma ltn_leq_trans : forall n m p : nat, m < n -> n <= p -> m < p. move=> n m p Hmn Hnp; rewrite -ltnS. diff --git a/mathcomp/ssrtest/elim.v b/mathcomp/ssrtest/elim.v index 5ab8f41..1adbb5e 100644 --- a/mathcomp/ssrtest/elim.v +++ b/mathcomp/ssrtest/elim.v @@ -1,5 +1,8 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat seq. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat seq. +Axiom daemon : False. Ltac myadmit := case: daemon. (* Ltac debugging feature: recursive elim + eq generation *) Lemma testL1 : forall A (s : seq A), s = s. @@ -154,18 +157,19 @@ Proof. move=> cf0 cf1 t; split => [] Ecf. elim: Ecf. match goal with |- forall cf2 cf3 : T, tr cf2 = cf3 -> - execr cf2 cf3 [::] => admit | _ => fail end. + execr cf2 cf3 [::] => myadmit | _ => fail end. match goal with |- forall (cf2 cf3 cf4 : T) (t0 : seq T), tr cf2 = cf4 -> exec cf4 cf3 t0 -> execr cf4 cf3 t0 -> - execr cf2 cf3 (cf4 :: t0) => admit | _ => fail end. + execr cf2 cf3 (cf4 :: t0) => myadmit | _ => fail end. elim: Ecf. match goal with |- forall cf2 : T, - tr cf0 = cf2 -> exec cf0 cf2 [::] => admit | _ => fail end. + tr cf0 = cf2 -> exec cf0 cf2 [::] => myadmit | _ => fail end. match goal with |- forall (cf2 cf3 : T) (t0 : seq T), execr cf0 cf3 t0 -> exec cf0 cf3 t0 -> tr cf3 = cf2 -> - exec cf0 cf2 (t0 ++ [:: cf3]) => admit | _ => fail end. + exec cf0 cf2 (t0 ++ [:: cf3]) => myadmit | _ => fail end. Qed. +From mathcomp Require Import seq div prime bigop. Lemma mem_primes : forall p n, diff --git a/mathcomp/ssrtest/elim2.v b/mathcomp/ssrtest/elim2.v index 344ee52..b3e764e 100644 --- a/mathcomp/ssrtest/elim2.v +++ b/mathcomp/ssrtest/elim2.v @@ -1,4 +1,7 @@ -Require Import ssreflect eqtype ssrbool ssrnat seq div fintype finfun path bigop. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import eqtype ssrbool ssrnat seq div fintype finfun path bigop. +Axiom daemon : False. Ltac myadmit := case: daemon. Lemma big_load R (K K' : R -> Type) idx op I r (P : pred I) F : let s := \big[op/idx]_(i <- r | P i) F i in @@ -13,9 +16,9 @@ Variables (idx : R) (op op' : R -> R -> R). Hypothesis Kid : K idx. -Ltac ASSERT1 := match goal with |- (K idx) => admit end. +Ltac ASSERT1 := match goal with |- (K idx) => myadmit end. Ltac ASSERT2 K := match goal with |- (forall x1 : R, R -> - forall y1 : R, R -> K x1 -> K y1 -> K (op x1 y1)) => admit end. + forall y1 : R, R -> K x1 -> K y1 -> K (op x1 y1)) => myadmit end. Lemma big_rec I r (P : pred I) F @@ -23,13 +26,13 @@ Lemma big_rec I r (P : pred I) F K (\big[op/idx]_(i <- r | P i) F i). Proof. elim/big_ind2: {-}_. - ASSERT1. ASSERT2 K. match goal with |- (forall i : I, is_true (P i) -> K (F i)) => admit end. Undo 4. + ASSERT1. ASSERT2 K. match goal with |- (forall i : I, is_true (P i) -> K (F i)) => myadmit end. Undo 4. elim/big_ind2: _ / {-}_. - ASSERT1. ASSERT2 K. match goal with |- (forall i : I, is_true (P i) -> K (F i)) => admit end. Undo 4. + ASSERT1. ASSERT2 K. match goal with |- (forall i : I, is_true (P i) -> K (F i)) => myadmit end. Undo 4. elim/big_rec2: (\big[op/idx]_(i <- r | P i) op idx (F i)) / (\big[op/idx]_(i <- r | P i) F i). - ASSERT1. match goal with |- (forall i : I, R -> forall y2 : R, is_true (P i) -> K y2 -> K (op (F i) y2)) => admit end. Undo 3. + ASSERT1. match goal with |- (forall i : I, R -> forall y2 : R, is_true (P i) -> K y2 -> K (op (F i) y2)) => myadmit end. Undo 3. elim/(big_load (phantom R)): _. Undo. diff --git a/mathcomp/ssrtest/elim_pattern.v b/mathcomp/ssrtest/elim_pattern.v index 51ab216..78abb5e 100644 --- a/mathcomp/ssrtest/elim_pattern.v +++ b/mathcomp/ssrtest/elim_pattern.v @@ -1,14 +1,17 @@ -Require Import ssreflect ssrbool eqtype ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool eqtype ssrnat. +Axiom daemon : False. Ltac myadmit := case: daemon. Lemma test x : (x == x) = (x + x.+1 == 2 * x + 1). case: (X in _ = X) / eqP => _. -match goal with |- (x == x) = true => admit end. -match goal with |- (x == x) = false => admit end. +match goal with |- (x == x) = true => myadmit end. +match goal with |- (x == x) = false => myadmit end. Qed. Lemma test1 x : (x == x) = (x + x.+1 == 2 * x + 1). elim: (x in RHS). -match goal with |- (x == x) = _ => admit end. -match goal with |- forall n, (x == x) = _ -> (x == x) = _ => admit end. +match goal with |- (x == x) = _ => myadmit end. +match goal with |- forall n, (x == x) = _ -> (x == x) = _ => myadmit end. Qed. diff --git a/mathcomp/ssrtest/first_n.v b/mathcomp/ssrtest/first_n.v index 175684a..e6af0b6 100644 --- a/mathcomp/ssrtest/first_n.v +++ b/mathcomp/ssrtest/first_n.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool. Lemma test : False -> (bool -> False -> True -> True) -> True. move=> F; let w := 2 in apply; last w first. diff --git a/mathcomp/ssrtest/gen_have.v b/mathcomp/ssrtest/gen_have.v index 757dba5..59f19d6 100644 --- a/mathcomp/ssrtest/gen_have.v +++ b/mathcomp/ssrtest/gen_have.v @@ -1,11 +1,14 @@ -Require Import ssreflect ssrfun ssrbool eqtype ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat. +Axiom daemon : False. Ltac myadmit := case: daemon. Axiom P : nat -> Prop. Lemma clear_test (b1 b2 : bool) : b2 = b2. Proof. -(* wlog gH : (b3 := b2) / b2 = b3. admit. *) +(* wlog gH : (b3 := b2) / b2 = b3. myadmit. *) gen have {b1} H, gH : (b3 := b2) (w := erefl 3) / b2 = b3. - admit. + myadmit. Fail exact (H b1). exact (H b2 (erefl _)). Qed. @@ -13,46 +16,46 @@ Qed. Lemma test1 n (ngt0 : 0 < n) : P n. gen have lt2le, /andP[H1 H2] : n ngt0 / (0 <= n) && (n != 0). - match goal with |- is_true((0 <= n) && (n != 0)) => admit end. + match goal with |- is_true((0 <= n) && (n != 0)) => myadmit end. Check (lt2le : forall n : nat, 0 < n -> (0 <= n) && (n != 0)). Check (H1 : 0 <= n). Check (H2 : n != 0). -admit. +myadmit. Qed. Lemma test2 n (ngt0 : 0 < n) : P n. gen have _, /andP[H1 H2] : n ngt0 / (0 <= n) && (n != 0). - match goal with |- is_true((0 <= n) && (n != 0)) => admit end. + match goal with |- is_true((0 <= n) && (n != 0)) => myadmit end. lazymatch goal with | lt2le : forall n : nat, is_true(0 < n) -> is_true((0 <= n) && (n != 0)) |- _ => fail "not cleared" | _ => idtac end. Check (H1 : 0 <= n). Check (H2 : n != 0). -admit. +myadmit. Qed. Lemma test3 n (ngt0 : 0 < n) : P n. gen have H : n ngt0 / (0 <= n) && (n != 0). - match goal with |- is_true((0 <= n) && (n != 0)) => admit end. + match goal with |- is_true((0 <= n) && (n != 0)) => myadmit end. Check (H : forall n : nat, 0 < n -> (0 <= n) && (n != 0)). -admit. +myadmit. Qed. Lemma test4 n (ngt0 : 0 < n) : P n. gen have : n ngt0 / (0 <= n) && (n != 0). - match goal with |- is_true((0 <= n) && (n != 0)) => admit end. + match goal with |- is_true((0 <= n) && (n != 0)) => myadmit end. move=> H. Check(H : forall n : nat, 0 < n -> (0 <= n) && (n != 0)). -admit. +myadmit. Qed. Lemma test4bis n (ngt0 : 0 < n) : P n. wlog suff : n ngt0 / (0 <= n) && (n != 0); last first. - match goal with |- is_true((0 <= n) && (n != 0)) => admit end. + match goal with |- is_true((0 <= n) && (n != 0)) => myadmit end. move=> H. Check(H : forall n : nat, 0 < n -> (0 <= n) && (n != 0)). -admit. +myadmit. Qed. Lemma test5 n (ngt0 : 0 < n) : P n. @@ -60,7 +63,7 @@ Fail gen have : / (0 <= n) && (n != 0). Abort. Lemma test6 n (ngt0 : 0 < n) : P n. -gen have : n ngt0 / (0 <= n) && (n != 0) by admit. +gen have : n ngt0 / (0 <= n) && (n != 0) by myadmit. Abort. Lemma test7 n (ngt0 : 0 < n) : P n. @@ -70,17 +73,17 @@ Abort. Lemma test3wlog2 n (ngt0 : 0 < n) : P n. gen have H : (m := n) ngt0 / (0 <= m) && (m != 0). match goal with - ngt0 : is_true(0 < m) |- is_true((0 <= m) && (m != 0)) => admit end. + ngt0 : is_true(0 < m) |- is_true((0 <= m) && (m != 0)) => myadmit end. Check (H : forall n : nat, 0 < n -> (0 <= n) && (n != 0)). -admit. +myadmit. Qed. Lemma test3wlog3 n (ngt0 : 0 < n) : P n. gen have H : {n} (m := n) (n := 0) ngt0 / (0 <= m) && (m != n). match goal with - ngt0 : is_true(n < m) |- is_true((0 <= m) && (m != n)) => admit end. + ngt0 : is_true(n < m) |- is_true((0 <= m) && (m != n)) => myadmit end. Check (H : forall m n : nat, n < m -> (0 <= m) && (m != n)). -admit. +myadmit. Qed. Lemma testw1 n (ngt0 : 0 < n) : n <= 0. @@ -88,13 +91,13 @@ wlog H : (z := 0) (m := n) ngt0 / m != 0. match goal with |- (forall z m, is_true(z < m) -> is_true(m != 0) -> is_true(m <= z)) -> - is_true(n <= 0) => admit end. + is_true(n <= 0) => myadmit end. Check(n : nat). Check(m : nat). Check(z : nat). Check(ngt0 : z < m). Check(H : m != 0). -admit. +myadmit. Qed. Lemma testw2 n (ngt0 : 0 < n) : n <= 0. @@ -115,9 +118,9 @@ wlog H : {n} (m := n) (z := (X in _ <= X)) ngt0 / m != z. |- (forall m z : nat, is_true(0 < z) -> is_true(m != z) -> is_true(m <= 0)) -> is_true(n <= 0) => idtac end. - admit. + myadmit. Fail Check n. -admit. +myadmit. Qed. Section Test. @@ -126,8 +129,8 @@ Definition addx y := y + x. Lemma testw3 (m n : nat) (ngt0 : 0 < n) : n <= addx x. wlog H : (n0 := n) (y := x) (@twoy := (id _ as X in _ <= X)) / twoy = 2 * y. - admit. -admit. + myadmit. +myadmit. Qed. @@ -140,15 +143,15 @@ wlog H : (y := x) (@twoy := (X in _ <= X)) / twoy = 2 * y. |- (forall y : nat, let twoy := y + y in twoy = 2 * y -> is_true(n + y <= twoy)) -> - is_true(n + x <= twox) => admit end. + is_true(n + x <= twox) => myadmit end. Restart. wlog H : (y := x) (@twoy := (id _ as X in _ <= X)) / twoy = 2 * y. match goal with |- (forall y : nat, let twoy := twox in twoy = 2 * y -> is_true(n + y <= twoy)) -> - is_true(n + x <= twox) => admit end. -admit. + is_true(n + x <= twox) => myadmit end. +myadmit. Qed. End Test. @@ -156,5 +159,5 @@ End Test. Lemma test_in n k (def_k : k = 0) (ngtk : k < n) : P n. rewrite -(add0n n) in {def_k k ngtk} (m := k) (def_m := def_k) (ngtm := ngtk). rewrite def_m add0n in {ngtm} (e := erefl 0 ) (ngt0 := ngtm) => {def_m}. -admit. +myadmit. Qed. diff --git a/mathcomp/ssrtest/gen_pattern.v b/mathcomp/ssrtest/gen_pattern.v index de57e8d..e5af827 100644 --- a/mathcomp/ssrtest/gen_pattern.v +++ b/mathcomp/ssrtest/gen_pattern.v @@ -1,4 +1,6 @@ -Require Import ssreflect ssrbool ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrnat. Notation "( a 'in' c )" := (a + c) (only parsing) : myscope. Delimit Scope myscope with myscope. diff --git a/mathcomp/ssrtest/have_TC.v b/mathcomp/ssrtest/have_TC.v index 3204c42..de28520 100644 --- a/mathcomp/ssrtest/have_TC.v +++ b/mathcomp/ssrtest/have_TC.v @@ -1,4 +1,6 @@ -Require Import ssreflect. +Require Import mathcomp.ssreflect.ssreflect. + +Axiom daemon : False. Ltac myadmit := case: daemon. Class foo (T : Type) := { n : nat }. Instance five : foo nat := {| n := 5 |}. @@ -14,7 +16,7 @@ have titi : bar _ 5. reflexivity. have titi2 : bar _ 5 := . Fail reflexivity. - by admit. + by myadmit. have totoc (H : bar _ 5) : 3 = 3 := eq_refl. move/totoc: nat => _. exact I. @@ -26,10 +28,10 @@ Lemma a' : True. set toto := bar _ 8. have titi : bar _ 5. Fail reflexivity. - by admit. + by myadmit. have titi2 : bar _ 5 := . Fail reflexivity. - by admit. + by myadmit. have totoc (H : bar _ 5) : 3 = 3 := eq_refl. move/totoc: nat => _. exact I. diff --git a/mathcomp/ssrtest/have_transp.v b/mathcomp/ssrtest/have_transp.v index f1e3203..3eba582 100644 --- a/mathcomp/ssrtest/have_transp.v +++ b/mathcomp/ssrtest/have_transp.v @@ -1,4 +1,6 @@ -Require Import ssreflect ssrfun ssrbool eqtype ssrnat seq choice fintype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat seq choice fintype. Lemma test1 n : n >= 0. diff --git a/mathcomp/ssrtest/have_view_idiom.v b/mathcomp/ssrtest/have_view_idiom.v index d42a3ac..6faae97 100644 --- a/mathcomp/ssrtest/have_view_idiom.v +++ b/mathcomp/ssrtest/have_view_idiom.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool. Lemma test (a b : bool) (pab : a && b) : b. have {pab} /= /andP [pa -> //] /= : true && (a && b) := pab. diff --git a/mathcomp/ssrtest/havesuff.v b/mathcomp/ssrtest/havesuff.v index c497773..b15728d 100644 --- a/mathcomp/ssrtest/havesuff.v +++ b/mathcomp/ssrtest/havesuff.v @@ -1,5 +1,6 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect. +Require Import mathcomp.ssreflect.ssreflect. + Variables P G : Prop. diff --git a/mathcomp/ssrtest/if_isnt.v b/mathcomp/ssrtest/if_isnt.v index 619df45..58812d5 100644 --- a/mathcomp/ssrtest/if_isnt.v +++ b/mathcomp/ssrtest/if_isnt.v @@ -1,4 +1,5 @@ -Require Import ssreflect. +Require Import mathcomp.ssreflect.ssreflect. + Definition unopt (x : option bool) := if x isn't Some x then false else x. diff --git a/mathcomp/ssrtest/indetLHS.v b/mathcomp/ssrtest/indetLHS.v index f9d42ff..3f5e7f0 100644 --- a/mathcomp/ssrtest/indetLHS.v +++ b/mathcomp/ssrtest/indetLHS.v @@ -1,4 +1,6 @@ -Require Import ssreflect ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrnat. Goal 5 = 3. Fail (rewrite -(addnK _ 5)). Abort.
\ No newline at end of file diff --git a/mathcomp/ssrtest/intro_beta.v b/mathcomp/ssrtest/intro_beta.v index 6ede976..7a1b0e7 100644 --- a/mathcomp/ssrtest/intro_beta.v +++ b/mathcomp/ssrtest/intro_beta.v @@ -1,4 +1,5 @@ -Require Import ssreflect. +Require Import mathcomp.ssreflect.ssreflect. + Axiom T : Type. @@ -9,5 +10,4 @@ Axiom P : T -> T -> Prop. Lemma foo : C (fun x => forall y, let z := x in P y x). move=> a b. match goal with |- (let y := _ in _) => idtac end. -admit. -Qed. +Admitted. diff --git a/mathcomp/ssrtest/intro_noop.v b/mathcomp/ssrtest/intro_noop.v index 91a87b5..be0bea7 100644 --- a/mathcomp/ssrtest/intro_noop.v +++ b/mathcomp/ssrtest/intro_noop.v @@ -1,4 +1,7 @@ -Require Import ssreflect ssrbool. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool. +Axiom daemon : False. Ltac myadmit := case: daemon. Lemma v : True -> bool -> bool. Proof. by []. Qed. @@ -16,8 +19,8 @@ have - -> : a = (id a) by []. have --> : a = (id a) by []. have - - _ : a = (id a) by []. have -{1}-> : a = (id a) by []. - by admit. + by myadmit. move: a. case: b => -[] //. -by admit. +by myadmit. Qed. diff --git a/mathcomp/ssrtest/ipatalternation.v b/mathcomp/ssrtest/ipatalternation.v index 9796648..eb29fd7 100644 --- a/mathcomp/ssrtest/ipatalternation.v +++ b/mathcomp/ssrtest/ipatalternation.v @@ -1,5 +1,6 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect. +Require Import mathcomp.ssreflect.ssreflect. + Lemma test1 : Prop -> Prop -> Prop -> Prop -> Prop -> True = False -> Prop -> True \/ True. by move=> A /= /= /= B C {A} {B} ? _ {C} {1}-> *; right. diff --git a/mathcomp/ssrtest/ltac_have.v b/mathcomp/ssrtest/ltac_have.v index 3ed274d..c106b42 100644 --- a/mathcomp/ssrtest/ltac_have.v +++ b/mathcomp/ssrtest/ltac_have.v @@ -1,4 +1,6 @@ -Require Import ssreflect ssrbool ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrnat. Ltac SUFF1 h t := suff h x (p := x < 0) : t. Ltac SUFF2 h t := suff h x (p := x < 0) : t by apply h. diff --git a/mathcomp/ssrtest/ltac_in.v b/mathcomp/ssrtest/ltac_in.v index c9f15dd..4cc0f9c 100644 --- a/mathcomp/ssrtest/ltac_in.v +++ b/mathcomp/ssrtest/ltac_in.v @@ -1,4 +1,6 @@ -Require Import ssreflect ssrbool eqtype ssrnat ssrfun. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool eqtype ssrnat ssrfun. Set Implicit Arguments. Unset Strict Implicit. diff --git a/mathcomp/ssrtest/move_after.v b/mathcomp/ssrtest/move_after.v index 9289193..d62926d 100644 --- a/mathcomp/ssrtest/move_after.v +++ b/mathcomp/ssrtest/move_after.v @@ -1,4 +1,5 @@ -Require Import ssreflect. +Require Import mathcomp.ssreflect.ssreflect. + Goal True -> True -> True. move=> H1 H2. diff --git a/mathcomp/ssrtest/multiview.v b/mathcomp/ssrtest/multiview.v index 53b3b4e..6a4f35b 100644 --- a/mathcomp/ssrtest/multiview.v +++ b/mathcomp/ssrtest/multiview.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrnat. Goal forall m n p, n <= p -> m <= n -> m <= p. by move=> m n p le_n_p /leq_trans; apply. @@ -45,6 +47,7 @@ Lemma test7: forall A B, (A -> B) -> A -> B. move=> A B A_to_B a; apply A_to_B in a; exact: a. Qed. +From mathcomp Require Import ssrfun eqtype ssrnat div seq choice fintype finfun finset. Lemma test8 (T : finType) (A B : {set T}) x (Ax : x \in A) (_ : B = A) : x \in B. diff --git a/mathcomp/ssrtest/occarrow.v b/mathcomp/ssrtest/occarrow.v index deaee0c..8efa4bc 100644 --- a/mathcomp/ssrtest/occarrow.v +++ b/mathcomp/ssrtest/occarrow.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect eqtype ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import eqtype ssrnat. Lemma test1 : forall n m : nat, n = m -> m * m + n * n = n * n + n * n. move=> n m E; have [{2}-> _] : n * n = m * n /\ True by move: E => {1}<-. diff --git a/mathcomp/ssrtest/patnoX.v b/mathcomp/ssrtest/patnoX.v index 9cde676..75dce69 100644 --- a/mathcomp/ssrtest/patnoX.v +++ b/mathcomp/ssrtest/patnoX.v @@ -1,4 +1,6 @@ -Require Import ssreflect ssrbool. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool. Goal forall x, x && true = x. move=> x. Fail (rewrite [X in _ && _]andbT). diff --git a/mathcomp/ssrtest/rewpatterns.v b/mathcomp/ssrtest/rewpatterns.v index 88c2a2f..33c1903 100644 --- a/mathcomp/ssrtest/rewpatterns.v +++ b/mathcomp/ssrtest/rewpatterns.v @@ -1,6 +1,8 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool ssrfun eqtype ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrfun eqtype ssrnat. Lemma test1 : forall x y (f : nat -> nat), f (x + y).+1 = f (y + x.+1). by move=> x y f; rewrite [_.+1](addnC x.+1). @@ -89,18 +91,16 @@ rewrite [x.+1 + y as X in f X _]addnC. match goal with |- (x + y).+1 + f (y + x.+1) (z + (y + x.+1)) = 0 + (0 + 0) => idtac end. Admitted. +From mathcomp Require Import fintype ssrnat finset fingroup gproduct. Goal (forall (gT : finGroupType) (G H: {set gT}) (Z : {group gT}), G = Z). move=> gT G H K. suff: G \x H = K. case/dprodP=> {G H} [[G H -> -> defK _ _]]. -admit. -admit. -Qed. +Admitted. Goal (exists x : 'I_3, x > 0). apply: (ex_intro _ (@Ordinal _ 2 _)). -admit. -Qed. +Admitted. Goal (forall y, 1 < y < 2 -> exists x : 'I_3, x > 0). move=> y; case/andP=> y_gt1 y_lt2; apply: (ex_intro _ (@Ordinal _ y _)). @@ -111,8 +111,7 @@ Qed. Goal (forall x y : nat, forall P : nat -> Prop, x = y -> True). move=> x y P E. have: P x -> P y by suff: x = y by move=> ?; congr (P _). -by admit. -Qed. +Admitted. Goal forall a : bool, a -> true && a || false && a. by move=> a ?; rewrite [true && _]/= [_ && a]/= orbC [_ || _]//=. diff --git a/mathcomp/ssrtest/set_lamda.v b/mathcomp/ssrtest/set_lamda.v index 51b8e61..432e5d3 100644 --- a/mathcomp/ssrtest/set_lamda.v +++ b/mathcomp/ssrtest/set_lamda.v @@ -1,4 +1,6 @@ -Require Import ssreflect ssrbool eqtype ssrnat ssrfun. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool eqtype ssrnat ssrfun. Set Implicit Arguments. Unset Strict Implicit. diff --git a/mathcomp/ssrtest/set_pattern.v b/mathcomp/ssrtest/set_pattern.v index 0a98267..50dc262 100644 --- a/mathcomp/ssrtest/set_pattern.v +++ b/mathcomp/ssrtest/set_pattern.v @@ -1,4 +1,6 @@ -Require Import ssreflect. +Require Import mathcomp.ssreflect.ssreflect. + +Axiom daemon : False. Ltac myadmit := case: daemon. Ltac T1 x := match goal with |- _ => set t := (x in X in _ = X) end. Ltac T2 x := first [set t := (x in RHS)]. @@ -6,6 +8,7 @@ Ltac T3 x := first [set t := (x in Y in _ = Y)|idtac]. Ltac T4 x := set t := (x in RHS);idtac. Ltac T5 x := match goal with |- _ => set t := (x in RHS) | |- _ => idtac end. +From mathcomp Require Import ssrbool ssrnat. Lemma foo x y : x.+1 = y + x.+1. @@ -31,7 +34,7 @@ rewrite [RHS]addnC. match goal with |- x.+1 = x.+1 + y => rewrite -[RHS]addnC end. rewrite -[in RHS](@subnK 1 x.+1) //. match goal with |- x.+1 = y + (x.+1 - 1 + 1) => rewrite subnK // end. -have H : x.+1 = y by admit. +have H : x.+1 = y by myadmit. set t := _.+1 in H |- *. match goal with H : t = y |- t = y + t => rewrite /t {t} in H * end. set t := (_.+1 in X in _ + X) in H |- *. diff --git a/mathcomp/ssrtest/ssrsyntax1.v b/mathcomp/ssrtest/ssrsyntax1.v index 64e78da..3a5a731 100644 --- a/mathcomp/ssrtest/ssrsyntax1.v +++ b/mathcomp/ssrtest/ssrsyntax1.v @@ -1,5 +1,5 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require ssreflect. +Require mathcomp.ssreflect.ssreflect. Require Import Arith. Goal (forall a b, a + b = b + a). diff --git a/mathcomp/ssrtest/ssrsyntax2.v b/mathcomp/ssrtest/ssrsyntax2.v index 29985b9..c82458d 100644 --- a/mathcomp/ssrtest/ssrsyntax2.v +++ b/mathcomp/ssrtest/ssrsyntax2.v @@ -1,4 +1,5 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) +From mathcomp Require Import ssrsyntax1. Require Import Arith. diff --git a/mathcomp/ssrtest/tc.v b/mathcomp/ssrtest/tc.v index b3f7d3f..7235a9b 100644 --- a/mathcomp/ssrtest/tc.v +++ b/mathcomp/ssrtest/tc.v @@ -1,4 +1,5 @@ -Require Import ssreflect. +Require Import mathcomp.ssreflect.ssreflect. + Class foo (A : Type) : Type := mkFoo { val : A }. Instance foo_pair {A B} {f1 : foo A} {f2 : foo B} : foo (A * B) | 2 := diff --git a/mathcomp/ssrtest/testmx.v b/mathcomp/ssrtest/testmx.v index 931cbad..2704a30 100644 --- a/mathcomp/ssrtest/testmx.v +++ b/mathcomp/ssrtest/testmx.v @@ -1,5 +1,8 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrfun ssrbool eqtype ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype ssrnat. +From mathcomp Require Import ssralg matrix. Set Implicit Arguments. diff --git a/mathcomp/ssrtest/typeof.v b/mathcomp/ssrtest/typeof.v index 8ad81a3..1e1ee8f 100644 --- a/mathcomp/ssrtest/typeof.v +++ b/mathcomp/ssrtest/typeof.v @@ -1,4 +1,5 @@ -Require Import ssreflect. +Require Import mathcomp.ssreflect.ssreflect. + Ltac mycut x := let tx := type of x in cut tx. diff --git a/mathcomp/ssrtest/unkeyed.v b/mathcomp/ssrtest/unkeyed.v index f6b2021..85b224d 100644 --- a/mathcomp/ssrtest/unkeyed.v +++ b/mathcomp/ssrtest/unkeyed.v @@ -1,4 +1,6 @@ -Require Import ssreflect ssrfun ssrbool eqtype. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrfun ssrbool eqtype. Set Implicit Arguments. Unset Strict Implicit. diff --git a/mathcomp/ssrtest/view_case.v b/mathcomp/ssrtest/view_case.v index f6de3df..577182a 100644 --- a/mathcomp/ssrtest/view_case.v +++ b/mathcomp/ssrtest/view_case.v @@ -1,4 +1,6 @@ -Require Import ssreflect ssrbool ssrnat eqtype seq fintype zmodp. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrnat eqtype seq fintype zmodp. Axiom P : forall T, seq T -> Prop. diff --git a/mathcomp/ssrtest/wlog_suff.v b/mathcomp/ssrtest/wlog_suff.v index 4e1c86d..a48e770 100644 --- a/mathcomp/ssrtest/wlog_suff.v +++ b/mathcomp/ssrtest/wlog_suff.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect ssrbool. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool. Lemma test b : b || ~~b. wlog _ : b / b = true. diff --git a/mathcomp/ssrtest/wlogletin.v b/mathcomp/ssrtest/wlogletin.v index 4d20321..1ab1c7c 100644 --- a/mathcomp/ssrtest/wlogletin.v +++ b/mathcomp/ssrtest/wlogletin.v @@ -1,5 +1,7 @@ (* (c) Copyright Microsoft Corporation and Inria. All rights reserved. *) -Require Import ssreflect eqtype ssrbool. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import eqtype ssrbool. Variable T : Type. Variables P : T -> Prop. diff --git a/mathcomp/ssrtest/wlong_intro.v b/mathcomp/ssrtest/wlong_intro.v index 97e378a..977b5eb 100644 --- a/mathcomp/ssrtest/wlong_intro.v +++ b/mathcomp/ssrtest/wlong_intro.v @@ -1,4 +1,6 @@ -Require Import ssreflect ssrbool ssrnat. +Require Import mathcomp.ssreflect.ssreflect. +From mathcomp +Require Import ssrbool ssrnat. Goal (forall x y : nat, True). move=> x y. |
