From 46029938cd88d8e7e7bb5e4ea576ea3d05a2f86d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 15:57:36 +0200 Subject: [PATCH 001/282] Nat: add comments and todos Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 79 +++++++++++++++++++++++++------------- 1 file changed, 52 insertions(+), 27 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index db413c74543..d701fd83a29 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1,34 +1,31 @@ -(* -*- mode: coq; mode: visual-line -*- *) -Require Import Basics. -Require Export Basics.Nat. -Require Import Basics.Classes. +Require Import Basics.Overture Basics.Tactics Basics.PathGroupoids + Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat Basics.Classes. +Export Basics.Nat. Local Set Universe Minimization ToSet. - Local Unset Elimination Schemes. -(** * Theorems about the natural numbers *) - -(** Many of these definitions and proofs have been ported from the coq stdlib. *) - -(** Some results are prefixed with [nat_] and some are not. Should we be more consistent? *) +(** * Natural Numbers *) (** We want to close the trunc_scope so that notations from there don't conflict here. *) Local Close Scope trunc_scope. Local Open Scope nat_scope. +(** TODO: Some results are prefixed with [nat_] and some are not. Should we be more consistent? *) + (** ** Basic operations on naturals *) (** It is common to call [S] [succ] so we add it as a parsing only notation. *) Notation succ := S (only parsing). -(** The predecessor of a natural number. *) +(** [pred n] is the predecessor of a natural number [n]. When [n] is [0] we return [0]. *) Definition pred n : nat := match n with | 0 => n | S n' => n' end. +(** TODO: rename to [nat_add]. *) (** Addition of natural numbers *) Fixpoint add n m : nat := match n with @@ -38,8 +35,10 @@ Fixpoint add n m : nat := Notation "n + m" := (add n m) : nat_scope. +(** TODO: remove *) Definition double n : nat := n + n. +(** TODO: rename to [nat_mul]. *) Fixpoint mul n m : nat := match n with | 0 => 0 @@ -48,6 +47,7 @@ Fixpoint mul n m : nat := Notation "n * m" := (mul n m) : nat_scope. +(** TODO: rename [nat_sub]. *) (** Truncated subtraction: [n - m] is [0] if [n <= m] *) Fixpoint sub n m : nat := match n, m with @@ -57,8 +57,8 @@ Fixpoint sub n m : nat := Notation "n - m" := (sub n m) : nat_scope. -(** ** Minimum, maximum *) - +(** TODO: rename [nat_max]. *) +(** The [max n m] of two natural numbers [n] and [m]. *) Fixpoint max n m := match n, m with | 0 , _ => m @@ -66,6 +66,8 @@ Fixpoint max n m := | S n' , S m' => (max n' m').+1 end. +(** TODO: rename [nat_min]. *) +(** The [min n m] of two natural numbers [n] and [m]. *) Fixpoint min n m := match n, m with | 0 , _ => 0 @@ -73,15 +75,15 @@ Fixpoint min n m := | S n' , S m' => S (min n' m') end. -(** ** Power *) - +(** TODO: rename to [nat_pow]. *) +(** Exponentiation of natural numbers. *) Fixpoint pow n m := match m with | 0 => 1 | S m' => n * (pow n m') end. -(** ** Euclidean division *) +(** *** Euclidean division *) (** This division is linear and tail-recursive. In [divmod], [y] is the predecessor of the actual divisor, and [u] is [y] sub the real remainder. *) @@ -110,7 +112,7 @@ Definition modulo x y : nat := Infix "/" := div : nat_scope. Infix "mod" := modulo : nat_scope. -(** ** Greatest common divisor *) +(** *** Greatest common divisor *) (** We use Euclid algorithm, which is normally not structural, but Coq is now clever enough to accept this (behind modulo there is a subtraction, which now preserves being a subterm) *) @@ -120,11 +122,11 @@ Fixpoint gcd a b := | S a' => gcd (b mod a'.+1) a'.+1 end. -(** ** Square *) +(** *** Square *) Definition square n : nat := n * n. -(** ** Square root *) +(** *** Square root *) (** The following square root function is linear (and tail-recursive). With Peano representation, we can't do better. For faster algorithm, @@ -196,70 +198,87 @@ Fixpoint log2_iter k p q r : nat := Definition log2 n : nat := log2_iter (pred n) 0 1 0. +(** ** Properties of Successors *) + +(** TODO: remove these *) Local Definition ap_S := @ap _ _ S. Local Definition ap_nat := @ap nat. #[export] Hint Resolve ap_S : core. #[export] Hint Resolve ap_nat : core. -Theorem pred_Sn : forall n:nat, n = pred (S n). +(** TODO: remove, this is trivial/by definition *) +Definition pred_Sn : forall n:nat, n = pred (S n). Proof. auto. Defined. (** Injectivity of successor *) - Definition path_nat_S n m (H : S n = S m) : n = m := ap pred H. #[export] Hint Immediate path_nat_S : core. -Theorem not_eq_S : forall n m:nat, n <> m -> S n <> S m. +(** TODO: rename to [neq_S] *) +(** TODO: avoid auto in proof. *) +Definition not_eq_S n m : n <> m -> S n <> S m. Proof. auto. Defined. #[export] Hint Resolve not_eq_S : core. -(** TODO: keep or remove? *) +(** TODO: remove *) Definition IsSucc (n: nat) : Type0 := match n with | O => Empty | S p => Unit end. +(** TODO: rename to [neq_O_S] *) (** Zero is not the successor of a number *) - -Theorem not_eq_O_S : forall n:nat, 0 <> S n. +Definition not_eq_O_S : forall n, 0 <> S n. Proof. discriminate. Defined. #[export] Hint Resolve not_eq_O_S : core. +(** TODO: rename to [neq_n_Sn] *) +(** TODO: prove above using this *) +(** TODO: remove auto. *) Theorem not_eq_n_Sn : forall n:nat, n <> S n. Proof. simple_induction' n; auto. Defined. #[export] Hint Resolve not_eq_n_Sn : core. +(** TODO: remove *) Local Definition ap011_add := @ap011 _ _ _ add. Local Definition ap011_nat := @ap011 nat nat. #[export] Hint Resolve ap011_add : core. #[export] Hint Resolve ap011_nat : core. +(** TODO: rename [nat_add_zero_r] *) +(** TODO: reverse direction *) Lemma add_n_O : forall (n : nat), n = n + 0. Proof. simple_induction' n; simpl; auto. Defined. #[export] Hint Resolve add_n_O : core. +(** TODO: rename [nat_add_zero_l] *) Lemma add_O_n : forall (n : nat), 0 + n = n. Proof. auto. Defined. +(** TODO: rename [nat_add_succ_r] *) +(** TODO: reverse direction *) Lemma add_n_Sm : forall n m:nat, S (n + m) = n + S m. Proof. simple_induction' n; simpl; auto. Defined. #[export] Hint Resolve add_n_Sm: core. +(** TODO: rename [nat_add_succ_l] *) +(** TODO: reverse direction *) +(** TODO: remove auto *) Lemma add_Sn_m : forall n m:nat, S n + m = S (n + m). Proof. auto. @@ -267,15 +286,20 @@ Defined. (** Multiplication *) +(** TODO: remove? *) Local Definition ap011_mul := @ap011 _ _ _ mul. #[export] Hint Resolve ap011_mul : core. +(** TODO: rename [nat_mul_zero_r] *) +(** TODO: reverse direction *) Lemma mul_n_O : forall n:nat, 0 = n * 0. Proof. simple_induction' n; simpl; auto. Defined. #[export] Hint Resolve mul_n_O : core. +(** TODO: rename [nat_mul_succ_r] *) +(** TODO: reverse direction *) Lemma mul_n_Sm : forall n m:nat, n * m + n = n * S m. Proof. intros; simple_induction n p H; simpl; auto. @@ -286,6 +310,7 @@ Defined. (** Standard associated names *) +(** TODO: remove? *) Notation mul_0_r_reverse := mul_n_O (only parsing). Notation mul_succ_r_reverse := mul_n_Sm (only parsing). @@ -541,7 +566,7 @@ Proof. intros n m; revert n; simple_induction m m IHm; auto. intros [] p. 1: inversion p. - cbn; by apply ap_S, IHm, leq_S_n. + cbn; by apply (ap S), IHm, leq_S_n. Defined. Theorem max_r : forall n m : nat, n <= m -> max n m = m. @@ -559,7 +584,7 @@ Proof. simple_induction n n IHn; auto. intros [] p. 1: inversion p. - cbn; by apply ap_S, IHn, leq_S_n. + cbn; by apply (ap S), IHn, leq_S_n. Defined. Theorem min_r : forall n m : nat, m <= n -> min n m = m. From 7129614cba044351946d62269bc5a98bf5945770 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 16:00:28 +0200 Subject: [PATCH 002/282] Nat: rename add -> nat_add Signed-off-by: Ali Caglayan --- .../Classes/implementations/peano_naturals.v | 6 +++--- theories/Spaces/Finite/Finite.v | 2 +- theories/Spaces/Nat/Core.v | 19 +++++++++---------- theories/Spaces/Pos/Core.v | 4 ++-- 4 files changed, 15 insertions(+), 16 deletions(-) diff --git a/theories/Classes/implementations/peano_naturals.v b/theories/Classes/implementations/peano_naturals.v index cfb2e947874..8a464ae6dd0 100644 --- a/theories/Classes/implementations/peano_naturals.v +++ b/theories/Classes/implementations/peano_naturals.v @@ -28,17 +28,17 @@ Defined. Global Instance nat_0: Zero@{N} nat := 0%nat. Global Instance nat_1: One@{N} nat := 1%nat. -Global Instance nat_plus: Plus@{N} nat := Nat.Core.add. +Global Instance nat_plus: Plus@{N} nat := Nat.Core.nat_add. Notation mul := Nat.Core.mul. Global Instance nat_mult: Mult@{N} nat := Nat.Core.mul. Ltac simpl_nat := - change (@plus nat _) with Nat.Core.add; + change (@plus nat _) with Nat.Core.nat_add; change (@mult nat _) with Nat.Core.mul; simpl; - change Nat.Core.add with (@plus nat Nat.Core.add); + change Nat.Core.nat_add with (@plus nat Nat.Core.nat_add); change Nat.Core.mul with (@mult nat Nat.Core.mul). (** [0 + a =N= a] *) diff --git a/theories/Spaces/Finite/Finite.v b/theories/Spaces/Finite/Finite.v index 87d0afe7004..c6bef58056a 100644 --- a/theories/Spaces/Finite/Finite.v +++ b/theories/Spaces/Finite/Finite.v @@ -300,7 +300,7 @@ Proof. refine (fcard_sum _ _ @ _). simpl. refine (_ @ nat_add_comm _ _). - refine (ap011 add _ _). + refine (ap011 nat_add _ _). + apply IH. + apply fcard_equiv', prod_unit_l. Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index d701fd83a29..2ef582381e3 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -25,15 +25,14 @@ Definition pred n : nat := | S n' => n' end. -(** TODO: rename to [nat_add]. *) (** Addition of natural numbers *) -Fixpoint add n m : nat := +Fixpoint nat_add n m : nat := match n with | 0 => m - | S n' => S (add n' m) + | S n' => S (nat_add n' m) end. -Notation "n + m" := (add n m) : nat_scope. +Notation "n + m" := (nat_add n m) : nat_scope. (** TODO: remove *) Definition double n : nat := n + n. @@ -249,10 +248,10 @@ Defined. #[export] Hint Resolve not_eq_n_Sn : core. (** TODO: remove *) -Local Definition ap011_add := @ap011 _ _ _ add. -Local Definition ap011_nat := @ap011 nat nat. -#[export] Hint Resolve ap011_add : core. -#[export] Hint Resolve ap011_nat : core. +(* Local Definition ap011_add := @ap011 _ _ _ add. *) +(* Local Definition ap011_nat := @ap011 nat nat. *) +(* #[export] Hint Resolve ap011_add : core. *) +(* #[export] Hint Resolve ap011_nat : core. *) (** TODO: rename [nat_add_zero_r] *) (** TODO: reverse direction *) @@ -647,12 +646,12 @@ Proof. by apply path_nat_S. Defined. -Global Instance isinj_nat_add_l@{} k : IsInjective (add k). +Global Instance isinj_nat_add_l@{} k : IsInjective (nat_add k). Proof. simple_induction k k Ik; exact _. Defined. -Definition isinj_nat_add_r@{} k : IsInjective (fun x => add x k). +Definition isinj_nat_add_r@{} k : IsInjective (fun x =>nat_add x k). Proof. intros x y H. rewrite 2 (nat_add_comm _ k) in H. diff --git a/theories/Spaces/Pos/Core.v b/theories/Spaces/Pos/Core.v index 9222c2bbcfb..618ff236d45 100644 --- a/theories/Spaces/Pos/Core.v +++ b/theories/Spaces/Pos/Core.v @@ -420,8 +420,8 @@ Definition pos_to_nat : Pos -> nat. Proof. intro p. induction p. + exact (S O). - + exact (add IHp IHp). - + exact (S (add IHp IHp)). + + exact (nat_add IHp IHp). + + exact (S (nat_add IHp IHp)). Defined. Number Notation Pos pos_of_number_uint pos_to_number_uint : positive_scope. From 2061fde014401cdf2034138563a346b6e1a7cc0a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 16:36:01 +0200 Subject: [PATCH 003/282] Nat: rename mul -> nat_mul Signed-off-by: Ali Caglayan --- theories/Classes/implementations/peano_naturals.v | 8 ++++---- theories/Classes/interfaces/abstract_algebra.v | 2 +- theories/Spaces/Finite/Finite.v | 2 +- theories/Spaces/Nat/Core.v | 10 +++++----- 4 files changed, 11 insertions(+), 11 deletions(-) diff --git a/theories/Classes/implementations/peano_naturals.v b/theories/Classes/implementations/peano_naturals.v index 8a464ae6dd0..c76ba85016c 100644 --- a/theories/Classes/implementations/peano_naturals.v +++ b/theories/Classes/implementations/peano_naturals.v @@ -30,16 +30,16 @@ Global Instance nat_1: One@{N} nat := 1%nat. Global Instance nat_plus: Plus@{N} nat := Nat.Core.nat_add. -Notation mul := Nat.Core.mul. +Notation mul := Nat.Core.nat_mul. -Global Instance nat_mult: Mult@{N} nat := Nat.Core.mul. +Global Instance nat_mult: Mult@{N} nat := Nat.Core.nat_mul. Ltac simpl_nat := change (@plus nat _) with Nat.Core.nat_add; - change (@mult nat _) with Nat.Core.mul; + change (@mult nat _) with Nat.Core.nat_mul; simpl; change Nat.Core.nat_add with (@plus nat Nat.Core.nat_add); - change Nat.Core.mul with (@mult nat Nat.Core.mul). + change Nat.Core.nat_mul with (@mult nat Nat.Core.nat_mul). (** [0 + a =N= a] *) Local Instance add_0_l : LeftIdentity@{N N} (plus : Plus nat) 0 := fun _ => idpath. diff --git a/theories/Classes/interfaces/abstract_algebra.v b/theories/Classes/interfaces/abstract_algebra.v index 0f1ecf8c514..674a5404216 100644 --- a/theories/Classes/interfaces/abstract_algebra.v +++ b/theories/Classes/interfaces/abstract_algebra.v @@ -203,7 +203,7 @@ Section upper_classes. := field_characteristic : forall n : nat, Nat.Core.lt 0 n -> iff@{j j j} (forall m : nat, not@{j} (paths@{Set} n - (Nat.Core.mul k m))) + (nat_mul k m))) (@apart A Aap (nat_iter n (1 +) 0) 0). End upper_classes. diff --git a/theories/Spaces/Finite/Finite.v b/theories/Spaces/Finite/Finite.v index c6bef58056a..5880524bad8 100644 --- a/theories/Spaces/Finite/Finite.v +++ b/theories/Spaces/Finite/Finite.v @@ -344,7 +344,7 @@ Proof. - reflexivity. - refine (fcard_equiv (equiv_sum_ind (fun (_:Fin n.+1) => Y))^-1 @ _). refine (fcard_prod _ _ @ _). - apply (ap011 mul). + apply (ap011 nat_mul). + assumption. + refine (fcard_equiv (@Unit_ind (fun (_:Unit) => Y))^-1). Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 2ef582381e3..cd4433be350 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -37,14 +37,14 @@ Notation "n + m" := (nat_add n m) : nat_scope. (** TODO: remove *) Definition double n : nat := n + n. -(** TODO: rename to [nat_mul]. *) -Fixpoint mul n m : nat := +(** Multiplication of natural numbers *) +Fixpoint nat_mul n m : nat := match n with | 0 => 0 - | S n' => m + (mul n' m) + | S n' => m + (nat_mul n' m) end. -Notation "n * m" := (mul n m) : nat_scope. +Notation "n * m" := (nat_mul n m) : nat_scope. (** TODO: rename [nat_sub]. *) (** Truncated subtraction: [n - m] is [0] if [n <= m] *) @@ -286,7 +286,7 @@ Defined. (** Multiplication *) (** TODO: remove? *) -Local Definition ap011_mul := @ap011 _ _ _ mul. +Local Definition ap011_mul := @ap011 _ _ _ nat_mul. #[export] Hint Resolve ap011_mul : core. (** TODO: rename [nat_mul_zero_r] *) From 9924bb22fafa1df5b700931418af6aaaed6dd21c Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 16:36:47 +0200 Subject: [PATCH 004/282] Nat: remove double Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 3 --- 1 file changed, 3 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index cd4433be350..abb25131109 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -34,9 +34,6 @@ Fixpoint nat_add n m : nat := Notation "n + m" := (nat_add n m) : nat_scope. -(** TODO: remove *) -Definition double n : nat := n + n. - (** Multiplication of natural numbers *) Fixpoint nat_mul n m : nat := match n with From 8d40b53e45c077a45688f816da9ad37afcadbf20 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 16:39:20 +0200 Subject: [PATCH 005/282] Nat: rename sub -> nat_sub Signed-off-by: Ali Caglayan --- theories/Classes/implementations/peano_naturals.v | 2 +- theories/Spaces/Nat/Core.v | 7 +++---- 2 files changed, 4 insertions(+), 5 deletions(-) diff --git a/theories/Classes/implementations/peano_naturals.v b/theories/Classes/implementations/peano_naturals.v index c76ba85016c..124404d1301 100644 --- a/theories/Classes/implementations/peano_naturals.v +++ b/theories/Classes/implementations/peano_naturals.v @@ -672,7 +672,7 @@ intros;apply toR_unique, _. Qed. Global Existing Instance nat_naturals. -Global Instance nat_cut_minus: CutMinus@{N} nat := Nat.Core.sub. +Global Instance nat_cut_minus: CutMinus@{N} nat := Nat.Core.nat_sub. Lemma plus_minus : forall a b, cut_minus (a + b) b =N= a. Proof. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index abb25131109..06308f33900 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -43,15 +43,14 @@ Fixpoint nat_mul n m : nat := Notation "n * m" := (nat_mul n m) : nat_scope. -(** TODO: rename [nat_sub]. *) (** Truncated subtraction: [n - m] is [0] if [n <= m] *) -Fixpoint sub n m : nat := +Fixpoint nat_sub n m : nat := match n, m with - | S n' , S m' => sub n' m' + | S n' , S m' => nat_sub n' m' | _ , _ => n end. -Notation "n - m" := (sub n m) : nat_scope. +Notation "n - m" := (nat_sub n m) : nat_scope. (** TODO: rename [nat_max]. *) (** The [max n m] of two natural numbers [n] and [m]. *) From 65a4a41c1d540e97faca664f44a6e78c13c30882 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 16:42:35 +0200 Subject: [PATCH 006/282] Nat: rename max -> nat_max Signed-off-by: Ali Caglayan --- .../Classes/implementations/peano_naturals.v | 4 +-- theories/Classes/interfaces/cauchy.v | 2 +- theories/Spaces/Nat/Core.v | 33 +++++++++---------- 3 files changed, 19 insertions(+), 20 deletions(-) diff --git a/theories/Classes/implementations/peano_naturals.v b/theories/Classes/implementations/peano_naturals.v index 124404d1301..5ddbc9a47a2 100644 --- a/theories/Classes/implementations/peano_naturals.v +++ b/theories/Classes/implementations/peano_naturals.v @@ -574,7 +574,7 @@ Definition nat_full@{} := ltac:(first[exact nat_full'@{Ularge Ularge}| exact nat_full'@{}]). Local Existing Instance nat_full. -Lemma le_nat_max_l n m : n <= Nat.Core.max n m. +Lemma le_nat_max_l n m : n <= Nat.Core.nat_max n m. Proof. revert m. induction n as [|n' IHn]; @@ -582,7 +582,7 @@ Proof. - apply zero_least. - apply le_S_S. exact (IHn m'). Qed. -Lemma le_nat_max_r n m : m <= Nat.Core.max n m. +Lemma le_nat_max_r n m : m <= Nat.Core.nat_max n m. Proof. revert m. induction n as [|n' IHn]; diff --git a/theories/Classes/interfaces/cauchy.v b/theories/Classes/interfaces/cauchy.v index 556eac1d5f2..47df617e642 100644 --- a/theories/Classes/interfaces/cauchy.v +++ b/theories/Classes/interfaces/cauchy.v @@ -53,7 +53,7 @@ Section cauchy. assert (lim_close := is_limit x (epsilon / 2)); strip_truncations. destruct lim_close as [N isclose']. - set (n := Nat.Core.max (M (epsilon / 2)) N). + set (n := Nat.Core.nat_max (M (epsilon / 2)) N). assert (leNn := le_nat_max_r (M (epsilon / 2)) N : N ≤ n). assert (isclose := isclose' n leNn). clear isclose'. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 06308f33900..438fda871be 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -52,13 +52,12 @@ Fixpoint nat_sub n m : nat := Notation "n - m" := (nat_sub n m) : nat_scope. -(** TODO: rename [nat_max]. *) -(** The [max n m] of two natural numbers [n] and [m]. *) -Fixpoint max n m := +(** The maximum [nat_max n m] of two natural numbers [n] and [m]. *) +Fixpoint nat_max n m := match n, m with | 0 , _ => m | S n' , 0 => n'.+1 - | S n' , S m' => (max n' m').+1 + | S n' , S m' => (nat_max n' m').+1 end. (** TODO: rename [nat_min]. *) @@ -527,36 +526,36 @@ Defined. (** Maximum and minimum : definitions and specifications *) -Lemma max_n_n n : max n n = n. +Lemma nat_max_n_n n : nat_max n n = n. Proof. simple_induction' n; cbn; auto. Defined. -#[export] Hint Resolve max_n_n : core. +#[export] Hint Resolve nat_max_n_n : core. -Lemma max_Sn_n n : max (S n) n = S n. +Lemma nat_max_Sn_n n : nat_max (S n) n = S n. Proof. simple_induction' n; cbn; auto. Defined. -#[export] Hint Resolve max_Sn_n : core. +#[export] Hint Resolve nat_max_Sn_n : core. -Lemma max_comm n m : max n m = max m n. +Lemma nat_max_comm n m : nat_max n m = nat_max m n. Proof. revert m; simple_induction' n; destruct m; cbn; auto. Defined. -Lemma max_0_n n : max 0 n = n. +Lemma nat_max_0_n n : nat_max 0 n = n. Proof. auto. Defined. -#[export] Hint Resolve max_0_n : core. +#[export] Hint Resolve nat_max_0_n : core. -Lemma max_n_0 n : max n 0 = n. +Lemma nat_max_n_0 n : nat_max n 0 = n. Proof. - by rewrite max_comm. + by rewrite nat_max_comm. Defined. -#[export] Hint Resolve max_n_0 : core. +#[export] Hint Resolve nat_max_n_0 : core. -Theorem max_l : forall n m, m <= n -> max n m = n. +Theorem nat_max_l : forall n m, m <= n -> nat_max n m = n. Proof. intros n m; revert n; simple_induction m m IHm; auto. intros [] p. @@ -564,9 +563,9 @@ Proof. cbn; by apply (ap S), IHm, leq_S_n. Defined. -Theorem max_r : forall n m : nat, n <= m -> max n m = m. +Theorem nat_max_r : forall n m : nat, n <= m -> nat_max n m = m. Proof. - intros; rewrite max_comm; by apply max_l. + intros; rewrite nat_max_comm; by apply nat_max_l. Defined. Lemma min_comm : forall n m, min n m = min m n. From cb2f76decb608735c39728a2c555ea779661e122 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 16:46:33 +0200 Subject: [PATCH 007/282] Nat: rename min -> nat_min Signed-off-by: Ali Caglayan --- theories/Spaces/List/Theory.v | 6 +++--- theories/Spaces/Nat/Core.v | 15 +++++++-------- 2 files changed, 10 insertions(+), 11 deletions(-) diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index 5c18da7f420..5ac8ade7f58 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -600,11 +600,11 @@ Defined. (** The length of a [take n] is the minimum of [n] and the length of the original list. *) Definition length_take@{i|} {A : Type@{i}} (n : nat) (l : list A) - : length (take n l) = min n (length l). + : length (take n l) = nat_min n (length l). Proof. induction l as [|a l IHl] in n |- * using list_ind@{i i}. { rewrite take_nil. - rewrite min_r. + rewrite nat_min_r. 1: reflexivity. cbn; exact _. } destruct n. @@ -661,7 +661,7 @@ Proof. rewrite length_app@{i}. rewrite length_take. rewrite length_drop. - rewrite min_l. + rewrite nat_min_l. - rewrite nataddsub_assoc. 2: exact H. rewrite <- predeqminus1. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 438fda871be..ef9053208ab 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -60,13 +60,12 @@ Fixpoint nat_max n m := | S n' , S m' => (nat_max n' m').+1 end. -(** TODO: rename [nat_min]. *) -(** The [min n m] of two natural numbers [n] and [m]. *) -Fixpoint min n m := +(** The minimum [nat_min n m] of two natural numbers [n] and [m]. *) +Fixpoint nat_min n m := match n, m with | 0 , _ => 0 | S n' , 0 => 0 - | S n' , S m' => S (min n' m') + | S n' , S m' => S (nat_min n' m') end. (** TODO: rename to [nat_pow]. *) @@ -568,12 +567,12 @@ Proof. intros; rewrite nat_max_comm; by apply nat_max_l. Defined. -Lemma min_comm : forall n m, min n m = min m n. +Lemma nat_min_comm : forall n m, nat_min n m = nat_min m n. Proof. simple_induction' n; destruct m; cbn; auto. Defined. -Theorem min_l : forall n m : nat, n <= m -> min n m = n. +Theorem nat_min_l : forall n m : nat, n <= m -> nat_min n m = n. Proof. simple_induction n n IHn; auto. intros [] p. @@ -581,9 +580,9 @@ Proof. cbn; by apply (ap S), IHn, leq_S_n. Defined. -Theorem min_r : forall n m : nat, m <= n -> min n m = m. +Theorem nat_min_r : forall n m : nat, m <= n -> nat_min n m = m. Proof. - intros; rewrite min_comm; by apply min_l. + intros; rewrite nat_min_comm; by apply nat_min_l. Defined. (** [n]th iteration of the function [f : A -> A]. We have definitional equalities [nat_iter 0 f x = x] and [nat_iter n.+1 f x = f (nat_iter n f x)]. We make this a notation, so it doesn't add a universe variable for the universe containing [A]. *) From 21e2a2d4ce4f6f486f67ba44fe8c6294bc674e49 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 16:48:29 +0200 Subject: [PATCH 008/282] Nat: rename succ -> nat_succ Signed-off-by: Ali Caglayan --- theories/Homotopy/SuccessorStructure.v | 2 +- theories/Spaces/Nat/Core.v | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/theories/Homotopy/SuccessorStructure.v b/theories/Homotopy/SuccessorStructure.v index a7f26cd70e3..66e8997b752 100644 --- a/theories/Homotopy/SuccessorStructure.v +++ b/theories/Homotopy/SuccessorStructure.v @@ -27,7 +27,7 @@ Arguments ss_succ {_} _. Notation "x .+1" := (ss_succ x) : succ_scope. (** Successor structure of naturals *) -Definition NatSucc : SuccStr := Build_SuccStr nat Nat.Core.succ. +Definition NatSucc : SuccStr := Build_SuccStr nat nat_succ. (** Successor structure of integers *) Definition BinIntSucc : SuccStr := Build_SuccStr Int int_succ. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ef9053208ab..ae51a4f1ff9 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -15,8 +15,8 @@ Local Open Scope nat_scope. (** ** Basic operations on naturals *) -(** It is common to call [S] [succ] so we add it as a parsing only notation. *) -Notation succ := S (only parsing). +(** [nat_succ n] is the successor of a natural number [n]. *) +Notation nat_succ := S (only parsing). (** [pred n] is the predecessor of a natural number [n]. When [n] is [0] we return [0]. *) Definition pred n : nat := From 32bd59c9051c026af0a0dfc927e15ede6475e8b3 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 16:53:18 +0200 Subject: [PATCH 009/282] Nat: rename pred -> nat_pred Signed-off-by: Ali Caglayan --- theories/Algebra/Universal/Operation.v | 2 +- theories/Spaces/Finite/FinSeq.v | 16 +++++++-------- theories/Spaces/Int.v | 4 ++-- theories/Spaces/List/Theory.v | 4 ++-- theories/Spaces/Nat/Arithmetic.v | 28 +++++++++++++------------- theories/Spaces/Nat/Core.v | 12 +++++------ 6 files changed, 33 insertions(+), 33 deletions(-) diff --git a/theories/Algebra/Universal/Operation.v b/theories/Algebra/Universal/Operation.v index e3c27f4ad2f..901ee9092b1 100644 --- a/theories/Algebra/Universal/Operation.v +++ b/theories/Algebra/Universal/Operation.v @@ -33,7 +33,7 @@ Monomorphic Definition head_dom {σ} (A : Carriers σ) {n : nat} of an operation domain [a : forall i, A (ss i)]. *) Monomorphic Definition tail_dom' {σ} (A : Carriers σ) (n : nat) - : forall (ss : FinSeq n (Sort σ)) (a : forall i, A (ss i)) (i : Fin (pred n)), + : forall (ss : FinSeq n (Sort σ)) (a : forall i, A (ss i)) (i : Fin (nat_pred n)), A (fstail' n ss i) := match n with | 0 => fun ss _ i => Empty_rec i diff --git a/theories/Spaces/Finite/FinSeq.v b/theories/Spaces/Finite/FinSeq.v index 330061197d3..eed66355d50 100644 --- a/theories/Spaces/Finite/FinSeq.v +++ b/theories/Spaces/Finite/FinSeq.v @@ -27,9 +27,9 @@ Defined. (** Add an element in the end of a finite sequence, [fscons'] and [fscons]. *) -Definition fscons' {A : Type} (n : nat) (a : A) (v : FinSeq (pred n) A) +Definition fscons' {A : Type} (n : nat) (a : A) (v : FinSeq (nat_pred n) A) : FinSeq n A - := fun i => fin_rec (fun n => FinSeq (pred n) A -> A) + := fun i => fin_rec (fun n => FinSeq (nat_pred n) A -> A) (fun _ _ => a) (fun n' i _ v => v i) i v. Definition fscons {A : Type} {n : nat} : A -> FinSeq n A -> FinSeq n.+1 A @@ -46,7 +46,7 @@ Definition fshead' {A} (n : nat) : 0 < n -> FinSeq n A -> A Definition fshead {A} {n : nat} : FinSeq n.+1 A -> A := fshead' n.+1 _. -Definition compute_fshead' {A} n (N : n > 0) (a : A) (v : FinSeq (pred n) A) +Definition compute_fshead' {A} n (N : n > 0) (a : A) (v : FinSeq (nat_pred n) A) : fshead' n N (fscons' n a v) = a. Proof. destruct n; [elim (not_lt_n_n _ N)|]. @@ -61,7 +61,7 @@ Defined. (** If the sequence is non-empty, then remove the first element. *) -Definition fstail' {A} (n : nat) : FinSeq n A -> FinSeq (pred n) A +Definition fstail' {A} (n : nat) : FinSeq n A -> FinSeq (nat_pred n) A := match n with | 0 => fun _ => Empty_rec | n'.+1 => fun v i => v (fsucc i) @@ -71,7 +71,7 @@ Definition fstail' {A} (n : nat) : FinSeq n A -> FinSeq (pred n) A Definition fstail {A} {n : nat} : FinSeq n.+1 A -> FinSeq n A := fstail' n.+1. -Definition compute_fstail' {A} n (a : A) (v : FinSeq (pred n) A) +Definition compute_fstail' {A} n (a : A) (v : FinSeq (nat_pred n) A) : fstail' n (fscons' n a v) == v. Proof. intro i. @@ -109,7 +109,7 @@ Defined. a path between [fscons] finite sequences. They cooperate nicely with [path_expand_fscons'] and [path_expand_fscons]. *) -Definition path_fscons' {A} n {a1 a2 : A} {v1 v2 : FinSeq (pred n) A} +Definition path_fscons' {A} n {a1 a2 : A} {v1 v2 : FinSeq (nat_pred n) A} (p : a1 = a2) (q : forall i, v1 i = v2 i) (i : Fin n) : fscons' n a1 v1 i = fscons' n a2 v2 i. Proof. @@ -120,7 +120,7 @@ Proof. Defined. Definition compute_path_fscons' {A} (n : nat) - (a : A) (v : FinSeq (pred n) A) (i : Fin n) + (a : A) (v : FinSeq (nat_pred n) A) (i : Fin n) : path_fscons' n (idpath a) (fun j => idpath (v j)) i = idpath. Proof. induction i using fin_ind; unfold path_fscons'. @@ -153,7 +153,7 @@ Defined. [path_expand_fscons] with [path_fscons]. *) Lemma path_expand_fscons_fscons' {A : Type} (n : nat) - (N : n > 0) (a : A) (v : FinSeq (pred n) A) (i : Fin n) + (N : n > 0) (a : A) (v : FinSeq (nat_pred n) A) (i : Fin n) : path_expand_fscons' n i N (fscons' n a v) = path_fscons' n (compute_fshead' n N a v) (compute_fstail' n a v) i. Proof. diff --git a/theories/Spaces/Int.v b/theories/Spaces/Int.v index c38e248c7c0..a352b33d5a8 100644 --- a/theories/Spaces/Int.v +++ b/theories/Spaces/Int.v @@ -45,9 +45,9 @@ Definition int_to_number_int (n : Int) : Numeral.int := Definition int_of_number_int (d : Numeral.int) := match d with | Numeral.IntDec (Decimal.Pos d) => int_of_nat (Nat.of_uint d) - | Numeral.IntDec (Decimal.Neg d) => negS (pred (Nat.of_uint d)) + | Numeral.IntDec (Decimal.Neg d) => negS (nat_pred (Nat.of_uint d)) | Numeral.IntHex (Hexadecimal.Pos u) => int_of_nat (Nat.of_hex_uint u) - | Numeral.IntHex (Hexadecimal.Neg u) => negS (pred (Nat.of_hex_uint u)) + | Numeral.IntHex (Hexadecimal.Neg u) => negS (nat_pred (Nat.of_hex_uint u)) end. Number Notation Int int_of_number_int int_to_number_int : int_scope. diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index 5ac8ade7f58..1a8cd44420d 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -469,7 +469,7 @@ Proof. Defined. (** The [nth i] element where [pred (length l) = i] is the last element of the list. *) -Definition nth_last {A : Type} (l : list A) (i : nat) (p : pred (length l) = i) +Definition nth_last {A : Type} (l : list A) (i : nat) (p : nat_pred (length l) = i) : nth l i = last l. Proof. destruct p. @@ -655,7 +655,7 @@ Defined. (** The length of a [remove n] is the length of the original list minus one. *) Definition length_remove@{i|} {A : Type@{i}} (n : nat) (l : list A) (H : (n < length l)%nat) - : length (remove n l) = pred (length l)%nat. + : length (remove n l) = nat_pred (length l)%nat. Proof. unfold remove. rewrite length_app@{i}. diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 2bb6df21370..ee0396c0d42 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -577,21 +577,21 @@ Defined. #[export] Hint Resolve nat_add_sub_eq : nat. -Proposition predeqminus1 { n : nat } : n - 1 = pred n. +Proposition predeqminus1 { n : nat } : n - 1 = nat_pred n. Proof. simple_induction' n. - reflexivity. - apply sub_n_0. Defined. -Proposition predn_leq_n (n : nat) : pred n <= n. +Proposition predn_leq_n (n : nat) : nat_pred n <= n. Proof. case n; [ apply leq_n | intro; apply leq_S; apply leq_n]. Defined. #[export] Hint Resolve predn_leq_n : nat. -Proposition S_predn (n i: nat) : (i < n) -> S(pred n) = n. +Proposition S_predn (n i: nat) : (i < n) -> S(nat_pred n) = n. Proof. simple_induction' n; intros l. - contradiction (not_lt_n_0 i). @@ -604,20 +604,20 @@ Defined. #[export] Hint Resolve S_predn : nat. #[export] Hint Resolve leq_n_pred : nat. -Proposition pred_equiv (k n : nat) : k < n -> k < S (pred n). +Proposition pred_equiv (k n : nat) : k < n -> k < S (nat_pred n). Proof. intro ineq; destruct n. - contradiction (not_lt_n_0 _ _). - assumption. Defined. -Proposition n_leq_pred_Sn (n : nat) : n <= S (pred n). +Proposition n_leq_pred_Sn (n : nat) : n <= S (nat_pred n). Proof. destruct n; auto. Defined. Proposition leq_implies_pred_lt (i n k : nat) - : (n > i) -> n <= k -> pred n < k. + : (n > i) -> n <= k -> nat_pred n < k. Proof. intro ineq; destruct n. - contradiction (not_lt_n_0 i). @@ -625,14 +625,14 @@ Proof. Defined. Proposition pred_lt_implies_leq (n k : nat) - : pred n < k -> n <= k. + : nat_pred n < k -> n <= k. Proof. intro l; destruct n. - apply leq_0_n. - assumption. Defined. -Proposition lt_implies_pred_geq (i j : nat) : i < j -> i <= pred j. +Proposition lt_implies_pred_geq (i j : nat) : i < j -> i <= nat_pred j. Proof. intro l; apply leq_n_pred in l; assumption. Defined. @@ -640,7 +640,7 @@ Defined. #[export] Hint Resolve lt_implies_pred_geq : nat. Proposition j_geq_0_lt_implies_pred_geq (i j k : nat) - : i < j -> k.+1 <= j -> k <= pred j. + : i < j -> k.+1 <= j -> k <= nat_pred j. Proof. intros l ineq. destruct j. @@ -651,11 +651,11 @@ Defined. #[export] Hint Resolve lt_implies_pred_geq : nat. Proposition pred_gt_implies_lt (i j : nat) - : i < pred j -> i.+1 < j. + : i < nat_pred j -> i.+1 < j. Proof. intros ineq. assert (H := leq_S_n' _ _ ineq). assert (i < j) as X. { - apply (@mixed_trans2 _ (pred j) _); + apply (@mixed_trans2 _ (nat_pred j) _); [assumption | apply predn_leq_n]. } destruct (symmetric_paths _ _ (S_predn _ _ X)) in H. @@ -663,7 +663,7 @@ Proof. Defined. Proposition pred_preserves_lt {i n: nat} (p : i < n) m - : (n < m) -> (pred n < pred m). + : (n < m) -> (nat_pred n < nat_pred m). Proof. intro l. apply leq_S_n. destruct (symmetric_paths _ _ (S_predn n i _)). @@ -888,9 +888,9 @@ Proof. (destruct (nat_add_comm n 1)). destruct (symmetric_paths _ _ (subsubadd m n 1)). destruct (S_predn (m - n) 0 (lt_sub_gt_0 _ _ ineq)); simpl; - destruct (symmetric_paths _ _ (sub_n_0 (pred (m - n)))). + destruct (symmetric_paths _ _ (sub_n_0 (nat_pred (m - n)))). assert (0 < m - n) as dp by exact (lt_sub_gt_0 _ _ ineq). - assert (pred (m - n) < m) as sh by + assert (nat_pred (m - n) < m) as sh by ( unfold "<"; destruct (symmetric_paths _ _ (S_predn _ 0 _)); exact sub_less). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ae51a4f1ff9..d1ba0ea3b9d 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -18,8 +18,8 @@ Local Open Scope nat_scope. (** [nat_succ n] is the successor of a natural number [n]. *) Notation nat_succ := S (only parsing). -(** [pred n] is the predecessor of a natural number [n]. When [n] is [0] we return [0]. *) -Definition pred n : nat := +(** [nat_pred n] is the predecessor of a natural number [n]. When [n] is [0] we return [0]. *) +Definition nat_pred n : nat := match n with | 0 => n | S n' => n' @@ -189,7 +189,7 @@ Fixpoint log2_iter k p q r : nat := end end. -Definition log2 n : nat := log2_iter (pred n) 0 1 0. +Definition log2 n : nat := log2_iter (nat_pred n) 0 1 0. (** ** Properties of Successors *) @@ -200,13 +200,13 @@ Local Definition ap_nat := @ap nat. #[export] Hint Resolve ap_nat : core. (** TODO: remove, this is trivial/by definition *) -Definition pred_Sn : forall n:nat, n = pred (S n). +Definition pred_Sn : forall n:nat, n = nat_pred (S n). Proof. auto. Defined. (** Injectivity of successor *) -Definition path_nat_S n m (H : S n = S m) : n = m := ap pred H. +Definition path_nat_S n m (H : S n = S m) : n = m := ap nat_pred H. #[export] Hint Immediate path_nat_S : core. (** TODO: rename to [neq_S] *) @@ -353,7 +353,7 @@ Defined. Global Instance transitive_leq : Transitive leq := @leq_trans. -Lemma leq_n_pred n m : leq n m -> leq (pred n) (pred m). +Lemma leq_n_pred n m : leq n m -> leq (nat_pred n) (nat_pred m). Proof. induction 1; auto. destruct m; simpl; auto. From 0342c09c66d2ce376e4cbc2bc0829a53c41ee68f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 16:55:51 +0200 Subject: [PATCH 010/282] Nat: rename pow -> nat_pow Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index d1ba0ea3b9d..8b33fea490b 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -68,12 +68,12 @@ Fixpoint nat_min n m := | S n' , S m' => S (nat_min n' m') end. -(** TODO: rename to [nat_pow]. *) +(** TODO: redefine in terms of nat_iter *) (** Exponentiation of natural numbers. *) -Fixpoint pow n m := +Fixpoint nat_pow n m := match m with | 0 => 1 - | S m' => n * (pow n m') + | S m' => n * (nat_pow n m') end. (** *** Euclidean division *) From f72821de0b2da30c5e0e484d484e19e801b15960 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 16:57:11 +0200 Subject: [PATCH 011/282] Nat: remove log2 Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 45 -------------------------------------- 1 file changed, 45 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 8b33fea490b..3af552c3100 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -146,51 +146,6 @@ Fixpoint sqrt_iter k p q r : nat := Definition sqrt n : nat := sqrt_iter n 0 0 0. -(** ** Log2 *) - -(** This base-2 logarithm is linear and tail-recursive. - - In [log2_iter], we maintain the logarithm [p] of the counter [q], - while [r] is the distance between [q] and the next power of 2, - more precisely [q + S r = 2^(S p)] and [r<2^p]. At each - recursive call, [q] goes up while [r] goes down. When [r] - is 0, we know that [q] has almost reached a power of 2, - and we increase [p] at the next call, while resetting [r] - to [q]. - - Graphically (numbers are [q], stars are [r]) : - -<< - 10 - 9 - 8 - 7 * - 6 * - 5 ... - 4 - 3 * - 2 * - 1 * * -0 * * * ->> - - We stop when [k], the global downward counter reaches 0. - At that moment, [q] is the number we're considering (since - [k+q] is invariant), and [p] its logarithm. -*) - -Fixpoint log2_iter k p q r : nat := - match k with - | O => p - | S k' => - match r with - | O => log2_iter k' (S p) (S q) q - | S r' => log2_iter k' p (S q) r' - end - end. - -Definition log2 n : nat := log2_iter (nat_pred n) 0 1 0. - (** ** Properties of Successors *) (** TODO: remove these *) From 4058c83beafdb5cc763d4b31b8723a4681f23fc2 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 16:58:48 +0200 Subject: [PATCH 012/282] Nat: remove sqrt Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 27 --------------------------- 1 file changed, 27 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 3af552c3100..e5b36f51a80 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -119,33 +119,6 @@ Fixpoint gcd a b := Definition square n : nat := n * n. -(** *** Square root *) - -(** The following square root function is linear (and tail-recursive). - With Peano representation, we can't do better. For faster algorithm, - see Psqrt/Zsqrt/Nsqrt... - - We search the square root of n = k + p^2 + (q - r) - with q = 2p and 0<=r<=q. We start with p=q=r=0, hence - looking for the square root of n = k. Then we progressively - decrease k and r. When k = S k' and r=0, it means we can use (S p) - as new sqrt candidate, since (S k')+p^2+2p = k'+(S p)^2. - When k reaches 0, we have found the biggest p^2 square contained - in n, hence the square root of n is p. -*) - -Fixpoint sqrt_iter k p q r : nat := - match k with - | O => p - | S k' => - match r with - | O => sqrt_iter k' p.+1 q.+2 q.+2 - | S r' => sqrt_iter k' p q r' - end - end. - -Definition sqrt n : nat := sqrt_iter n 0 0 0. - (** ** Properties of Successors *) (** TODO: remove these *) From d440522cb32eca9d6153c405fd1a3df9e06e0bf7 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 16:59:36 +0200 Subject: [PATCH 013/282] Nat: remove square Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 ---- 1 file changed, 4 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index e5b36f51a80..17debc127d5 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -115,10 +115,6 @@ Fixpoint gcd a b := | S a' => gcd (b mod a'.+1) a'.+1 end. -(** *** Square *) - -Definition square n : nat := n * n. - (** ** Properties of Successors *) (** TODO: remove these *) From fe0d83a50bd2e1ba22796867fe7e2d11400f3c7a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 17:00:32 +0200 Subject: [PATCH 014/282] Nat: remove IsSucc Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 7 ------- 1 file changed, 7 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 17debc127d5..93c4635f913 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -141,13 +141,6 @@ Proof. Defined. #[export] Hint Resolve not_eq_S : core. -(** TODO: remove *) -Definition IsSucc (n: nat) : Type0 := - match n with - | O => Empty - | S p => Unit - end. - (** TODO: rename to [neq_O_S] *) (** Zero is not the successor of a number *) Definition not_eq_O_S : forall n, 0 <> S n. From e592ac3145701a75b2efcb9bdd792bd28554e3c6 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 17:11:38 +0200 Subject: [PATCH 015/282] Nat: rename pred_Sn -> nat_pred_succ Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 7 ++----- 2 files changed, 3 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index ee0396c0d42..f66302bcd52 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -599,7 +599,7 @@ Proof. Defined. #[export] Hint Rewrite S_predn : nat. -#[export] Hint Rewrite <- pred_Sn : nat. +#[export] Hint Rewrite <- nat_pred_succ : nat. #[export] Hint Resolve S_predn : nat. #[export] Hint Resolve leq_n_pred : nat. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 93c4635f913..80cbca26761 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -123,11 +123,8 @@ Local Definition ap_nat := @ap nat. #[export] Hint Resolve ap_S : core. #[export] Hint Resolve ap_nat : core. -(** TODO: remove, this is trivial/by definition *) -Definition pred_Sn : forall n:nat, n = nat_pred (S n). -Proof. - auto. -Defined. +Definition nat_pred_succ n : nat_pred (nat_succ n) = n + := idpath. (** Injectivity of successor *) Definition path_nat_S n m (H : S n = S m) : n = m := ap nat_pred H. From e1a13c337c47e9505b73d3870ad8dae43c8ce645 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 17:12:03 +0200 Subject: [PATCH 016/282] Nat: move nat_exp next to nat_pow and add todo Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 80cbca26761..0599069cb93 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -76,6 +76,14 @@ Fixpoint nat_pow n m := | S m' => n * (nat_pow n m') end. +(** TODO: merge witth nat_pow (order of arguments needs adjusting). *) +(** Exponentiation *) +Fixpoint nat_exp (n m : nat) : nat + := match m with + | 0 => 1 + | S m => nat_exp n m * n + end. + (** *** Euclidean division *) (** This division is linear and tail-recursive. In [divmod], [y] is the predecessor of the actual divisor, and [u] is [y] sub the real remainder. *) @@ -574,14 +582,6 @@ Proof. nrapply mul_n_Sm. Defined. -(** ** Exponentiation *) - -Fixpoint nat_exp (n m : nat) : nat - := match m with - | 0 => 1 - | S m => nat_exp n m * n - end. - (** ** Factorials *) Fixpoint factorial (n : nat) : nat From fc38327f900c97cbc5882c5e20b867d173b11722 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 17:15:55 +0200 Subject: [PATCH 017/282] Nat: move factorial Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 0599069cb93..cd9f52745a8 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -123,6 +123,14 @@ Fixpoint gcd a b := | S a' => gcd (b mod a'.+1) a'.+1 end. +(** *** Factorials *) + +Fixpoint factorial (n : nat) : nat + := match n with + | 0 => 1 + | S n => S n * factorial n + end. + (** ** Properties of Successors *) (** TODO: remove these *) @@ -582,14 +590,6 @@ Proof. nrapply mul_n_Sm. Defined. -(** ** Factorials *) - -Fixpoint factorial (n : nat) : nat - := match n with - | 0 => 1 - | S n => S n * factorial n - end. - (** ** Natural number ordering *) (** ** Theorems about natural number ordering *) From 1dde79de18c24a482680154970b6cd604668a209 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 17:29:55 +0200 Subject: [PATCH 018/282] Nat: redefine a few operations in terms of nat_iter Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 38 ++++++++++++++------------------------ 1 file changed, 14 insertions(+), 24 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index cd9f52745a8..244eb43fae9 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -15,6 +15,14 @@ Local Open Scope nat_scope. (** ** Basic operations on naturals *) +(** [n]th iteration of the function [f : A -> A]. We have definitional equalities [nat_iter 0 f x = x] and [nat_iter n.+1 f x = f (nat_iter n f x)]. We make this a notation, so it doesn't add a universe variable for the universe containing [A]. *) +Notation nat_iter n f x + := ((fix F (m : nat) + := match m with + | 0 => x + | m'.+1 => f (F m') + end) n). + (** [nat_succ n] is the successor of a natural number [n]. *) Notation nat_succ := S (only parsing). @@ -26,20 +34,14 @@ Definition nat_pred n : nat := end. (** Addition of natural numbers *) -Fixpoint nat_add n m : nat := - match n with - | 0 => m - | S n' => S (nat_add n' m) - end. +Definition nat_add n m : nat + := nat_iter n nat_succ m. Notation "n + m" := (nat_add n m) : nat_scope. (** Multiplication of natural numbers *) -Fixpoint nat_mul n m : nat := - match n with - | 0 => 0 - | S n' => m + (nat_mul n' m) - end. +Definition nat_mul n m : nat + := nat_iter n (nat_add m) 0. Notation "n * m" := (nat_mul n m) : nat_scope. @@ -68,13 +70,9 @@ Fixpoint nat_min n m := | S n' , S m' => S (nat_min n' m') end. -(** TODO: redefine in terms of nat_iter *) (** Exponentiation of natural numbers. *) -Fixpoint nat_pow n m := - match m with - | 0 => 1 - | S m' => n * (nat_pow n m') - end. +Definition nat_pow n m := + nat_iter m (nat_mul n) 1. (** TODO: merge witth nat_pow (order of arguments needs adjusting). *) (** Exponentiation *) @@ -515,14 +513,6 @@ Proof. intros; rewrite nat_min_comm; by apply nat_min_l. Defined. -(** [n]th iteration of the function [f : A -> A]. We have definitional equalities [nat_iter 0 f x = x] and [nat_iter n.+1 f x = f (nat_iter n f x)]. We make this a notation, so it doesn't add a universe variable for the universe containing [A]. *) -Notation nat_iter n f x - := ((fix F (m : nat) - := match m with - | 0 => x - | m'.+1 => f (F m') - end) n). - Lemma nat_iter_succ_r n {A} (f : A -> A) (x : A) : nat_iter (S n) f x = nat_iter n f (f x). Proof. From 71543d819bcf911e93949fd08792579af7ad460d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 17:32:23 +0200 Subject: [PATCH 019/282] Nat: improve comments Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 34 +++++++++++++++++++++------------- 1 file changed, 21 insertions(+), 13 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 244eb43fae9..ec611ccbf2b 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -15,6 +15,8 @@ Local Open Scope nat_scope. (** ** Basic operations on naturals *) +(** *** Iteration *) + (** [n]th iteration of the function [f : A -> A]. We have definitional equalities [nat_iter 0 f x = x] and [nat_iter n.+1 f x = f (nat_iter n f x)]. We make this a notation, so it doesn't add a universe variable for the universe containing [A]. *) Notation nat_iter n f x := ((fix F (m : nat) @@ -23,6 +25,8 @@ Notation nat_iter n f x | m'.+1 => f (F m') end) n). +(** *** Successor and Predecessor *) + (** [nat_succ n] is the successor of a natural number [n]. *) Notation nat_succ := S (only parsing). @@ -33,6 +37,8 @@ Definition nat_pred n : nat := | S n' => n' end. +(** *** Arithmetic Operations *) + (** Addition of natural numbers *) Definition nat_add n m : nat := nat_iter n nat_succ m. @@ -54,6 +60,20 @@ Fixpoint nat_sub n m : nat := Notation "n - m" := (nat_sub n m) : nat_scope. +(** Exponentiation of natural numbers. *) +Definition nat_pow n m := + nat_iter m (nat_mul n) 1. + +(** TODO: merge witth nat_pow (order of arguments needs adjusting). *) +(** Exponentiation *) +Fixpoint nat_exp (n m : nat) : nat + := match m with + | 0 => 1 + | S m => nat_exp n m * n + end. + +(** *** Maximum and Minimum *) + (** The maximum [nat_max n m] of two natural numbers [n] and [m]. *) Fixpoint nat_max n m := match n, m with @@ -70,18 +90,6 @@ Fixpoint nat_min n m := | S n' , S m' => S (nat_min n' m') end. -(** Exponentiation of natural numbers. *) -Definition nat_pow n m := - nat_iter m (nat_mul n) 1. - -(** TODO: merge witth nat_pow (order of arguments needs adjusting). *) -(** Exponentiation *) -Fixpoint nat_exp (n m : nat) : nat - := match m with - | 0 => 1 - | S m => nat_exp n m * n - end. - (** *** Euclidean division *) (** This division is linear and tail-recursive. In [divmod], [y] is the predecessor of the actual divisor, and [u] is [y] sub the real remainder. *) @@ -121,7 +129,7 @@ Fixpoint gcd a b := | S a' => gcd (b mod a'.+1) a'.+1 end. -(** *** Factorials *) +(** *** Factorial *) Fixpoint factorial (n : nat) : nat := match n with From d303d796b712b14c712bb7385ec1d38285a82912 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 17:34:59 +0200 Subject: [PATCH 020/282] Nat: move properties of nat_iter to own section Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 48 ++++++++++++++++++++------------------ 1 file changed, 25 insertions(+), 23 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ec611ccbf2b..a93867d7bc8 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -137,6 +137,31 @@ Fixpoint factorial (n : nat) : nat | S n => S n * factorial n end. +(** ** Properties of [nat_iter]. *) + +Lemma nat_iter_succ_r n {A} (f : A -> A) (x : A) + : nat_iter (S n) f x = nat_iter n f (f x). +Proof. + simple_induction n n IHn; simpl; trivial. + exact (ap f IHn). +Defined. + +Theorem nat_iter_add (n m : nat) {A} (f : A -> A) (x : A) + : nat_iter (n + m) f x = nat_iter n f (nat_iter m f x). +Proof. + simple_induction n n IHn; simpl; trivial. + exact (ap f IHn). +Defined. + +(** Preservation of invariants : if [f : A -> A] preserves the invariant [P], then the iterates of [f] also preserve it. *) +Theorem nat_iter_invariant (n : nat) {A} (f : A -> A) (P : A -> Type) + : (forall x, P x -> P (f x)) -> forall x, P x -> P (nat_iter n f x). +Proof. + simple_induction n n IHn; simpl; trivial. + intros Hf x Hx. + apply Hf, IHn; trivial. +Defined. + (** ** Properties of Successors *) (** TODO: remove these *) @@ -521,29 +546,6 @@ Proof. intros; rewrite nat_min_comm; by apply nat_min_l. Defined. -Lemma nat_iter_succ_r n {A} (f : A -> A) (x : A) - : nat_iter (S n) f x = nat_iter n f (f x). -Proof. - simple_induction n n IHn; simpl; trivial. - exact (ap f IHn). -Defined. - -Theorem nat_iter_add (n m : nat) {A} (f : A -> A) (x : A) - : nat_iter (n + m) f x = nat_iter n f (nat_iter m f x). -Proof. - simple_induction n n IHn; simpl; trivial. - exact (ap f IHn). -Defined. - -(** Preservation of invariants : if [f : A -> A] preserves the invariant [P], then the iterates of [f] also preserve it. *) -Theorem nat_iter_invariant (n : nat) {A} (f : A -> A) (P : A -> Type) - : (forall x, P x -> P (f x)) -> forall x, P x -> P (nat_iter n f x). -Proof. - simple_induction n n IHn; simpl; trivial. - intros Hf x Hx. - apply Hf, IHn; trivial. -Defined. - (** ** Arithmetic *) Lemma nat_add_n_Sm (n m : nat) : (n + m).+1 = n + m.+1. From c96b9a8fa33061a02213389ec8e8078a0e7fc845 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 17:59:40 +0200 Subject: [PATCH 021/282] Nat: rename add_n_O -> nat_add_zero_r Signed-off-by: Ali Caglayan --- theories/Colimits/Sequential.v | 5 +++++ theories/Spaces/List/Theory.v | 2 +- theories/Spaces/Nat/Arithmetic.v | 14 +++++++------- theories/Spaces/Nat/Core.v | 29 +++++++++++++---------------- 4 files changed, 26 insertions(+), 24 deletions(-) diff --git a/theories/Colimits/Sequential.v b/theories/Colimits/Sequential.v index a42de66a1ab..62f9073dc15 100644 --- a/theories/Colimits/Sequential.v +++ b/theories/Colimits/Sequential.v @@ -208,6 +208,11 @@ Proof. destruct p; rewrite !concat_1p, concat_p1; reflexivity. Defined. +Local Lemma add_n_O : forall (n : nat), n = n + 0. +Proof. + simple_induction' n; simpl; auto. +Defined. + Global Instance isequiv_colim_shift_seq_to_colim_seq `{Funext} A n : IsEquiv (colim_shift_seq_to_colim_seq A n). Proof. diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index 1a8cd44420d..4c836e056f5 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -235,7 +235,7 @@ Definition length_reverse_acc@{i|} {A : Type@{i}} (acc l : list A) : length (reverse_acc acc l) = (length acc + length l)%nat. Proof. induction l as [|x l IHl] in acc |- * using list_ind@{i i}. - - apply add_n_O. + - symmetry; apply nat_add_zero_r. - lhs nrapply IHl. apply nat_add_n_Sm. Defined. diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index f66302bcd52..5814c21b11e 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -200,7 +200,7 @@ Proof. + reflexivity. + assumption. - simple_induction' n. - + simpl. destruct (add_n_O m); reflexivity. + + simpl. destruct (nat_add_zero_r m)^; reflexivity. + simpl. destruct (add_n_Sm m n). assumption. Defined. @@ -328,7 +328,7 @@ Proposition natminuspluseq (n m : nat) Proof. revert m; simple_induction n n IHn. - intros. destruct m; [reflexivity |]. simpl. - apply (ap S), symmetric_paths, add_n_O. + apply (ap S), symmetric_paths, (nat_add_zero_r _)^. - intros m l. destruct m. + contradiction (not_leq_Sn_0 n). + simpl. apply leq_S_n, IHn in l. @@ -391,7 +391,7 @@ Proposition nataddpreservesleq { n m k : nat } Proof. intro l. simple_induction k k IHk. - - destruct (add_n_O n), (add_n_O m); exact l. + - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; exact l. - destruct (nat_add_n_Sm n k), (nat_add_n_Sm m k); apply leq_S_n'; exact IHk. Defined. @@ -415,7 +415,7 @@ Proof. change (n + k).+1 with (n.+1 + k). generalize (n.+1). intros n' l. simple_induction k k IHk. - - destruct (add_n_O n'), (add_n_O m); exact l. + - destruct (nat_add_zero_r n')^, (nat_add_zero_r m)^; exact l. - destruct (nat_add_n_Sm n' k), (nat_add_n_Sm m k); apply leq_S_n'; exact IHk. Defined. @@ -432,7 +432,7 @@ Proposition nataddreflectslt { n m k : nat } : n + k < m + k -> n < m. Proof. simple_induction k k IHk. - - destruct (add_n_O n), (add_n_O m); trivial. + - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; trivial. - intro l. destruct (nat_add_n_Sm n k), (nat_add_n_Sm m k) in l. apply leq_S_n, IHk in l; exact l. Defined. @@ -530,7 +530,7 @@ Proposition nataddsub_comm_ineq (n m k : nat) : (n + k) - m <= (n - m) + k. Proof. simple_induction k k IHk. - - destruct (add_n_O n), (add_n_O (n - m)); constructor. + - destruct (nat_add_zero_r n)^, (nat_add_zero_r (n - m))^; constructor. - destruct (add_n_Sm n k). refine (leq_trans (nataddsub_comm_ineq_lemma (n+k) m) _). destruct (add_n_Sm (n - m) k). @@ -792,7 +792,7 @@ Defined. #[export] Hint Immediate add_n_sub_n_eq : nat. #[export] Hint Immediate add_n_sub_n_eq' : nat. -#[export] Hint Rewrite <- add_n_O : nat. +#[export] Hint Rewrite -> nat_add_zero_r : nat. #[export] Hint Rewrite -> add_O_n : nat. #[export] Hint Rewrite -> add_n_sub_n_eq : nat. #[export] Hint Rewrite -> add_n_sub_n_eq' : nat. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index a93867d7bc8..89223bbd659 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -170,16 +170,16 @@ Local Definition ap_nat := @ap nat. #[export] Hint Resolve ap_S : core. #[export] Hint Resolve ap_nat : core. -Definition nat_pred_succ n : nat_pred (nat_succ n) = n +Definition nat_pred_succ@{} n : nat_pred (nat_succ n) = n := idpath. (** Injectivity of successor *) -Definition path_nat_S n m (H : S n = S m) : n = m := ap nat_pred H. +Definition path_nat_S@{} n m (H : S n = S m) : n = m := ap nat_pred H. #[export] Hint Immediate path_nat_S : core. (** TODO: rename to [neq_S] *) (** TODO: avoid auto in proof. *) -Definition not_eq_S n m : n <> m -> S n <> S m. +Definition not_eq_S@{} n m : n <> m -> S n <> S m. Proof. auto. Defined. @@ -187,7 +187,7 @@ Defined. (** TODO: rename to [neq_O_S] *) (** Zero is not the successor of a number *) -Definition not_eq_O_S : forall n, 0 <> S n. +Definition not_eq_O_S@{} : forall n, 0 <> S n. Proof. discriminate. Defined. @@ -196,25 +196,22 @@ Defined. (** TODO: rename to [neq_n_Sn] *) (** TODO: prove above using this *) (** TODO: remove auto. *) -Theorem not_eq_n_Sn : forall n:nat, n <> S n. +Theorem not_eq_n_Sn@{} n : n <> S n. Proof. simple_induction' n; auto. Defined. #[export] Hint Resolve not_eq_n_Sn : core. -(** TODO: remove *) -(* Local Definition ap011_add := @ap011 _ _ _ add. *) -(* Local Definition ap011_nat := @ap011 nat nat. *) -(* #[export] Hint Resolve ap011_add : core. *) -(* #[export] Hint Resolve ap011_nat : core. *) +(** *** Properties of Addition *) -(** TODO: rename [nat_add_zero_r] *) -(** TODO: reverse direction *) -Lemma add_n_O : forall (n : nat), n = n + 0. +Definition nat_add_zero_r@{} n : n + 0 = n. Proof. - simple_induction' n; simpl; auto. + induction n as [|n IHn]. + - reflexivity. + - apply (ap nat_succ). + exact IHn. Defined. -#[export] Hint Resolve add_n_O : core. +#[export] Hint Resolve nat_add_zero_r : core. (** TODO: rename [nat_add_zero_l] *) Lemma add_O_n : forall (n : nat), 0 + n = n. @@ -558,7 +555,7 @@ Defined. Definition nat_add_comm (n m : nat) : n + m = m + n. Proof. simple_induction n n IHn; simpl. - - exact (add_n_O m). + - exact (nat_add_zero_r m)^. - transitivity (m + n).+1. + apply ap, IHn. + apply nat_add_n_Sm. From e8b8f590fb94a1315644bf9b2d328d55fe73f18d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 18:02:27 +0200 Subject: [PATCH 022/282] Nat: rename add_O_n -> nat_add_zero_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 9 +++------ 2 files changed, 4 insertions(+), 7 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 5814c21b11e..db8335ed929 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -793,7 +793,7 @@ Defined. #[export] Hint Immediate add_n_sub_n_eq' : nat. #[export] Hint Rewrite -> nat_add_zero_r : nat. -#[export] Hint Rewrite -> add_O_n : nat. +#[export] Hint Rewrite -> nat_add_zero_l : nat. #[export] Hint Rewrite -> add_n_sub_n_eq : nat. #[export] Hint Rewrite -> add_n_sub_n_eq' : nat. #[export] Hint Rewrite -> nataddsub_assoc : nat. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 89223bbd659..52f20a3c159 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -204,6 +204,9 @@ Defined. (** *** Properties of Addition *) +Definition nat_add_zero_l@{} n : 0 + n = n + := idpath. + Definition nat_add_zero_r@{} n : n + 0 = n. Proof. induction n as [|n IHn]. @@ -213,12 +216,6 @@ Proof. Defined. #[export] Hint Resolve nat_add_zero_r : core. -(** TODO: rename [nat_add_zero_l] *) -Lemma add_O_n : forall (n : nat), 0 + n = n. -Proof. - auto. -Defined. - (** TODO: rename [nat_add_succ_r] *) (** TODO: reverse direction *) Lemma add_n_Sm : forall n m:nat, S (n + m) = n + S m. From 0e1df358d4993b2ece183d85737e1ebec6da12c7 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 18:05:59 +0200 Subject: [PATCH 023/282] Nat: rename add_n_Sm -> nat_add_succ_r Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 6 +++--- theories/Spaces/Nat/Core.v | 10 ++++------ 2 files changed, 7 insertions(+), 9 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index db8335ed929..4a375f2c843 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -201,7 +201,7 @@ Proof. + assumption. - simple_induction' n. + simpl. destruct (nat_add_zero_r m)^; reflexivity. - + simpl. destruct (add_n_Sm m n). assumption. + + simpl. destruct (nat_add_succ_r m n)^. assumption. Defined. Proposition add_n_sub_n_eq' (m n : nat) : n + m - n = m. @@ -531,9 +531,9 @@ Proposition nataddsub_comm_ineq (n m k : nat) Proof. simple_induction k k IHk. - destruct (nat_add_zero_r n)^, (nat_add_zero_r (n - m))^; constructor. - - destruct (add_n_Sm n k). + - destruct (nat_add_succ_r n k)^. refine (leq_trans (nataddsub_comm_ineq_lemma (n+k) m) _). - destruct (add_n_Sm (n - m) k). + destruct (nat_add_succ_r (n - m) k)^. now apply leq_S_n'. Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 52f20a3c159..20f89f25fd2 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -216,13 +216,11 @@ Proof. Defined. #[export] Hint Resolve nat_add_zero_r : core. -(** TODO: rename [nat_add_succ_r] *) -(** TODO: reverse direction *) -Lemma add_n_Sm : forall n m:nat, S (n + m) = n + S m. +Lemma nat_add_succ_r n m : n + m.+1 = (n + m).+1. Proof. simple_induction' n; simpl; auto. Defined. -#[export] Hint Resolve add_n_Sm: core. +#[export] Hint Resolve nat_add_succ_r: core. (** TODO: rename [nat_add_succ_l] *) (** TODO: reverse direction *) @@ -251,7 +249,7 @@ Defined. Lemma mul_n_Sm : forall n m:nat, n * m + n = n * S m. Proof. intros; simple_induction n p H; simpl; auto. - destruct H; rewrite <- add_n_Sm; apply ap. + destruct H; rewrite nat_add_succ_r; apply ap. pattern m at 1 3; elim m; simpl; auto. Defined. #[export] Hint Resolve mul_n_Sm: core. @@ -632,6 +630,6 @@ Proof. refine (trunc_index_add_succ _ _ @ _). exact (ap trunc_S IH). } refine (_ @ ap nat_to_trunc_index _). - 2: exact (ap _ (add_Sn_m _ _)^ @ add_n_Sm _ _). + 2: exact (nat_add_succ_r _ _ @ ap _ (add_Sn_m _ _))^. reflexivity. Defined. From 5cc747af62083f68b1427be05dbf88ca8a1d9bc9 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 18:09:24 +0200 Subject: [PATCH 024/282] Nat: rename add_Sn_m -> nat_add_succ_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 15 +++++---------- 1 file changed, 5 insertions(+), 10 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 20f89f25fd2..dee0bcf7aa1 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -216,20 +216,15 @@ Proof. Defined. #[export] Hint Resolve nat_add_zero_r : core. -Lemma nat_add_succ_r n m : n + m.+1 = (n + m).+1. +Definition nat_add_succ_l n m : n.+1 + m = (n + m).+1 + := idpath. + +Definition nat_add_succ_r n m : n + m.+1 = (n + m).+1. Proof. simple_induction' n; simpl; auto. Defined. #[export] Hint Resolve nat_add_succ_r: core. -(** TODO: rename [nat_add_succ_l] *) -(** TODO: reverse direction *) -(** TODO: remove auto *) -Lemma add_Sn_m : forall n m:nat, S n + m = S (n + m). -Proof. - auto. -Defined. - (** Multiplication *) (** TODO: remove? *) @@ -630,6 +625,6 @@ Proof. refine (trunc_index_add_succ _ _ @ _). exact (ap trunc_S IH). } refine (_ @ ap nat_to_trunc_index _). - 2: exact (nat_add_succ_r _ _ @ ap _ (add_Sn_m _ _))^. + 2: exact (nat_add_succ_r _ _)^. reflexivity. Defined. From 3a5c8877cf45416d219e8e4957f3df68474f8744 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 18:18:48 +0200 Subject: [PATCH 025/282] Nat: remove nat_add_n_Sm (redundant) and move nat_add_comm Signed-off-by: Ali Caglayan --- theories/Colimits/Sequential.v | 7 +++++++ theories/Spaces/List/Theory.v | 7 ++++--- theories/Spaces/Nat/Arithmetic.v | 14 +++++++------- theories/Spaces/Nat/Core.v | 29 +++++++++++------------------ 4 files changed, 29 insertions(+), 28 deletions(-) diff --git a/theories/Colimits/Sequential.v b/theories/Colimits/Sequential.v index 62f9073dc15..2178dd652a0 100644 --- a/theories/Colimits/Sequential.v +++ b/theories/Colimits/Sequential.v @@ -213,6 +213,13 @@ Proof. simple_induction' n; simpl; auto. Defined. +Local Lemma nat_add_n_Sm (n m : nat) : (n + m).+1 = n + m.+1. +Proof. + simple_induction' n; simpl. + - reflexivity. + - apply ap; assumption. +Defined. + Global Instance isequiv_colim_shift_seq_to_colim_seq `{Funext} A n : IsEquiv (colim_shift_seq_to_colim_seq A n). Proof. diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index 4c836e056f5..960fc14f5d3 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -234,10 +234,11 @@ Defined. Definition length_reverse_acc@{i|} {A : Type@{i}} (acc l : list A) : length (reverse_acc acc l) = (length acc + length l)%nat. Proof. + symmetry. induction l as [|x l IHl] in acc |- * using list_ind@{i i}. - - symmetry; apply nat_add_zero_r. - - lhs nrapply IHl. - apply nat_add_n_Sm. + - apply nat_add_zero_r. + - rhs_V nrapply IHl. + apply nat_add_succ_r. Defined. (** The length of [reverse] is the same as the length of the original list. *) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 4a375f2c843..93d1194301e 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -42,12 +42,12 @@ Proof. revert m k; simple_induction n n IHn. - reflexivity. - intros m k. change (n.+1 + (m + k)) with (n + (m + k)).+1. - apply (transitive_paths _ _ _ (nat_add_n_Sm _ _)). + apply (transitive_paths _ _ _ (nat_add_succ_r _ _)^). change (m + k).+1 with (m.+1 + k); change (n.+1 + m) with (n + m).+1. apply (transitive_paths _ _ _ (IHn m.+1 k)). apply (ap (fun zz => zz + k)). - apply symmetric_paths, nat_add_n_Sm. + apply symmetric_paths, (nat_add_succ_r _ _)^. Defined. Proposition not_lt_implies_geq {n m : nat} : ~(n < m) -> m <= n. @@ -332,7 +332,7 @@ Proof. - intros m l. destruct m. + contradiction (not_leq_Sn_0 n). + simpl. apply leq_S_n, IHn in l. - destruct (nat_add_n_Sm (m - n) n). + destruct (nat_add_succ_r (m - n) n)^. destruct (symmetric_paths _ _ l). reflexivity. Defined. @@ -392,7 +392,7 @@ Proof. intro l. simple_induction k k IHk. - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; exact l. - - destruct (nat_add_n_Sm n k), (nat_add_n_Sm m k); + - destruct (nat_add_succ_r n k)^, (nat_add_succ_r m k)^; apply leq_S_n'; exact IHk. Defined. @@ -416,7 +416,7 @@ Proof. generalize (n.+1). intros n' l. simple_induction k k IHk. - destruct (nat_add_zero_r n')^, (nat_add_zero_r m)^; exact l. - - destruct (nat_add_n_Sm n' k), (nat_add_n_Sm m k); + - destruct (nat_add_succ_r n' k)^, (nat_add_succ_r m k)^; apply leq_S_n'; exact IHk. Defined. @@ -433,7 +433,7 @@ Proposition nataddreflectslt { n m k : nat } Proof. simple_induction k k IHk. - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; trivial. - - intro l. destruct (nat_add_n_Sm n k), (nat_add_n_Sm m k) in l. + - intro l. destruct (nat_add_succ_r n k)^, (nat_add_succ_r m k)^ in l. apply leq_S_n, IHk in l; exact l. Defined. @@ -499,7 +499,7 @@ Proof. | contradiction (not_lt_n_0 k _) | ]. simpl "-". apply leq_S_n in l. - destruct (symmetric_paths _ _ (nat_add_n_Sm n (m - k))). + destruct (nat_add_succ_r n (m - k)). destruct (nataddsub_assoc_lemma l). apply (IHn m.+1 k). apply leq_S. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index dee0bcf7aa1..c539a32262a 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -216,15 +216,24 @@ Proof. Defined. #[export] Hint Resolve nat_add_zero_r : core. -Definition nat_add_succ_l n m : n.+1 + m = (n + m).+1 +Definition nat_add_succ_l@{} n m : n.+1 + m = (n + m).+1 := idpath. -Definition nat_add_succ_r n m : n + m.+1 = (n + m).+1. +Definition nat_add_succ_r@{} n m : n + m.+1 = (n + m).+1. Proof. simple_induction' n; simpl; auto. Defined. #[export] Hint Resolve nat_add_succ_r: core. +Definition nat_add_comm@{} n m : n + m = m + n. +Proof. + induction n. + - exact (nat_add_zero_r m)^. + - rhs nrapply nat_add_succ_r. + apply (ap nat_succ). + exact IHn. +Defined. + (** Multiplication *) (** TODO: remove? *) @@ -535,22 +544,6 @@ Defined. (** ** Arithmetic *) -Lemma nat_add_n_Sm (n m : nat) : (n + m).+1 = n + m.+1. -Proof. - simple_induction' n; simpl. - - reflexivity. - - apply ap; assumption. -Defined. - -Definition nat_add_comm (n m : nat) : n + m = m + n. -Proof. - simple_induction n n IHn; simpl. - - exact (nat_add_zero_r m)^. - - transitivity (m + n).+1. - + apply ap, IHn. - + apply nat_add_n_Sm. -Defined. - Global Instance isinj_S : IsInjective S. Proof. intros x y p. From 84e09599e43898761e4be013f45fcbdfef30e0d8 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 18:24:55 +0200 Subject: [PATCH 026/282] Nat: rename/move/reprove assoc_nat_add -> nat_add_assoc Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 14 -------------- theories/Spaces/Nat/Core.v | 14 ++++++++++++++ 2 files changed, 14 insertions(+), 14 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 93d1194301e..237302df283 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -36,20 +36,6 @@ Local Definition transitive_paths_nat (n m k : nat) #[export] Hint Resolve leq_trans : nat. #[export] Hint Resolve leq_antisym : nat. -Proposition assoc_nat_add (n m k : nat) - : n + (m + k) = (n + m) + k. -Proof. - revert m k; simple_induction n n IHn. - - reflexivity. - - intros m k. change (n.+1 + (m + k)) with (n + (m + k)).+1. - apply (transitive_paths _ _ _ (nat_add_succ_r _ _)^). - change (m + k).+1 with (m.+1 + k); - change (n.+1 + m) with (n + m).+1. - apply (transitive_paths _ _ _ (IHn m.+1 k)). - apply (ap (fun zz => zz + k)). - apply symmetric_paths, (nat_add_succ_r _ _)^. -Defined. - Proposition not_lt_implies_geq {n m : nat} : ~(n < m) -> m <= n. Proof. intros not_lt. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index c539a32262a..3968b8798e4 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -204,9 +204,11 @@ Defined. (** *** Properties of Addition *) +(** [0] is the left identity of addition. *) Definition nat_add_zero_l@{} n : 0 + n = n := idpath. +(** [0] is the right identity of addition. *) Definition nat_add_zero_r@{} n : n + 0 = n. Proof. induction n as [|n IHn]. @@ -216,15 +218,18 @@ Proof. Defined. #[export] Hint Resolve nat_add_zero_r : core. +(** Adding a successor on the left is the same as adding and then taking the successor. *) Definition nat_add_succ_l@{} n m : n.+1 + m = (n + m).+1 := idpath. +(** Adding a successor on the right is the same as adding and then taking the successor. *) Definition nat_add_succ_r@{} n m : n + m.+1 = (n + m).+1. Proof. simple_induction' n; simpl; auto. Defined. #[export] Hint Resolve nat_add_succ_r: core. +(** Addition of natural numbers is commutative. *) Definition nat_add_comm@{} n m : n + m = m + n. Proof. induction n. @@ -234,6 +239,15 @@ Proof. exact IHn. Defined. +(** Addition of natural numbers is associative. *) +Definition nat_add_assoc n m k : n + (m + k) = (n + m) + k. +Proof. + induction n as [|n IHn]. + - reflexivity. + - nrapply (ap nat_succ). + exact IHn. +Defined. + (** Multiplication *) (** TODO: remove? *) From e58740f233dd016433c5d902c7f6dee31d586dd3 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 18:29:11 +0200 Subject: [PATCH 027/282] Nat: rename mul_n_O -> nat_mul_zero_r Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 18 ++++++++++-------- 1 file changed, 10 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 3968b8798e4..e2e62855596 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -248,19 +248,22 @@ Proof. exact IHn. Defined. -(** Multiplication *) +(** *** Properties of Multiplication *) (** TODO: remove? *) Local Definition ap011_mul := @ap011 _ _ _ nat_mul. #[export] Hint Resolve ap011_mul : core. -(** TODO: rename [nat_mul_zero_r] *) -(** TODO: reverse direction *) -Lemma mul_n_O : forall n:nat, 0 = n * 0. +(** Multiplication by [0] on the left is [0]. *) +Definition nat_mul_zero_l n : 0 * n = 0 + := idpath. + +(** Multiplicaiton by [0] on the right is [0]. *) +Definition nat_mul_zero_r n : n * 0 = 0. Proof. - simple_induction' n; simpl; auto. + by induction n. Defined. -#[export] Hint Resolve mul_n_O : core. +#[export] Hint Resolve nat_mul_zero_r : core. (** TODO: rename [nat_mul_succ_r] *) (** TODO: reverse direction *) @@ -275,7 +278,6 @@ Defined. (** Standard associated names *) (** TODO: remove? *) -Notation mul_0_r_reverse := mul_n_O (only parsing). Notation mul_succ_r_reverse := mul_n_Sm (only parsing). (** ** Equality of natural numbers *) @@ -579,7 +581,7 @@ Defined. Definition nat_mul_comm@{} (x y : nat) : x * y = y * x. Proof. induction x as [|x IHx] in y |- * using nat_rect@{Set}. - - apply mul_n_O. + - apply (nat_mul_zero_r _)^. - simpl; rewrite nat_add_comm, IHx. nrapply mul_n_Sm. Defined. From e96b323c9de50b20ac19d012bdb7487ba7ab36b6 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 18:44:09 +0200 Subject: [PATCH 028/282] Nat: remove unused mul hints Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 ---- 1 file changed, 4 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index e2e62855596..2986a794f43 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -250,10 +250,6 @@ Defined. (** *** Properties of Multiplication *) -(** TODO: remove? *) -Local Definition ap011_mul := @ap011 _ _ _ nat_mul. -#[export] Hint Resolve ap011_mul : core. - (** Multiplication by [0] on the left is [0]. *) Definition nat_mul_zero_l n : 0 * n = 0 := idpath. From 68d8b67b3f42637883647544d67fd955a356b8ea Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 18:52:52 +0200 Subject: [PATCH 029/282] Nat: rename mul_n_Sm -> nat_mul_succ_r Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 32 ++++++++++++++++---------------- 1 file changed, 16 insertions(+), 16 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 2986a794f43..fe18f30f3d9 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -261,20 +261,20 @@ Proof. Defined. #[export] Hint Resolve nat_mul_zero_r : core. -(** TODO: rename [nat_mul_succ_r] *) -(** TODO: reverse direction *) -Lemma mul_n_Sm : forall n m:nat, n * m + n = n * S m. +Definition nat_mul_succ_l n m : n.+1 * m = m + n * m + := idpath. + +Definition nat_mul_succ_r n m : n * m.+1 = n * m + n. Proof. - intros; simple_induction n p H; simpl; auto. - destruct H; rewrite nat_add_succ_r; apply ap. - pattern m at 1 3; elim m; simpl; auto. + induction n as [|n IHn]. + - reflexivity. + - rhs nrapply nat_add_succ_r. + nrapply (ap nat_succ). + rhs_V nrapply nat_add_assoc. + nrapply (ap (nat_add m)). + exact IHn. Defined. -#[export] Hint Resolve mul_n_Sm: core. - -(** Standard associated names *) - -(** TODO: remove? *) -Notation mul_succ_r_reverse := mul_n_Sm (only parsing). +#[export] Hint Resolve nat_mul_succ_r : core. (** ** Equality of natural numbers *) @@ -576,10 +576,10 @@ Defined. Definition nat_mul_comm@{} (x y : nat) : x * y = y * x. Proof. - induction x as [|x IHx] in y |- * using nat_rect@{Set}. - - apply (nat_mul_zero_r _)^. - - simpl; rewrite nat_add_comm, IHx. - nrapply mul_n_Sm. + induction y as [|y IHy]. + - apply (nat_mul_zero_r _). + - simpl; rewrite nat_add_comm, <- IHy. + nrapply nat_mul_succ_r. Defined. (** ** Natural number ordering *) From a66e52394aef09e470fb40ed7c328f4d03db73de Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 18:58:51 +0200 Subject: [PATCH 030/282] Nat: move/reprove commutativity of nat_mul Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 29 ++++++++++++++++------------- 1 file changed, 16 insertions(+), 13 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index fe18f30f3d9..97e6a4a33ef 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -240,7 +240,7 @@ Proof. Defined. (** Addition of natural numbers is associative. *) -Definition nat_add_assoc n m k : n + (m + k) = (n + m) + k. +Definition nat_add_assoc@{} n m k : n + (m + k) = (n + m) + k. Proof. induction n as [|n IHn]. - reflexivity. @@ -251,20 +251,20 @@ Defined. (** *** Properties of Multiplication *) (** Multiplication by [0] on the left is [0]. *) -Definition nat_mul_zero_l n : 0 * n = 0 +Definition nat_mul_zero_l@{} n : 0 * n = 0 := idpath. (** Multiplicaiton by [0] on the right is [0]. *) -Definition nat_mul_zero_r n : n * 0 = 0. +Definition nat_mul_zero_r@{} n : n * 0 = 0. Proof. by induction n. Defined. #[export] Hint Resolve nat_mul_zero_r : core. -Definition nat_mul_succ_l n m : n.+1 * m = m + n * m +Definition nat_mul_succ_l@{} n m : n.+1 * m = m + n * m := idpath. -Definition nat_mul_succ_r n m : n * m.+1 = n * m + n. +Definition nat_mul_succ_r@{} n m : n * m.+1 = n * m + n. Proof. induction n as [|n IHn]. - reflexivity. @@ -276,6 +276,17 @@ Proof. Defined. #[export] Hint Resolve nat_mul_succ_r : core. +(** Multiplication of natural numbers is commutative. *) +Definition nat_mul_comm@{} n m : n * m = m * n. +Proof. + induction m as [|m IHm]; simpl. + - nrapply nat_mul_zero_r. + - lhs nrapply nat_mul_succ_r. + lhs nrapply nat_add_comm. + snrapply (ap (nat_add n)). + exact IHm. +Defined. + (** ** Equality of natural numbers *) (** *** Boolean equality and its properties *) @@ -574,14 +585,6 @@ Proof. exact (isinj_nat_add_l k _ _ H). Defined. -Definition nat_mul_comm@{} (x y : nat) : x * y = y * x. -Proof. - induction y as [|y IHy]. - - apply (nat_mul_zero_r _). - - simpl; rewrite nat_add_comm, <- IHy. - nrapply nat_mul_succ_r. -Defined. - (** ** Natural number ordering *) (** ** Theorems about natural number ordering *) From afab41695ef885a414ae36b08834bfcfd8b04805 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 19:08:58 +0200 Subject: [PATCH 031/282] Nat: distributivity of multiplication over addition Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 97e6a4a33ef..4f1dfb846fe 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -287,6 +287,28 @@ Proof. exact IHm. Defined. +(** Multiplication of natural numbers distributes over addition on the left. *) +Definition nat_dist_l@{} n m k : n * (m + k) = n * m + n * k. +Proof. + induction n as [|n IHn]; simpl. + - reflexivity. + - lhs_V nrapply nat_add_assoc. + rhs_V nrapply nat_add_assoc. + nrapply (ap (nat_add m)). + lhs nrapply nat_add_comm. + rewrite IHn. + lhs_V nrapply nat_add_assoc. + nrapply (ap (nat_add (n * m))). + nrapply nat_add_comm. +Defined. + +(** Multiplication of natural numbers distributes over addition on the right. *) +Definition nat_dist_r@{} n m k : (n + m) * k = n * k + m * k. +Proof. + rewrite 3 (nat_mul_comm _ k). + nrapply nat_dist_l. +Defined. + (** ** Equality of natural numbers *) (** *** Boolean equality and its properties *) From eed88a44e35be89d9b18b554b1eab2d34632f0e4 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 19:10:50 +0200 Subject: [PATCH 032/282] Nat: multiplication is associative Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 10 ++++++++++ 1 file changed, 10 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 4f1dfb846fe..4d46ac12079 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -309,6 +309,16 @@ Proof. nrapply nat_dist_l. Defined. +(** Multiplication of natural numbers is associative. *) +Definition nat_mul_assoc@{} n m k : n * (m * k) = n * m * k. +Proof. + induction n as [|n IHn]; simpl. + - reflexivity. + - rhs nrapply nat_dist_r. + nrapply (ap (nat_add (m * k))). + exact IHn. +Defined. + (** ** Equality of natural numbers *) (** *** Boolean equality and its properties *) From 5d398e4b1ec81055de0c4aa84782eaa5f1eca415 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 19:11:57 +0200 Subject: [PATCH 033/282] Nat: move decidability and hset proof of nat Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 42 ++++++++++++++++++-------------------- 1 file changed, 20 insertions(+), 22 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 4d46ac12079..a113f869df9 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -202,7 +202,25 @@ Proof. Defined. #[export] Hint Resolve not_eq_n_Sn : core. -(** *** Properties of Addition *) +(** ** Truncatedness of Natural Numbers *) + +(** [nat] has decidable paths. *) +Global Instance decidable_paths_nat@{} : DecidablePaths nat. +Proof. + intros n; simple_induction n n IHn; + intros m; destruct m. + - exact (inl idpath). + - exact (inr (not_eq_O_S m)). + - exact (inr (fun p => not_eq_O_S n p^)). + - destruct (IHn m) as [p|q]. + + exact (inl (ap S p)). + + exact (inr (fun p => q (path_nat_S _ _ p))). +Defined. + +(** [nat] is therefore a hset. *) +Global Instance ishset_nat : IsHSet nat := _. + +(** ** Properties of Addition *) (** [0] is the left identity of addition. *) Definition nat_add_zero_l@{} n : 0 + n = n @@ -248,7 +266,7 @@ Proof. exact IHn. Defined. -(** *** Properties of Multiplication *) +(** ** Properties of Multiplication *) (** Multiplication by [0] on the left is [0]. *) Definition nat_mul_zero_l@{} n : 0 * n = 0 @@ -319,26 +337,6 @@ Proof. exact IHn. Defined. -(** ** Equality of natural numbers *) - -(** *** Boolean equality and its properties *) - -(** [nat] has decidable paths *) -Global Instance decidable_paths_nat@{} : DecidablePaths nat. -Proof. - intros n; simple_induction n n IHn; - intros m; destruct m. - - exact (inl idpath). - - exact (inr (not_eq_O_S m)). - - exact (inr (fun p => not_eq_O_S n p^)). - - destruct (IHn m) as [p|q]. - + exact (inl (ap S p)). - + exact (inr (fun p => q (path_nat_S _ _ p))). -Defined. - -(** And is therefore a HSet *) -Global Instance hset_nat : IsHSet nat := _. - (** ** Inequality of natural numbers *) Cumulative Inductive leq (n : nat) : nat -> Type0 := From 0b079136092a395489860d3c458e97e94b965ad0 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 16 Jul 2024 20:11:37 +0200 Subject: [PATCH 034/282] fix HoTTBook.v Signed-off-by: Ali Caglayan --- contrib/HoTTBook.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/contrib/HoTTBook.v b/contrib/HoTTBook.v index c3252c9336f..44bc5a0846d 100644 --- a/contrib/HoTTBook.v +++ b/contrib/HoTTBook.v @@ -345,7 +345,7 @@ Definition Book_3_1_3 := @HoTT.Types.Empty.istrunc_Empty (-2). (* ================================================== thm:nat-set *) (** Example 3.1.4 *) -Definition Book_3_1_4 := @HoTT.Spaces.Nat.Core.hset_nat. +Definition Book_3_1_4 := @HoTT.Spaces.Nat.Core.ishset_nat. (* ================================================== thm:isset-prod *) (** Example 3.1.5 *) From cda5b88432d95ad74293426f55f7acdef0c7f554 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 17 Jul 2024 14:01:44 +0200 Subject: [PATCH 035/282] revise definition of normal subgroup Signed-off-by: Ali Caglayan --- theories/Algebra/AbGroups/AbelianGroup.v | 8 ++-- theories/Algebra/Groups/Kernel.v | 12 ++--- theories/Algebra/Groups/Subgroup.v | 56 ++++++++++++++++++------ 3 files changed, 51 insertions(+), 25 deletions(-) diff --git a/theories/Algebra/AbGroups/AbelianGroup.v b/theories/Algebra/AbGroups/AbelianGroup.v index 7c9a818bbed..8d557283b4d 100644 --- a/theories/Algebra/AbGroups/AbelianGroup.v +++ b/theories/Algebra/AbGroups/AbelianGroup.v @@ -74,11 +74,9 @@ Coercion abgroup_subgroup : Subgroup >-> AbGroup. Global Instance isnormal_ab_subgroup (G : AbGroup) (H : Subgroup G) : IsNormalSubgroup H. Proof. - intros x y; unfold in_cosetL, in_cosetR. - refine (_ oE equiv_subgroup_inverse _ _). - rewrite negate_sg_op. - rewrite negate_involutive. - by rewrite (commutativity (-y) x). + intros x y; unfold in_cosetL, in_cosetR; intros h. + rewrite ab_comm, grp_assoc, grp_inv_l, grp_unit_l. + exact h. Defined. (** ** Quotients of abelian groups *) diff --git a/theories/Algebra/Groups/Kernel.v b/theories/Algebra/Groups/Kernel.v index e4e2d0b25b8..99e00f7d446 100644 --- a/theories/Algebra/Groups/Kernel.v +++ b/theories/Algebra/Groups/Kernel.v @@ -17,13 +17,13 @@ Proof. refine (grp_homo_op _ _ _ @ ap011 _ p _ @ _). 1: apply grp_homo_inv. rewrite q; apply right_inverse. - - intros x y; cbn. + - intros x y; cbn; intros p. rewrite 2 grp_homo_op. - rewrite 2 grp_homo_inv. - refine (_^-1 oE grp_moveL_M1). - refine (_ oE equiv_path_inverse _ _). - apply grp_moveR_1M. - Defined. + rewrite grp_homo_inv. + apply grp_moveL_1M^-1. + rewrite p. + apply grp_unit_r. +Defined. (** ** Corecursion principle for group kernels *) diff --git a/theories/Algebra/Groups/Subgroup.v b/theories/Algebra/Groups/Subgroup.v index 259cefa77b6..2188cca1a84 100644 --- a/theories/Algebra/Groups/Subgroup.v +++ b/theories/Algebra/Groups/Subgroup.v @@ -337,9 +337,9 @@ Proof. all: by intro. Defined. -(** A subgroup is normal if being in a left coset is equivalent to being in a right coset represented by the same element. *) +(** A normal subgroup is a subgroup closed under conjugation. *) Class IsNormalSubgroup {G : Group} (N : Subgroup G) - := isnormal : forall {x y}, in_cosetL N x y <~> in_cosetR N x y. + := isnormal : forall {x y}, N x -> N (y * x * -y). Record NormalSubgroup (G : Group) := { normalsubgroup_subgroup : Subgroup G ; @@ -348,26 +348,54 @@ Record NormalSubgroup (G : Group) := { Coercion normalsubgroup_subgroup : NormalSubgroup >-> Subgroup. Global Existing Instance normalsubgroup_isnormal. + + +(** Left and right cosets are equivalent in normal subgroups. *) +Definition equiv_isnormalsubgroup_in_cosetL_in_cosetR {G : Group} + (N : NormalSubgroup G) + : forall x y, in_cosetL N x y <~> in_cosetR N x y. +Proof. + intros x y. + rapply equiv_iff_hprop; cbn. + - intros n. + nrefine (transport N (grp_unit_r _) _). + nrefine (transport (fun z => N (x * -y * z)) (grp_inv_r x) _). + nrefine (transport N (grp_assoc _ _ _) _). + nrefine (transport (fun z => N (x * z)) (grp_assoc _ _ _)^ _). + nrefine (transport N (grp_assoc _ _ _)^ _). + apply isnormal. + apply subgroup_in_inv'. + nrefine (transport N (grp_inv_op _ _)^ _). + nrefine (transport (fun z => N (- x * z)) (grp_inv_inv _)^ _). + exact n. + - intros n. + nrefine (transport N (grp_unit_r _) _). + nrefine (transport (fun z => N ((- x * y) * z)) (grp_inv_r (-x)) _). + nrefine (transport N (grp_assoc _ _ _) _). + nrefine (transport (fun z => N (- x * z)) (grp_assoc _ _ _)^ _). + nrefine (transport N (grp_assoc _ _ _)^ _). + apply isnormal. + apply subgroup_in_inv'. + nrefine (transport N (grp_inv_op _ _)^ _). + nrefine (transport (fun z => N (z * -y)) (grp_inv_inv _)^ _). + exact n. +Defined. (* Inverses are then respected *) Definition in_cosetL_inverse {G : Group} {N : NormalSubgroup G} : forall x y : G, in_cosetL N (-x) (-y) <~> in_cosetL N x y. Proof. intros x y. - unfold in_cosetL. - rewrite negate_involutive. - symmetry; apply isnormal. + refine (_ oE equiv_isnormalsubgroup_in_cosetL_in_cosetR _ _ _); cbn. + by rewrite negate_involutive. Defined. Definition in_cosetR_inverse {G : Group} {N : NormalSubgroup G} : forall x y : G, in_cosetR N (-x) (-y) <~> in_cosetR N x y. Proof. intros x y. - refine (_ oE (in_cosetR_unit _ _)^-1). - refine (_ oE isnormal^-1). - refine (_ oE in_cosetL_unit _ _). - refine (_ oE isnormal). - by rewrite negate_involutive. + refine (_ oE equiv_isnormalsubgroup_in_cosetL_in_cosetR _ _ _); cbn. + by rewrite grp_inv_inv. Defined. (** This lets us prove that left and right coset relations are congruences. *) @@ -380,11 +408,11 @@ Proof. rewrite negate_sg_op, <- simple_associativity. apply symmetric_in_cosetL; cbn. rewrite simple_associativity. - apply isnormal; cbn. + apply equiv_isnormalsubgroup_in_cosetL_in_cosetR; cbn. rewrite <- simple_associativity. apply subgroup_in_op. 1: assumption. - by apply isnormal, symmetric_in_cosetL. + by apply equiv_isnormalsubgroup_in_cosetL_in_cosetR, symmetric_in_cosetL. Defined. Definition in_cosetR_cong {G : Group} {N : NormalSubgroup G} @@ -396,11 +424,11 @@ Proof. rewrite negate_sg_op, simple_associativity. apply symmetric_in_cosetR; cbn. rewrite <- simple_associativity. - apply isnormal; cbn. + apply equiv_isnormalsubgroup_in_cosetL_in_cosetR; cbn. rewrite simple_associativity. apply subgroup_in_op. 2: assumption. - by apply isnormal, symmetric_in_cosetR. + by apply equiv_isnormalsubgroup_in_cosetL_in_cosetR, symmetric_in_cosetR. Defined. (** The property of being the trivial subgroup is useful. *) From 22a445a6f03fca459fb495dcf486f850b7101579 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sat, 20 Jul 2024 14:00:56 +0200 Subject: [PATCH 036/282] better looking terms in presentation test Signed-off-by: Ali Caglayan --- test/Algebra/Groups/Presentation.v | 1 + 1 file changed, 1 insertion(+) diff --git a/test/Algebra/Groups/Presentation.v b/test/Algebra/Groups/Presentation.v index 6baf11c6d15..58b1125a4ae 100644 --- a/test/Algebra/Groups/Presentation.v +++ b/test/Algebra/Groups/Presentation.v @@ -1,3 +1,4 @@ +From HoTT Require Import Basics Spaces.Finite.Fin Spaces.Finite.FinSeq. From HoTT.Algebra.Groups Require Import Group Presentation. Local Open Scope mc_scope. From 21a1cdf39459a535c295beed834a399b4294d03b Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 23 Jul 2024 11:39:11 +0200 Subject: [PATCH 037/282] introduce Coeq_ind_hprop and use it Signed-off-by: Ali Caglayan --- theories/Algebra/AbGroups/Abelianization.v | 5 +++-- theories/Algebra/Groups/FreeGroup.v | 26 ++++++++++------------ theories/Algebra/Groups/GroupCoeq.v | 3 +-- theories/Colimits/Coeq.v | 11 +++++++++ 4 files changed, 27 insertions(+), 18 deletions(-) diff --git a/theories/Algebra/AbGroups/Abelianization.v b/theories/Algebra/AbGroups/Abelianization.v index e30e3127cd1..a66f446c405 100644 --- a/theories/Algebra/AbGroups/Abelianization.v +++ b/theories/Algebra/AbGroups/Abelianization.v @@ -145,8 +145,9 @@ Section Abel. (a : forall x, P (abel_in x)) : forall (x : Abel), P x. Proof. - srapply (Abel_ind _ a). - intros; apply path_ishprop. + srapply Trunc_ind. + srapply Coeq_ind_hprop. + exact a. Defined. (** And its recursion version. *) diff --git a/theories/Algebra/Groups/FreeGroup.v b/theories/Algebra/Groups/FreeGroup.v index c6817071a1e..81e2ba0cb2a 100644 --- a/theories/Algebra/Groups/FreeGroup.v +++ b/theories/Algebra/Groups/FreeGroup.v @@ -91,8 +91,7 @@ Section Reduction. nrapply (freegroup_tau _ a). } intros [[c b] d]. revert y. - snrapply Coeq_ind. (* TODO: create and use Coeq_ind_hprop *) - 2: intro; rapply path_ishprop. + srapply Coeq_ind_hprop. intro a. change (freegroup_eta ((c ++ [b] ++ [b^] ++ d) ++ a) = freegroup_eta ((c ++ d) ++ a)). @@ -198,9 +197,9 @@ Section Reduction. Proof. intros x y z. strip_truncations. - revert x; snrapply Coeq_ind; intro x; [ | apply path_ishprop]. - revert y; snrapply Coeq_ind; intro y; [ | apply path_ishprop]. - revert z; snrapply Coeq_ind; intro z; [ | apply path_ishprop]. + revert x; srapply Coeq_ind_hprop; intro x. + revert y; srapply Coeq_ind_hprop; intro y. + revert z; srapply Coeq_ind_hprop; intro z. nrapply (ap (tr o coeq)). nrapply app_assoc. Defined. @@ -209,7 +208,7 @@ Section Reduction. Global Instance leftidentity_freegroup_type : LeftIdentity sg_op mon_unit. Proof. rapply Trunc_ind. - srapply Coeq_ind; intro x; [ | apply path_ishprop]. + srapply Coeq_ind_hprop; intros x. reflexivity. Defined. @@ -217,7 +216,7 @@ Section Reduction. Global Instance rightidentity_freegroup_type : RightIdentity sg_op mon_unit. Proof. rapply Trunc_ind. - srapply Coeq_ind; intro x; [ | apply path_ishprop]. + srapply Coeq_ind_hprop; intros x. apply (ap tr), ap. nrapply app_nil. Defined. @@ -226,7 +225,7 @@ Section Reduction. Global Instance leftinverse_freegroup_type : LeftInverse sg_op negate mon_unit. Proof. rapply Trunc_ind. - srapply Coeq_ind; intro x; [ | apply path_ishprop]. + srapply Coeq_ind_hprop; intro x. apply word_concat_Vw. Defined. @@ -234,7 +233,7 @@ Section Reduction. Global Instance rightinverse_freegroup_type : RightInverse sg_op negate mon_unit. Proof. rapply Trunc_ind. - srapply Coeq_ind; intro x; [ | apply path_ishprop]. + srapply Coeq_ind_hprop; intro x. apply word_concat_wV. Defined. @@ -313,8 +312,8 @@ Section Reduction. intros [[b a] c]. apply words_rec_coh. } intros x y; strip_truncations. - revert x; snrapply Coeq_ind; hnf; intro x; [ | apply path_ishprop ]. - revert y; snrapply Coeq_ind; hnf; intro y; [ | apply path_ishprop ]. + revert x; srapply Coeq_ind_hprop; intro x. + revert y; srapply Coeq_ind_hprop; intro y. simpl. apply words_rec_pp. Defined. @@ -334,9 +333,8 @@ Section Reduction. : forall x, P x. Proof. rapply Trunc_ind. - snrapply Coeq_ind. - - exact H1. - - intro; apply path_ishprop. + srapply Coeq_ind_hprop. + exact H1. Defined. Definition FreeGroup_ind_hprop (P : FreeGroup -> Type) diff --git a/theories/Algebra/Groups/GroupCoeq.v b/theories/Algebra/Groups/GroupCoeq.v index 42f009caf9b..3986a7c7484 100644 --- a/theories/Algebra/Groups/GroupCoeq.v +++ b/theories/Algebra/Groups/GroupCoeq.v @@ -34,8 +34,7 @@ Proof. refine (p _ @ _). revert x. rapply Trunc_ind. - srapply Coeq_ind. - 2: intros; apply path_ishprop. + srapply Coeq_ind_hprop. intros w. hnf. induction w. 1: apply ap, grp_homo_unit. diff --git a/theories/Colimits/Coeq.v b/theories/Colimits/Coeq.v index 0d3b2904ae4..30fedec3730 100644 --- a/theories/Colimits/Coeq.v +++ b/theories/Colimits/Coeq.v @@ -53,6 +53,17 @@ Proof. rapply GraphQuotient_rec_beta_gqglue. Defined. +Definition Coeq_ind_hprop {B A f g} (P : @Coeq B A f g -> Type) + `{forall x, IsHProp (P x)} + (i : forall a, P (coeq a)) + : forall x, P x. +Proof. + snrapply Coeq_ind. + 1: exact i. + intros b. + rapply path_ishprop. +Defined. + (** ** Universal property *) (** See Colimits/CoeqUnivProp.v for a similar universal property without [Funext]. *) From 8f8acb40a844a4d201ec09c9eca71a896b280737 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 24 Jul 2024 16:17:54 +0200 Subject: [PATCH 038/282] Update test/Algebra/Groups/Presentation.v --- test/Algebra/Groups/Presentation.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/test/Algebra/Groups/Presentation.v b/test/Algebra/Groups/Presentation.v index 58b1125a4ae..c76df930aec 100644 --- a/test/Algebra/Groups/Presentation.v +++ b/test/Algebra/Groups/Presentation.v @@ -1,5 +1,5 @@ From HoTT Require Import Basics Spaces.Finite.Fin Spaces.Finite.FinSeq. -From HoTT.Algebra.Groups Require Import Group Presentation. +From HoTT.Algebra.Groups Require Import Group Presentation FreeGroup. Local Open Scope mc_scope. Local Open Scope mc_mult_scope. From 8552b1bf5a48d7ad566075c6046e949c4e6fbec6 Mon Sep 17 00:00:00 2001 From: Dan Christensen Date: Thu, 25 Jul 2024 18:23:16 -0400 Subject: [PATCH 039/282] In a normal subgroup, N(x * y) -> N(y * x) --- theories/Algebra/AbGroups/AbelianGroup.v | 2 +- theories/Algebra/Groups/Subgroup.v | 51 ++++++++++++------------ 2 files changed, 26 insertions(+), 27 deletions(-) diff --git a/theories/Algebra/AbGroups/AbelianGroup.v b/theories/Algebra/AbGroups/AbelianGroup.v index 8d557283b4d..ed07aa7ba34 100644 --- a/theories/Algebra/AbGroups/AbelianGroup.v +++ b/theories/Algebra/AbGroups/AbelianGroup.v @@ -74,7 +74,7 @@ Coercion abgroup_subgroup : Subgroup >-> AbGroup. Global Instance isnormal_ab_subgroup (G : AbGroup) (H : Subgroup G) : IsNormalSubgroup H. Proof. - intros x y; unfold in_cosetL, in_cosetR; intros h. + intros x y h. rewrite ab_comm, grp_assoc, grp_inv_l, grp_unit_l. exact h. Defined. diff --git a/theories/Algebra/Groups/Subgroup.v b/theories/Algebra/Groups/Subgroup.v index 2188cca1a84..4a391b392fc 100644 --- a/theories/Algebra/Groups/Subgroup.v +++ b/theories/Algebra/Groups/Subgroup.v @@ -348,40 +348,39 @@ Record NormalSubgroup (G : Group) := { Coercion normalsubgroup_subgroup : NormalSubgroup >-> Subgroup. Global Existing Instance normalsubgroup_isnormal. - - + +(** A product [x * y] is in a normal subgroup if and only if [y * x] is. *) +Definition symmetric_in_normalsubgroup {G : Group} + (N : NormalSubgroup G) + : forall x y, N (x * y) -> N (y * x). +Proof. + intros x y n. + refine (transport N _ (isnormal (y:=y) n)). + lhs nrapply (ap (.* -y) (grp_assoc y x y)). + lhs_V nrapply grp_assoc. + lhs nrapply (ap _ (grp_inv_r y)). + apply grp_unit_r. +Defined. + +Definition equiv_symmetric_in_normalsubgroup {G : Group} + (N : NormalSubgroup G) + : forall x y, N (x * y) <~> N (y * x). +Proof. + intros x y. + rapply equiv_iff_hprop. + all: apply symmetric_in_normalsubgroup. +Defined. + (** Left and right cosets are equivalent in normal subgroups. *) Definition equiv_isnormalsubgroup_in_cosetL_in_cosetR {G : Group} (N : NormalSubgroup G) : forall x y, in_cosetL N x y <~> in_cosetR N x y. Proof. intros x y. - rapply equiv_iff_hprop; cbn. - - intros n. - nrefine (transport N (grp_unit_r _) _). - nrefine (transport (fun z => N (x * -y * z)) (grp_inv_r x) _). - nrefine (transport N (grp_assoc _ _ _) _). - nrefine (transport (fun z => N (x * z)) (grp_assoc _ _ _)^ _). - nrefine (transport N (grp_assoc _ _ _)^ _). - apply isnormal. - apply subgroup_in_inv'. - nrefine (transport N (grp_inv_op _ _)^ _). - nrefine (transport (fun z => N (- x * z)) (grp_inv_inv _)^ _). - exact n. - - intros n. - nrefine (transport N (grp_unit_r _) _). - nrefine (transport (fun z => N ((- x * y) * z)) (grp_inv_r (-x)) _). - nrefine (transport N (grp_assoc _ _ _) _). - nrefine (transport (fun z => N (- x * z)) (grp_assoc _ _ _)^ _). - nrefine (transport N (grp_assoc _ _ _)^ _). - apply isnormal. - apply subgroup_in_inv'. - nrefine (transport N (grp_inv_op _ _)^ _). - nrefine (transport (fun z => N (z * -y)) (grp_inv_inv _)^ _). - exact n. + exact (equiv_in_cosetR_symm _ _ oE equiv_symmetric_in_normalsubgroup _ _ _). Defined. -(* Inverses are then respected *) +(** Inverses are then respected *) Definition in_cosetL_inverse {G : Group} {N : NormalSubgroup G} : forall x y : G, in_cosetL N (-x) (-y) <~> in_cosetL N x y. Proof. From 202cc74dea672d418bc94ec51929d0ba3bfb779a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 12:52:15 +0200 Subject: [PATCH 040/282] rename equiv_isnormalsubgroup_in_cosetL_in_cosetR Signed-off-by: Ali Caglayan --- theories/Algebra/Groups/Subgroup.v | 14 +++++++------- 1 file changed, 7 insertions(+), 7 deletions(-) diff --git a/theories/Algebra/Groups/Subgroup.v b/theories/Algebra/Groups/Subgroup.v index 4a391b392fc..6065ae8ca7d 100644 --- a/theories/Algebra/Groups/Subgroup.v +++ b/theories/Algebra/Groups/Subgroup.v @@ -372,7 +372,7 @@ Proof. Defined. (** Left and right cosets are equivalent in normal subgroups. *) -Definition equiv_isnormalsubgroup_in_cosetL_in_cosetR {G : Group} +Definition equiv_in_cosetL_in_cosetR_normalsubgroup {G : Group} (N : NormalSubgroup G) : forall x y, in_cosetL N x y <~> in_cosetR N x y. Proof. @@ -385,7 +385,7 @@ Definition in_cosetL_inverse {G : Group} {N : NormalSubgroup G} : forall x y : G, in_cosetL N (-x) (-y) <~> in_cosetL N x y. Proof. intros x y. - refine (_ oE equiv_isnormalsubgroup_in_cosetL_in_cosetR _ _ _); cbn. + refine (_ oE equiv_in_cosetL_in_cosetR_normalsubgroup _ _ _); cbn. by rewrite negate_involutive. Defined. @@ -393,7 +393,7 @@ Definition in_cosetR_inverse {G : Group} {N : NormalSubgroup G} : forall x y : G, in_cosetR N (-x) (-y) <~> in_cosetR N x y. Proof. intros x y. - refine (_ oE equiv_isnormalsubgroup_in_cosetL_in_cosetR _ _ _); cbn. + refine (_ oE equiv_in_cosetL_in_cosetR_normalsubgroup _ _ _); cbn. by rewrite grp_inv_inv. Defined. @@ -407,11 +407,11 @@ Proof. rewrite negate_sg_op, <- simple_associativity. apply symmetric_in_cosetL; cbn. rewrite simple_associativity. - apply equiv_isnormalsubgroup_in_cosetL_in_cosetR; cbn. + apply equiv_in_cosetL_in_cosetR_normalsubgroup; cbn. rewrite <- simple_associativity. apply subgroup_in_op. 1: assumption. - by apply equiv_isnormalsubgroup_in_cosetL_in_cosetR, symmetric_in_cosetL. + by apply equiv_in_cosetL_in_cosetR_normalsubgroup, symmetric_in_cosetL. Defined. Definition in_cosetR_cong {G : Group} {N : NormalSubgroup G} @@ -423,11 +423,11 @@ Proof. rewrite negate_sg_op, simple_associativity. apply symmetric_in_cosetR; cbn. rewrite <- simple_associativity. - apply equiv_isnormalsubgroup_in_cosetL_in_cosetR; cbn. + apply equiv_in_cosetL_in_cosetR_normalsubgroup; cbn. rewrite simple_associativity. apply subgroup_in_op. 2: assumption. - by apply equiv_isnormalsubgroup_in_cosetL_in_cosetR, symmetric_in_cosetR. + by apply equiv_in_cosetL_in_cosetR_normalsubgroup, symmetric_in_cosetR. Defined. (** The property of being the trivial subgroup is useful. *) From cc687e37328afd1a51c52e46c73bc25514b4c90a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 13:09:25 +0200 Subject: [PATCH 041/282] change definition of normal subgroup to forall x y, N(x*y) -> N(y*x) Signed-off-by: Ali Caglayan --- theories/Algebra/AbGroups/AbelianGroup.v | 3 +- theories/Algebra/Groups/Kernel.v | 21 +++++---- theories/Algebra/Groups/Subgroup.v | 58 ++++++++---------------- 3 files changed, 33 insertions(+), 49 deletions(-) diff --git a/theories/Algebra/AbGroups/AbelianGroup.v b/theories/Algebra/AbGroups/AbelianGroup.v index ed07aa7ba34..8d5bc5f232c 100644 --- a/theories/Algebra/AbGroups/AbelianGroup.v +++ b/theories/Algebra/AbGroups/AbelianGroup.v @@ -75,8 +75,7 @@ Global Instance isnormal_ab_subgroup (G : AbGroup) (H : Subgroup G) : IsNormalSubgroup H. Proof. intros x y h. - rewrite ab_comm, grp_assoc, grp_inv_l, grp_unit_l. - exact h. + by rewrite ab_comm. Defined. (** ** Quotients of abelian groups *) diff --git a/theories/Algebra/Groups/Kernel.v b/theories/Algebra/Groups/Kernel.v index 99e00f7d446..83011e6d8f0 100644 --- a/theories/Algebra/Groups/Kernel.v +++ b/theories/Algebra/Groups/Kernel.v @@ -14,15 +14,20 @@ Proof. - srapply (Build_Subgroup' (fun x => f x = group_unit)). 1: apply grp_homo_unit. intros x y p q; cbn in p, q; cbn. - refine (grp_homo_op _ _ _ @ ap011 _ p _ @ _). - 1: apply grp_homo_inv. - rewrite q; apply right_inverse. + lhs nrapply grp_homo_op. + rhs_V nrapply (grp_inv_r mon_unit). + nrapply (ap011 (.*.) p). + lhs nrapply grp_homo_inv. + nrapply (ap (-)). + exact q. - intros x y; cbn; intros p. - rewrite 2 grp_homo_op. - rewrite grp_homo_inv. - apply grp_moveL_1M^-1. - rewrite p. - apply grp_unit_r. + lhs nrapply grp_homo_op. + nrapply grp_moveR_Mg. + rhs nrapply grp_unit_r. + rhs_V nrapply grp_unit_l. + nrapply grp_moveL_gV. + lhs_V nrapply grp_homo_op. + exact p. Defined. (** ** Corecursion principle for group kernels *) diff --git a/theories/Algebra/Groups/Subgroup.v b/theories/Algebra/Groups/Subgroup.v index 6065ae8ca7d..6727be719dc 100644 --- a/theories/Algebra/Groups/Subgroup.v +++ b/theories/Algebra/Groups/Subgroup.v @@ -339,7 +339,7 @@ Defined. (** A normal subgroup is a subgroup closed under conjugation. *) Class IsNormalSubgroup {G : Group} (N : Subgroup G) - := isnormal : forall {x y}, N x -> N (y * x * -y). + := isnormal : forall {x y}, N (x * y) -> N (y * x). Record NormalSubgroup (G : Group) := { normalsubgroup_subgroup : Subgroup G ; @@ -349,50 +349,32 @@ Record NormalSubgroup (G : Group) := { Coercion normalsubgroup_subgroup : NormalSubgroup >-> Subgroup. Global Existing Instance normalsubgroup_isnormal. -(** A product [x * y] is in a normal subgroup if and only if [y * x] is. *) -Definition symmetric_in_normalsubgroup {G : Group} - (N : NormalSubgroup G) - : forall x y, N (x * y) -> N (y * x). -Proof. - intros x y n. - refine (transport N _ (isnormal (y:=y) n)). - lhs nrapply (ap (.* -y) (grp_assoc y x y)). - lhs_V nrapply grp_assoc. - lhs nrapply (ap _ (grp_inv_r y)). - apply grp_unit_r. -Defined. - Definition equiv_symmetric_in_normalsubgroup {G : Group} (N : NormalSubgroup G) : forall x y, N (x * y) <~> N (y * x). Proof. intros x y. rapply equiv_iff_hprop. - all: apply symmetric_in_normalsubgroup. + all: apply isnormal. Defined. (** Left and right cosets are equivalent in normal subgroups. *) Definition equiv_in_cosetL_in_cosetR_normalsubgroup {G : Group} - (N : NormalSubgroup G) - : forall x y, in_cosetL N x y <~> in_cosetR N x y. -Proof. - intros x y. - exact (equiv_in_cosetR_symm _ _ oE equiv_symmetric_in_normalsubgroup _ _ _). -Defined. + (N : NormalSubgroup G) (x y : G) + : in_cosetL N x y <~> in_cosetR N x y + := equiv_in_cosetR_symm _ _ oE equiv_symmetric_in_normalsubgroup _ _ _. (** Inverses are then respected *) -Definition in_cosetL_inverse {G : Group} {N : NormalSubgroup G} - : forall x y : G, in_cosetL N (-x) (-y) <~> in_cosetL N x y. +Definition in_cosetL_inverse {G : Group} {N : NormalSubgroup G} (x y : G) + : in_cosetL N (-x) (-y) <~> in_cosetL N x y. Proof. - intros x y. refine (_ oE equiv_in_cosetL_in_cosetR_normalsubgroup _ _ _); cbn. by rewrite negate_involutive. Defined. -Definition in_cosetR_inverse {G : Group} {N : NormalSubgroup G} - : forall x y : G, in_cosetR N (-x) (-y) <~> in_cosetR N x y. +Definition in_cosetR_inverse {G : Group} {N : NormalSubgroup G} (x y : G) + : in_cosetR N (-x) (-y) <~> in_cosetR N x y. Proof. - intros x y. refine (_ oE equiv_in_cosetL_in_cosetR_normalsubgroup _ _ _); cbn. by rewrite grp_inv_inv. Defined. @@ -405,13 +387,12 @@ Proof. cbn; intros p q. (** rewrite goal before applying subgroup_op *) rewrite negate_sg_op, <- simple_associativity. - apply symmetric_in_cosetL; cbn. - rewrite simple_associativity. - apply equiv_in_cosetL_in_cosetR_normalsubgroup; cbn. - rewrite <- simple_associativity. + apply isnormal. + rewrite simple_associativity, <- simple_associativity. apply subgroup_in_op. - 1: assumption. - by apply equiv_in_cosetL_in_cosetR_normalsubgroup, symmetric_in_cosetL. + 1: exact p. + apply isnormal. + exact q. Defined. Definition in_cosetR_cong {G : Group} {N : NormalSubgroup G} @@ -421,13 +402,12 @@ Proof. cbn; intros p q. (** rewrite goal before applying subgroup_op *) rewrite negate_sg_op, simple_associativity. - apply symmetric_in_cosetR; cbn. - rewrite <- simple_associativity. - apply equiv_in_cosetL_in_cosetR_normalsubgroup; cbn. - rewrite simple_associativity. + apply isnormal. + rewrite <- simple_associativity, simple_associativity. apply subgroup_in_op. - 2: assumption. - by apply equiv_in_cosetL_in_cosetR_normalsubgroup, symmetric_in_cosetR. + 2: exact q. + apply isnormal. + exact p. Defined. (** The property of being the trivial subgroup is useful. *) From 93c60ce7e6085b4885b6fbff011ded5b797956b1 Mon Sep 17 00:00:00 2001 From: Dan Christensen Date: Tue, 30 Jul 2024 12:32:13 -0400 Subject: [PATCH 042/282] Groups: add and use some movement lemmas involving homomorphisms --- theories/Algebra/Groups/Group.v | 13 +++++++++++++ theories/Algebra/Groups/Kernel.v | 23 ++++++++--------------- 2 files changed, 21 insertions(+), 15 deletions(-) diff --git a/theories/Algebra/Groups/Group.v b/theories/Algebra/Groups/Group.v index 39eca032333..ab279839b7b 100644 --- a/theories/Algebra/Groups/Group.v +++ b/theories/Algebra/Groups/Group.v @@ -922,3 +922,16 @@ Proof. refine (ap f _). apply C. Defined. + +(** The group movement lemmas can be extended to when there is a homomorphism involved. For now, we only include these two. *) +Definition grp_homo_moveL_1V {A B : Group} (f : GroupHomomorphism A B) (x y : A) + : f (x * y) = group_unit <~> (f x = - f y) + := grp_moveL_1V oE equiv_concat_l (grp_homo_op f x y)^ _. + +Definition grp_homo_moveL_1M {A B : Group} (f : GroupHomomorphism A B) (x y : A) + : f (x * -y) = group_unit <~> (f x = f y). +Proof. + refine (grp_moveL_1M oE equiv_concat_l _^ _). + lhs nrapply grp_homo_op. + apply ap, grp_homo_inv. +Defined. diff --git a/theories/Algebra/Groups/Kernel.v b/theories/Algebra/Groups/Kernel.v index 83011e6d8f0..2d418803558 100644 --- a/theories/Algebra/Groups/Kernel.v +++ b/theories/Algebra/Groups/Kernel.v @@ -11,23 +11,16 @@ Local Open Scope mc_mult_scope. Definition grp_kernel {A B : Group} (f : GroupHomomorphism A B) : NormalSubgroup A. Proof. snrapply Build_NormalSubgroup. - - srapply (Build_Subgroup' (fun x => f x = group_unit)). + - srapply (Build_Subgroup' (fun x => f x = group_unit)); cbn beta. 1: apply grp_homo_unit. - intros x y p q; cbn in p, q; cbn. - lhs nrapply grp_homo_op. - rhs_V nrapply (grp_inv_r mon_unit). - nrapply (ap011 (.*.) p). - lhs nrapply grp_homo_inv. - nrapply (ap (-)). - exact q. + intros x y p q. + apply (grp_homo_moveL_1M _ _ _)^-1. + exact (p @ q^). - intros x y; cbn; intros p. - lhs nrapply grp_homo_op. - nrapply grp_moveR_Mg. - rhs nrapply grp_unit_r. - rhs_V nrapply grp_unit_l. - nrapply grp_moveL_gV. - lhs_V nrapply grp_homo_op. - exact p. + apply (grp_homo_moveL_1V _ _ _)^-1. + lhs_V nrapply grp_inv_inv. + apply (ap (-)). + exact ((grp_homo_moveL_1V f x y) p)^. Defined. (** ** Corecursion principle for group kernels *) From dea660a04d12fb1a13a3edcbcccfd11045d85c8a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 18:57:03 +0200 Subject: [PATCH 043/282] show that IsNormalSubgroup is equivalent to the classical definition Signed-off-by: Ali Caglayan --- theories/Algebra/Groups/Subgroup.v | 45 ++++++++++++++++++++++++++++++ 1 file changed, 45 insertions(+) diff --git a/theories/Algebra/Groups/Subgroup.v b/theories/Algebra/Groups/Subgroup.v index 5bc0e69ce27..7bab540f3a3 100644 --- a/theories/Algebra/Groups/Subgroup.v +++ b/theories/Algebra/Groups/Subgroup.v @@ -346,6 +346,8 @@ Record NormalSubgroup (G : Group) := { normalsubgroup_isnormal : IsNormalSubgroup normalsubgroup_subgroup ; }. +Arguments Build_NormalSubgroup G N _ : rename. + Coercion normalsubgroup_subgroup : NormalSubgroup >-> Subgroup. Global Existing Instance normalsubgroup_isnormal. @@ -358,6 +360,49 @@ Proof. all: apply isnormal. Defined. +(** Our definiiton of normal subgroup implies the usual definition of invariance under conjugation. *) +Definition isnormal_conjugate {G : Group} (N : NormalSubgroup G) {x y : G} + : N x -> N (y * x * -y). +Proof. + intros n. + apply isnormal. + nrefine (transport N (grp_assoc _ _ _)^ _). + nrefine (transport (fun y => N (y * x)) (grp_inv_l _)^ _). + nrefine (transport N (grp_unit_l _)^ _). + exact n. +Defined. + +(** We can show a subgroup is normal if it is invariant under conjugation. *) +Definition Build_IsNormalSubgroup' (G : Group) (N : Subgroup G) + (isnormal : forall x y, N x -> N (y * x * -y)) + : IsNormalSubgroup N. +Proof. + intros x y n. + nrefine (transport N (grp_unit_r _) _). + nrefine (transport (fun z => N (_ * z)) (grp_inv_r y) _). + nrefine (transport N (grp_assoc _ _ _)^ _). + nrefine (transport (fun z => N (z * _)) (grp_assoc _ _ _) _). + by apply isnormal. +Defined. + +(** Under funext, being a normal subgroup is a hprop. *) +Global Instance ishprop_isnormalsubgroup `{Funext} {G : Group} (N : Subgroup G) + : IsHProp (IsNormalSubgroup N). +Proof. + unfold IsNormalSubgroup; exact _. +Defined. + +(** Our definition of normal subgroup and the usual definition are therefore equivalent. *) +Definition equiv_isnormal_conjugate `{Funext} {G : Group} (N : Subgroup G) + : IsNormalSubgroup N <~> (forall x y, N x -> N (y * x * -y)). +Proof. + rapply equiv_iff_hprop. + - intros is_normal x y. + exact (isnormal_conjugate (Build_NormalSubgroup G N is_normal)). + - intros is_normal'. + by snrapply Build_IsNormalSubgroup'. +Defined. + (** Left and right cosets are equivalent in normal subgroups. *) Definition equiv_in_cosetL_in_cosetR_normalsubgroup {G : Group} (N : NormalSubgroup G) (x y : G) From 5e1940fdd3fcfa1afaca1f230e01f397c26c66bf Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 19:01:44 +0200 Subject: [PATCH 044/282] ci: python -> python3 Signed-off-by: Ali Caglayan --- .github/workflows/ci.yml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index be897505df4..0fafd68f5f1 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -133,7 +133,7 @@ jobs: endGroup echo "::remove-matcher owner=coq-problem-matcher::" # remove problem matcher installed by Coq docker action, so we don't get duplicate warning annotations sudo apt-get -o Acquire::Retries=30 update -q - sudo apt-get -o Acquire::Retries=30 install python -y --allow-unauthenticated + sudo apt-get -o Acquire::Retries=30 install python3 -y --allow-unauthenticated etc/coq-scripts/github/reportify-coq.sh --errors ${{ matrix.extra-gh-reportify }} make TIMED=1 -j2 --output-sync - name: Revert permissions @@ -375,7 +375,7 @@ jobs: ocaml_version: ${{ env.ocaml-version }} custom_script: | sudo apt-get update - sudo apt-get install -y time python lua5.1 + sudo apt-get install -y time python3 lua5.1 startGroup "Workaround permission issue" # https://github.com/coq-community/docker-coq-action#permissions sudo chown -R coq:coq . endGroup From faab9404c6f328a1fd7f785ae914e7b9bac5325e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 19:19:33 +0200 Subject: [PATCH 045/282] ci: try to create a virtual python environment for alectryon job Signed-off-by: Ali Caglayan --- .github/workflows/ci.yml | 7 ++++++- 1 file changed, 6 insertions(+), 1 deletion(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 0fafd68f5f1..39c35eca0b3 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -203,7 +203,12 @@ jobs: opam install -y coq-serapi sudo apt-get -o Acquire::Retries=30 update -q sudo apt-get -o Acquire::Retries=30 install python3-pip autoconf -y --allow-unauthenticated - python3 -m pip install --user --upgrade pygments dominate beautifulsoup4 docutils==0.17.1 + # Create and activate a virtual environment + python3 -m venv myenv + source myenv/bin/activate + # Install the required Python packages in the virtual environment + python -m pip install --upgrade pip + python -m pip install pygments dominate beautifulsoup4 docutils==0.17.1 startGroup "Workaround permission issue" # https://github.com/coq-community/docker-coq-action#permissions sudo chown -R coq:coq . endGroup From d2bed00a42d75504230419b493fe9886e0925bc2 Mon Sep 17 00:00:00 2001 From: Dan Christensen Date: Tue, 30 Jul 2024 15:45:51 -0400 Subject: [PATCH 046/282] ci: move workaround for permissions earlier --- .github/workflows/ci.yml | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 39c35eca0b3..5e8efa773c9 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -203,15 +203,15 @@ jobs: opam install -y coq-serapi sudo apt-get -o Acquire::Retries=30 update -q sudo apt-get -o Acquire::Retries=30 install python3-pip autoconf -y --allow-unauthenticated + startGroup "Workaround permission issue" # https://github.com/coq-community/docker-coq-action#permissions + sudo chown -R coq:coq . + endGroup # Create and activate a virtual environment python3 -m venv myenv source myenv/bin/activate # Install the required Python packages in the virtual environment python -m pip install --upgrade pip python -m pip install pygments dominate beautifulsoup4 docutils==0.17.1 - startGroup "Workaround permission issue" # https://github.com/coq-community/docker-coq-action#permissions - sudo chown -R coq:coq . - endGroup echo "::remove-matcher owner=coq-problem-matcher::" # remove problem matcher installed by Coq docker action, so we don't get duplicate warning annotations make alectryon ALECTRYON_EXTRAFLAGS=--traceback - name: Revert permissions From 7c4f68576100a93843ba653c9b5139b3dbde0ceb Mon Sep 17 00:00:00 2001 From: Dan Christensen Date: Tue, 30 Jul 2024 19:47:05 -0400 Subject: [PATCH 047/282] Install python-is-python3 Co-authored-by: Jason Gross --- .github/workflows/ci.yml | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 5e8efa773c9..25fc56ac472 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -133,7 +133,7 @@ jobs: endGroup echo "::remove-matcher owner=coq-problem-matcher::" # remove problem matcher installed by Coq docker action, so we don't get duplicate warning annotations sudo apt-get -o Acquire::Retries=30 update -q - sudo apt-get -o Acquire::Retries=30 install python3 -y --allow-unauthenticated + sudo apt-get -o Acquire::Retries=30 install python3 python-is-python3 -y --allow-unauthenticated etc/coq-scripts/github/reportify-coq.sh --errors ${{ matrix.extra-gh-reportify }} make TIMED=1 -j2 --output-sync - name: Revert permissions @@ -380,7 +380,7 @@ jobs: ocaml_version: ${{ env.ocaml-version }} custom_script: | sudo apt-get update - sudo apt-get install -y time python3 lua5.1 + sudo apt-get install -y time python3 python-is-python3 lua5.1 startGroup "Workaround permission issue" # https://github.com/coq-community/docker-coq-action#permissions sudo chown -R coq:coq . endGroup From e5ea9cfffde80b12ac5e37b297ceac4810bab2e5 Mon Sep 17 00:00:00 2001 From: Dan Christensen Date: Tue, 30 Jul 2024 19:48:16 -0400 Subject: [PATCH 048/282] ci: install python3-venv --- .github/workflows/ci.yml | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/.github/workflows/ci.yml b/.github/workflows/ci.yml index 25fc56ac472..f447f0e6d9e 100644 --- a/.github/workflows/ci.yml +++ b/.github/workflows/ci.yml @@ -202,7 +202,7 @@ jobs: custom_script: | opam install -y coq-serapi sudo apt-get -o Acquire::Retries=30 update -q - sudo apt-get -o Acquire::Retries=30 install python3-pip autoconf -y --allow-unauthenticated + sudo apt-get -o Acquire::Retries=30 install python3-pip python3-venv autoconf -y --allow-unauthenticated startGroup "Workaround permission issue" # https://github.com/coq-community/docker-coq-action#permissions sudo chown -R coq:coq . endGroup From c98c6d387c3359516411b216786400ff26d521b8 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 14:53:38 +0200 Subject: [PATCH 049/282] Fix binding of type_scope to Sortclass We fix the order in which we bind type_scope to Sortclass allowing targets of function types to automatically be interpreted as types. This allows us to remove %type annotations throughout the library. Signed-off-by: Ali Caglayan --- contrib/HoTTBookExercises.v | 2 +- theories/Basics/Notations.v | 6 ++++-- theories/Spaces/BAut/Cantor.v | 2 +- theories/Spaces/Nat/Arithmetic.v | 4 ++-- theories/Spaces/No/Addition.v | 2 +- theories/Tactics.v | 4 ++-- 6 files changed, 11 insertions(+), 9 deletions(-) diff --git a/contrib/HoTTBookExercises.v b/contrib/HoTTBookExercises.v index a80996429c5..9595407da5c 100644 --- a/contrib/HoTTBookExercises.v +++ b/contrib/HoTTBookExercises.v @@ -1169,7 +1169,7 @@ End Book_4_5. Section Book_4_6_i. Definition is_qinv {A B : Type} (f : A -> B) - := { g : B -> A & ((f o g == idmap) * (g o f == idmap))%type }. + := { g : B -> A & (f o g == idmap) * (g o f == idmap) }. Definition qinv (A B : Type) := { f : A -> B & is_qinv f }. Definition qinv_id A : qinv A A diff --git a/theories/Basics/Notations.v b/theories/Basics/Notations.v index 3aa4ff17b2c..5acb4fcf3ec 100644 --- a/theories/Basics/Notations.v +++ b/theories/Basics/Notations.v @@ -1,3 +1,7 @@ +(** [type_scope] must be declared and bound early on so that later reserved notations register correctly. *) +Declare Scope type_scope. +Bind Scope type_scope with Sortclass. + (** To reserve this notation, we must first bootstrap, and preserve the underlying [forall] notation *) Notation "'forall' x .. y , P" := (forall x , .. (forall y, P) ..) (at level 200, x binder, y binder, right associativity). Reserved Notation "'exists' x .. y , p" @@ -265,7 +269,6 @@ Reserved Infix ":::" (at level 60, right associativity). (** We define various scopes and open them in the order we expect to use them. *) Declare Scope core_scope. Declare Scope function_scope. -Declare Scope type_scope. Declare Scope equiv_scope. Declare Scope path_scope. Declare Scope fibration_scope. @@ -290,4 +293,3 @@ Global Open Scope type_scope. Global Open Scope core_scope. Bind Scope function_scope with Funclass. -Bind Scope type_scope with Sortclass. diff --git a/theories/Spaces/BAut/Cantor.v b/theories/Spaces/BAut/Cantor.v index 47e26b80e93..c624110fc27 100644 --- a/theories/Spaces/BAut/Cantor.v +++ b/theories/Spaces/BAut/Cantor.v @@ -20,7 +20,7 @@ Section Assumptions. Proof. intros Z. (** Here is the important part of this definition. *) - exists (Z + Cantor)%type. + exists (Z + Cantor). (** The rest is just a proof that [Z+Cantor] is again equivalent to [Cantor], using [cantor_fold] and the assumption that [Z] is equivalent to [Cantor]. *) pose (e := Z.2); simpl in e; clearbody e. strip_truncations. diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 2bb6df21370..76a9f0bf4f5 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -93,7 +93,7 @@ Proof. Defined. Proposition diseq_implies_lt (n m : nat) - : n <> m -> sum (n < m) (n > m). + : n <> m -> (n < m) + (n > m). Proof. intros diseq. destruct (dec (n < m)) as [| a]; [ now left |]. @@ -271,7 +271,7 @@ Proof. Defined. Definition nleqSm_dichot {n m : nat} - : (n <= m.+1) -> sum (n <= m) (n = m.+1). + : (n <= m.+1) -> (n <= m) + (n = m.+1). Proof. revert m; simple_induction n n IHn. - intro. left. exact (leq_0_n m). diff --git a/theories/Spaces/No/Addition.v b/theories/Spaces/No/Addition.v index a20fe702477..94187db35a8 100644 --- a/theories/Spaces/No/Addition.v +++ b/theories/Spaces/No/Addition.v @@ -298,7 +298,7 @@ Section Addition. refine (path_No_easy _ _ _ _ eL eR _ _ _ _); intros; repeat match goal with - | [ H : (?A + ?B)%type |- _ ] => destruct H + | [ H : (?A + ?B) |- _ ] => destruct H end; repeat match goal with | [ |- context[@equiv_fun ?A ?B ?e ?v] ] diff --git a/theories/Tactics.v b/theories/Tactics.v index a0714555fe5..0518bd4438c 100644 --- a/theories/Tactics.v +++ b/theories/Tactics.v @@ -49,7 +49,7 @@ Lemma path_forall_recr_beta' `{Funext} A B x0 P f g e Px g (@path_forall _ _ _ _ _ e) Px - = @transport ((forall a, B a) * B x0)%type + = @transport ((forall a, B a) * B x0) (fun x => P (fst x) (snd x)) (f, f x0) (g, g x0) @@ -140,7 +140,7 @@ Ltac transport_path_forall_hammer := (** An example showing that it works *) Lemma path_forall_2_beta' `{Funext} A B x0 x1 P f g e Px : @transport (forall a : A, B a) (fun f => P (f x0) (f x1)) f g (@path_forall _ _ _ _ _ e) Px - = @transport (B x0 * B x1)%type (fun x => P (fst x) (snd x)) (f x0, f x1) (g x0, g x1) (path_prod' (e x0) (e x1)) Px. + = @transport (B x0 * B x1) (fun x => P (fst x) (snd x)) (f x0, f x1) (g x0, g x1) (path_prod' (e x0) (e x1)) Px. Proof. transport_path_forall_hammer. repeat match goal with From 42647dca5b1d0412d77b8c5c1c476090be7c0594 Mon Sep 17 00:00:00 2001 From: Dan Christensen Date: Wed, 31 Jul 2024 17:49:06 -0400 Subject: [PATCH 050/282] Categories: remove some %type annotations --- theories/Categories/Category/Prod.v | 6 +++--- theories/Categories/Category/Sigma/Univalent.v | 8 ++++---- theories/Categories/Category/Sum.v | 2 +- theories/Categories/ExponentialLaws/Tactics.v | 4 ++-- theories/Categories/Functor/Sum.v | 2 +- theories/Categories/NatCategory.v | 2 +- 6 files changed, 12 insertions(+), 12 deletions(-) diff --git a/theories/Categories/Category/Prod.v b/theories/Categories/Category/Prod.v index cc26ff69f9f..b1a31b81051 100644 --- a/theories/Categories/Category/Prod.v +++ b/theories/Categories/Category/Prod.v @@ -17,9 +17,9 @@ Section prod. Definition prod : PreCategory. refine (@Build_PreCategory - (C * D)%type - (fun s d => (morphism C (fst s) (fst d) - * morphism D (snd s) (snd d))%type) + (C * D) + (fun s d => morphism C (fst s) (fst d) + * morphism D (snd s) (snd d)) (fun x => (identity (fst x), identity (snd x))) (fun s d d' m2 m1 => (fst m2 o fst m1, snd m2 o snd m1)) _ diff --git a/theories/Categories/Category/Sigma/Univalent.v b/theories/Categories/Category/Sigma/Univalent.v index bed82289eb2..eb86e42e55c 100644 --- a/theories/Categories/Category/Sigma/Univalent.v +++ b/theories/Categories/Category/Sigma/Univalent.v @@ -50,7 +50,7 @@ Section onmorphisms. Variable A : PreCategory. Variable Pmor : forall s d, morphism A s d -> Type. - Local Notation mor s d := { m : _ | Pmor s d m }%type. + Local Notation mor s d := { m : _ | Pmor s d m }. Context `(HPmor : forall s d, IsHSet (mor s d)). Variable Pidentity : forall x, @Pmor x x (@identity A _). @@ -148,11 +148,11 @@ Section on_both. Variable A : PreCategory. Variable Pobj : A -> Type. - Local Notation obj := { x : _ | Pobj x }%type (only parsing). + Local Notation obj := { x : _ | Pobj x } (only parsing). Variable Pmor : forall s d : obj, morphism A s.1 d.1 -> Type. - Local Notation mor s d := { m : _ | Pmor s d m }%type (only parsing). + Local Notation mor s d := { m : _ | Pmor s d m } (only parsing). Context `(HPmor : forall s d, IsHSet (mor s d)). Variable Pidentity : forall x, @Pmor x x (@identity A _). @@ -283,7 +283,7 @@ Section on_both. Local Definition equiv_iso_A'_eisretr_helper {s d} (x : {e : @Isomorphic A s.1 d.1 - | Pmor_iso_T s d e e^-1 (@left_inverse _ _ _ e e) right_inverse }%type) + | Pmor_iso_T s d e e^-1 (@left_inverse _ _ _ e e) right_inverse }) : transport (fun e : @Isomorphic A s.1 d.1 => Pmor_iso_T s d e e^-1 left_inverse right_inverse) diff --git a/theories/Categories/Category/Sum.v b/theories/Categories/Category/Sum.v index 6608f7f4a63..89cab02edf5 100644 --- a/theories/Categories/Category/Sum.v +++ b/theories/Categories/Category/Sum.v @@ -43,7 +43,7 @@ End internals. Definition sum (C D : PreCategory) : PreCategory. Proof. refine (@Build_PreCategory - (C + D)%type + (C + D) (sum_morphism C D) (sum_identity C D) (sum_compose C D) diff --git a/theories/Categories/ExponentialLaws/Tactics.v b/theories/Categories/ExponentialLaws/Tactics.v index d280489b41c..e78835c799b 100644 --- a/theories/Categories/ExponentialLaws/Tactics.v +++ b/theories/Categories/ExponentialLaws/Tactics.v @@ -21,10 +21,10 @@ Ltac exp_laws_misc_t' := Ltac exp_laws_simplify_types' := idtac; match goal with - | [ H : (_ + _)%type |- _ ] => destruct H + | [ H : (_ + _) |- _ ] => destruct H | [ H : Unit |- _ ] => destruct H | [ H : Empty |- _ ] => destruct H - | [ H : (_ * _)%type |- _ ] => destruct H + | [ H : (_ * _) |- _ ] => destruct H | [ |- _ = _ :> Functor _ _ ] => progress path_functor | [ |- _ = _ :> NaturalTransformation _ _ ] => progress path_natural_transformation | [ |- _ = _ :> prod _ _ ] => apply path_prod diff --git a/theories/Categories/Functor/Sum.v b/theories/Categories/Functor/Sum.v index fc951597a8d..ab041edf94e 100644 --- a/theories/Categories/Functor/Sum.v +++ b/theories/Categories/Functor/Sum.v @@ -89,7 +89,7 @@ Section swap_functor. transport_path_forall_hammer. by repeat match goal with | [ H : Empty |- _ ] => destruct H - | [ H : (_ + _)%type |- _ ] => destruct H + | [ H : (_ + _) |- _ ] => destruct H | _ => progress hnf in * end. Qed. diff --git a/theories/Categories/NatCategory.v b/theories/Categories/NatCategory.v index 686f091a0f7..0a2e08aa550 100644 --- a/theories/Categories/NatCategory.v +++ b/theories/Categories/NatCategory.v @@ -18,7 +18,7 @@ Module Export Core. match n with | 0 => Empty | 1 => Unit - | S n' => (CardinalityRepresentative n' + Unit)%type + | S n' => CardinalityRepresentative n' + Unit end. Coercion CardinalityRepresentative : nat >-> Sortclass. From 73807f685805d7accacedcbd18741ef42ae9ea61 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 02:01:18 +0200 Subject: [PATCH 051/282] Update theories/Spaces/Nat/Core.v --- theories/Spaces/Nat/Core.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index a113f869df9..cd8efbaab66 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -608,7 +608,7 @@ Proof. simple_induction k k Ik; exact _. Defined. -Definition isinj_nat_add_r@{} k : IsInjective (fun x =>nat_add x k). +Definition isinj_nat_add_r@{} k : IsInjective (fun x => nat_add x k). Proof. intros x y H. rewrite 2 (nat_add_comm _ k) in H. From 5e8a988a6f3f4595963b89e62d25d2fbe6b1233e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 02:09:58 +0200 Subject: [PATCH 052/282] Nat: mention Power in nat_pow comment Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index cd8efbaab66..b80466ef7d7 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -60,7 +60,7 @@ Fixpoint nat_sub n m : nat := Notation "n - m" := (nat_sub n m) : nat_scope. -(** Exponentiation of natural numbers. *) +(** Powers or exponentiation of natural numbers. *) Definition nat_pow n m := nat_iter m (nat_mul n) 1. From 7c87eebd5b0a066728bb9ff2ea363312496a9cc4 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:22:37 +0200 Subject: [PATCH 053/282] Add TODO for Sequential.v Signed-off-by: Ali Caglayan --- theories/Colimits/Sequential.v | 2 ++ 1 file changed, 2 insertions(+) diff --git a/theories/Colimits/Sequential.v b/theories/Colimits/Sequential.v index 2178dd652a0..8a5fd5ba830 100644 --- a/theories/Colimits/Sequential.v +++ b/theories/Colimits/Sequential.v @@ -208,6 +208,8 @@ Proof. destruct p; rewrite !concat_1p, concat_p1; reflexivity. Defined. +(** TODO: The following lemmas exist in Nat/Core.v however the proofs in this file expect them to be proven in the following way. At some point we should revise the proof and adapt them to the new versions. *) + Local Lemma add_n_O : forall (n : nat), n = n + 0. Proof. simple_induction' n; simpl; auto. From 46f4d75b10fb6398a17f76ad9985edb02f5ee87c Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:25:59 +0200 Subject: [PATCH 054/282] Fix capitalisation in documentation headers Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index b80466ef7d7..a23ccc0d0e1 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -5,7 +5,7 @@ Export Basics.Nat. Local Set Universe Minimization ToSet. Local Unset Elimination Schemes. -(** * Natural Numbers *) +(** * Natural numbers *) (** We want to close the trunc_scope so that notations from there don't conflict here. *) Local Close Scope trunc_scope. @@ -25,7 +25,7 @@ Notation nat_iter n f x | m'.+1 => f (F m') end) n). -(** *** Successor and Predecessor *) +(** *** Successor and predecessor *) (** [nat_succ n] is the successor of a natural number [n]. *) Notation nat_succ := S (only parsing). @@ -37,7 +37,7 @@ Definition nat_pred n : nat := | S n' => n' end. -(** *** Arithmetic Operations *) +(** *** Arithmetic operations *) (** Addition of natural numbers *) Definition nat_add n m : nat @@ -72,7 +72,7 @@ Fixpoint nat_exp (n m : nat) : nat | S m => nat_exp n m * n end. -(** *** Maximum and Minimum *) +(** *** Maximum and minimum *) (** The maximum [nat_max n m] of two natural numbers [n] and [m]. *) Fixpoint nat_max n m := @@ -162,7 +162,7 @@ Proof. apply Hf, IHn; trivial. Defined. -(** ** Properties of Successors *) +(** ** Properties of successors *) (** TODO: remove these *) Local Definition ap_S := @ap _ _ S. @@ -202,7 +202,7 @@ Proof. Defined. #[export] Hint Resolve not_eq_n_Sn : core. -(** ** Truncatedness of Natural Numbers *) +(** ** Truncatedness of natural numbers *) (** [nat] has decidable paths. *) Global Instance decidable_paths_nat@{} : DecidablePaths nat. @@ -220,7 +220,7 @@ Defined. (** [nat] is therefore a hset. *) Global Instance ishset_nat : IsHSet nat := _. -(** ** Properties of Addition *) +(** ** Properties of addition *) (** [0] is the left identity of addition. *) Definition nat_add_zero_l@{} n : 0 + n = n @@ -266,7 +266,7 @@ Proof. exact IHn. Defined. -(** ** Properties of Multiplication *) +(** ** Properties of multiplication *) (** Multiplication by [0] on the left is [0]. *) Definition nat_mul_zero_l@{} n : 0 * n = 0 From ee6c0d72eb84836540b48c1fea55740753052588 Mon Sep 17 00:00:00 2001 From: Dan Christensen Date: Thu, 1 Aug 2024 10:24:44 -0400 Subject: [PATCH 055/282] Colimits/Sequential: use lemmas from Nat/Core instead of custom versions --- theories/Colimits/Sequential.v | 38 ++++++++++------------------------ 1 file changed, 11 insertions(+), 27 deletions(-) diff --git a/theories/Colimits/Sequential.v b/theories/Colimits/Sequential.v index 8a5fd5ba830..96393ff8f5f 100644 --- a/theories/Colimits/Sequential.v +++ b/theories/Colimits/Sequential.v @@ -187,8 +187,8 @@ Proof. destruct p; reflexivity. Defined. -Local Definition J {X Y Z} {x1 x2 : X} {y} {I : forall x, Y x -> Z} (p : x1 = x2) - : I x2 y = I x1 (coe (ap Y p^) y). +Local Definition J {X Y Z} {x1 x2 : X} {y} {I : forall x, Y x -> Z} (p : x2 = x1) + : I x2 y = I x1 (coe (ap Y p) y). Proof. destruct p; reflexivity. Defined. @@ -199,52 +199,36 @@ Proof. destruct p; reflexivity. Defined. -Local Definition L {X Y Z} {x1 x2 : X} {y} {F G} {I : forall x, Y x -> Z} {p : x1 = x2} +Local Definition L {X Y Z} {x1 x2 : X} {y} {F G} {I : forall x, Y x -> Z} {p : x2 = x1} (Q : forall x y, I (F x) (G x y) = I x y) : Q x2 y @ J p = - J (ap F p) @ (ap (I (F x1)) (K F G p^ @ ap10 (ap coe (ap (ap Y) (ap_V F p))) (G x2 y))^ @ - Q x1 (coe (ap Y p^) y)). + J (ap F p) @ (ap (I (F x1)) (K F G p)^ @ + Q x1 (coe (ap Y p) y)). Proof. destruct p; rewrite !concat_1p, concat_p1; reflexivity. Defined. -(** TODO: The following lemmas exist in Nat/Core.v however the proofs in this file expect them to be proven in the following way. At some point we should revise the proof and adapt them to the new versions. *) - -Local Lemma add_n_O : forall (n : nat), n = n + 0. -Proof. - simple_induction' n; simpl; auto. -Defined. - -Local Lemma nat_add_n_Sm (n m : nat) : (n + m).+1 = n + m.+1. -Proof. - simple_induction' n; simpl. - - reflexivity. - - apply ap; assumption. -Defined. - Global Instance isequiv_colim_shift_seq_to_colim_seq `{Funext} A n : IsEquiv (colim_shift_seq_to_colim_seq A n). Proof. induction n as [ | n e]; srapply isequiv_homotopic'. - srapply equiv_functor_colimit; srapply Build_diagram_equiv. + srapply Build_DiagramMap. - * exact (fun k => coe (ap A (add_n_O k)^)). - * intros k l p a; destruct p; srapply (K S (fun n a => a^+) (add_n_O k)^ @ _). - srapply (ap10 (ap coe (ap (ap _) (ap_V _ _)))). + * exact (fun k => coe (ap A (nat_add_zero_r k))). + * intros k l p a; destruct p. srapply (K S (fun n a => a^+) _). + exact _. - symmetry; srapply seq_colimit_uniq. - + intros k a; exact (J (add_n_O k)). + + intros k a; exact (J (nat_add_zero_r k)). + intros k a; rewrite !Colimit_rec_beta_colimp; srapply (L (glue A)). - transitivity (Colimit (succ_seq (shift_seq A n))). + srapply equiv_functor_colimit; srapply Build_diagram_equiv. * srapply Build_DiagramMap. - { exact (fun k => coe (ap A (nat_add_n_Sm k n)^)). } - { intros k l p a; destruct p; rapply (K S (fun n a => a^+) (nat_add_n_Sm k n)^ @ _). - srapply (ap10 (ap coe (ap (ap _) (ap_V _ _)))). } + { exact (fun k => coe (ap A (nat_add_succ_r k n))). } + { intros k l p a; destruct p; rapply (K S (fun n a => a^+) (nat_add_succ_r k n)). } * exact _. + srefine (transitivity (equiv_colim_succ_seq_to_colim_seq _) (Build_Equiv _ _ _ e)). - symmetry; srapply seq_colimit_uniq. - + intros k a; exact (J (nat_add_n_Sm k n)). + + intros k a; exact (J (nat_add_succ_r k n)). + intros k a; rewrite Colimit_rec_beta_colimp; simpl. rewrite 2(ap_compose' _ _ (glue _ k a)), Colimit_rec_beta_colimp, 2ap_pp. rewrite colim_succ_seq_to_colim_seq_ap_inj, colim_shift_seq_to_colim_seq_ap_inj. From 6850fbd76af3c246a0017373d900410a0ef2f82a Mon Sep 17 00:00:00 2001 From: Dan Christensen Date: Thu, 1 Aug 2024 10:25:23 -0400 Subject: [PATCH 056/282] Colimits/Sequential: get rid of rewrites in one proof --- theories/Colimits/Sequential.v | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/theories/Colimits/Sequential.v b/theories/Colimits/Sequential.v index 96393ff8f5f..65906d9b35b 100644 --- a/theories/Colimits/Sequential.v +++ b/theories/Colimits/Sequential.v @@ -205,7 +205,9 @@ Local Definition L {X Y Z} {x1 x2 : X} {y} {F G} {I : forall x, Y x -> Z} {p : x J (ap F p) @ (ap (I (F x1)) (K F G p)^ @ Q x1 (coe (ap Y p) y)). Proof. - destruct p; rewrite !concat_1p, concat_p1; reflexivity. + destruct p; cbn. + apply equiv_p1_1q. + symmetry; apply concat_1p. Defined. Global Instance isequiv_colim_shift_seq_to_colim_seq `{Funext} A n From b8bebda62c020fb7b2cdb10098c076711ec1c1b1 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 15:07:44 +0200 Subject: [PATCH 057/282] Nat: remove cumulative flag from leq since it has no universes Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index a23ccc0d0e1..66daa1ceda3 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -339,7 +339,7 @@ Defined. (** ** Inequality of natural numbers *) -Cumulative Inductive leq (n : nat) : nat -> Type0 := +Inductive leq (n : nat) : nat -> Type0 := | leq_n : leq n n | leq_S : forall m, leq n m -> leq n (S m). From 59e32e9a384896d7b3687e8bec6b424963930825 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 15:42:39 +0200 Subject: [PATCH 058/282] Nat: remove unused hints Hint databases require much more care than afforded here. The current situation seems to be adding every lemma etc as a hint. This will have to be rethought if we want a auto solver for nat lemmas, but in the mean while we should avoid using auto and rely on typeclasses instead. Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 124 ------------------------------- 1 file changed, 124 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 442782ac95d..9b7e55057ff 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -6,23 +6,6 @@ Local Set Universe Minimization ToSet. Local Close Scope trunc_scope. Local Open Scope nat_scope. -Ltac nat_absurd_trivial := - unfold ">" in *; unfold "<" in *; - match goal with - | [ H : ?n.+1 <= 0 |- _ ] => contradiction (not_leq_Sn_0 n H) - | [ H : ?n.+1 <= ?n |- _ ] => contradiction (not_lt_n_n n H) - | [ H1 : ?k.+1 <= ?n |- _ ] => - match goal with - | [ H2 : ?n <= ?k |- _] => - contradiction (not_leq_Sn_n k (@leq_trans _ _ _ H1 H2)) - end - end. - -#[export] Hint Resolve not_lt_n_n : nat. -#[export] Hint Resolve not_lt_n_0 : nat. -#[export] Hint Resolve not_leq_Sn_0 : nat. -#[export] Hint Extern 2 => nat_absurd_trivial : nat. - (** This is defined so that it can be added to the [nat] auto hint database. *) Local Definition symmetric_paths_nat (n m : nat) : n = m -> m = n := @symmetric_paths nat n m. @@ -31,10 +14,7 @@ Local Definition transitive_paths_nat (n m k : nat) : n = m -> m = k -> n = k := @transitive_paths nat n m k. #[export] Hint Resolve symmetric_paths_nat | 5 : nat. -#[export] Hint Resolve transitive_paths_nat : nat. #[export] Hint Resolve leq_0_n : nat. -#[export] Hint Resolve leq_trans : nat. -#[export] Hint Resolve leq_antisym : nat. Proposition not_lt_implies_geq {n m : nat} : ~(n < m) -> m <= n. Proof. @@ -68,8 +48,6 @@ Ltac convert_to_positive := | [|- ~ (?n <= ?m) ] => apply lt_implies_not_geq end. -#[export] Hint Extern 2 => convert_to_positive : nat. - (** Because of the inductive definition of [<=], one can destruct the proof of [n <= m] and get a judgemental identification between [n] and [m] rather than a propositional one, which may be preferable to the following lemma. *) Proposition leq_split {n m : nat} : (n <= m) -> sum (n < m) (n = m). Proof. @@ -97,8 +75,6 @@ Proof. contradiction (not_lt_n_n m). Defined. -#[export] Hint Resolve lt_implies_diseq : nat. - (** This lemma is just for convenience in the case where the user forgets to unfold the definition of [<]. *) Proposition n_lt_Sn (n : nat) : n < n.+1. Proof. @@ -111,20 +87,6 @@ Proof. now apply leq_S_n, leq_S. Defined. -Ltac easy_eq_to_ineq := - match goal with - | [ H : ?x = ?n |- ?x <= ?n ] => destruct H; constructor - | [ H : ?x.+1 = ?n |- ?x <= ?n ] => rewrite <- H; constructor; - constructor - | [ H : ?x.+1 = ?n |- ?x < ?n ] => rewrite <- H; apply leq_n - | [ H : ?x.+2 = ?n |- ?x <= ?n ] => rewrite <- H; apply leq_S'; - apply leq_S'; apply leq_n - | [ H : ?x.+2 = ?n |- ?x < ?n ] => rewrite <- H; apply leq_S_n'; - apply leq_S - end. - -#[export] Hint Extern 3 => easy_eq_to_ineq : nat. - Proposition mixed_trans1 (n m k : nat) : n <= m -> m < k -> n < k. Proof. @@ -150,9 +112,6 @@ Proof. intros l j. apply (@leq_trans (n.+1) m k); trivial. Defined. -#[export] Hint Resolve mixed_trans1 : nat. -#[export] Hint Resolve mixed_trans2 : nat. - Proposition sub_n_n (n : nat) : n - n = 0. Proof. simple_induction n n IHn. @@ -165,20 +124,6 @@ Proof. destruct n; reflexivity. Defined. -Ltac rewrite_subn0 := - match goal with - | [ |- context [ ?n - 0 ] ] => rewrite -> sub_n_0 - end. - -Ltac rewrite_subnn := - match goal with - | [ |- context [ ?n - ?n ] ] => rewrite -> sub_n_n - end. - -#[export] Hint Rewrite -> sub_n_0 : nat. -#[export] Hint Rewrite -> sub_n_n : nat. -#[export] Hint Resolve sub_n_0 : nat. - Proposition add_n_sub_n_eq (m n : nat) : m + n - n = m. Proof. destruct m. @@ -339,9 +284,6 @@ Proof. apply natminuspluseq. assumption. Defined. -#[export] Hint Rewrite -> natminuspluseq : nat. -#[export] Hint Rewrite -> natminuspluseq' : nat. - Lemma equiv_leq_add n m : leq n m <~> exists k, k + n = m. Proof. @@ -360,18 +302,6 @@ Proof. apply leq_add. Defined. -#[export] Hint Resolve leq_S_n' : nat. - -Ltac leq_S_n_in_hypotheses := - match goal with - | [ H : ?n.+1 <= ?m.+1 |- _ ] => apply leq_S_n in H - | [ H : ?n < ?m.+1 |- _ ] => apply leq_S_n in H - | [ H : ?m.+1 > ?n |- _ ] => apply leq_S_n in H - | [ H : ?m.+1 >= ?n.+1 |- _ ] => apply leq_S_n in H - end. - -#[export] Hint Extern 4 => leq_S_n_in_hypotheses : nat. - Proposition nataddpreservesleq { n m k : nat } : n <= m -> n + k <= m + k. Proof. @@ -382,8 +312,6 @@ Proof. apply leq_S_n'; exact IHk. Defined. -#[export] Hint Resolve nataddpreservesleq : nat. - Proposition nataddpreservesleq' { n m k : nat } : n <= m -> k + n <= k + m. Proof. @@ -392,8 +320,6 @@ Proof. exact nataddpreservesleq. Defined. -#[export] Hint Resolve nataddpreservesleq' : nat. - Proposition nataddpreserveslt { n m k : nat } : n < m -> n + k < m + k. Proof. @@ -551,8 +477,6 @@ Proof. destruct (symmetric_paths _ _ (nataddsub_assoc n H)); done. Defined. -#[export] Hint Resolve nataddsub_assoc_implication : nat. - Proposition nat_add_sub_eq (n : nat) {k: nat} : (k <= n) -> k + (n - k) = n. Proof. @@ -561,8 +485,6 @@ Proof. destruct (nat_add_comm n k); exact (add_n_sub_n_eq _ _). Defined. -#[export] Hint Resolve nat_add_sub_eq : nat. - Proposition predeqminus1 { n : nat } : n - 1 = nat_pred n. Proof. simple_induction' n. @@ -575,8 +497,6 @@ Proof. case n; [ apply leq_n | intro; apply leq_S; apply leq_n]. Defined. -#[export] Hint Resolve predn_leq_n : nat. - Proposition S_predn (n i: nat) : (i < n) -> S(nat_pred n) = n. Proof. simple_induction' n; intros l. @@ -584,12 +504,6 @@ Proof. - reflexivity. Defined. -#[export] Hint Rewrite S_predn : nat. -#[export] Hint Rewrite <- nat_pred_succ : nat. - -#[export] Hint Resolve S_predn : nat. -#[export] Hint Resolve leq_n_pred : nat. - Proposition pred_equiv (k n : nat) : k < n -> k < S (nat_pred n). Proof. intro ineq; destruct n. @@ -623,8 +537,6 @@ Proof. intro l; apply leq_n_pred in l; assumption. Defined. -#[export] Hint Resolve lt_implies_pred_geq : nat. - Proposition j_geq_0_lt_implies_pred_geq (i j k : nat) : i < j -> k.+1 <= j -> k <= nat_pred j. Proof. @@ -634,8 +546,6 @@ Proof. - now simpl; apply leq_S_n. Defined. -#[export] Hint Resolve lt_implies_pred_geq : nat. - Proposition pred_gt_implies_lt (i j : nat) : i < nat_pred j -> i.+1 < j. Proof. @@ -673,8 +583,6 @@ Proof. apply leq_n_pred, IHk. exact l. Defined. -#[export] Hint Resolve natsubpreservesleq : nat. - Proposition sub_less { n k : nat } : n - k <= n. Proof. revert k. @@ -686,9 +594,6 @@ Proof. [ apply IHn | apply leq_S, leq_n]. Defined. -#[export] Hint Resolve sub_less : nat. -#[export] Hint Resolve leq_S_n' : nat. - Proposition sub_less_strict { n k : nat } : 0 < n -> 0 < k -> n - k < n. Proof. @@ -710,8 +615,6 @@ Proof. exact (nataddreflectslt q'). Defined. -#[export] Hint Resolve natpmswap1 : nat. - Proposition natpmswap2 (k m n : nat) : n <= k -> k - n <= m -> k <= n + m. Proof. @@ -722,8 +625,6 @@ Proof. assumption. Defined. -#[export] Hint Resolve natpmswap2 : nat. - Proposition natpmswap3 (k m n : nat) : k <= n -> m <= n - k -> k + m <= n. Proof. @@ -734,8 +635,6 @@ Proof. assumption. Defined. -#[export] Hint Resolve natpmswap3 : nat. - Proposition natpmswap4 (k m n : nat) : k - n < m -> k < n + m. Proof. @@ -745,8 +644,6 @@ Proof. (nat_sub_add_ineq _ _)). Defined. -#[export] Hint Resolve natpmswap4 : nat. - Proposition n_leq_m_n_leq_plus_m_k (n m k : nat) : n <= m -> n <= m + k. Proof. @@ -767,27 +664,6 @@ Proof. * exact l'. Defined. -#[export] Hint Resolve nat_add_bifunctor : nat. -#[export] Hint Resolve nataddpreserveslt : nat. -#[export] Hint Resolve nataddpreservesleq' : nat. -#[export] Hint Resolve nataddpreserveslt' : nat. -#[export] Hint Resolve natineq0eq0 : nat. -#[export] Hint Resolve n_leq_add_n_k : nat. -#[export] Hint Resolve n_leq_m_n_leq_plus_m_k : nat. - -#[export] Hint Immediate add_n_sub_n_eq : nat. -#[export] Hint Immediate add_n_sub_n_eq' : nat. - -#[export] Hint Rewrite -> nat_add_zero_r : nat. -#[export] Hint Rewrite -> nat_add_zero_l : nat. -#[export] Hint Rewrite -> add_n_sub_n_eq : nat. -#[export] Hint Rewrite -> add_n_sub_n_eq' : nat. -#[export] Hint Rewrite -> nataddsub_assoc : nat. - -Ltac autorewrite_or_fail := progress ltac:(autorewrite with nat). - -#[export] Hint Extern 7 => autorewrite_or_fail : nat. - Proposition strong_induction (P : nat -> Type) : (forall n : nat, (forall m : nat, (m < n) -> P m) -> P n) -> forall n : nat, P n. From e5c743e597a7daa82c672dd98aa29ca912da7be1 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 15:45:04 +0200 Subject: [PATCH 059/282] Nat: avoid use of `auto with nat` This database doesn't appear to be constructed with much thought. We don't even declare this database which means it is either a default Coq hintdb or is declared on the fly. We explicitly remove all hints and uses of it. Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 32 +++++++------------------------- 1 file changed, 7 insertions(+), 25 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 9b7e55057ff..63028000bb0 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -13,9 +13,6 @@ Local Definition symmetric_paths_nat (n m : nat) Local Definition transitive_paths_nat (n m k : nat) : n = m -> m = k -> n = k := @transitive_paths nat n m k. -#[export] Hint Resolve symmetric_paths_nat | 5 : nat. -#[export] Hint Resolve leq_0_n : nat. - Proposition not_lt_implies_geq {n m : nat} : ~(n < m) -> m <= n. Proof. intros not_lt. @@ -94,18 +91,6 @@ Proof. apply (@leq_trans (n.+1) (m.+1) k); trivial. Defined. -Ltac leq_trans_resolve := - match goal with - | [ H : ?n <= ?m |- ?n <= ?k ] => apply (leq_trans H) - | [ H : ?k <= ?m |- ?n <= ?k ] => refine (leq_trans _ H) - | [ H : ?n <= ?m |- ?n < ?k ] => apply (mixed_trans1 _ _ _ H) - | [ H : ?m <= ?k |- ?n < ?k ] => refine (leq_trans _ H) - | [ H : ?m < ?k |- ?n < ?k ] => refine (mixed_trans1 _ _ _ _ H) - | [ H : ?n < ?m |- ?n < ?k ] => apply (leq_trans H) - end. - -#[export] Hint Extern 2 => leq_trans_resolve : nat. - Proposition mixed_trans2 (n m k : nat) : n < m -> m <= k -> n < k. Proof. @@ -152,11 +137,11 @@ Proof. intro H. apply leq_S, leq_S_n in H; exact H. Defined. -#[export] Hint Resolve n_lt_m_n_leq_m : nat. - Proposition lt_trans {n m k : nat} : n < m -> m < k -> n < k. Proof. - eauto with nat. + intros H1 H2. + nrapply n_lt_m_n_leq_m. + exact (mixed_trans1 _ _ _ H1 H2). Defined. Proposition not_both_less (n m : nat) : n < m -> ~(m < n). @@ -194,12 +179,8 @@ Proof. + change (m.+1 + k) with (m + k).+1; apply IHn. Defined. -#[export] Hint Resolve subsubadd : nat. - -Proposition subsubadd' (n m k : nat) : n - m - k = n - (m + k). -Proof. - auto with nat. -Defined. +Definition subsubadd' (n m k : nat) : n - m - k = n - (m + k) + := (subsubadd n m k)^. Definition nleqSm_dichot {n m : nat} : (n <= m.+1) -> (n <= m) + (n = m.+1). @@ -229,7 +210,8 @@ Defined. Proposition sub_leq_0_converse (n m : nat) : n - m = 0 -> n <= m. Proof. revert m; simple_induction n n IHn. - - auto with nat. + - simpl. intros m eq. + apply leq_0_n. - intros m eq. destruct m. + simpl in eq. apply symmetric_paths in eq. contradiction (not_eq_O_S n eq). From ede18e3407bf215b538aa85b6c218345d540468f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 15:49:18 +0200 Subject: [PATCH 060/282] Nat: remove symmetric_paths_nat Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 7 ------- 1 file changed, 7 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 63028000bb0..ab5a04bbf83 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -6,13 +6,6 @@ Local Set Universe Minimization ToSet. Local Close Scope trunc_scope. Local Open Scope nat_scope. -(** This is defined so that it can be added to the [nat] auto hint database. *) -Local Definition symmetric_paths_nat (n m : nat) - : n = m -> m = n := @symmetric_paths nat n m. - -Local Definition transitive_paths_nat (n m k : nat) - : n = m -> m = k -> n = k := @transitive_paths nat n m k. - Proposition not_lt_implies_geq {n m : nat} : ~(n < m) -> m <= n. Proof. intros not_lt. From db186b6344d905d4cd5799107de64dd8d5af0e02 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 15:50:08 +0200 Subject: [PATCH 061/282] Nat: remove unused ltac Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 8 -------- 1 file changed, 8 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index ab5a04bbf83..056329e9993 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -30,14 +30,6 @@ Proof. contradiction (not_lt_n_n n); by refine (leq_trans _ ineq2). Defined. -Ltac convert_to_positive := - match goal with - | [ H : ~ (?n < ?m) |- _] => apply not_lt_implies_geq in H - | [ H : ~ (?n <= ?m) |- _] => apply not_leq_implies_gt in H - | [|- ~ (?n < ?m) ] => apply leq_implies_not_gt - | [|- ~ (?n <= ?m) ] => apply lt_implies_not_geq - end. - (** Because of the inductive definition of [<=], one can destruct the proof of [n <= m] and get a judgemental identification between [n] and [m] rather than a propositional one, which may be preferable to the following lemma. *) Proposition leq_split {n m : nat} : (n <= m) -> sum (n < m) (n = m). Proof. From f15c7b9aa2f1829ab65899eb87957e2172f15d8a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 16:59:17 +0200 Subject: [PATCH 062/282] Nat: remove auto hints for nat Signed-off-by: Ali Caglayan --- theories/Classes/implementations/ne_list.v | 14 ++-- theories/Spaces/List/Theory.v | 6 +- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 82 ++++++++++------------ 4 files changed, 53 insertions(+), 51 deletions(-) diff --git a/theories/Classes/implementations/ne_list.v b/theories/Classes/implementations/ne_list.v index 337c443d116..439e86f4610 100644 --- a/theories/Classes/implementations/ne_list.v +++ b/theories/Classes/implementations/ne_list.v @@ -1,7 +1,7 @@ Require Import HoTT.Utf8Minimal HoTT.Spaces.List.Core - HoTT.Basics.Overture + HoTT.Basics.Overture Basics.Tactics HoTT.Spaces.Nat.Core. Local Open Scope nat_scope. @@ -124,11 +124,13 @@ Lemma tails_are_shorter {T} (y x: ne_list T): InList x (to_list (tails y)) → leq (length (to_list x)) (length (to_list y)). Proof. - induction y; cbn. - - intros [[] | C]. - + constructor. - + elim C. - - intros [[] | C]; auto. + induction y; cbn. + - intros [[] | C]. + + constructor. + + elim C. + - intros [[] | C]. + + exact _. + + by apply leq_S, IHy. Qed. Fixpoint map {A B} (f: A → B) (l: ne_list A): ne_list B := diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index 725e1bed0bd..f69d5846b08 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -971,7 +971,11 @@ Proof. - destruct l2 as [|y l2]. + discriminate. + intros [Hx Hl1] [Hy Hl2]. - split; auto. + split. + * by apply Hf. + * apply IHl1; trivial. + apply path_nat_S. + exact p. Defined. (** The left fold of [f] on a list [l] for which [for_all Q l] satisfies [P] if [P] and [Q] imply [P] composed with [f]. *) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 056329e9993..604b1e323be 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -480,7 +480,7 @@ Defined. Proposition n_leq_pred_Sn (n : nat) : n <= S (nat_pred n). Proof. - destruct n; auto. + destruct n; exact _. Defined. Proposition leq_implies_pred_lt (i n k : nat) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 66daa1ceda3..be1ef723a7b 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -164,26 +164,19 @@ Defined. (** ** Properties of successors *) -(** TODO: remove these *) -Local Definition ap_S := @ap _ _ S. -Local Definition ap_nat := @ap nat. -#[export] Hint Resolve ap_S : core. -#[export] Hint Resolve ap_nat : core. - Definition nat_pred_succ@{} n : nat_pred (nat_succ n) = n := idpath. (** Injectivity of successor *) Definition path_nat_S@{} n m (H : S n = S m) : n = m := ap nat_pred H. -#[export] Hint Immediate path_nat_S : core. (** TODO: rename to [neq_S] *) -(** TODO: avoid auto in proof. *) Definition not_eq_S@{} n m : n <> m -> S n <> S m. Proof. - auto. + intros np q. + apply np. + exact (path_nat_S _ _ q). Defined. -#[export] Hint Resolve not_eq_S : core. (** TODO: rename to [neq_O_S] *) (** Zero is not the successor of a number *) @@ -191,16 +184,14 @@ Definition not_eq_O_S@{} : forall n, 0 <> S n. Proof. discriminate. Defined. -#[export] Hint Resolve not_eq_O_S : core. (** TODO: rename to [neq_n_Sn] *) -(** TODO: prove above using this *) -(** TODO: remove auto. *) Theorem not_eq_n_Sn@{} n : n <> S n. Proof. - simple_induction' n; auto. + simple_induction' n. + - apply not_eq_O_S. + - by apply not_eq_S. Defined. -#[export] Hint Resolve not_eq_n_Sn : core. (** ** Truncatedness of natural numbers *) @@ -234,7 +225,6 @@ Proof. - apply (ap nat_succ). exact IHn. Defined. -#[export] Hint Resolve nat_add_zero_r : core. (** Adding a successor on the left is the same as adding and then taking the successor. *) Definition nat_add_succ_l@{} n m : n.+1 + m = (n + m).+1 @@ -243,9 +233,10 @@ Definition nat_add_succ_l@{} n m : n.+1 + m = (n + m).+1 (** Adding a successor on the right is the same as adding and then taking the successor. *) Definition nat_add_succ_r@{} n m : n + m.+1 = (n + m).+1. Proof. - simple_induction' n; simpl; auto. + simple_induction' n; simpl. + 1: reflexivity. + exact (ap S IH). Defined. -#[export] Hint Resolve nat_add_succ_r: core. (** Addition of natural numbers is commutative. *) Definition nat_add_comm@{} n m : n + m = m + n. @@ -277,7 +268,6 @@ Definition nat_mul_zero_r@{} n : n * 0 = 0. Proof. by induction n. Defined. -#[export] Hint Resolve nat_mul_zero_r : core. Definition nat_mul_succ_l@{} n m : n.+1 * m = m + n * m := idpath. @@ -292,7 +282,6 @@ Proof. nrapply (ap (nat_add m)). exact IHn. Defined. -#[export] Hint Resolve nat_mul_succ_r : core. (** Multiplication of natural numbers is commutative. *) Definition nat_mul_comm@{} n m : n * m = m * n. @@ -348,7 +337,6 @@ Scheme leq_rect := Induction for leq Sort Type. Scheme leq_rec := Minimality for leq Sort Type. Notation "n <= m" := (leq n m) : nat_scope. -#[export] Hint Constructors leq : core. Existing Class leq. Global Existing Instances leq_n leq_S. @@ -358,15 +346,15 @@ Global Instance reflexive_leq : Reflexive leq := leq_n. Lemma leq_trans {x y z} : x <= y -> y <= z -> x <= z. Proof. - induction 2; auto. + induction 2; exact _. Defined. Global Instance transitive_leq : Transitive leq := @leq_trans. Lemma leq_n_pred n m : leq n m -> leq (nat_pred n) (nat_pred m). Proof. - induction 1; auto. - destruct m; simpl; auto. + induction 1; try exact _. + destruct m; simpl; exact _. Defined. Lemma leq_S_n : forall n m, n.+1 <= m.+1 -> n <= m. @@ -377,7 +365,7 @@ Defined. Lemma leq_S_n' n m : n <= m -> n.+1 <= m.+1. Proof. - induction 1; auto. + induction 1; exact _. Defined. Global Existing Instance leq_S_n' | 100. @@ -440,7 +428,7 @@ Defined. Global Instance leq_0_n n : 0 <= n | 10. Proof. - simple_induction' n; auto. + simple_induction' n; exact _. Defined. Lemma not_leq_Sn_0 n : ~ (n.+1 <= 0). @@ -482,7 +470,7 @@ Definition lt n m : Type0 := leq (S n) m. (** We declare it as an existing class so typeclass search is performed on its goals. *) Existing Class lt. -#[export] Hint Unfold lt : core typeclass_instances. +#[export] Hint Unfold lt : typeclass_instances. Infix "<" := lt : nat_scope. (** We add a typeclass instance for unfolding the definition so lemmas about [leq] can be used. *) Global Instance lt_is_leq n m : leq n.+1 m -> lt n m | 100 := idmap. @@ -499,7 +487,7 @@ Global Instance decidable_lt n m : Decidable (lt n m) := _. Definition ge n m := leq m n. Existing Class ge. -#[export] Hint Unfold ge : core typeclass_instances. +#[export] Hint Unfold ge : typeclass_instances. Infix ">=" := ge : nat_scope. Global Instance ge_is_leq n m : leq m n -> ge n m | 100 := idmap. @@ -509,7 +497,7 @@ Global Instance decidable_ge n m : Decidable (ge n m) := _. Definition gt n m := lt m n. Existing Class gt. -#[export] Hint Unfold gt : core typeclass_instances. +#[export] Hint Unfold gt : typeclass_instances. Infix ">" := gt : nat_scope. Global Instance gt_is_leq n m : leq m.+1 n -> gt n m | 100 := idmap. @@ -537,36 +525,42 @@ Defined. Lemma nat_max_n_n n : nat_max n n = n. Proof. - simple_induction' n; cbn; auto. + simple_induction' n; cbn. + 1: reflexivity. + exact (ap S IH). Defined. -#[export] Hint Resolve nat_max_n_n : core. Lemma nat_max_Sn_n n : nat_max (S n) n = S n. Proof. - simple_induction' n; cbn; auto. + simple_induction' n; cbn. + 1: reflexivity. + exact (ap S IH). Defined. -#[export] Hint Resolve nat_max_Sn_n : core. -Lemma nat_max_comm n m : nat_max n m = nat_max m n. +Lemma nat_max_0_n n : nat_max 0 n = n. Proof. - revert m; simple_induction' n; destruct m; cbn; auto. + reflexivity. Defined. -Lemma nat_max_0_n n : nat_max 0 n = n. +Lemma nat_max_n_0 n : nat_max n 0 = n. Proof. - auto. + induction n as [|n IHn]; reflexivity. Defined. -#[export] Hint Resolve nat_max_0_n : core. -Lemma nat_max_n_0 n : nat_max n 0 = n. +Lemma nat_max_comm n m : nat_max n m = nat_max m n. Proof. - by rewrite nat_max_comm. + induction m as [|m IHm] in n |- *. + - nrapply nat_max_n_0. + - destruct n. + + reflexivity. + + exact (ap S (IHm n)). Defined. -#[export] Hint Resolve nat_max_n_0 : core. Theorem nat_max_l : forall n m, m <= n -> nat_max n m = n. Proof. - intros n m; revert n; simple_induction m m IHm; auto. + intros n m; revert n; simple_induction m m IHm. + { intros n H. + nrapply nat_max_n_0. } intros [] p. 1: inversion p. cbn; by apply (ap S), IHm, leq_S_n. @@ -579,7 +573,9 @@ Defined. Lemma nat_min_comm : forall n m, nat_min n m = nat_min m n. Proof. - simple_induction' n; destruct m; cbn; auto. + simple_induction' n; destruct m; cbn. + 1-3: reflexivity. + exact (ap S (IH _)). Defined. Theorem nat_min_l : forall n m : nat, n <= m -> nat_min n m = n. From adfa10e277e69fb08df0d000fc6e1cb65eb7fc8c Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:05:08 +0200 Subject: [PATCH 063/282] Nat: rename not_eq_S -> neq_nat_succ Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index be1ef723a7b..aca18f1aabc 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -170,8 +170,8 @@ Definition nat_pred_succ@{} n : nat_pred (nat_succ n) = n (** Injectivity of successor *) Definition path_nat_S@{} n m (H : S n = S m) : n = m := ap nat_pred H. -(** TODO: rename to [neq_S] *) -Definition not_eq_S@{} n m : n <> m -> S n <> S m. +(** Inequality of sucessors is implied with inequality of the arguments. *) +Definition neq_nat_succ@{} n m : n <> m -> S n <> S m. Proof. intros np q. apply np. @@ -190,7 +190,7 @@ Theorem not_eq_n_Sn@{} n : n <> S n. Proof. simple_induction' n. - apply not_eq_O_S. - - by apply not_eq_S. + - by apply neq_nat_succ. Defined. (** ** Truncatedness of natural numbers *) From e0ccb906a888aaa999a9ccd3f719a8e6c395882a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:07:56 +0200 Subject: [PATCH 064/282] Nat: rename path_nat_S -> path_nat_succ Signed-off-by: Ali Caglayan --- theories/Spaces/List/Theory.v | 10 +++++----- theories/Spaces/Nat/Core.v | 10 +++++----- 2 files changed, 10 insertions(+), 10 deletions(-) diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index f69d5846b08..79db96fab8c 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -194,7 +194,7 @@ Proof. - destruct l2. + inversion p. + cbn; f_ap. - by apply IHl1, path_nat_S. + by apply IHl1, path_nat_succ. Defined. (** An element of a [list_map2] is the result of applying the function to some elements of the original lists. *) @@ -215,7 +215,7 @@ Proof. destruct H as [q | i]. 1: exact (y; z; (q, (inl idpath, inl idpath))). destruct (IHl1 l2 x i) as [y' [z' [q [r s]]]]. - 1: apply path_nat_S, p. + 1: apply path_nat_succ, p. exact (y'; z'; (q, (inr r, inr s))). Defined. @@ -454,7 +454,7 @@ Proof. unshelve erewrite nth'_cons. 1: apply leq_S_n, H''. apply IHl1. - by apply path_nat_S. + by apply path_nat_succ. Defined. (** The [nth'] element of a [repeat] is the repeated value. *) @@ -485,7 +485,7 @@ Proof. f_ap. - exact (H 0%nat _). - snrapply IHl. - 1: by apply path_nat_S. + 1: by apply path_nat_succ. intros n Hn. snrefine ((nth'_cons l n a Hn _)^ @ _). 1: apply leq_S_n', Hn. @@ -974,7 +974,7 @@ Proof. split. * by apply Hf. * apply IHl1; trivial. - apply path_nat_S. + apply path_nat_succ. exact p. Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index aca18f1aabc..0e4e3d28489 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -168,14 +168,14 @@ Definition nat_pred_succ@{} n : nat_pred (nat_succ n) = n := idpath. (** Injectivity of successor *) -Definition path_nat_S@{} n m (H : S n = S m) : n = m := ap nat_pred H. +Definition path_nat_succ@{} n m (H : S n = S m) : n = m := ap nat_pred H. (** Inequality of sucessors is implied with inequality of the arguments. *) Definition neq_nat_succ@{} n m : n <> m -> S n <> S m. Proof. intros np q. apply np. - exact (path_nat_S _ _ q). + exact (path_nat_succ _ _ q). Defined. (** TODO: rename to [neq_O_S] *) @@ -205,7 +205,7 @@ Proof. - exact (inr (fun p => not_eq_O_S n p^)). - destruct (IHn m) as [p|q]. + exact (inl (ap S p)). - + exact (inr (fun p => q (path_nat_S _ _ p))). + + exact (inr (fun p => q (path_nat_succ _ _ p))). Defined. (** [nat] is therefore a hset. *) @@ -402,7 +402,7 @@ Proof. destruct r. contradiction (not_leq_Sn_n _ p). + intros m' q r. - pose (r' := path_nat_S _ _ r). + pose (r' := path_nat_succ _ _ r). destruct r'. assert (t : idpath = r) by apply path_ishprop. destruct t. @@ -596,7 +596,7 @@ Defined. Global Instance isinj_S : IsInjective S. Proof. intros x y p. - by apply path_nat_S. + by apply path_nat_succ. Defined. Global Instance isinj_nat_add_l@{} k : IsInjective (nat_add k). From 69db57d5ae5c036214516ebb0fa91f335701aa35 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:11:34 +0200 Subject: [PATCH 065/282] Nat: rename not_eq_O_S -> neq_nat_zero_succ Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 9 ++++----- 2 files changed, 5 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 604b1e323be..37e2c00404c 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -199,7 +199,7 @@ Proof. apply leq_0_n. - intros m eq. destruct m. + simpl in eq. apply symmetric_paths in eq. - contradiction (not_eq_O_S n eq). + contradiction (neq_nat_zero_succ n eq). + simpl in eq. apply leq_S_n', IHn, eq. Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 0e4e3d28489..f66cfb3cf80 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -178,9 +178,8 @@ Proof. exact (path_nat_succ _ _ q). Defined. -(** TODO: rename to [neq_O_S] *) (** Zero is not the successor of a number *) -Definition not_eq_O_S@{} : forall n, 0 <> S n. +Definition neq_nat_zero_succ@{} n : 0 <> S n. Proof. discriminate. Defined. @@ -189,7 +188,7 @@ Defined. Theorem not_eq_n_Sn@{} n : n <> S n. Proof. simple_induction' n. - - apply not_eq_O_S. + - apply neq_nat_zero_succ. - by apply neq_nat_succ. Defined. @@ -201,8 +200,8 @@ Proof. intros n; simple_induction n n IHn; intros m; destruct m. - exact (inl idpath). - - exact (inr (not_eq_O_S m)). - - exact (inr (fun p => not_eq_O_S n p^)). + - exact (inr (neq_nat_zero_succ m)). + - exact (inr (fun p => neq_nat_zero_succ n p^)). - destruct (IHn m) as [p|q]. + exact (inl (ap S p)). + exact (inr (fun p => q (path_nat_succ _ _ p))). From 15e6d2ef7b860380c90b33ee6de0ad02433483f5 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:14:40 +0200 Subject: [PATCH 066/282] Nat: rename not_eq_n_Sn -> neq_nat_succ' Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index f66cfb3cf80..8db03ce8463 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -167,7 +167,7 @@ Defined. Definition nat_pred_succ@{} n : nat_pred (nat_succ n) = n := idpath. -(** Injectivity of successor *) +(** Injectivity of successor. *) Definition path_nat_succ@{} n m (H : S n = S m) : n = m := ap nat_pred H. (** Inequality of sucessors is implied with inequality of the arguments. *) @@ -178,14 +178,14 @@ Proof. exact (path_nat_succ _ _ q). Defined. -(** Zero is not the successor of a number *) +(** Zero is not the successor of a number. *) Definition neq_nat_zero_succ@{} n : 0 <> S n. Proof. discriminate. Defined. -(** TODO: rename to [neq_n_Sn] *) -Theorem not_eq_n_Sn@{} n : n <> S n. +(** A natural number cannot be equal to its own successor. *) +Theorem neq_nat_succ'@{} n : n <> S n. Proof. simple_induction' n. - apply neq_nat_zero_succ. From 93ab0e8a1895ef24ab9e0be7bf51d0962637f9e8 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:26:27 +0200 Subject: [PATCH 067/282] Nat: move comparison predicates to one place Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 94 +++++++++++++++++++++----------------- 1 file changed, 53 insertions(+), 41 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 8db03ce8463..7c0693ddb03 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -137,6 +137,59 @@ Fixpoint factorial (n : nat) : nat | S n => S n * factorial n end. +(** ** Comparison Predicates *) + +(** *** Less Than or Equal To *) + +Inductive leq (n : nat) : nat -> Type0 := +| leq_n : leq n n +| leq_S : forall m, leq n m -> leq n (S m). + +Scheme leq_ind := Induction for leq Sort Type. +Scheme leq_rect := Induction for leq Sort Type. +Scheme leq_rec := Minimality for leq Sort Type. + +Notation "n <= m" := (leq n m) : nat_scope. + +Existing Class leq. +Global Existing Instances leq_n leq_S. + +Notation leq_refl := leq_n (only parsing). + +(** *** Less Than *) + +(** We define the less-than relation [lt] in terms of [leq] *) +Definition lt n m : Type0 := leq (S n) m. + +(** We declare it as an existing class so typeclass search is performed on its goals. *) +Existing Class lt. +#[export] Hint Unfold lt : typeclass_instances. +Infix "<" := lt : nat_scope. +Global Instance lt_is_leq n m : leq n.+1 m -> lt n m | 100 := idmap. + +(*** Greater Than or Equal To *) + +Definition gt n m := lt m n. +Existing Class gt. +#[export] Hint Unfold gt : typeclass_instances. +Infix ">" := gt : nat_scope. +Global Instance gt_is_leq n m : leq m.+1 n -> gt n m | 100 := idmap. + +(** *** Greater Than *) + +Definition ge n m := leq m n. +Existing Class ge. +#[export] Hint Unfold ge : typeclass_instances. +Infix ">=" := ge : nat_scope. +Global Instance ge_is_leq n m : leq m n -> ge n m | 100 := idmap. + +(** *** Combined Comparison Predicates *) + +Notation "x <= y <= z" := (x <= y /\ y <= z) : nat_scope. +Notation "x <= y < z" := (x <= y /\ y < z) : nat_scope. +Notation "x < y < z" := (x < y /\ y < z) : nat_scope. +Notation "x < y <= z" := (x < y /\ y <= z) : nat_scope. + (** ** Properties of [nat_iter]. *) Lemma nat_iter_succ_r n {A} (f : A -> A) (x : A) @@ -327,20 +380,6 @@ Defined. (** ** Inequality of natural numbers *) -Inductive leq (n : nat) : nat -> Type0 := -| leq_n : leq n n -| leq_S : forall m, leq n m -> leq n (S m). - -Scheme leq_ind := Induction for leq Sort Type. -Scheme leq_rect := Induction for leq Sort Type. -Scheme leq_rec := Minimality for leq Sort Type. - -Notation "n <= m" := (leq n m) : nat_scope. - -Existing Class leq. -Global Existing Instances leq_n leq_S. - -Notation leq_refl := leq_n (only parsing). Global Instance reflexive_leq : Reflexive leq := leq_n. Lemma leq_trans {x y z} : x <= y -> y <= z -> x <= z. @@ -464,16 +503,6 @@ Proof. apply leq_S, leq_add. Defined. -(** We define the less-than relation [lt] in terms of [leq] *) -Definition lt n m : Type0 := leq (S n) m. - -(** We declare it as an existing class so typeclass search is performed on its goals. *) -Existing Class lt. -#[export] Hint Unfold lt : typeclass_instances. -Infix "<" := lt : nat_scope. -(** We add a typeclass instance for unfolding the definition so lemmas about [leq] can be used. *) -Global Instance lt_is_leq n m : leq n.+1 m -> lt n m | 100 := idmap. - (** We should also give them their various typeclass instances *) Global Instance transitive_lt : Transitive lt. Proof. @@ -484,31 +513,14 @@ Defined. Global Instance decidable_lt n m : Decidable (lt n m) := _. -Definition ge n m := leq m n. -Existing Class ge. -#[export] Hint Unfold ge : typeclass_instances. -Infix ">=" := ge : nat_scope. -Global Instance ge_is_leq n m : leq m n -> ge n m | 100 := idmap. - Global Instance reflexive_ge : Reflexive ge := leq_n. Global Instance transitive_ge : Transitive ge := fun x y z p q => leq_trans q p. Global Instance decidable_ge n m : Decidable (ge n m) := _. -Definition gt n m := lt m n. -Existing Class gt. -#[export] Hint Unfold gt : typeclass_instances. -Infix ">" := gt : nat_scope. -Global Instance gt_is_leq n m : leq m.+1 n -> gt n m | 100 := idmap. - Global Instance transitive_gt : Transitive gt := fun x y z p q => transitive_lt _ _ _ q p. Global Instance decidable_gt n m : Decidable (gt n m) := _. -Notation "x <= y <= z" := (x <= y /\ y <= z) : nat_scope. -Notation "x <= y < z" := (x <= y /\ y < z) : nat_scope. -Notation "x < y < z" := (x < y /\ y < z) : nat_scope. -Notation "x < y <= z" := (x < y /\ y <= z) : nat_scope. - (** Principle of double induction *) Theorem nat_double_ind (R : nat -> nat -> Type) From 9df68ceaa0250defb93e5da9bace5ee21224cce9 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:37:39 +0200 Subject: [PATCH 068/282] Nat: move basic properties of subtraction to Nat.Core Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 34 ---------------------------- theories/Spaces/Nat/Core.v | 38 ++++++++++++++++++++++++++++++++ 2 files changed, 38 insertions(+), 34 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 37e2c00404c..48e3c6b8d64 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -82,17 +82,6 @@ Proof. intros l j. apply (@leq_trans (n.+1) m k); trivial. Defined. -Proposition sub_n_n (n : nat) : n - n = 0. -Proof. - simple_induction n n IHn. - - reflexivity. - - simpl; exact IHn. -Defined. - -Proposition sub_n_0 (n : nat) : n - 0 = n. -Proof. - destruct n; reflexivity. -Defined. Proposition add_n_sub_n_eq (m n : nat) : m + n - n = m. Proof. @@ -155,18 +144,6 @@ Proof. - intro. contradiction (not_leq_Sn_0 n). Defined. -Proposition subsubadd (n m k : nat) : n - (m + k) = n - m - k. -Proof. - revert m k; simple_induction n n IHn. - - reflexivity. - - intro m; destruct m; intro k. - + change (0 + k) with k; reflexivity. - + change (m.+1 + k) with (m + k).+1; apply IHn. -Defined. - -Definition subsubadd' (n m k : nat) : n - m - k = n - (m + k) - := (subsubadd n m k)^. - Definition nleqSm_dichot {n m : nat} : (n <= m.+1) -> (n <= m) + (n = m.+1). Proof. @@ -180,17 +157,6 @@ Proof. * right. apply ap; exact b. Defined. -Proposition sub_leq_0 (n m : nat) : n <= m -> n - m = 0. -Proof. - intro l; induction l. - - exact (sub_n_n n). - - change (m.+1) with (1 + m). destruct n. - + reflexivity. - + destruct (nat_add_comm m 1). - destruct (symmetric_paths _ _ (subsubadd n.+1 m 1)). - destruct (symmetric_paths _ _ IHl). - reflexivity. -Defined. Proposition sub_leq_0_converse (n m : nat) : n - m = 0 -> n <= m. Proof. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 7c0693ddb03..04fde5355fd 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -378,6 +378,44 @@ Proof. exact IHn. Defined. +(** ** Properties of Subtraction *) + +Proposition sub_n_0 (n : nat) : n - 0 = n. +Proof. + destruct n; reflexivity. +Defined. + +Proposition sub_n_n (n : nat) : n - n = 0. +Proof. + simple_induction n n IHn. + - reflexivity. + - simpl; exact IHn. +Defined. + +Proposition subsubadd (n m k : nat) : n - (m + k) = n - m - k. +Proof. + revert m k; simple_induction n n IHn. + - reflexivity. + - intro m; destruct m; intro k. + + change (0 + k) with k; reflexivity. + + change (m.+1 + k) with (m + k).+1; apply IHn. +Defined. + +Definition subsubadd' (n m k : nat) : n - m - k = n - (m + k) + := (subsubadd n m k)^. + +Proposition sub_leq_0 (n m : nat) : n <= m -> n - m = 0. +Proof. + intro l; induction l. + - exact (sub_n_n n). + - change (m.+1) with (1 + m). destruct n. + + reflexivity. + + destruct (nat_add_comm m 1). + destruct (symmetric_paths _ _ (subsubadd n.+1 m 1)). + destruct (symmetric_paths _ _ IHl). + reflexivity. +Defined. + (** ** Inequality of natural numbers *) Global Instance reflexive_leq : Reflexive leq := leq_n. From f50a5489dde6ff540a1cb952580b97c2a6fe6d4e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:44:19 +0200 Subject: [PATCH 069/282] Nat: add comments describing nat_sub lemmas and include todos Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 8 ++++---- theories/Spaces/Nat/Core.v | 21 +++++++++++++++------ 2 files changed, 19 insertions(+), 10 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 48e3c6b8d64..497bbe4dc37 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -207,7 +207,7 @@ Proof. - destruct (symmetric_paths _ _ (natminuspluseq _ _ l)); constructor. - apply n_lt_m_n_leq_m in g. - now destruct (symmetric_paths _ _ (sub_leq_0 n m _)). + now destruct (symmetric_paths _ _ (sub_leq_0 _)). Defined. Proposition natminuspluseq' (n m : nat) @@ -389,7 +389,7 @@ Proof. destruct (symmetric_paths _ _ (add_n_sub_n_eq n m)). apply leq_refl; done. - apply n_lt_m_n_leq_m in gt. - destruct (symmetric_paths _ _ (sub_leq_0 n m _)). + destruct (symmetric_paths _ _ (sub_leq_0 _)). assumption. Defined. @@ -652,9 +652,9 @@ Proof. simple_induction k k IHk. - destruct (symmetric_paths _ _ (sub_n_0 n)); constructor. - destruct (@leq_dichot n k) as [l | g]. - + destruct (symmetric_paths _ _ (sub_leq_0 _ _ _)) in IHk. + + destruct (symmetric_paths _ _ (sub_leq_0 _)) in IHk. apply leq_S in l. - destruct (symmetric_paths _ _ (sub_leq_0 _ _ _)). exact IHk. + destruct (symmetric_paths _ _ (sub_leq_0 _)). exact IHk. + change k.+1 with (1 + k). destruct (nat_add_comm k 1). destruct (symmetric_paths _ _ (subsubadd n k 1)). destruct (symmetric_paths _ _ (@predeqminus1 (n - k))). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 04fde5355fd..ca90fb3e217 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -380,19 +380,25 @@ Defined. (** ** Properties of Subtraction *) -Proposition sub_n_0 (n : nat) : n - 0 = n. +(** TODO: rename [nat_sub_zero_r] *) +(** Subtracting [0] from a number is the number itself. *) +Definition sub_n_0 (n : nat) : n - 0 = n. Proof. - destruct n; reflexivity. + destruct n; reflexivity. Defined. -Proposition sub_n_n (n : nat) : n - n = 0. +(** TODO: rename [nat_sub_cancel_r] *) +(** Subtracting a number from itself is [0]. *) +Definition sub_n_n (n : nat) : n - n = 0. Proof. simple_induction n n IHn. - reflexivity. - simpl; exact IHn. Defined. -Proposition subsubadd (n m k : nat) : n - (m + k) = n - m - k. +(** TODO: rename [nat_sub_add] *) +(** Subtracting an addition is the same as subtracting the two numbers separately. *) +Definition subsubadd n m k : n - (m + k) = n - m - k. Proof. revert m k; simple_induction n n IHn. - reflexivity. @@ -401,10 +407,13 @@ Proof. + change (m.+1 + k) with (m + k).+1; apply IHn. Defined. -Definition subsubadd' (n m k : nat) : n - m - k = n - (m + k) +(** TODO: remove *) +Definition subsubadd' n m k : n - m - k = n - (m + k) := (subsubadd n m k)^. -Proposition sub_leq_0 (n m : nat) : n <= m -> n - m = 0. +(** TODO: rename [nat_sub_leq] *) +(** Subtracting a larger number from a smaller number is [0]. *) +Definition sub_leq_0 {n m} : n <= m -> n - m = 0. Proof. intro l; induction l. - exact (sub_n_n n). From b262118305a731817a648a0fabb33714a84987a7 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:46:10 +0200 Subject: [PATCH 070/282] Nat: add nat_sub_zero_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 3 +++ 1 file changed, 3 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ca90fb3e217..8dcab664761 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -380,6 +380,9 @@ Defined. (** ** Properties of Subtraction *) +(** Subtracting a number from [0] is [0]. *) +Definition nat_sub_zero_l@{} n : 0 - n = 0 := idpath. + (** TODO: rename [nat_sub_zero_r] *) (** Subtracting [0] from a number is the number itself. *) Definition sub_n_0 (n : nat) : n - 0 = n. From 72d1f66d6d21e3a028591103b38523b9be326907 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:46:51 +0200 Subject: [PATCH 071/282] Nat: rename nat_sub_n_0 -> nat_sub_zero_r Signed-off-by: Ali Caglayan --- theories/Spaces/List/Theory.v | 2 +- theories/Spaces/Nat/Arithmetic.v | 12 ++++++------ theories/Spaces/Nat/Core.v | 3 +-- 3 files changed, 8 insertions(+), 9 deletions(-) diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index 79db96fab8c..dc2ac79bc49 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -821,7 +821,7 @@ Proof. induction i as [|i IHi] in n, H |- *. - induction n. 1: destruct (not_leq_Sn_0 _ H). - cbn; by rewrite sub_n_0. + cbn; by rewrite nat_sub_zero_r. - induction n as [|n IHn]. 1: destruct (not_leq_Sn_0 _ H). by apply IHi, leq_S_n. diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 497bbe4dc37..47c7ed93b99 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -422,7 +422,7 @@ Proposition predeqminus1 { n : nat } : n - 1 = nat_pred n. Proof. simple_induction' n. - reflexivity. - - apply sub_n_0. + - apply nat_sub_zero_r. Defined. Proposition predn_leq_n (n : nat) : nat_pred n <= n. @@ -505,8 +505,8 @@ Proposition natsubpreservesleq { n m k : nat } : n <= m -> n - k <= m - k. Proof. simple_induction k k IHk. - - destruct (symmetric_paths _ _ (sub_n_0 n)), - (symmetric_paths _ _ (sub_n_0 m)); done. + - destruct (symmetric_paths _ _ (nat_sub_zero_r n)), + (symmetric_paths _ _ (nat_sub_zero_r m)); done. - intro l. change (k.+1) with (1 + k). destruct (nat_add_comm k 1). destruct (symmetric_paths _ _ (subsubadd n k 1)). @@ -650,7 +650,7 @@ Lemma increasing_geq_minus (n k : nat) : increasing_geq n (n - k). Proof. simple_induction k k IHk. - - destruct (symmetric_paths _ _ (sub_n_0 n)); constructor. + - destruct (symmetric_paths _ _ (nat_sub_zero_r n)); constructor. - destruct (@leq_dichot n k) as [l | g]. + destruct (symmetric_paths _ _ (sub_leq_0 _)) in IHk. apply leq_S in l. @@ -676,14 +676,14 @@ Defined. Lemma ineq_sub (n m : nat) : n <= m -> m - (m - n) = n. Proof. revert m; simple_induction n n IHn. - - intros. destruct (symmetric_paths _ _ (sub_n_0 m)), + - intros. destruct (symmetric_paths _ _ (nat_sub_zero_r m)), (symmetric_paths _ _ (sub_n_n m)); reflexivity. - intros m ineq. change (m - n.+1) with (m - (1 + n)). (destruct (nat_add_comm n 1)). destruct (symmetric_paths _ _ (subsubadd m n 1)). destruct (S_predn (m - n) 0 (lt_sub_gt_0 _ _ ineq)); simpl; - destruct (symmetric_paths _ _ (sub_n_0 (nat_pred (m - n)))). + destruct (symmetric_paths _ _ (nat_sub_zero_r (nat_pred (m - n)))). assert (0 < m - n) as dp by exact (lt_sub_gt_0 _ _ ineq). assert (nat_pred (m - n) < m) as sh by ( unfold "<"; diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 8dcab664761..6f4faf5ddd4 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -383,9 +383,8 @@ Defined. (** Subtracting a number from [0] is [0]. *) Definition nat_sub_zero_l@{} n : 0 - n = 0 := idpath. -(** TODO: rename [nat_sub_zero_r] *) (** Subtracting [0] from a number is the number itself. *) -Definition sub_n_0 (n : nat) : n - 0 = n. +Definition nat_sub_zero_r (n : nat) : n - 0 = n. Proof. destruct n; reflexivity. Defined. From 27edd1342d6a6bf73e35be20f1036bd477531a2a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:49:36 +0200 Subject: [PATCH 072/282] Nat: rename sub_n_n -> nat_sub_cancel Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 5 ++--- 2 files changed, 3 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 47c7ed93b99..660b1a150b8 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -677,7 +677,7 @@ Lemma ineq_sub (n m : nat) : n <= m -> m - (m - n) = n. Proof. revert m; simple_induction n n IHn. - intros. destruct (symmetric_paths _ _ (nat_sub_zero_r m)), - (symmetric_paths _ _ (sub_n_n m)); + (symmetric_paths _ _ (nat_sub_cancel m)); reflexivity. - intros m ineq. change (m - n.+1) with (m - (1 + n)). (destruct (nat_add_comm n 1)). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 6f4faf5ddd4..a624b27e01c 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -389,9 +389,8 @@ Proof. destruct n; reflexivity. Defined. -(** TODO: rename [nat_sub_cancel_r] *) (** Subtracting a number from itself is [0]. *) -Definition sub_n_n (n : nat) : n - n = 0. +Definition nat_sub_cancel (n : nat) : n - n = 0. Proof. simple_induction n n IHn. - reflexivity. @@ -418,7 +417,7 @@ Definition subsubadd' n m k : n - m - k = n - (m + k) Definition sub_leq_0 {n m} : n <= m -> n - m = 0. Proof. intro l; induction l. - - exact (sub_n_n n). + - exact (nat_sub_cancel n). - change (m.+1) with (1 + m). destruct n. + reflexivity. + destruct (nat_add_comm m 1). From d1efca84b8565e13cdaff7fe24618bcc205ecc3f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:51:34 +0200 Subject: [PATCH 073/282] Nat: rename subsubadd -> nat_sub_add Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 8 ++++---- theories/Spaces/Nat/Core.v | 7 +++---- 2 files changed, 7 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 660b1a150b8..7f0443255a4 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -509,8 +509,8 @@ Proof. (symmetric_paths _ _ (nat_sub_zero_r m)); done. - intro l. change (k.+1) with (1 + k). destruct (nat_add_comm k 1). - destruct (symmetric_paths _ _ (subsubadd n k 1)). - destruct (symmetric_paths _ _ (subsubadd m k 1)). + destruct (symmetric_paths _ _ (nat_sub_add n k 1)). + destruct (symmetric_paths _ _ (nat_sub_add m k 1)). destruct (symmetric_paths _ _ (@predeqminus1 (n -k))). destruct (symmetric_paths _ _ (@predeqminus1 (m -k))). apply leq_n_pred, IHk. exact l. @@ -656,7 +656,7 @@ Proof. apply leq_S in l. destruct (symmetric_paths _ _ (sub_leq_0 _)). exact IHk. + change k.+1 with (1 + k). destruct (nat_add_comm k 1). - destruct (symmetric_paths _ _ (subsubadd n k 1)). + destruct (symmetric_paths _ _ (nat_sub_add n k 1)). destruct (symmetric_paths _ _ (@predeqminus1 (n - k))). apply increasing_geq_S. unfold ">", "<" in *. @@ -681,7 +681,7 @@ Proof. reflexivity. - intros m ineq. change (m - n.+1) with (m - (1 + n)). (destruct (nat_add_comm n 1)). - destruct (symmetric_paths _ _ (subsubadd m n 1)). + destruct (symmetric_paths _ _ (nat_sub_add m n 1)). destruct (S_predn (m - n) 0 (lt_sub_gt_0 _ _ ineq)); simpl; destruct (symmetric_paths _ _ (nat_sub_zero_r (nat_pred (m - n)))). assert (0 < m - n) as dp by exact (lt_sub_gt_0 _ _ ineq). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index a624b27e01c..11e52ad7c96 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -397,9 +397,8 @@ Proof. - simpl; exact IHn. Defined. -(** TODO: rename [nat_sub_add] *) (** Subtracting an addition is the same as subtracting the two numbers separately. *) -Definition subsubadd n m k : n - (m + k) = n - m - k. +Definition nat_sub_add n m k : n - (m + k) = n - m - k. Proof. revert m k; simple_induction n n IHn. - reflexivity. @@ -410,7 +409,7 @@ Defined. (** TODO: remove *) Definition subsubadd' n m k : n - m - k = n - (m + k) - := (subsubadd n m k)^. + := (nat_sub_add n m k)^. (** TODO: rename [nat_sub_leq] *) (** Subtracting a larger number from a smaller number is [0]. *) @@ -421,7 +420,7 @@ Proof. - change (m.+1) with (1 + m). destruct n. + reflexivity. + destruct (nat_add_comm m 1). - destruct (symmetric_paths _ _ (subsubadd n.+1 m 1)). + destruct (symmetric_paths _ _ (nat_sub_add n.+1 m 1)). destruct (symmetric_paths _ _ IHl). reflexivity. Defined. From d7952fae8b1fbb99f7c4ba5e6e29fe7b108e2562 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:53:29 +0200 Subject: [PATCH 074/282] Nat: remove subsubadd' Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 ---- 1 file changed, 4 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 11e52ad7c96..14c6579745a 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -407,10 +407,6 @@ Proof. + change (m.+1 + k) with (m + k).+1; apply IHn. Defined. -(** TODO: remove *) -Definition subsubadd' n m k : n - m - k = n - (m + k) - := (nat_sub_add n m k)^. - (** TODO: rename [nat_sub_leq] *) (** Subtracting a larger number from a smaller number is [0]. *) Definition sub_leq_0 {n m} : n <= m -> n - m = 0. From 5251125908b816ad8783392b0bc2b6c66eca70ce Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:55:45 +0200 Subject: [PATCH 075/282] Nat: improve some sub proofs Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 18 +++++++++--------- 1 file changed, 9 insertions(+), 9 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 14c6579745a..480412a2a46 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -384,28 +384,28 @@ Defined. Definition nat_sub_zero_l@{} n : 0 - n = 0 := idpath. (** Subtracting [0] from a number is the number itself. *) -Definition nat_sub_zero_r (n : nat) : n - 0 = n. +Definition nat_sub_zero_r@{} (n : nat) : n - 0 = n. Proof. destruct n; reflexivity. Defined. (** Subtracting a number from itself is [0]. *) -Definition nat_sub_cancel (n : nat) : n - n = 0. +Definition nat_sub_cancel@{} (n : nat) : n - n = 0. Proof. simple_induction n n IHn. - reflexivity. - - simpl; exact IHn. + - exact IHn. Defined. (** Subtracting an addition is the same as subtracting the two numbers separately. *) -Definition nat_sub_add n m k : n - (m + k) = n - m - k. +Definition nat_sub_add@{} n m k : n - (m + k) = n - m - k. Proof. - revert m k; simple_induction n n IHn. + induction n as [|n IHn] in m, k |- *. - reflexivity. - - intro m; destruct m; intro k. - + change (0 + k) with k; reflexivity. - + change (m.+1 + k) with (m + k).+1; apply IHn. -Defined. + - destruct m. + + reflexivity. + + nrapply IHn. +Defined. (** TODO: rename [nat_sub_leq] *) (** Subtracting a larger number from a smaller number is [0]. *) From a09f36499ffeb04c2a37e5e7b24febe37017f936 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 17:56:29 +0200 Subject: [PATCH 076/282] Nat: rename sub_leq_0 -> nat_sub_leq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 10 +++++----- theories/Spaces/Nat/Core.v | 3 +-- 2 files changed, 6 insertions(+), 7 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 7f0443255a4..e23fc5d63bc 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -173,7 +173,7 @@ Proposition sub_gt_0_lt (n m : nat) : n - m > 0 -> m < n. Proof. intro ineq. destruct (@leq_dichot n m) as [n_leq_m |]; [ | assumption]. - apply sub_leq_0 in n_leq_m. + apply nat_sub_leq in n_leq_m. contradiction (not_lt_n_n 0). now rewrite n_leq_m in ineq. Defined. @@ -207,7 +207,7 @@ Proof. - destruct (symmetric_paths _ _ (natminuspluseq _ _ l)); constructor. - apply n_lt_m_n_leq_m in g. - now destruct (symmetric_paths _ _ (sub_leq_0 _)). + now destruct (symmetric_paths _ _ (nat_sub_leq _)). Defined. Proposition natminuspluseq' (n m : nat) @@ -389,7 +389,7 @@ Proof. destruct (symmetric_paths _ _ (add_n_sub_n_eq n m)). apply leq_refl; done. - apply n_lt_m_n_leq_m in gt. - destruct (symmetric_paths _ _ (sub_leq_0 _)). + destruct (symmetric_paths _ _ (nat_sub_leq _)). assumption. Defined. @@ -652,9 +652,9 @@ Proof. simple_induction k k IHk. - destruct (symmetric_paths _ _ (nat_sub_zero_r n)); constructor. - destruct (@leq_dichot n k) as [l | g]. - + destruct (symmetric_paths _ _ (sub_leq_0 _)) in IHk. + + destruct (symmetric_paths _ _ (nat_sub_leq _)) in IHk. apply leq_S in l. - destruct (symmetric_paths _ _ (sub_leq_0 _)). exact IHk. + destruct (symmetric_paths _ _ (nat_sub_leq _)). exact IHk. + change k.+1 with (1 + k). destruct (nat_add_comm k 1). destruct (symmetric_paths _ _ (nat_sub_add n k 1)). destruct (symmetric_paths _ _ (@predeqminus1 (n - k))). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 480412a2a46..ef486ccf999 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -407,9 +407,8 @@ Proof. + nrapply IHn. Defined. -(** TODO: rename [nat_sub_leq] *) (** Subtracting a larger number from a smaller number is [0]. *) -Definition sub_leq_0 {n m} : n <= m -> n - m = 0. +Definition nat_sub_leq {n m} : n <= m -> n - m = 0. Proof. intro l; induction l. - exact (nat_sub_cancel n). From 5afcd4cb0108c331fef0adb747c603443e2d4da3 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 18:07:22 +0200 Subject: [PATCH 077/282] Nat: add nat_sub_comm_r and improve proof of nat_sub_leq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 18 ++++++++++++------ 1 file changed, 12 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ef486ccf999..de7c2b43006 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -407,17 +407,23 @@ Proof. + nrapply IHn. Defined. +(** The order in which two numbers are subtracted does not matter. *) +Definition nat_sub_comm_r@{} n m k : n - m - k = n - k - m. +Proof. + lhs_V nrapply nat_sub_add. + rewrite nat_add_comm. + nrapply nat_sub_add. +Defined. + (** Subtracting a larger number from a smaller number is [0]. *) Definition nat_sub_leq {n m} : n <= m -> n - m = 0. Proof. intro l; induction l. - exact (nat_sub_cancel n). - - change (m.+1) with (1 + m). destruct n. - + reflexivity. - + destruct (nat_add_comm m 1). - destruct (symmetric_paths _ _ (nat_sub_add n.+1 m 1)). - destruct (symmetric_paths _ _ IHl). - reflexivity. + - change (m.+1) with (1 + m). + lhs nrapply nat_sub_add. + lhs nrapply nat_sub_comm_r. + by destruct IHl^. Defined. (** ** Inequality of natural numbers *) From adb8cf8e651143caddf1044b9e81612e605e1f86 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 18:45:48 +0200 Subject: [PATCH 078/282] Nat: organise properties of leq and other comparison predicates Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 166 +++++++++++++++++++++---------------- 1 file changed, 95 insertions(+), 71 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index de7c2b43006..afa5638e3de 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -378,83 +378,54 @@ Proof. exact IHn. Defined. -(** ** Properties of Subtraction *) - -(** Subtracting a number from [0] is [0]. *) -Definition nat_sub_zero_l@{} n : 0 - n = 0 := idpath. - -(** Subtracting [0] from a number is the number itself. *) -Definition nat_sub_zero_r@{} (n : nat) : n - 0 = n. -Proof. - destruct n; reflexivity. -Defined. - -(** Subtracting a number from itself is [0]. *) -Definition nat_sub_cancel@{} (n : nat) : n - n = 0. -Proof. - simple_induction n n IHn. - - reflexivity. - - exact IHn. -Defined. +(** ** Basic Properties of Comparison Predicates *) -(** Subtracting an addition is the same as subtracting the two numbers separately. *) -Definition nat_sub_add@{} n m k : n - (m + k) = n - m - k. -Proof. - induction n as [|n IHn] in m, k |- *. - - reflexivity. - - destruct m. - + reflexivity. - + nrapply IHn. -Defined. +(** *** Basic Properties of [leq] *) -(** The order in which two numbers are subtracted does not matter. *) -Definition nat_sub_comm_r@{} n m k : n - m - k = n - k - m. -Proof. - lhs_V nrapply nat_sub_add. - rewrite nat_add_comm. - nrapply nat_sub_add. -Defined. +(** [leq] is reflexive by definition. *) +Global Instance reflexive_leq : Reflexive leq := leq_n. -(** Subtracting a larger number from a smaller number is [0]. *) -Definition nat_sub_leq {n m} : n <= m -> n - m = 0. +(** Being less than or equal to is a transitive relation. *) +Definition leq_trans {x y z} : x <= y -> y <= z -> x <= z. Proof. - intro l; induction l. - - exact (nat_sub_cancel n). - - change (m.+1) with (1 + m). - lhs nrapply nat_sub_add. - lhs nrapply nat_sub_comm_r. - by destruct IHl^. + intros H1 H2; induction H2; exact _. Defined. -(** ** Inequality of natural numbers *) - -Global Instance reflexive_leq : Reflexive leq := leq_n. +(** [leq] is transtiive. *) +Global Instance transitive_leq : Transitive leq := @leq_trans. -Lemma leq_trans {x y z} : x <= y -> y <= z -> x <= z. +(** TODO: rename to [leq_zero] *) +(** [0] is less than or equal to any natural number. *) +Definition leq_0_n n : 0 <= n. Proof. - induction 2; exact _. + simple_induction' n; exact _. Defined. +Global Existing Instance leq_0_n | 10. -Global Instance transitive_leq : Transitive leq := @leq_trans. - -Lemma leq_n_pred n m : leq n m -> leq (nat_pred n) (nat_pred m). +(** TODO: rename to [leq_pred] *) +(** A predecessor is less than or equal to a predecessor if the original number is less than or equal. *) +Definition leq_n_pred n m : leq n m -> leq (nat_pred n) (nat_pred m). Proof. - induction 1; try exact _. - destruct m; simpl; exact _. + intros H; induction H. + 1: exact _. + destruct m; exact _. Defined. +(** TODO: is this needed? *) Lemma leq_S_n : forall n m, n.+1 <= m.+1 -> n <= m. Proof. intros n m. apply leq_n_pred. Defined. +(** TODO: possibly rename to [leq_succ] depending on above. *) Lemma leq_S_n' n m : n <= m -> n.+1 <= m.+1. Proof. induction 1; exact _. Defined. Global Existing Instance leq_S_n' | 100. +(** TODO: use lemmas about negating predicate *) Lemma not_leq_Sn_n n : ~ (n.+1 <= n). Proof. simple_induction n n IHn. @@ -464,6 +435,13 @@ Proof. by apply IHn, leq_S_n. Defined. +Lemma not_leq_Sn_0 n : ~ (n.+1 <= 0). +Proof. + intros p. + apply (fun x => leq_trans x (leq_0_n n)) in p. + contradiction (not_leq_Sn_n _ p). +Defined. + (** A general form for injectivity of this constructor *) Definition leq_n_inj_gen n k (p : n <= k) (r : n = k) : p = r # leq_n n. Proof. @@ -512,18 +490,6 @@ Proof. rapply leq_S_inj. Defined. -Global Instance leq_0_n n : 0 <= n | 10. -Proof. - simple_induction' n; exact _. -Defined. - -Lemma not_leq_Sn_0 n : ~ (n.+1 <= 0). -Proof. - intros p. - apply (fun x => leq_trans x (leq_0_n n)) in p. - contradiction (not_leq_Sn_n _ p). -Defined. - Definition equiv_leq_S_n n m : n.+1 <= m.+1 <~> n <= m. Proof. srapply equiv_iff_hprop. @@ -544,14 +510,8 @@ Proof. apply equiv_leq_S_n. Defined. -Fixpoint leq_add n m : n <= (m + n). -Proof. - destruct m. - 1: apply leq_n. - apply leq_S, leq_add. -Defined. +(** *** Basic Properties of [lt] *) -(** We should also give them their various typeclass instances *) Global Instance transitive_lt : Transitive lt. Proof. hnf; unfold lt in *. @@ -559,16 +519,80 @@ Proof. rapply leq_trans. Defined. +Global Instance ishprop_lt n m : IsHProp (n < m) := _. Global Instance decidable_lt n m : Decidable (lt n m) := _. +(** *** Basic Properties of [ge] *) + Global Instance reflexive_ge : Reflexive ge := leq_n. Global Instance transitive_ge : Transitive ge := fun x y z p q => leq_trans q p. +Global Instance ishprop_ge n m : IsHProp (ge n m) := _. Global Instance decidable_ge n m : Decidable (ge n m) := _. +(** *** Basic Properties of [gt] *) + Global Instance transitive_gt : Transitive gt := fun x y z p q => transitive_lt _ _ _ q p. +Global Instance ishprop_gt n m : IsHProp (gt n m) := _. Global Instance decidable_gt n m : Decidable (gt n m) := _. +(** ** Properties of Subtraction *) + +(** Subtracting a number from [0] is [0]. *) +Definition nat_sub_zero_l@{} n : 0 - n = 0 := idpath. + +(** Subtracting [0] from a number is the number itself. *) +Definition nat_sub_zero_r@{} (n : nat) : n - 0 = n. +Proof. + destruct n; reflexivity. +Defined. + +(** Subtracting a number from itself is [0]. *) +Definition nat_sub_cancel@{} (n : nat) : n - n = 0. +Proof. + simple_induction n n IHn. + - reflexivity. + - exact IHn. +Defined. + +(** Subtracting an addition is the same as subtracting the two numbers separately. *) +Definition nat_sub_add@{} n m k : n - (m + k) = n - m - k. +Proof. + induction n as [|n IHn] in m, k |- *. + - reflexivity. + - destruct m. + + reflexivity. + + nrapply IHn. +Defined. + +(** The order in which two numbers are subtracted does not matter. *) +Definition nat_sub_comm_r@{} n m k : n - m - k = n - k - m. +Proof. + lhs_V nrapply nat_sub_add. + rewrite nat_add_comm. + nrapply nat_sub_add. +Defined. + +(** Subtracting a larger number from a smaller number is [0]. *) +Definition nat_sub_leq {n m} : n <= m -> n - m = 0. +Proof. + intro l; induction l. + - exact (nat_sub_cancel n). + - change (m.+1) with (1 + m). + lhs nrapply nat_sub_add. + lhs nrapply nat_sub_comm_r. + by destruct IHl^. +Defined. + +(** ** Inequality of natural numbers *) + +Fixpoint leq_add n m : n <= (m + n). +Proof. + destruct m. + 1: apply leq_n. + apply leq_S, leq_add. +Defined. + (** Principle of double induction *) Theorem nat_double_ind (R : nat -> nat -> Type) From 15789f2caf5e4f09eb0a0a87eef72bbc0890e848 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 19:25:38 +0200 Subject: [PATCH 079/282] Nat: improvements to max/min lemmas Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 120 ++++++++++++++++++++----------------- 1 file changed, 64 insertions(+), 56 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index afa5638e3de..44df65fd57c 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -147,7 +147,7 @@ Inductive leq (n : nat) : nat -> Type0 := Scheme leq_ind := Induction for leq Sort Type. Scheme leq_rect := Induction for leq Sort Type. -Scheme leq_rec := Minimality for leq Sort Type. +Scheme leq_rec := Induction for leq Sort Type. Notation "n <= m" := (leq n m) : nat_scope. @@ -584,94 +584,102 @@ Proof. by destruct IHl^. Defined. -(** ** Inequality of natural numbers *) +(** ** Properties of Maximum and Minimum *) -Fixpoint leq_add n m : n <= (m + n). -Proof. - destruct m. - 1: apply leq_n. - apply leq_S, leq_add. -Defined. +(** *** Properties of Maxima *) -(** Principle of double induction *) - -Theorem nat_double_ind (R : nat -> nat -> Type) - (H1 : forall n, R 0 n) (H2 : forall n, R (S n) 0) - (H3 : forall n m, R n m -> R (S n) (S m)) - : forall n m:nat, R n m. -Proof. - simple_induction' n; auto. - destruct m; auto. -Defined. - -(** Maximum and minimum : definitions and specifications *) - -Lemma nat_max_n_n n : nat_max n n = n. +(** TODO: rename to [nat_max_idem] *) +(** [nat_max] is idempotent. *) +Definition nat_max_n_n@{} n : nat_max n n = n. Proof. simple_induction' n; cbn. 1: reflexivity. exact (ap S IH). Defined. -Lemma nat_max_Sn_n n : nat_max (S n) n = S n. +(** [nat_max] is commutative. *) +Definition nat_max_comm@{} n m : nat_max n m = nat_max m n. +Proof. + induction n as [|n IHn] in m |- *; destruct m; cbn. + 1-3: reflexivity. + exact (ap S (IHn _)). +Defined. + +(** TODO: rename to [nat_max_succ_l] *) +Definition nat_max_Sn_n@{} n : nat_max (S n) n = S n. Proof. simple_induction' n; cbn. 1: reflexivity. exact (ap S IH). Defined. -Lemma nat_max_0_n n : nat_max 0 n = n. -Proof. - reflexivity. -Defined. +(** TODO: rename to [nat_max_zero_l] *) +(** [0] is the left identity of [nat_max]. *) +Definition nat_max_0_n@{} n : nat_max 0 n = n := idpath. +Notation nat_max_zero_l := nat_max_0_n. -Lemma nat_max_n_0 n : nat_max n 0 = n. -Proof. - induction n as [|n IHn]; reflexivity. -Defined. +(** TODO: rename to [nat_max_zero_r] *) +(** [0] is the right identity of [nat_max]. *) +Definition nat_max_n_0@{} n : nat_max n 0 = n + := nat_max_comm _ _ @ nat_max_zero_l _. +Notation nat_max_zero_r := nat_max_n_0. -Lemma nat_max_comm n m : nat_max n m = nat_max m n. +Definition nat_max_l@{} {n m} : m <= n -> nat_max n m = n. Proof. - induction m as [|m IHm] in n |- *. - - nrapply nat_max_n_0. - - destruct n. - + reflexivity. - + exact (ap S (IHm n)). -Defined. - -Theorem nat_max_l : forall n m, m <= n -> nat_max n m = n. -Proof. - intros n m; revert n; simple_induction m m IHm. - { intros n H. - nrapply nat_max_n_0. } - intros [] p. - 1: inversion p. + intros H. + induction m as [|m IHm] in n, H |- *. + 1: nrapply nat_max_n_0. + destruct n. + 1: inversion H. cbn; by apply (ap S), IHm, leq_S_n. Defined. -Theorem nat_max_r : forall n m : nat, n <= m -> nat_max n m = m. -Proof. - intros; rewrite nat_max_comm; by apply nat_max_l. -Defined. +Definition nat_max_r {n m} : n <= m -> nat_max n m = m + := fun _ => nat_max_comm _ _ @ nat_max_l _. + +(** Properties of Minima *) -Lemma nat_min_comm : forall n m, nat_min n m = nat_min m n. +Definition nat_min_comm n m : nat_min n m = nat_min m n. Proof. - simple_induction' n; destruct m; cbn. + induction n as [|n IHn] in m |- *; destruct m; cbn. 1-3: reflexivity. - exact (ap S (IH _)). + exact (ap S (IHn _)). Defined. -Theorem nat_min_l : forall n m : nat, n <= m -> nat_min n m = n. +Definition nat_min_zero_l n : nat_min 0 n = 0 := idpath. +Definition nat_min_zero_r n : nat_min n 0 = 0:= + nat_min_comm _ _ @ nat_min_zero_l _. + +Definition nat_min_l {n m} : n <= m -> nat_min n m = n. Proof. + revert n m. simple_induction n n IHn; auto. intros [] p. 1: inversion p. cbn; by apply (ap S), IHn, leq_S_n. Defined. -Theorem nat_min_r : forall n m : nat, m <= n -> nat_min n m = m. +Definition nat_min_r {n m} : m <= n -> nat_min n m = m + := fun _ => nat_min_comm _ _ @ nat_min_l _. + +(** ** Inequality of natural numbers *) + +Fixpoint leq_add n m : n <= (m + n). +Proof. + destruct m. + 1: apply leq_n. + apply leq_S, leq_add. +Defined. + +(** Principle of double induction *) + +Theorem nat_double_ind (R : nat -> nat -> Type) + (H1 : forall n, R 0 n) (H2 : forall n, R (S n) 0) + (H3 : forall n m, R n m -> R (S n) (S m)) + : forall n m:nat, R n m. Proof. - intros; rewrite nat_min_comm; by apply nat_min_l. + simple_induction' n; auto. + destruct m; auto. Defined. (** ** Arithmetic *) From a11eb6890d36e918aaf15ada08cd327151a2d056 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 19:26:00 +0200 Subject: [PATCH 080/282] Nat: remove unused "double induction" lemma This doesn't seem correct anyway. Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 11 ----------- 1 file changed, 11 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 44df65fd57c..f306fa08578 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -671,17 +671,6 @@ Proof. apply leq_S, leq_add. Defined. -(** Principle of double induction *) - -Theorem nat_double_ind (R : nat -> nat -> Type) - (H1 : forall n, R 0 n) (H2 : forall n, R (S n) 0) - (H3 : forall n m, R n m -> R (S n) (S m)) - : forall n m:nat, R n m. -Proof. - simple_induction' n; auto. - destruct m; auto. -Defined. - (** ** Arithmetic *) Global Instance isinj_S : IsInjective S. From fa1441634f1ac72fddbd72cf4a07819e6de5faa4 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 19:54:38 +0200 Subject: [PATCH 081/282] Nat: rename leq_0_n -> leq_zero Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 14 +++++++------- theories/Spaces/Nat/Core.v | 11 +++++------ 2 files changed, 12 insertions(+), 13 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index e23fc5d63bc..34f8fc712e2 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -126,7 +126,7 @@ Defined. Proposition n_leq_add_n_k (n m : nat) : n <= n + m. Proof. simple_induction n n IHn. - - apply leq_0_n. + - apply leq_zero. - simpl; apply leq_S_n', IHn. Defined. @@ -148,7 +148,7 @@ Definition nleqSm_dichot {n m : nat} : (n <= m.+1) -> (n <= m) + (n = m.+1). Proof. revert m; simple_induction n n IHn. - - intro. left. exact (leq_0_n m). + - intro. left. exact (leq_zero m). - destruct m. + intro l. apply leq_S_n, natineq0eq0 in l. right; apply ap; exact l. @@ -162,7 +162,7 @@ Proposition sub_leq_0_converse (n m : nat) : n - m = 0 -> n <= m. Proof. revert m; simple_induction n n IHn. - simpl. intros m eq. - apply leq_0_n. + apply leq_zero. - intros m eq. destruct m. + simpl in eq. apply symmetric_paths in eq. contradiction (neq_nat_zero_succ n eq). @@ -286,7 +286,7 @@ Proposition nataddreflectsleq { n m k : nat } : n + k <= m + k -> n <= m. Proof. destruct n. - - intros ?; apply leq_0_n. + - intros ?; apply leq_zero. - intro a. change (n.+1 + k) with (n + k).+1 in a. now apply (@nataddreflectslt n m k). Defined. @@ -399,7 +399,7 @@ Proof. revert m i; simple_induction n n IHn. - intros m i l. simpl in l. contradiction (not_lt_n_0 _ _). - intros m i l. destruct m. - + apply leq_0_n. + + apply leq_zero. + apply leq_S_n'. simpl in l. apply (IHn m i l). Defined. @@ -461,7 +461,7 @@ Proposition pred_lt_implies_leq (n k : nat) : nat_pred n < k -> n <= k. Proof. intro l; destruct n. - - apply leq_0_n. + - apply leq_zero. - assumption. Defined. @@ -520,7 +520,7 @@ Proposition sub_less { n k : nat } : n - k <= n. Proof. revert k. simple_induction n n IHn. - - intros; apply leq_0_n. + - intros; apply leq_zero. - destruct k. + apply leq_n. + simpl; apply (@leq_trans _ n _); diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index f306fa08578..f95e2e72b44 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -394,13 +394,12 @@ Defined. (** [leq] is transtiive. *) Global Instance transitive_leq : Transitive leq := @leq_trans. -(** TODO: rename to [leq_zero] *) (** [0] is less than or equal to any natural number. *) -Definition leq_0_n n : 0 <= n. +Definition leq_zero n : 0 <= n. Proof. simple_induction' n; exact _. Defined. -Global Existing Instance leq_0_n | 10. +Global Existing Instance leq_zero | 10. (** TODO: rename to [leq_pred] *) (** A predecessor is less than or equal to a predecessor if the original number is less than or equal. *) @@ -438,7 +437,7 @@ Defined. Lemma not_leq_Sn_0 n : ~ (n.+1 <= 0). Proof. intros p. - apply (fun x => leq_trans x (leq_0_n n)) in p. + apply (fun x => leq_trans x (leq_zero n)) in p. contradiction (not_leq_Sn_n _ p). Defined. @@ -708,13 +707,13 @@ Defined. Definition not_lt_n_n n : ~ (n < n) := not_leq_Sn_n n. -Definition leq_1_Sn {n} : 1 <= n.+1 := leq_S_n' 0 n (leq_0_n _). +Definition leq_1_Sn {n} : 1 <= n.+1 := leq_S_n' 0 n (leq_zero _). Fixpoint leq_dichot {m} {n} : (m <= n) + (m > n). Proof. simple_induction' m; simple_induction' n. - left; reflexivity. - - left; apply leq_0_n. + - left; apply leq_zero. - right; unfold lt; apply leq_1_Sn. - assert ((m <= n) + (n < m)) as X by apply leq_dichot. destruct X as [leqmn|ltnm]. From 1c74106db793bfbddf73a95de529341350062c8c Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 19:57:55 +0200 Subject: [PATCH 082/282] Nat: rename leq_n_pred -> leq_pred Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 4 ++-- theories/Spaces/Nat/Core.v | 5 ++--- 2 files changed, 4 insertions(+), 5 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 34f8fc712e2..34912e60ca4 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -467,7 +467,7 @@ Defined. Proposition lt_implies_pred_geq (i j : nat) : i < j -> i <= nat_pred j. Proof. - intro l; apply leq_n_pred in l; assumption. + intro l; apply leq_pred in l; assumption. Defined. Proposition j_geq_0_lt_implies_pred_geq (i j k : nat) @@ -513,7 +513,7 @@ Proof. destruct (symmetric_paths _ _ (nat_sub_add m k 1)). destruct (symmetric_paths _ _ (@predeqminus1 (n -k))). destruct (symmetric_paths _ _ (@predeqminus1 (m -k))). - apply leq_n_pred, IHk. exact l. + apply leq_pred, IHk. exact l. Defined. Proposition sub_less { n k : nat } : n - k <= n. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index f95e2e72b44..a64aff05060 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -401,9 +401,8 @@ Proof. Defined. Global Existing Instance leq_zero | 10. -(** TODO: rename to [leq_pred] *) (** A predecessor is less than or equal to a predecessor if the original number is less than or equal. *) -Definition leq_n_pred n m : leq n m -> leq (nat_pred n) (nat_pred m). +Definition leq_pred n m : leq n m -> leq (nat_pred n) (nat_pred m). Proof. intros H; induction H. 1: exact _. @@ -414,7 +413,7 @@ Defined. Lemma leq_S_n : forall n m, n.+1 <= m.+1 -> n <= m. Proof. intros n m. - apply leq_n_pred. + apply leq_pred. Defined. (** TODO: possibly rename to [leq_succ] depending on above. *) From cf8f6a2033296313cd58f6b737471d9a98eb7f0a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 20:01:10 +0200 Subject: [PATCH 083/282] Nat: adjust comment about renaming leq_S_n Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index a64aff05060..a22121361fa 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -409,14 +409,15 @@ Proof. destruct m; exact _. Defined. -(** TODO: is this needed? *) -Lemma leq_S_n : forall n m, n.+1 <= m.+1 -> n <= m. +(** TODO: rename leq_S_n -> leq_succ' *) +Definition leq_S_n : forall n m, n.+1 <= m.+1 -> n <= m. Proof. intros n m. apply leq_pred. Defined. (** TODO: possibly rename to [leq_succ] depending on above. *) +(** TODO: rename leq_S_n' -> leq_succ *) Lemma leq_S_n' n m : n <= m -> n.+1 <= m.+1. Proof. induction 1; exact _. From b79ce7025889e7f548115fdbb616567ad4c538a0 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 20:01:55 +0200 Subject: [PATCH 084/282] Nat: rename leq_S_n' -> leq_succ Signed-off-by: Ali Caglayan --- theories/Spaces/Finite/FinNat.v | 6 +++--- theories/Spaces/Finite/Finite.v | 2 +- theories/Spaces/List/Theory.v | 6 +++--- theories/Spaces/Nat/Arithmetic.v | 30 +++++++++++++++--------------- theories/Spaces/Nat/Core.v | 12 +++++------- 5 files changed, 27 insertions(+), 29 deletions(-) diff --git a/theories/Spaces/Finite/FinNat.v b/theories/Spaces/Finite/FinNat.v index 6314ae40233..164ca3d7f92 100644 --- a/theories/Spaces/Finite/FinNat.v +++ b/theories/Spaces/Finite/FinNat.v @@ -18,7 +18,7 @@ Proof. Defined. Definition succ_finnat {n : nat} (u : FinNat n) : FinNat n.+1 - := (u.1.+1; leq_S_n' u.1.+1 n u.2). + := (u.1.+1; leq_succ u.2). Lemma path_succ_finnat {n : nat} (u : FinNat n) (h : u.1.+1 < n.+1) : succ_finnat u = (u.1.+1; h). @@ -78,9 +78,9 @@ Proof. refine (_ @ transport (fun p => transport _ p (s n u _) = s n u (finnat_ind P z s u)) - (hset_path2 1 (path_succ_finnat u (leq_S_n' _ _ u.2))) 1). + (hset_path2 1 (path_succ_finnat u (leq_succ u.2))) 1). destruct u as [u1 u2]. - assert (u2 = leq_S_n u1.+1 n (leq_S_n' u1.+1 n u2)) as p. + assert (u2 = leq_S_n u1.+1 n (leq_succ u2)) as p. - apply path_ishprop. - simpl. by induction p. Defined. diff --git a/theories/Spaces/Finite/Finite.v b/theories/Spaces/Finite/Finite.v index 5880524bad8..928a9df8587 100644 --- a/theories/Spaces/Finite/Finite.v +++ b/theories/Spaces/Finite/Finite.v @@ -676,7 +676,7 @@ Proof. { unfold h; apply moveR_equiv_V; symmetry. apply fin_transpose_last_with_last. } rewrite q; exact tt. } - apply leq_S_n'. + apply leq_succ. exact (IHn m (unfunctor_sum_l h Ha) (mapinO_unfunctor_sum_l (Tr (-1)) h Ha Hb)). Qed. diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index dc2ac79bc49..e5a72dd03d6 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -394,12 +394,12 @@ Proof. snrapply paths_ind_r@{i i}. snrefine (exist@{i i} _ 0%nat _). snrefine (exist _ _ idpath). - apply leq_S_n'. + apply leq_succ. exact _. - destruct (IHl i) as [n [H H']]. snrefine (exist@{i i} _ n.+1%nat _). snrefine (_; _); cbn. - 1: apply leq_S_n', H. + 1: apply leq_succ, H. refine (_ @ H'). apply nth'_cons. Defined. @@ -488,7 +488,7 @@ Proof. 1: by apply path_nat_succ. intros n Hn. snrefine ((nth'_cons l n a Hn _)^ @ _). - 1: apply leq_S_n', Hn. + 1: apply leq_succ, Hn. lhs nrapply H. nrapply nth'_cons. Defined. diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 34912e60ca4..44a3575342c 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -35,7 +35,7 @@ Proposition leq_split {n m : nat} : (n <= m) -> sum (n < m) (n = m). Proof. intro l. induction l. - now right. - - left. exact (leq_S_n' _ _ l). + - left. exact (leq_succ l). Defined. Proposition diseq_implies_lt (n m : nat) @@ -47,7 +47,7 @@ Proof. [ | assumption]. destruct n_leq_m; [ now contradiction diseq - | contradiction a; now apply leq_S_n']. + | contradiction a; now apply leq_succ]. Defined. Proposition lt_implies_diseq (n m : nat) @@ -72,7 +72,7 @@ Defined. Proposition mixed_trans1 (n m k : nat) : n <= m -> m < k -> n < k. Proof. - intros l j. apply leq_S_n' in l. + intros l j. apply leq_succ in l. apply (@leq_trans (n.+1) (m.+1) k); trivial. Defined. @@ -127,7 +127,7 @@ Proposition n_leq_add_n_k (n m : nat) : n <= n + m. Proof. simple_induction n n IHn. - apply leq_zero. - - simpl; apply leq_S_n', IHn. + - simpl; apply leq_succ, IHn. Defined. Proposition n_leq_add_n_k' (n m : nat) : n <= m + n. @@ -153,7 +153,7 @@ Proof. + intro l. apply leq_S_n, natineq0eq0 in l. right; apply ap; exact l. + intro l. apply leq_S_n, IHn in l; destruct l as [a | b]. - * left. apply leq_S_n'; exact a. + * left. apply leq_succ; exact a. * right. apply ap; exact b. Defined. @@ -166,7 +166,7 @@ Proof. - intros m eq. destruct m. + simpl in eq. apply symmetric_paths in eq. contradiction (neq_nat_zero_succ n eq). - + simpl in eq. apply leq_S_n', IHn, eq. + + simpl in eq. apply leq_succ, IHn, eq. Defined. Proposition sub_gt_0_lt (n m : nat) : n - m > 0 -> m < n. @@ -242,7 +242,7 @@ Proof. simple_induction k k IHk. - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; exact l. - destruct (nat_add_succ_r n k)^, (nat_add_succ_r m k)^; - apply leq_S_n'; exact IHk. + apply leq_succ; exact IHk. Defined. Proposition nataddpreservesleq' { n m k : nat } @@ -262,7 +262,7 @@ Proof. simple_induction k k IHk. - destruct (nat_add_zero_r n')^, (nat_add_zero_r m)^; exact l. - destruct (nat_add_succ_r n' k)^, (nat_add_succ_r m k)^; - apply leq_S_n'; exact IHk. + apply leq_succ; exact IHk. Defined. Proposition nataddpreserveslt' { n m k : nat } @@ -379,7 +379,7 @@ Proof. - destruct (nat_add_succ_r n k)^. refine (leq_trans (nataddsub_comm_ineq_lemma (n+k) m) _). destruct (nat_add_succ_r (n - m) k)^. - now apply leq_S_n'. + now apply leq_succ. Defined. Proposition nat_sub_add_ineq (n m : nat) : n <= n - m + m. @@ -400,7 +400,7 @@ Proof. - intros m i l. simpl in l. contradiction (not_lt_n_0 _ _). - intros m i l. destruct m. + apply leq_zero. - + apply leq_S_n'. simpl in l. apply (IHn m i l). + + apply leq_succ. simpl in l. apply (IHn m i l). Defined. Proposition nataddsub_assoc_implication (n : nat) {m k z : nat} @@ -483,7 +483,7 @@ Proposition pred_gt_implies_lt (i j : nat) : i < nat_pred j -> i.+1 < j. Proof. intros ineq. - assert (H := leq_S_n' _ _ ineq). assert (i < j) as X. { + assert (H := leq_succ ineq). assert (i < j) as X. { apply (@mixed_trans2 _ (nat_pred j) _); [assumption | apply predn_leq_n]. } @@ -534,7 +534,7 @@ Proof. unfold "<". destruct k, n; try (contradiction (not_lt_n_0 _ _)). - simpl; apply leq_S_n', sub_less. + simpl; apply leq_succ, sub_less. Defined. Proposition natpmswap1 (k m n : nat) @@ -592,7 +592,7 @@ Proof. - intros n' m; destruct m. + intros. contradiction (not_leq_Sn_0 n). + intros m' l l'. apply leq_S_n in l. simpl. - apply leq_S_n', IHn. + apply leq_succ, IHn. * exact l. * exact l'. Defined. @@ -608,7 +608,7 @@ Proof. - intros m l. apply leq_S_n in l. destruct l as [ | n]. + apply a; intros ? ?; now apply IHn. - + now apply (IHn m), leq_S_n'. + + now apply (IHn m), leq_succ. } intro n. apply (X (n.+1) n), (leq_n n.+1). Defined. @@ -724,6 +724,6 @@ Proof. destruct (@leq_dichot m n) as [m_leq_n | m_gt_n]. - apply symmetry. destruct m_leq_n. + apply reflexivity. - + apply A. apply leq_S_n'. assumption. + + apply A. apply leq_succ. assumption. - apply A, m_gt_n. Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index a22121361fa..de550ca3bd9 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -416,13 +416,11 @@ Proof. apply leq_pred. Defined. -(** TODO: possibly rename to [leq_succ] depending on above. *) -(** TODO: rename leq_S_n' -> leq_succ *) -Lemma leq_S_n' n m : n <= m -> n.+1 <= m.+1. +Definition leq_succ {n m} : n <= m -> n.+1 <= m.+1. Proof. induction 1; exact _. Defined. -Global Existing Instance leq_S_n' | 100. +Global Existing Instance leq_succ | 100. (** TODO: use lemmas about negating predicate *) Lemma not_leq_Sn_n n : ~ (n.+1 <= n). @@ -707,7 +705,7 @@ Defined. Definition not_lt_n_n n : ~ (n < n) := not_leq_Sn_n n. -Definition leq_1_Sn {n} : 1 <= n.+1 := leq_S_n' 0 n (leq_zero _). +Definition leq_1_Sn {n} : 1 <= n.+1 := leq_succ (leq_zero _). Fixpoint leq_dichot {m} {n} : (m <= n) + (m > n). Proof. @@ -717,8 +715,8 @@ Proof. - right; unfold lt; apply leq_1_Sn. - assert ((m <= n) + (n < m)) as X by apply leq_dichot. destruct X as [leqmn|ltnm]. - + left; apply leq_S_n'; assumption. - + right; apply leq_S_n'; assumption. + + left; apply leq_succ; assumption. + + right; apply leq_succ; assumption. Defined. Lemma not_lt_n_0 n : ~ (n < 0). From bdff85f76da37bff7e17a599592e7d8514b14522 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 20:08:37 +0200 Subject: [PATCH 085/282] Nat: rename nat_max_n_n -> nat_max_idem Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index de550ca3bd9..955f5229cae 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -585,9 +585,8 @@ Defined. (** *** Properties of Maxima *) -(** TODO: rename to [nat_max_idem] *) (** [nat_max] is idempotent. *) -Definition nat_max_n_n@{} n : nat_max n n = n. +Definition nat_max_idem@{} n : nat_max n n = n. Proof. simple_induction' n; cbn. 1: reflexivity. From 412124edbf157d2b957ee95fd816b8fc1477ab7e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 20:09:21 +0200 Subject: [PATCH 086/282] Nat: rename nat_max_0_n -> nat_max_zero_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 955f5229cae..8a6f75c3183 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -609,10 +609,8 @@ Proof. exact (ap S IH). Defined. -(** TODO: rename to [nat_max_zero_l] *) (** [0] is the left identity of [nat_max]. *) -Definition nat_max_0_n@{} n : nat_max 0 n = n := idpath. -Notation nat_max_zero_l := nat_max_0_n. +Definition nat_max_zero_l@{} n : nat_max 0 n = n := idpath. (** TODO: rename to [nat_max_zero_r] *) (** [0] is the right identity of [nat_max]. *) From 9525e0584d58aa3faabcc6656c3c77ff41b5840f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 20:10:03 +0200 Subject: [PATCH 087/282] Nat: rename nat_max_n_0 -> nat_max_zero_r Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 8a6f75c3183..2f1a7237a5f 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -612,17 +612,15 @@ Defined. (** [0] is the left identity of [nat_max]. *) Definition nat_max_zero_l@{} n : nat_max 0 n = n := idpath. -(** TODO: rename to [nat_max_zero_r] *) (** [0] is the right identity of [nat_max]. *) -Definition nat_max_n_0@{} n : nat_max n 0 = n +Definition nat_max_zero_r@{} n : nat_max n 0 = n := nat_max_comm _ _ @ nat_max_zero_l _. -Notation nat_max_zero_r := nat_max_n_0. Definition nat_max_l@{} {n m} : m <= n -> nat_max n m = n. Proof. intros H. induction m as [|m IHm] in n, H |- *. - 1: nrapply nat_max_n_0. + 1: nrapply nat_max_zero_r. destruct n. 1: inversion H. cbn; by apply (ap S), IHm, leq_S_n. From b4c67afa100a44a9a22c6a5a4f8a9629a6513b2b Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 20:12:52 +0200 Subject: [PATCH 088/282] Nat: add more comments about max and min Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 2f1a7237a5f..b3cab3163c3 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -616,6 +616,7 @@ Definition nat_max_zero_l@{} n : nat_max 0 n = n := idpath. Definition nat_max_zero_r@{} n : nat_max n 0 = n := nat_max_comm _ _ @ nat_max_zero_l _. +(** [nat_max n m] is [n] if [m <= n]. *) Definition nat_max_l@{} {n m} : m <= n -> nat_max n m = n. Proof. intros H. @@ -626,11 +627,13 @@ Proof. cbn; by apply (ap S), IHm, leq_S_n. Defined. +(** [nat_max n m] is [m] if [n <= m]. *) Definition nat_max_r {n m} : n <= m -> nat_max n m = m := fun _ => nat_max_comm _ _ @ nat_max_l _. (** Properties of Minima *) +(** [nat_min] is commutative. *) Definition nat_min_comm n m : nat_min n m = nat_min m n. Proof. induction n as [|n IHn] in m |- *; destruct m; cbn. @@ -638,10 +641,14 @@ Proof. exact (ap S (IHn _)). Defined. +(** [nat_min] of [0] and [n] is [0]. *) Definition nat_min_zero_l n : nat_min 0 n = 0 := idpath. + +(** [nat_min] of [n] and [0] is [0]. *) Definition nat_min_zero_r n : nat_min n 0 = 0:= nat_min_comm _ _ @ nat_min_zero_l _. +(** [nat_min n m] is [n] if [n <= m]. *) Definition nat_min_l {n m} : n <= m -> nat_min n m = n. Proof. revert n m. @@ -651,6 +658,7 @@ Proof. cbn; by apply (ap S), IHn, leq_S_n. Defined. +(** [nat_min n m] is [m] if [m <= n]. *) Definition nat_min_r {n m} : m <= n -> nat_min n m = m := fun _ => nat_min_comm _ _ @ nat_min_l _. From 7340b9cebc7a0c68a9a35e85d133fc802669bad4 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 20:13:31 +0200 Subject: [PATCH 089/282] Nat: add nat_min_idem Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index b3cab3163c3..2360f81ec34 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -633,6 +633,14 @@ Definition nat_max_r {n m} : n <= m -> nat_max n m = m (** Properties of Minima *) +(** [nat_min] is idempotent. *) +Definition nat_min_idem n : nat_min n n = n. +Proof. + simple_induction' n; cbn. + 1: reflexivity. + exact (ap S IH). +Defined. + (** [nat_min] is commutative. *) Definition nat_min_comm n m : nat_min n m = nat_min m n. Proof. From 7b82c10e2e3b6cdaae3d7d298484f19ff5c1bfa7 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 20:15:31 +0200 Subject: [PATCH 090/282] Nat: rename nat_max_Sn_n -> nat_max_succ_l and add nat_max_succ_r Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 8 ++++++-- 1 file changed, 6 insertions(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 2360f81ec34..9c9610c8346 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -601,14 +601,18 @@ Proof. exact (ap S (IHn _)). Defined. -(** TODO: rename to [nat_max_succ_l] *) -Definition nat_max_Sn_n@{} n : nat_max (S n) n = S n. +(** The maximum of [n.+1] and [n] is [n.+1]. *) +Definition nat_max_succ_l@{} n : nat_max n.+1 n = n.+1. Proof. simple_induction' n; cbn. 1: reflexivity. exact (ap S IH). Defined. +(** The maximum of [n] and [n.+1] is [n.+1]. *) +Definition nat_max_succ_r@{} n : nat_max n n.+1 = n.+1 + := nat_max_comm _ _ @ nat_max_succ_l _. + (** [0] is the left identity of [nat_max]. *) Definition nat_max_zero_l@{} n : nat_max 0 n = n := idpath. From 40f2f03f3747da2684934256136875720f7276c1 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 20:17:36 +0200 Subject: [PATCH 091/282] Nat: add nat_max_assoc and nat_min_assoc Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 22 ++++++++++++++++++++++ 1 file changed, 22 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 9c9610c8346..fd396b34df1 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -635,6 +635,17 @@ Defined. Definition nat_max_r {n m} : n <= m -> nat_max n m = m := fun _ => nat_max_comm _ _ @ nat_max_l _. +(** [nat_max n m] is associative. *) +Definition nat_max_assoc@{} n m k + : nat_max n (nat_max m k) = nat_max (nat_max n m) k. +Proof. + induction n as [|n IHn] in m, k |- *. + 1: reflexivity. + destruct m, k. + 1-3: reflexivity. + by apply (ap S), IHn. +Defined. + (** Properties of Minima *) (** [nat_min] is idempotent. *) @@ -674,6 +685,17 @@ Defined. Definition nat_min_r {n m} : m <= n -> nat_min n m = m := fun _ => nat_min_comm _ _ @ nat_min_l _. +(** [nat_min n m] is associative. *) +Definition nat_min_assoc n m k + : nat_min n (nat_min m k) = nat_min (nat_min n m) k. +Proof. + induction n as [|n IHn] in m, k |- *. + 1: reflexivity. + destruct m, k. + 1-3: reflexivity. + by apply (ap S), IHn. +Defined. + (** ** Inequality of natural numbers *) Fixpoint leq_add n m : n <= (m + n). From 42101581ebb6f3391e0bee7f31906877de36879e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 22:28:54 +0200 Subject: [PATCH 092/282] Nat: leq is an antisymmetric relation Signed-off-by: Ali Caglayan --- theories/Basics/Classes.v | 6 ++++++ theories/Classes/interfaces/canonical_names.v | 4 ---- theories/Spaces/Nat/Core.v | 12 ++++++------ 3 files changed, 12 insertions(+), 10 deletions(-) diff --git a/theories/Basics/Classes.v b/theories/Basics/Classes.v index 61f87f77530..9a56fcbb173 100644 --- a/theories/Basics/Classes.v +++ b/theories/Basics/Classes.v @@ -37,3 +37,9 @@ Proof. apply (injective (g o f)). exact (ap g p). Defined. + +(** ** Antisymmetric Relations *) + +Class AntiSymmetric {A : Type} (R : A -> A -> Type) : Type + := antisymmetry : forall x y, R x y -> R y x -> x = y. +Arguments antisymmetry {A} R {AntiSymmetric} x y _ _. diff --git a/theories/Classes/interfaces/canonical_names.v b/theories/Classes/interfaces/canonical_names.v index ec68bbbc61b..fb45c33e6a6 100644 --- a/theories/Classes/interfaces/canonical_names.v +++ b/theories/Classes/interfaces/canonical_names.v @@ -271,10 +271,6 @@ Class CoTransitive `(R : Relation A) : Type := cotransitive : forall x y, R x y -> forall z, hor (R x z) (R z y). Arguments cotransitive {A R CoTransitive x y} _ _. -Class AntiSymmetric `(R : Relation A) : Type - := antisymmetry: forall x y, R x y -> R y x -> x = y. -Arguments antisymmetry {A} _ {AntiSymmetric} _ _ _ _. - Class EquivRel `(R : Relation A) : Type := Build_EquivRel { EquivRel_Reflexive : Reflexive R ; EquivRel_Symmetric : Symmetric R ; diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index fd396b34df1..fe677ad843e 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -725,21 +725,21 @@ Proof. exact (isinj_nat_add_l k _ _ H). Defined. -(** ** Natural number ordering *) +(** ** More Theory of Comparison Predicates *) -(** ** Theorems about natural number ordering *) - -Lemma leq_antisym {x y} : x <= y -> y <= x -> x = y. +(** [<=] is an antisymmetric relation. *) +Definition leq_antisym {x y} : x <= y -> y <= x -> x = y. Proof. intros p q. destruct p. 1: reflexivity. destruct x; [inversion q|]. apply leq_S_n in q. - pose (r := leq_trans p q). - by apply not_leq_Sn_n in r. + contradiction (not_leq_Sn_n _ (leq_trans p q)). Defined. +Global Instance antisymmetric_leq : AntiSymmetric leq := @leq_antisym. + Definition not_lt_n_n n : ~ (n < n) := not_leq_Sn_n n. Definition leq_1_Sn {n} : 1 <= n.+1 := leq_succ (leq_zero _). From 15195c5f3a36e055e4469e6234aa0335f918024f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 22:35:53 +0200 Subject: [PATCH 093/282] Nat: clarify < is irreflexive Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 3 +++ 1 file changed, 3 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index fe677ad843e..687a744473a 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -740,8 +740,11 @@ Defined. Global Instance antisymmetric_leq : AntiSymmetric leq := @leq_antisym. +(** [<] is an irreflexive relation. *) Definition not_lt_n_n n : ~ (n < n) := not_leq_Sn_n n. +Global Instance irreflexive_lt : Irreflexive lt := not_lt_n_n. +(** [1] is less than or equal to the successor of any number. *) Definition leq_1_Sn {n} : 1 <= n.+1 := leq_succ (leq_zero _). Fixpoint leq_dichot {m} {n} : (m <= n) + (m > n). From f0974a5f5d99f64a7f72e490caeedeb91ac55411 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 22:37:26 +0200 Subject: [PATCH 094/282] Nat: rename leq_S_n -> leq_succ' Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 687a744473a..2af1d0781fd 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -410,7 +410,7 @@ Proof. Defined. (** TODO: rename leq_S_n -> leq_succ' *) -Definition leq_S_n : forall n m, n.+1 <= m.+1 -> n <= m. +Definition leq_succ' : forall n m, n.+1 <= m.+1 -> n <= m. Proof. intros n m. apply leq_pred. @@ -429,7 +429,7 @@ Proof. { intro p. inversion p. } intros p. - by apply IHn, leq_S_n. + by apply IHn, leq_succ'. Defined. Lemma not_leq_Sn_0 n : ~ (n.+1 <= 0). @@ -490,7 +490,7 @@ Defined. Definition equiv_leq_S_n n m : n.+1 <= m.+1 <~> n <= m. Proof. srapply equiv_iff_hprop. - apply leq_S_n. + apply leq_succ'. Defined. Global Instance decidable_leq n m : Decidable (n <= m). @@ -628,7 +628,7 @@ Proof. 1: nrapply nat_max_zero_r. destruct n. 1: inversion H. - cbn; by apply (ap S), IHm, leq_S_n. + cbn; by apply (ap S), IHm, leq_succ'. Defined. (** [nat_max n m] is [m] if [n <= m]. *) @@ -678,7 +678,7 @@ Proof. simple_induction n n IHn; auto. intros [] p. 1: inversion p. - cbn; by apply (ap S), IHn, leq_S_n. + cbn; by apply (ap S), IHn, leq_succ'. Defined. (** [nat_min n m] is [m] if [m <= n]. *) @@ -734,7 +734,7 @@ Proof. destruct p. 1: reflexivity. destruct x; [inversion q|]. - apply leq_S_n in q. + apply leq_succ' in q. contradiction (not_leq_Sn_n _ (leq_trans p q)). Defined. From 619ec62ec89b60b30c01380438f89428ffa91b2b Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 22:47:59 +0200 Subject: [PATCH 095/282] Nat: rename leq_S_n -> leq_succ' Signed-off-by: Ali Caglayan --- theories/BoundedSearch.v | 2 +- theories/Spaces/Finite/FinNat.v | 17 ++++++++--------- theories/Spaces/List/Theory.v | 18 +++++++++--------- theories/Spaces/Nat/Arithmetic.v | 28 ++++++++++++++-------------- theories/Spaces/Nat/Core.v | 13 +++++-------- 5 files changed, 37 insertions(+), 41 deletions(-) diff --git a/theories/BoundedSearch.v b/theories/BoundedSearch.v index 234b13955c6..19f66d29c71 100644 --- a/theories/BoundedSearch.v +++ b/theories/BoundedSearch.v @@ -61,7 +61,7 @@ Section bounded_search. destruct X as [leqSnm|ltmSn]. ++ assumption. ++ unfold gt, lt in ltmSn. - assert (m <= n) as X by rapply leq_S_n. + assert (m <= n) as X by rapply leq_succ'. destruct (n0 m X pm). * right. intros l q. assert ((l <= n) + (l > n)) as X by apply leq_dichot. diff --git a/theories/Spaces/Finite/FinNat.v b/theories/Spaces/Finite/FinNat.v index 164ca3d7f92..d30db097c26 100644 --- a/theories/Spaces/Finite/FinNat.v +++ b/theories/Spaces/Finite/FinNat.v @@ -55,7 +55,7 @@ Proof. - destruct u as [x h]. destruct x as [| x]. + exact (transport (P n.+1) (path_zero_finnat _ h) (z _)). - + refine (transport (P n.+1) (path_succ_finnat (x; leq_S_n _ _ h) _) _). + + refine (transport (P n.+1) (path_succ_finnat (x; leq_succ' h) _) _). apply s. apply IHn. Defined. @@ -80,9 +80,8 @@ Proof. (fun p => transport _ p (s n u _) = s n u (finnat_ind P z s u)) (hset_path2 1 (path_succ_finnat u (leq_succ u.2))) 1). destruct u as [u1 u2]. - assert (u2 = leq_S_n u1.+1 n (leq_succ u2)) as p. - - apply path_ishprop. - - simpl. by induction p. + assert (u2 = leq_succ' (leq_succ u2)) as p by apply path_ishprop. + simpl; by induction p. Defined. Monomorphic Definition is_bounded_fin_to_nat {n} (k : Fin n) @@ -106,7 +105,7 @@ Monomorphic Fixpoint finnat_to_fin {n : nat} : FinNat n -> Fin n | n.+1 => fun u => match u with | (0; _) => fin_zero - | (x.+1; h) => fsucc (finnat_to_fin (x; leq_S_n _ _ h)) + | (x.+1; h) => fsucc (finnat_to_fin (x; leq_succ' h)) end end. @@ -155,9 +154,9 @@ Proof. - elim (not_lt_n_0 _ u.2). - destruct u as [x h]. destruct x as [| x]; [reflexivity|]. - refine ((ap _ (ap _ (path_succ_finnat (x; leq_S_n _ _ h) h)))^ @ _). - refine (_ @ ap fsucc (IHn (x; leq_S_n _ _ h))). - by induction (path_finnat_to_fin_succ (incl_finnat (x; leq_S_n _ _ h))). + refine ((ap _ (ap _ (path_succ_finnat (x; leq_succ' h) h)))^ @ _). + refine (_ @ ap fsucc (IHn (x; leq_succ' h))). + by induction (path_finnat_to_fin_succ (incl_finnat (x; leq_succ' h))). Defined. Lemma path_finnat_to_fin_last (n : nat) @@ -178,7 +177,7 @@ Proof. destruct x as [| x]. + exact (ap pr1 (path_fin_to_finnat_fin_zero n)). + refine ((path_fin_to_finnat_fsucc _)..1 @ _). - exact (ap S (IHn (x; leq_S_n _ _ h))..1). + exact (ap S (IHn (x; leq_succ' h))..1). Defined. Lemma path_fin_to_finnat_to_fin {n : nat} (k : Fin n) diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index e5a72dd03d6..32763677a2a 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -336,7 +336,7 @@ Proof. destruct n. 1: by exists a. apply IHa. - apply leq_S_n. + apply leq_succ'. exact H. Defined. @@ -448,11 +448,11 @@ Proof. + destruct n. * reflexivity. * unshelve erewrite nth'_cons. - 1: apply leq_S_n, H. + 1: apply leq_succ', H. unshelve erewrite nth'_cons. - 1: apply leq_S_n, H'. + 1: apply leq_succ', H'. unshelve erewrite nth'_cons. - 1: apply leq_S_n, H''. + 1: apply leq_succ', H''. apply IHl1. by apply path_nat_succ. Defined. @@ -502,7 +502,7 @@ Proof. 1: destruct (not_leq_Sn_0 _ H). destruct n. 1: reflexivity. - by apply IHl, leq_S_n. + by apply IHl, leq_succ'. Defined. (** The [nth i] element where [pred (length l) = i] is the last element of the list. *) @@ -574,7 +574,7 @@ Proof. destruct n. 1: destruct (not_leq_Sn_0 _ H). cbn; apply IHl. - apply leq_S_n. + apply leq_succ'. exact H. Defined. @@ -632,7 +632,7 @@ Proof. destruct n. 1: destruct (not_leq_Sn_0 _ H). cbn; f_ap. - by apply IHl, leq_S_n. + by apply IHl, leq_succ'. Defined. (** The length of a [take n] is the minimum of [n] and the length of the original list. *) @@ -824,7 +824,7 @@ Proof. cbn; by rewrite nat_sub_zero_r. - induction n as [|n IHn]. 1: destruct (not_leq_Sn_0 _ H). - by apply IHi, leq_S_n. + by apply IHi, leq_succ'. Defined. (** The [nth] element of a [seq] is [i]. *) @@ -841,7 +841,7 @@ Proof. - apply not_lt_implies_geq in H'. destruct (leq_split H') as [H'' | H'']. { apply lt_implies_not_geq in H''. - apply leq_S_n in H. + apply leq_succ' in H. contradiction. } destruct H''. lhs nrapply nth_last. diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 44a3575342c..fdf8f0bdef2 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -66,7 +66,7 @@ Defined. Proposition leq_S' (n m : nat) : n.+1 <= m -> n <= m. Proof. intro l. - now apply leq_S_n, leq_S. + now apply leq_succ', leq_S. Defined. Proposition mixed_trans1 (n m k : nat) @@ -108,7 +108,7 @@ Defined. Proposition n_lt_m_n_leq_m { n m : nat } : n < m -> n <= m. Proof. - intro H. apply leq_S, leq_S_n in H; exact H. + intro H. apply leq_S, leq_succ' in H; exact H. Defined. Proposition lt_trans {n m k : nat} : n < m -> m < k -> n < k. @@ -150,9 +150,9 @@ Proof. revert m; simple_induction n n IHn. - intro. left. exact (leq_zero m). - destruct m. - + intro l. apply leq_S_n, natineq0eq0 in l. + + intro l. apply leq_succ', natineq0eq0 in l. right; apply ap; exact l. - + intro l. apply leq_S_n, IHn in l; destruct l as [a | b]. + + intro l. apply leq_succ', IHn in l; destruct l as [a | b]. * left. apply leq_succ; exact a. * right. apply ap; exact b. Defined. @@ -183,7 +183,7 @@ Proof. - intros m ineq. contradiction (not_lt_n_0 m). - destruct m. + simpl. easy. - + simpl. intro ineq. apply leq_S_n in ineq. + + simpl. intro ineq. apply leq_succ' in ineq. now apply IHn in ineq. Defined. @@ -195,7 +195,7 @@ Proof. apply (ap S), symmetric_paths, (nat_add_zero_r _)^. - intros m l. destruct m. + contradiction (not_leq_Sn_0 n). - + simpl. apply leq_S_n, IHn in l. + + simpl. apply leq_succ', IHn in l. destruct (nat_add_succ_r (m - n) n)^. destruct (symmetric_paths _ _ l). reflexivity. @@ -279,7 +279,7 @@ Proof. simple_induction k k IHk. - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; trivial. - intro l. destruct (nat_add_succ_r n k)^, (nat_add_succ_r m k)^ in l. - apply leq_S_n, IHk in l; exact l. + apply leq_succ', IHk in l; exact l. Defined. Proposition nataddreflectsleq { n m k : nat } @@ -324,7 +324,7 @@ Proof. - intros m l; simpl. destruct m; reflexivity. - destruct m. + simpl; intro g; contradiction (not_leq_Sn_0 _ g). - + intro l; apply leq_S_n in l. + + intro l; apply leq_succ' in l. change (m.+2 - k.+1) with (m.+1 - k). change (m.+1 - k.+1) with (m - k). exact (IHk _ l). @@ -343,7 +343,7 @@ Proof. | reflexivity | contradiction (not_lt_n_0 k _) | ]. - simpl "-". apply leq_S_n in l. + simpl "-". apply leq_succ' in l. destruct (nat_add_succ_r n (m - k)). destruct (nataddsub_assoc_lemma l). apply (IHn m.+1 k). @@ -476,7 +476,7 @@ Proof. intros l ineq. destruct j. - contradiction (not_lt_n_0 i). - - now simpl; apply leq_S_n. + - now simpl; apply leq_succ'. Defined. Proposition pred_gt_implies_lt (i j : nat) @@ -495,7 +495,7 @@ Proposition pred_preserves_lt {i n: nat} (p : i < n) m : (n < m) -> (nat_pred n < nat_pred m). Proof. intro l. - apply leq_S_n. destruct (symmetric_paths _ _ (S_predn n i _)). + apply leq_succ'. destruct (symmetric_paths _ _ (S_predn n i _)). set (k := transitive_lt i n m p l). destruct (symmetric_paths _ _ (S_predn m i _)). assumption. @@ -591,7 +591,7 @@ Proof. apply (leq_trans l'). exact (n_leq_add_n_k' m' m). - intros n' m; destruct m. + intros. contradiction (not_leq_Sn_0 n). - + intros m' l l'. apply leq_S_n in l. simpl. + + intros m' l l'. apply leq_succ' in l. simpl. apply leq_succ, IHn. * exact l. * exact l'. @@ -605,7 +605,7 @@ Proof. assert (forall n m: nat, m < n -> P m) as X. { simple_induction n n IHn. - intros m l. contradiction (not_lt_n_0 m). - - intros m l. apply leq_S_n in l. + - intros m l. apply leq_succ' in l. destruct l as [ | n]. + apply a; intros ? ?; now apply IHn. + now apply (IHn m), leq_succ. @@ -669,7 +669,7 @@ Proof. intro ineq. destruct n. - contradiction (not_lt_n_0 k). - - change (n.+1 - k.+1) with (n - k). apply leq_S_n in ineq. + - change (n.+1 - k.+1) with (n - k). apply leq_succ' in ineq. apply (nataddsub_assoc_lemma _). Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 2af1d0781fd..2c5376cebfe 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -402,26 +402,23 @@ Defined. Global Existing Instance leq_zero | 10. (** A predecessor is less than or equal to a predecessor if the original number is less than or equal. *) -Definition leq_pred n m : leq n m -> leq (nat_pred n) (nat_pred m). +Definition leq_pred {n m} : leq n m -> leq (nat_pred n) (nat_pred m). Proof. intros H; induction H. 1: exact _. destruct m; exact _. Defined. -(** TODO: rename leq_S_n -> leq_succ' *) -Definition leq_succ' : forall n m, n.+1 <= m.+1 -> n <= m. -Proof. - intros n m. - apply leq_pred. -Defined. - +(** A successor is less than or equal to a successor if the original numbers are less than or equal. *) Definition leq_succ {n m} : n <= m -> n.+1 <= m.+1. Proof. induction 1; exact _. Defined. Global Existing Instance leq_succ | 100. +(** The converse to [leq_succ] also holds. *) +Definition leq_succ' {n m} : n.+1 <= m.+1 -> n <= m := leq_pred. + (** TODO: use lemmas about negating predicate *) Lemma not_leq_Sn_n n : ~ (n.+1 <= n). Proof. From e280bd2e9e452859b4c41f8a78129dd760f2a1f7 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 22:49:32 +0200 Subject: [PATCH 096/282] Nat: move leq_add to more appropriate place Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 16 +++++++--------- 1 file changed, 7 insertions(+), 9 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 2c5376cebfe..c2995feb6d7 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -693,15 +693,6 @@ Proof. by apply (ap S), IHn. Defined. -(** ** Inequality of natural numbers *) - -Fixpoint leq_add n m : n <= (m + n). -Proof. - destruct m. - 1: apply leq_n. - apply leq_S, leq_add. -Defined. - (** ** Arithmetic *) Global Instance isinj_S : IsInjective S. @@ -724,6 +715,13 @@ Defined. (** ** More Theory of Comparison Predicates *) +Fixpoint leq_add n m : n <= (m + n). +Proof. + destruct m. + 1: apply leq_n. + apply leq_S, leq_add. +Defined. + (** [<=] is an antisymmetric relation. *) Definition leq_antisym {x y} : x <= y -> y <= x -> x = y. Proof. From 5854c0888338a938cfd2660bba8162fb3bf83c4d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 22:53:03 +0200 Subject: [PATCH 097/282] Nat: move and reprove injectivity of S, add Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 37 +++++++++++++++++-------------------- 1 file changed, 17 insertions(+), 20 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index c2995feb6d7..0199a526a92 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -222,6 +222,7 @@ Definition nat_pred_succ@{} n : nat_pred (nat_succ n) = n (** Injectivity of successor. *) Definition path_nat_succ@{} n m (H : S n = S m) : n = m := ap nat_pred H. +Global Instance isinj_succ : IsInjective nat_succ := path_nat_succ. (** Inequality of sucessors is implied with inequality of the arguments. *) Definition neq_nat_succ@{} n m : n <> m -> S n <> S m. @@ -309,6 +310,22 @@ Proof. exact IHn. Defined. +(** Addition on the left is injective. *) +Global Instance isinj_nat_add_l@{} k : IsInjective (nat_add k). +Proof. + simple_induction k k Ik; exact _. +Defined. + +(** Addition on the right is injective. *) +Definition isinj_nat_add_r@{} k : IsInjective (fun x => nat_add x k). +Proof. + intros x y H. + nrapply (isinj_nat_add_l k). + lhs nrapply nat_add_comm. + lhs nrapply H. + nrapply nat_add_comm. +Defined. + (** ** Properties of multiplication *) (** Multiplication by [0] on the left is [0]. *) @@ -693,26 +710,6 @@ Proof. by apply (ap S), IHn. Defined. -(** ** Arithmetic *) - -Global Instance isinj_S : IsInjective S. -Proof. - intros x y p. - by apply path_nat_succ. -Defined. - -Global Instance isinj_nat_add_l@{} k : IsInjective (nat_add k). -Proof. - simple_induction k k Ik; exact _. -Defined. - -Definition isinj_nat_add_r@{} k : IsInjective (fun x => nat_add x k). -Proof. - intros x y H. - rewrite 2 (nat_add_comm _ k) in H. - exact (isinj_nat_add_l k _ _ H). -Defined. - (** ** More Theory of Comparison Predicates *) Fixpoint leq_add n m : n <= (m + n). From 803ffad304ac23289322c5665aa586905562a2a8 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 22:55:21 +0200 Subject: [PATCH 098/282] Nat: move negation of comparison lemmas to Core.v Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 24 ------------------------ theories/Spaces/Nat/Core.v | 24 ++++++++++++++++++++++++ 2 files changed, 24 insertions(+), 24 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index fdf8f0bdef2..2638d8a37ed 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -6,30 +6,6 @@ Local Set Universe Minimization ToSet. Local Close Scope trunc_scope. Local Open Scope nat_scope. -Proposition not_lt_implies_geq {n m : nat} : ~(n < m) -> m <= n. -Proof. - intros not_lt. - destruct (@leq_dichot m n); [ assumption | contradiction]. -Defined. - -Proposition not_leq_implies_gt {n m : nat} : ~(n <= m) -> m < n. -Proof. - intros not_leq. - destruct (@leq_dichot n m); [ contradiction | assumption]. -Defined. - -Proposition lt_implies_not_geq {n m : nat} : (n < m) -> ~(m <= n). -Proof. - intros ineq1 ineq2. - contradiction (not_lt_n_n n). by apply (leq_trans ineq1). -Defined. - -Proposition leq_implies_not_gt {n m : nat} : (n <= m) -> ~(m < n). -Proof. - intros ineq1 ineq2. - contradiction (not_lt_n_n n); by refine (leq_trans _ ineq2). -Defined. - (** Because of the inductive definition of [<=], one can destruct the proof of [n <= m] and get a judgemental identification between [n] and [m] rather than a propositional one, which may be preferable to the following lemma. *) Proposition leq_split {n m : nat} : (n <= m) -> sum (n < m) (n = m). Proof. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 0199a526a92..794caae5855 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -756,6 +756,30 @@ Proof. apply not_leq_Sn_0. Defined. +Definition not_lt_implies_geq {n m : nat} : ~(n < m) -> m <= n. +Proof. + intros not_lt. + destruct (@leq_dichot m n); [ assumption | contradiction]. +Defined. + +Definition not_leq_implies_gt {n m : nat} : ~(n <= m) -> m < n. +Proof. + intros not_leq. + destruct (@leq_dichot n m); [ contradiction | assumption]. +Defined. + +Definition lt_implies_not_geq {n m : nat} : (n < m) -> ~(m <= n). +Proof. + intros ineq1 ineq2. + contradiction (not_lt_n_n n). by apply (leq_trans ineq1). +Defined. + +Definition leq_implies_not_gt {n m : nat} : (n <= m) -> ~(m < n). +Proof. + intros ineq1 ineq2. + contradiction (not_lt_n_n n); by refine (leq_trans _ ineq2). +Defined. + (** ** Arithmetic relations between [trunc_index] and [nat]. *) Lemma trunc_index_add_nat_add (n : nat) From 2b21cf22ac112ea661c3d955d86b0d7b7bf54999 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 22:59:33 +0200 Subject: [PATCH 099/282] Nat: rename leq_n -> leq_refl Signed-off-by: Ali Caglayan --- .../Classes/implementations/peano_naturals.v | 2 +- theories/Classes/interfaces/cauchy.v | 2 +- theories/Spaces/Finite/FinNat.v | 2 +- theories/Spaces/Nat/Arithmetic.v | 16 +++++++------- theories/Spaces/Nat/Core.v | 22 +++++++++---------- 5 files changed, 21 insertions(+), 23 deletions(-) diff --git a/theories/Classes/implementations/peano_naturals.v b/theories/Classes/implementations/peano_naturals.v index 5ddbc9a47a2..085611f76a5 100644 --- a/theories/Classes/implementations/peano_naturals.v +++ b/theories/Classes/implementations/peano_naturals.v @@ -274,7 +274,7 @@ Global Instance nat_lt: Lt@{N N} nat := Nat.Core.lt. Lemma le_plus : forall n k, n <= k + n. Proof. induction k. -- apply Nat.Core.leq_n. +- apply Nat.Core.leq_refl. - simpl_nat. constructor. assumption. Qed. diff --git a/theories/Classes/interfaces/cauchy.v b/theories/Classes/interfaces/cauchy.v index 47df617e642..9e4f9f35d3a 100644 --- a/theories/Classes/interfaces/cauchy.v +++ b/theories/Classes/interfaces/cauchy.v @@ -58,7 +58,7 @@ Section cauchy. assert (isclose := isclose' n leNn). clear isclose'. assert (leMn := le_nat_max_l (M (epsilon / 2)) N : M (epsilon / 2) ≤ n). - assert (leMM : M (epsilon / 2) ≤ M (epsilon / 2) ) by apply (Nat.Core.leq_n). + assert (leMM : M (epsilon / 2) ≤ M (epsilon / 2) ) by apply (Nat.Core.leq_refl). assert (x_close := cauchy_convergence x (epsilon/2) n (M (epsilon / 2)) leMn leMM). cbn in isclose, x_close. rewrite (@preserves_mult Q F _ _ _ _ _ _ _ _ _ _ _ _) in isclose, x_close. diff --git a/theories/Spaces/Finite/FinNat.v b/theories/Spaces/Finite/FinNat.v index d30db097c26..4b8cf5d39d5 100644 --- a/theories/Spaces/Finite/FinNat.v +++ b/theories/Spaces/Finite/FinNat.v @@ -36,7 +36,7 @@ Proof. Defined. Definition incl_finnat {n : nat} (u : FinNat n) : FinNat n.+1 - := (u.1; leq_trans u.2 (leq_S n n (leq_n n))). + := (u.1; leq_trans u.2 (leq_S n n (leq_refl n))). Lemma path_incl_finnat (n : nat) (u : FinNat n) (h : u.1 < n.+1) : incl_finnat u = (u.1; h). diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 2638d8a37ed..175cc18b7f6 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -36,7 +36,7 @@ Defined. (** This lemma is just for convenience in the case where the user forgets to unfold the definition of [<]. *) Proposition n_lt_Sn (n : nat) : n < n.+1. Proof. - exact (leq_n n.+1). + exact (leq_refl n.+1). Defined. Proposition leq_S' (n m : nat) : n.+1 <= m -> n <= m. @@ -109,7 +109,7 @@ Defined. Proposition n_leq_add_n_k' (n m : nat) : n <= m + n. Proof. simple_induction' m. - - exact(leq_n n). + - exact(leq_refl n). - simpl. apply leq_S. assumption. Defined. @@ -341,9 +341,9 @@ Proposition nataddsub_comm_ineq_lemma (n m : nat) Proof. revert m. simple_induction n n IHn. - - simple_induction m m IHm; [ apply leq_n | apply leq_S; apply leq_n ]. + - simple_induction m m IHm; exact _. - intro m; simple_induction m m IHm. - + apply leq_n. + + apply leq_refl. + apply IHn. Defined. @@ -403,7 +403,7 @@ Defined. Proposition predn_leq_n (n : nat) : nat_pred n <= n. Proof. - case n; [ apply leq_n | intro; apply leq_S; apply leq_n]. + destruct n; exact _. Defined. Proposition S_predn (n i: nat) : (i < n) -> S(nat_pred n) = n. @@ -498,9 +498,9 @@ Proof. simple_induction n n IHn. - intros; apply leq_zero. - destruct k. - + apply leq_n. + + apply leq_refl. + simpl; apply (@leq_trans _ n _); - [ apply IHn | apply leq_S, leq_n]. + [ apply IHn | apply leq_S, leq_refl]. Defined. Proposition sub_less_strict { n k : nat } @@ -586,7 +586,7 @@ Proof. + apply a; intros ? ?; now apply IHn. + now apply (IHn m), leq_succ. } - intro n. apply (X (n.+1) n), (leq_n n.+1). + intro n. apply (X (n.+1) n), (leq_refl n.+1). Defined. (** This inductive type is defined because it lets you loop from [i = 0] up to [i = n] by structural induction on a proof of [increasing_geq n 0]. With the existing [leq] type and the inductive structure of [n], it is easier and more natural to loop downwards from [i = n] to [i = 0], but harder to find the least natural number in the interval $[0,n]$ satisfying a given property. *) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 794caae5855..62d895608ad 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -142,7 +142,7 @@ Fixpoint factorial (n : nat) : nat (** *** Less Than or Equal To *) Inductive leq (n : nat) : nat -> Type0 := -| leq_n : leq n n +| leq_refl : leq n n | leq_S : forall m, leq n m -> leq n (S m). Scheme leq_ind := Induction for leq Sort Type. @@ -152,9 +152,7 @@ Scheme leq_rec := Induction for leq Sort Type. Notation "n <= m" := (leq n m) : nat_scope. Existing Class leq. -Global Existing Instances leq_n leq_S. - -Notation leq_refl := leq_n (only parsing). +Global Existing Instances leq_refl leq_S. (** *** Less Than *) @@ -400,7 +398,7 @@ Defined. (** *** Basic Properties of [leq] *) (** [leq] is reflexive by definition. *) -Global Instance reflexive_leq : Reflexive leq := leq_n. +Global Instance reflexive_leq : Reflexive leq := leq_refl. (** Being less than or equal to is a transitive relation. *) Definition leq_trans {x y z} : x <= y -> y <= z -> x <= z. @@ -454,7 +452,7 @@ Proof. Defined. (** A general form for injectivity of this constructor *) -Definition leq_n_inj_gen n k (p : n <= k) (r : n = k) : p = r # leq_n n. +Definition leq_refl_inj_gen n k (p : n <= k) (r : n = k) : p = r # leq_refl n. Proof. destruct p. + assert (c : idpath = r) by apply path_ishprop. @@ -465,8 +463,8 @@ Proof. Defined. (** Which we specialise to this lemma *) -Definition leq_n_inj n (p : n <= n) : p = leq_n n - := leq_n_inj_gen n n p idpath. +Definition leq_refl_inj n (p : n <= n) : p = leq_refl n + := leq_refl_inj_gen n n p idpath. Fixpoint leq_S_inj_gen n m k (p : n <= k) (q : n <= m) (r : m.+1 = k) : p = r # leq_S n m q. @@ -483,7 +481,7 @@ Proof. destruct t. cbn. apply ap. destruct q. - 1: apply leq_n_inj. + 1: apply leq_refl_inj. apply (leq_S_inj_gen n m _ p q idpath). Defined. @@ -496,7 +494,7 @@ Proof. intros p q; revert p. induction q. + intros y. - rapply leq_n_inj. + rapply leq_refl_inj. + intros y. rapply leq_S_inj. Defined. @@ -535,7 +533,7 @@ Global Instance decidable_lt n m : Decidable (lt n m) := _. (** *** Basic Properties of [ge] *) -Global Instance reflexive_ge : Reflexive ge := leq_n. +Global Instance reflexive_ge : Reflexive ge := leq_refl. Global Instance transitive_ge : Transitive ge := fun x y z p q => leq_trans q p. Global Instance ishprop_ge n m : IsHProp (ge n m) := _. Global Instance decidable_ge n m : Decidable (ge n m) := _. @@ -715,7 +713,7 @@ Defined. Fixpoint leq_add n m : n <= (m + n). Proof. destruct m. - 1: apply leq_n. + 1: apply leq_refl. apply leq_S, leq_add. Defined. From 88101abd4f12ff1d814808d55a9c3c5b553ae62f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 23:04:54 +0200 Subject: [PATCH 100/282] Nat: rename leq_S -> leq_succ_r Signed-off-by: Ali Caglayan --- theories/Algebra/AbGroups/AbelianGroup.v | 2 +- theories/Classes/implementations/ne_list.v | 2 +- theories/Spaces/Finite/FinNat.v | 4 ++-- theories/Spaces/List/Theory.v | 2 +- theories/Spaces/Nat/Arithmetic.v | 13 ++++++------ theories/Spaces/Nat/Core.v | 24 ++++++++++++---------- 6 files changed, 24 insertions(+), 23 deletions(-) diff --git a/theories/Algebra/AbGroups/AbelianGroup.v b/theories/Algebra/AbGroups/AbelianGroup.v index 20d56011d6d..4df356aabb3 100644 --- a/theories/Algebra/AbGroups/AbelianGroup.v +++ b/theories/Algebra/AbGroups/AbelianGroup.v @@ -303,7 +303,7 @@ Proof. - refine (f n _ + IHn _). intros k Hk. refine (f k _). - apply leq_S. + apply leq_succ_r. exact Hk. Defined. diff --git a/theories/Classes/implementations/ne_list.v b/theories/Classes/implementations/ne_list.v index 439e86f4610..d4183d4daee 100644 --- a/theories/Classes/implementations/ne_list.v +++ b/theories/Classes/implementations/ne_list.v @@ -130,7 +130,7 @@ Proof. + elim C. - intros [[] | C]. + exact _. - + by apply leq_S, IHy. + + by apply leq_succ_r, IHy. Qed. Fixpoint map {A B} (f: A → B) (l: ne_list A): ne_list B := diff --git a/theories/Spaces/Finite/FinNat.v b/theories/Spaces/Finite/FinNat.v index 4b8cf5d39d5..18e8434e664 100644 --- a/theories/Spaces/Finite/FinNat.v +++ b/theories/Spaces/Finite/FinNat.v @@ -36,7 +36,7 @@ Proof. Defined. Definition incl_finnat {n : nat} (u : FinNat n) : FinNat n.+1 - := (u.1; leq_trans u.2 (leq_S n n (leq_refl n))). + := (u.1; leq_trans u.2 (leq_succ_r (leq_refl n))). Lemma path_incl_finnat (n : nat) (u : FinNat n) (h : u.1 < n.+1) : incl_finnat u = (u.1; h). @@ -92,7 +92,7 @@ Proof. - destruct k as [k | []]. + apply (@leq_trans _ n _). * apply IHn. - * by apply leq_S. + * by apply leq_succ_r. + apply leq_refl. Defined. diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index 32763677a2a..b8fe4210f6c 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -764,7 +764,7 @@ Proof. { intros m. snrapply (functor_sigma idmap). intros k H. - exact (leq_S _ _ H). } + exact (leq_succ_r H). } induction n as [|n IHn]. 1: exact nil. nrefine ((n; _) :: list_map (f n) IHn). diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 175cc18b7f6..77e93d4e220 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -41,8 +41,7 @@ Defined. Proposition leq_S' (n m : nat) : n.+1 <= m -> n <= m. Proof. - intro l. - now apply leq_succ', leq_S. + intro l; apply leq_succ'; exact _. Defined. Proposition mixed_trans1 (n m k : nat) @@ -84,7 +83,7 @@ Defined. Proposition n_lt_m_n_leq_m { n m : nat } : n < m -> n <= m. Proof. - intro H. apply leq_S, leq_succ' in H; exact H. + intro H. apply leq_succ_r, leq_succ' in H; exact H. Defined. Proposition lt_trans {n m k : nat} : n < m -> m < k -> n < k. @@ -110,7 +109,7 @@ Proposition n_leq_add_n_k' (n m : nat) : n <= m + n. Proof. simple_induction' m. - exact(leq_refl n). - - simpl. apply leq_S. assumption. + - simpl. apply leq_succ_r. assumption. Defined. Proposition natineq0eq0 {n : nat} : n <= 0 -> n = 0. @@ -323,7 +322,7 @@ Proof. destruct (nat_add_succ_r n (m - k)). destruct (nataddsub_assoc_lemma l). apply (IHn m.+1 k). - apply leq_S. + apply leq_succ_r. assumption. Defined. @@ -500,7 +499,7 @@ Proof. - destruct k. + apply leq_refl. + simpl; apply (@leq_trans _ n _); - [ apply IHn | apply leq_S, leq_refl]. + [ apply IHn | apply leq_succ_r, leq_refl]. Defined. Proposition sub_less_strict { n k : nat } @@ -629,7 +628,7 @@ Proof. - destruct (symmetric_paths _ _ (nat_sub_zero_r n)); constructor. - destruct (@leq_dichot n k) as [l | g]. + destruct (symmetric_paths _ _ (nat_sub_leq _)) in IHk. - apply leq_S in l. + apply leq_succ_r in l. destruct (symmetric_paths _ _ (nat_sub_leq _)). exact IHk. + change k.+1 with (1 + k). destruct (nat_add_comm k 1). destruct (symmetric_paths _ _ (nat_sub_add n k 1)). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 62d895608ad..b7a47ae5462 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -143,7 +143,9 @@ Fixpoint factorial (n : nat) : nat Inductive leq (n : nat) : nat -> Type0 := | leq_refl : leq n n -| leq_S : forall m, leq n m -> leq n (S m). +| leq_succ_r m : leq n m -> leq n (S m). + +Arguments leq_succ_r {n m} _. Scheme leq_ind := Induction for leq Sort Type. Scheme leq_rect := Induction for leq Sort Type. @@ -152,7 +154,7 @@ Scheme leq_rec := Induction for leq Sort Type. Notation "n <= m" := (leq n m) : nat_scope. Existing Class leq. -Global Existing Instances leq_refl leq_S. +Global Existing Instances leq_refl leq_succ_r. (** *** Less Than *) @@ -466,8 +468,8 @@ Defined. Definition leq_refl_inj n (p : n <= n) : p = leq_refl n := leq_refl_inj_gen n n p idpath. -Fixpoint leq_S_inj_gen n m k (p : n <= k) (q : n <= m) (r : m.+1 = k) - : p = r # leq_S n m q. +Fixpoint leq_succ_r_inj_gen n m k (p : n <= k) (q : n <= m) (r : m.+1 = k) + : p = r # leq_succ_r q. Proof. revert m q r. destruct p. @@ -482,11 +484,11 @@ Proof. cbn. apply ap. destruct q. 1: apply leq_refl_inj. - apply (leq_S_inj_gen n m _ p q idpath). + apply (leq_succ_r_inj_gen n m _ p q idpath). Defined. -Definition leq_S_inj n m (p : n <= m.+1) (q : n <= m) : p = leq_S n m q - := leq_S_inj_gen n m m.+1 p q idpath. +Definition leq_succ_r_inj n m (p : n <= m.+1) (q : n <= m) : p = leq_succ_r q + := leq_succ_r_inj_gen n m m.+1 p q idpath. Global Instance ishprop_leq n m : IsHProp (n <= m). Proof. @@ -496,10 +498,10 @@ Proof. + intros y. rapply leq_refl_inj. + intros y. - rapply leq_S_inj. + rapply leq_succ_r_inj. Defined. -Definition equiv_leq_S_n n m : n.+1 <= m.+1 <~> n <= m. +Definition equiv_leq_succ n m : n.+1 <= m.+1 <~> n <= m. Proof. srapply equiv_iff_hprop. apply leq_succ'. @@ -516,7 +518,7 @@ Proof. + left; exact _. + rapply decidable_equiv'. symmetry. - apply equiv_leq_S_n. + apply equiv_leq_succ. Defined. (** *** Basic Properties of [lt] *) @@ -714,7 +716,7 @@ Fixpoint leq_add n m : n <= (m + n). Proof. destruct m. 1: apply leq_refl. - apply leq_S, leq_add. + apply leq_succ_r, leq_add. Defined. (** [<=] is an antisymmetric relation. *) From 79f3675a63e8784ff3d696033e52c34b61de34f5 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 23:07:45 +0200 Subject: [PATCH 101/282] Nat: clarify notation in comments Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index b7a47ae5462..dc800917e0a 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -27,7 +27,7 @@ Notation nat_iter n f x (** *** Successor and predecessor *) -(** [nat_succ n] is the successor of a natural number [n]. *) +(** [nat_succ n] is the successor of a natural number [n]. We defined a notation [x.+1] for it in Overture.v. *) Notation nat_succ := S (only parsing). (** [nat_pred n] is the predecessor of a natural number [n]. When [n] is [0] we return [0]. *) From bef190f504c17586b28f411562d8817d51b84d4e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 23:10:32 +0200 Subject: [PATCH 102/282] Nat: rename ge -> geq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 28 ++++++++++++++-------------- 1 file changed, 14 insertions(+), 14 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index dc800917e0a..5a5ebd43bfc 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -167,7 +167,15 @@ Existing Class lt. Infix "<" := lt : nat_scope. Global Instance lt_is_leq n m : leq n.+1 m -> lt n m | 100 := idmap. -(*** Greater Than or Equal To *) +(** *** Greater Than or Equal To *) + +Definition geq n m := leq m n. +Existing Class geq. +#[export] Hint Unfold geq : typeclass_instances. +Infix ">=" := geq : nat_scope. +Global Instance geq_is_leq n m : leq m n -> geq n m | 100 := idmap. + +(*** Greater Than *) Definition gt n m := lt m n. Existing Class gt. @@ -175,14 +183,6 @@ Existing Class gt. Infix ">" := gt : nat_scope. Global Instance gt_is_leq n m : leq m.+1 n -> gt n m | 100 := idmap. -(** *** Greater Than *) - -Definition ge n m := leq m n. -Existing Class ge. -#[export] Hint Unfold ge : typeclass_instances. -Infix ">=" := ge : nat_scope. -Global Instance ge_is_leq n m : leq m n -> ge n m | 100 := idmap. - (** *** Combined Comparison Predicates *) Notation "x <= y <= z" := (x <= y /\ y <= z) : nat_scope. @@ -533,12 +533,12 @@ Defined. Global Instance ishprop_lt n m : IsHProp (n < m) := _. Global Instance decidable_lt n m : Decidable (lt n m) := _. -(** *** Basic Properties of [ge] *) +(** *** Basic Properties of [geq] *) -Global Instance reflexive_ge : Reflexive ge := leq_refl. -Global Instance transitive_ge : Transitive ge := fun x y z p q => leq_trans q p. -Global Instance ishprop_ge n m : IsHProp (ge n m) := _. -Global Instance decidable_ge n m : Decidable (ge n m) := _. +Global Instance reflexive_geq : Reflexive geq := leq_refl. +Global Instance transitive_geq : Transitive geq := fun x y z p q => leq_trans q p. +Global Instance ishprop_geq n m : IsHProp (geq n m) := _. +Global Instance decidable_geq n m : Decidable (geq n m) := _. (** *** Basic Properties of [gt] *) From dcab26ca910b4a2fb637ec66ed3291801d9de467 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 28 Jul 2024 23:16:21 +0200 Subject: [PATCH 103/282] Nat: comments about things that need to be renamed Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 11 +++++++++++ 1 file changed, 11 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 5a5ebd43bfc..86a02cf0acd 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -436,6 +436,7 @@ Global Existing Instance leq_succ | 100. (** The converse to [leq_succ] also holds. *) Definition leq_succ' {n m} : n.+1 <= m.+1 -> n <= m := leq_pred. +(** TODO: rename *) (** TODO: use lemmas about negating predicate *) Lemma not_leq_Sn_n n : ~ (n.+1 <= n). Proof. @@ -446,6 +447,7 @@ Proof. by apply IHn, leq_succ'. Defined. +(** TODO: rename *) Lemma not_leq_Sn_0 n : ~ (n.+1 <= 0). Proof. intros p. @@ -712,6 +714,7 @@ Defined. (** ** More Theory of Comparison Predicates *) +(** TODO: rename to [leq_add_r'] *) Fixpoint leq_add n m : n <= (m + n). Proof. destruct m. @@ -732,13 +735,16 @@ Defined. Global Instance antisymmetric_leq : AntiSymmetric leq := @leq_antisym. +(** TODO: rename to [lt_irrefl] *) (** [<] is an irreflexive relation. *) Definition not_lt_n_n n : ~ (n < n) := not_leq_Sn_n n. Global Instance irreflexive_lt : Irreflexive lt := not_lt_n_n. +(** TODO: rename *) (** [1] is less than or equal to the successor of any number. *) Definition leq_1_Sn {n} : 1 <= n.+1 := leq_succ (leq_zero _). +(** TODO: rename *) Fixpoint leq_dichot {m} {n} : (m <= n) + (m > n). Proof. simple_induction' m; simple_induction' n. @@ -751,29 +757,34 @@ Proof. + right; apply leq_succ; assumption. Defined. +(** TODO: rename *) Lemma not_lt_n_0 n : ~ (n < 0). Proof. apply not_leq_Sn_0. Defined. +(** TODO: rename *) Definition not_lt_implies_geq {n m : nat} : ~(n < m) -> m <= n. Proof. intros not_lt. destruct (@leq_dichot m n); [ assumption | contradiction]. Defined. +(** TODO: rename *) Definition not_leq_implies_gt {n m : nat} : ~(n <= m) -> m < n. Proof. intros not_leq. destruct (@leq_dichot n m); [ contradiction | assumption]. Defined. +(** TODO: rename *) Definition lt_implies_not_geq {n m : nat} : (n < m) -> ~(m <= n). Proof. intros ineq1 ineq2. contradiction (not_lt_n_n n). by apply (leq_trans ineq1). Defined. +(** TODO: rename *) Definition leq_implies_not_gt {n m : nat} : (n <= m) -> ~(m < n). Proof. intros ineq1 ineq2. From e51a309635a4a11dec383e3301af38aa09b53911 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 02:14:12 +0200 Subject: [PATCH 104/282] Nat: replace nat_exp with nat_pow Signed-off-by: Ali Caglayan --- theories/Spaces/Finite/Finite.v | 9 +++++---- theories/Spaces/Nat/Core.v | 11 +---------- 2 files changed, 6 insertions(+), 14 deletions(-) diff --git a/theories/Spaces/Finite/Finite.v b/theories/Spaces/Finite/Finite.v index 928a9df8587..8f140a9afc2 100644 --- a/theories/Spaces/Finite/Finite.v +++ b/theories/Spaces/Finite/Finite.v @@ -333,20 +333,21 @@ Defined. #[local] Hint Extern 4 => progress (cbv beta iota) : typeclass_instances. Definition fcard_arrow `{Funext} X Y `{Finite X} `{Finite Y} -: fcard (X -> Y) = nat_exp (fcard Y) (fcard X). +: fcard (X -> Y) = nat_pow (fcard Y) (fcard X). Proof. assert (e := merely_equiv_fin X). strip_truncations. refine (fcard_equiv (functor_arrow e idmap)^-1 @ _). - refine (_ @ ap (fun x => nat_exp (fcard Y) x) (fcard_equiv e)). + refine (_ @ ap (fun x => nat_pow (fcard Y) x) (fcard_equiv e)). generalize (fcard X); intros n. induction n as [|n IH]. - reflexivity. - refine (fcard_equiv (equiv_sum_ind (fun (_:Fin n.+1) => Y))^-1 @ _). refine (fcard_prod _ _ @ _). + lhs nrapply nat_mul_comm. apply (ap011 nat_mul). - + assumption. + refine (fcard_equiv (@Unit_ind (fun (_:Unit) => Y))^-1). + + assumption. Defined. (** [fcard] still computes, despite the funext: *) @@ -487,7 +488,7 @@ Defined. (** The product of a finite constant family is the exponential by its cardinality. *) Definition finmult_const `{Funext} X `{Finite X} n -: finmult (fun x:X => n) = nat_exp n (fcard X). +: finmult (fun x:X => n) = nat_pow n (fcard X). Proof. refine (fcard_arrow X (Fin n)). Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 86a02cf0acd..b994f3109c4 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -61,16 +61,7 @@ Fixpoint nat_sub n m : nat := Notation "n - m" := (nat_sub n m) : nat_scope. (** Powers or exponentiation of natural numbers. *) -Definition nat_pow n m := - nat_iter m (nat_mul n) 1. - -(** TODO: merge witth nat_pow (order of arguments needs adjusting). *) -(** Exponentiation *) -Fixpoint nat_exp (n m : nat) : nat - := match m with - | 0 => 1 - | S m => nat_exp n m * n - end. +Definition nat_pow n m := nat_iter m (nat_mul n) 1. (** *** Maximum and minimum *) From a2b36a11526c775b1a227fdf522d63260420bf1a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 02:38:49 +0200 Subject: [PATCH 105/282] Nat: add properties about powers Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 77 +++++++++++++++++++++++++++++++++++++- 1 file changed, 75 insertions(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index b994f3109c4..97451f3675b 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -372,8 +372,9 @@ Defined. (** Multiplication of natural numbers distributes over addition on the right. *) Definition nat_dist_r@{} n m k : (n + m) * k = n * k + m * k. Proof. - rewrite 3 (nat_mul_comm _ k). - nrapply nat_dist_l. + lhs nrapply nat_mul_comm. + lhs nrapply nat_dist_l. + nrapply ap011; nrapply nat_mul_comm. Defined. (** Multiplication of natural numbers is associative. *) @@ -386,6 +387,14 @@ Proof. exact IHn. Defined. +(** Multiplication by [1] on the left is the identity. *) +Definition nat_mul_one_l@{} n : 1 * n = n + := nat_add_zero_r _. + +(** Multiplication by [1] on the right is the identity. *) +Definition nat_mul_one_r@{} n : n * 1 = n + := nat_mul_comm _ _ @ nat_mul_one_l _. + (** ** Basic Properties of Comparison Predicates *) (** *** Basic Properties of [leq] *) @@ -797,3 +806,67 @@ Proof. 2: exact (nat_add_succ_r _ _)^. reflexivity. Defined. + +(** ** Properties of Powers *) + +(** [0] to any power is [0] unless that power is [0] in which case it is [1]. *) +Definition nat_pow_zero_l@{} n : nat_pow 0 n = if dec (n = 0) then 1 else 0. +Proof. + destruct n; reflexivity. +Defined. + +(** Any number to the power of [0] is [1]. *) +Definition nat_pow_zero_r@{} n : nat_pow n 0 = 1 + := idpath. + +(** [1] to any power is [1]. *) +Definition nat_pow_one_l@{} n : nat_pow 1 n = 1. +Proof. + induction n as [|n IHn]; simpl. + 1: reflexivity. + lhs nrapply nat_add_zero_r. + exact IHn. +Defined. + +(** Any number to the power of [1] is itself. *) +Definition nat_pow_one_r@{} n : nat_pow n 1 = n + := nat_mul_one_r _. + +(** Exponentiation of natural numbers is distributive over addition on the left. *) +Definition nat_pow_add_r@{} n m k + : nat_pow n (m + k) = nat_pow n m * nat_pow n k. +Proof. + induction m as [|m IHm]; simpl. + - symmetry. + apply nat_add_zero_r. + - rhs_V nrapply nat_mul_assoc. + exact (ap _ IHm). +Defined. + +(** Exponentiation of natural numbers is distributive over multiplication on the right. *) +Definition nat_pow_mul_l@{} n m k + : nat_pow (n * m) k = nat_pow n k * nat_pow m k. +Proof. + induction k as [|k IHk]; simpl. + 1: reflexivity. + lhs_V nrapply nat_mul_assoc. + rhs_V nrapply nat_mul_assoc. + nrapply ap. + rhs nrapply nat_mul_comm. + rhs_V nrapply nat_mul_assoc. + nrapply ap. + rhs nrapply nat_mul_comm. + exact IHk. +Defined. + +(** Exponentiation of natural numbers is distributive over multiplication on the left. *) +Definition nat_pow_mul_r@{} n m k + : nat_pow n (m * k) = nat_pow (nat_pow n m) k. +Proof. + induction m as [|m IHm]; simpl. + - exact (nat_pow_one_l _)^. + - lhs nrapply nat_pow_add_r. + rhs nrapply nat_pow_mul_l. + nrapply ap. + exact IHm. +Defined. From eaa7c93c39d85d4a7505d28084aa6ebf3489126e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 13:33:48 +0200 Subject: [PATCH 106/282] Nat: remove n_lt_Sn Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 8 ++------ 1 file changed, 2 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 77e93d4e220..312bca5abaa 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -6,6 +6,8 @@ Local Set Universe Minimization ToSet. Local Close Scope trunc_scope. Local Open Scope nat_scope. +(** TODO: The results in this file are in the process of being moved over to Core.v *) + (** Because of the inductive definition of [<=], one can destruct the proof of [n <= m] and get a judgemental identification between [n] and [m] rather than a propositional one, which may be preferable to the following lemma. *) Proposition leq_split {n m : nat} : (n <= m) -> sum (n < m) (n = m). Proof. @@ -33,12 +35,6 @@ Proof. contradiction (not_lt_n_n m). Defined. -(** This lemma is just for convenience in the case where the user forgets to unfold the definition of [<]. *) -Proposition n_lt_Sn (n : nat) : n < n.+1. -Proof. - exact (leq_refl n.+1). -Defined. - Proposition leq_S' (n m : nat) : n.+1 <= m -> n <= m. Proof. intro l; apply leq_succ'; exact _. From f4babd7701b9abd2765c12e709d2863920795a2a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 13:38:05 +0200 Subject: [PATCH 107/282] Nat: move leq_S' to Core.v Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 7 +------ theories/Spaces/Nat/Core.v | 7 +++++++ 2 files changed, 8 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 312bca5abaa..5458471966d 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -35,11 +35,6 @@ Proof. contradiction (not_lt_n_n m). Defined. -Proposition leq_S' (n m : nat) : n.+1 <= m -> n <= m. -Proof. - intro l; apply leq_succ'; exact _. -Defined. - Proposition mixed_trans1 (n m k : nat) : n <= m -> m < k -> n < k. Proof. @@ -675,7 +670,7 @@ Proof. - intro a. induction a. + constructor. - + exact (leq_S' _ _ _). + + exact (leq_S' _). Defined. (** This tautology accepts a (potentially opaqued or QED'ed) proof of [n <= m], and returns a transparent proof which can be computed with (i.e., one can loop from n to m) *) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 97451f3675b..e73efd27c65 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -714,6 +714,13 @@ Defined. (** ** More Theory of Comparison Predicates *) +(** TODO: rename leq_S' -> leq_succ_l *) +(** [n.+1 <= m] implies [n <= m]. *) +Definition leq_S' {n m} : n.+1 <= m -> n <= m. +Proof. + intro l; apply leq_succ'; exact _. +Defined. + (** TODO: rename to [leq_add_r'] *) Fixpoint leq_add n m : n <= (m + n). Proof. From c418638f325510d1fa4fa3aaf677af1add31565b Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 13:38:28 +0200 Subject: [PATCH 108/282] Nat: leq_S' -> leq_succ_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 4 ++-- theories/Spaces/Nat/Core.v | 3 +-- 2 files changed, 3 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 5458471966d..e4a206af247 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -657,7 +657,7 @@ Proof. exact sub_less). destruct (symmetric_paths _ _ (ineq_sub' _ _ _)). destruct (symmetric_paths _ _ (S_predn _ 0 _)). - apply (ap S), IHn, leq_S', ineq. + apply (ap S), IHn, leq_succ_l, ineq. Defined. Proposition leq_equivalent (n m : nat) @@ -670,7 +670,7 @@ Proof. - intro a. induction a. + constructor. - + exact (leq_S' _). + + exact (leq_succ_l _). Defined. (** This tautology accepts a (potentially opaqued or QED'ed) proof of [n <= m], and returns a transparent proof which can be computed with (i.e., one can loop from n to m) *) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index e73efd27c65..4a90683ee10 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -714,9 +714,8 @@ Defined. (** ** More Theory of Comparison Predicates *) -(** TODO: rename leq_S' -> leq_succ_l *) (** [n.+1 <= m] implies [n <= m]. *) -Definition leq_S' {n m} : n.+1 <= m -> n <= m. +Definition leq_succ_l {n m} : n.+1 <= m -> n <= m. Proof. intro l; apply leq_succ'; exact _. Defined. From 44b646b8a5fdb02f5792999268ba628ee8b8eb58 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 13:53:03 +0200 Subject: [PATCH 109/282] Nat: move leq_split and strengthen to an equivalence Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 7 ------- theories/Spaces/Nat/Core.v | 21 +++++++++++++++++++++ 2 files changed, 21 insertions(+), 7 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index e4a206af247..4f66437e541 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,13 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** Because of the inductive definition of [<=], one can destruct the proof of [n <= m] and get a judgemental identification between [n] and [m] rather than a propositional one, which may be preferable to the following lemma. *) -Proposition leq_split {n m : nat} : (n <= m) -> sum (n < m) (n = m). -Proof. - intro l. induction l. - - now right. - - left. exact (leq_succ l). -Defined. Proposition diseq_implies_lt (n m : nat) : n <> m -> (n < m) + (n > m). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 4a90683ee10..382f61eb7c7 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -797,6 +797,27 @@ Proof. contradiction (not_lt_n_n n); by refine (leq_trans _ ineq2). Defined. +(** TODO: rename [equiv_leq_lt_or_eq] *) +(** [n <= m] is equivalent to [(n < m) + (n = m)]. Note that it is not immediately obvious that the latter type is a hprop, hence we have to explicitly show the back and forth maps are inverses of eachother. This is possible and justifies the name "less than or equal to". *) +Definition leq_split {n m : nat} : (n <= m) <~> ((n < m) + (n = m))%type. +Proof. + snrapply equiv_adjointify. + - intro l; induction l. + + now right. + + left; exact (leq_succ l). + - intros [l|p]. + + exact (leq_succ_l l). + + destruct p. + exact (leq_refl _). + - intros [l|p]. + + induction l. + 1: reflexivity. + snrapply (ap (inl)). + rapply path_ishprop. + + by destruct p. + - intro; rapply path_ishprop. +Defined. + (** ** Arithmetic relations between [trunc_index] and [nat]. *) Lemma trunc_index_add_nat_add (n : nat) From 6a859e8691dfa6f8ca9786c9e1302867e782d28f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 13:55:16 +0200 Subject: [PATCH 110/282] Nat: rename leq_split -> equiv_leq_lt_or_eq We take this oppurtunity to improve a proof in Theory.v too. Signed-off-by: Ali Caglayan --- theories/Spaces/List/Theory.v | 7 ++----- theories/Spaces/Nat/Core.v | 3 +-- 2 files changed, 3 insertions(+), 7 deletions(-) diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index b8fe4210f6c..e9ee1c6bf37 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -839,11 +839,8 @@ Proof. 1: by rewrite length_seq. by apply IHn. - apply not_lt_implies_geq in H'. - destruct (leq_split H') as [H'' | H'']. - { apply lt_implies_not_geq in H''. - apply leq_succ' in H. - contradiction. } - destruct H''. + apply leq_succ' in H. + destruct (leq_antisym H H'). lhs nrapply nth_last. { rewrite length_app. rewrite nat_add_comm. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 382f61eb7c7..2b99ced5ad1 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -797,9 +797,8 @@ Proof. contradiction (not_lt_n_n n); by refine (leq_trans _ ineq2). Defined. -(** TODO: rename [equiv_leq_lt_or_eq] *) (** [n <= m] is equivalent to [(n < m) + (n = m)]. Note that it is not immediately obvious that the latter type is a hprop, hence we have to explicitly show the back and forth maps are inverses of eachother. This is possible and justifies the name "less than or equal to". *) -Definition leq_split {n m : nat} : (n <= m) <~> ((n < m) + (n = m))%type. +Definition equiv_leq_lt_or_eq {n m : nat} : (n <= m) <~> ((n < m) + (n = m))%type. Proof. snrapply equiv_adjointify. - intro l; induction l. From 0417819e88826f3339a28697073eb84bb7649cd5 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 15:42:56 +0200 Subject: [PATCH 111/282] Nat: move diseq_implies_lt to Core.v Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 12 ------------ theories/Spaces/Nat/Core.v | 17 ++++++++++++++++- 2 files changed, 16 insertions(+), 13 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 4f66437e541..afe667c4292 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -9,18 +9,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -Proposition diseq_implies_lt (n m : nat) - : n <> m -> (n < m) + (n > m). -Proof. - intros diseq. - destruct (dec (n < m)) as [| a]; [ now left |]. - right. destruct (@leq_dichot n m) as [n_leq_m | gt]; - [ | assumption]. - destruct n_leq_m; - [ now contradiction diseq - | contradiction a; now apply leq_succ]. -Defined. - Proposition lt_implies_diseq (n m : nat) : n < m -> (n <> m). Proof. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 2b99ced5ad1..aea6e83f7f3 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -798,7 +798,7 @@ Proof. Defined. (** [n <= m] is equivalent to [(n < m) + (n = m)]. Note that it is not immediately obvious that the latter type is a hprop, hence we have to explicitly show the back and forth maps are inverses of eachother. This is possible and justifies the name "less than or equal to". *) -Definition equiv_leq_lt_or_eq {n m : nat} : (n <= m) <~> ((n < m) + (n = m))%type. +Definition equiv_leq_lt_or_eq {n m : nat} : (n <= m) <~> (n < m) + (n = m). Proof. snrapply equiv_adjointify. - intro l; induction l. @@ -817,6 +817,21 @@ Proof. - intro; rapply path_ishprop. Defined. +(** TODO: rename diseq_implies_lt -> neq_iff_lt_or_gt *) +Definition diseq_implies_lt {n m} : n <> m <-> (n < m) + (n > m). +Proof. + split. + - intros diseq. + destruct (dec (n < m)) as [| a]; [ now left |]. + apply not_lt_implies_geq in a. + apply equiv_leq_lt_or_eq in a. + destruct a as [lt | eq]. + 1: by right. + symmetry in eq. + contradiction. + - intros [H' | H'] nq; destruct nq; exact (not_lt_n_n _ H'). +Defined. + (** ** Arithmetic relations between [trunc_index] and [nat]. *) Lemma trunc_index_add_nat_add (n : nat) From f692badd236599499892d3b4e433a1d512edd695 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 15:43:26 +0200 Subject: [PATCH 112/282] Nat: rename diseq_implies_lt -> neq_iff_lt_or_gt Signed-off-by: Ali Caglayan --- theories/Algebra/Rings/KroneckerDelta.v | 2 +- theories/Spaces/Nat/Core.v | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/theories/Algebra/Rings/KroneckerDelta.v b/theories/Algebra/Rings/KroneckerDelta.v index 788a392ca48..b78e6822136 100644 --- a/theories/Algebra/Rings/KroneckerDelta.v +++ b/theories/Algebra/Rings/KroneckerDelta.v @@ -110,7 +110,7 @@ Proof. apply rng_mult_zero_l. - simpl; lhs nrapply ap. + nrapply IHn. - apply diseq_implies_lt in p. + apply neq_iff_lt_or_gt in p. destruct p; [assumption|]. contradiction (lt_implies_not_geq Hi). + rewrite (kronecker_delta_neq p). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index aea6e83f7f3..404d563dea4 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -817,8 +817,8 @@ Proof. - intro; rapply path_ishprop. Defined. -(** TODO: rename diseq_implies_lt -> neq_iff_lt_or_gt *) -Definition diseq_implies_lt {n m} : n <> m <-> (n < m) + (n > m). +(** The inequality of natural numbers is equivalent to [n < m] or [n > m]. This could be an equivalence however one of the sections requires funext since we are comparing two inequality proofs. It is therefore more useful to keep it as a biimplication. *) +Definition neq_iff_lt_or_gt {n m} : n <> m <-> (n < m) + (n > m). Proof. split. - intros diseq. From 26838bc8bcd77e287217368d18148f83e034158b Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 15:47:15 +0200 Subject: [PATCH 113/282] Nat: remove lt_implies_diseq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 7 ------- 1 file changed, 7 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index afe667c4292..cd5cf15b3df 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -9,13 +9,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -Proposition lt_implies_diseq (n m : nat) - : n < m -> (n <> m). -Proof. - intros ineq eq. rewrite eq in ineq. - contradiction (not_lt_n_n m). -Defined. - Proposition mixed_trans1 (n m k : nat) : n <= m -> m < k -> n < k. Proof. From 1cafaa56610d2bd1a3f78c9dd2ce55711e3fa892 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 15:52:02 +0200 Subject: [PATCH 114/282] Nat: move mixed_trans* to Core.v Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 18 ++---------------- theories/Spaces/Nat/Core.v | 8 ++++++++ 2 files changed, 10 insertions(+), 16 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index cd5cf15b3df..9d9f720a771 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -9,19 +9,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -Proposition mixed_trans1 (n m k : nat) - : n <= m -> m < k -> n < k. -Proof. - intros l j. apply leq_succ in l. - apply (@leq_trans (n.+1) (m.+1) k); trivial. -Defined. - -Proposition mixed_trans2 (n m k : nat) - : n < m -> m <= k -> n < k. -Proof. - intros l j. apply (@leq_trans (n.+1) m k); trivial. -Defined. - Proposition add_n_sub_n_eq (m n : nat) : m + n - n = m. Proof. @@ -55,7 +42,7 @@ Proposition lt_trans {n m k : nat} : n < m -> m < k -> n < k. Proof. intros H1 H2. nrapply n_lt_m_n_leq_m. - exact (mixed_trans1 _ _ _ H1 H2). + exact (mixed_trans1 H1 H2). Defined. Proposition not_both_less (n m : nat) : n < m -> ~(m < n). @@ -513,8 +500,7 @@ Proposition natpmswap4 (k m n : nat) Proof. intro l; apply (@nataddpreserveslt (k - n) m n) in l. destruct (nat_add_comm m n). - now apply (mixed_trans1 k (k - n + n) (m + n) - (nat_sub_add_ineq _ _)). + now rapply (mixed_trans1 (nat_sub_add_ineq _ _)). Defined. Proposition n_leq_m_n_leq_plus_m_k (n m k : nat) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 404d563dea4..eaf4ae354d9 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -832,6 +832,14 @@ Proof. - intros [H' | H'] nq; destruct nq; exact (not_lt_n_n _ H'). Defined. +(** TODO: rename mixed_trans1 -> leq_lt_trans *) +Definition mixed_trans1 {n m k} : n <= m -> m < k -> n < k + := fun leq lt => leq_trans (leq_succ leq) lt. + +(** TODO: rename mixed_trans2 -> lt_leq_trans *) +Definition mixed_trans2 {n m k} : n < m -> m <= k -> n < k + := fun lt leq => leq_trans lt leq. + (** ** Arithmetic relations between [trunc_index] and [nat]. *) Lemma trunc_index_add_nat_add (n : nat) From 43943878285d2960f4be2b31db556f7a0cecaddd Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 15:52:42 +0200 Subject: [PATCH 115/282] Nat: rename mixed_trans1 -> leq_lt_trans, mixed_trans2 -> lt_leq_trans Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 8 ++++---- theories/Spaces/Nat/Core.v | 6 ++---- 2 files changed, 6 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 9d9f720a771..f4ceb8c6094 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -42,7 +42,7 @@ Proposition lt_trans {n m k : nat} : n < m -> m < k -> n < k. Proof. intros H1 H2. nrapply n_lt_m_n_leq_m. - exact (mixed_trans1 H1 H2). + exact (leq_lt_trans H1 H2). Defined. Proposition not_both_less (n m : nat) : n < m -> ~(m < n). @@ -411,7 +411,7 @@ Proposition pred_gt_implies_lt (i j : nat) Proof. intros ineq. assert (H := leq_succ ineq). assert (i < j) as X. { - apply (@mixed_trans2 _ (nat_pred j) _); + apply (@lt_leq_trans _ (nat_pred j) _); [assumption | apply predn_leq_n]. } destruct (symmetric_paths _ _ (S_predn _ _ X)) in H. @@ -500,7 +500,7 @@ Proposition natpmswap4 (k m n : nat) Proof. intro l; apply (@nataddpreserveslt (k - n) m n) in l. destruct (nat_add_comm m n). - now rapply (mixed_trans1 (nat_sub_add_ineq _ _)). + now rapply (leq_lt_trans (nat_sub_add_ineq _ _)). Defined. Proposition n_leq_m_n_leq_plus_m_k (n m k : nat) @@ -639,7 +639,7 @@ Proof. intro ineq. destruct (@leq_dichot n m) as [l | g]. - exact l. - - contradiction (not_lt_n_n m (@mixed_trans2 _ _ _ g ineq)). + - contradiction (not_lt_n_n m (lt_leq_trans g ineq)). Defined. Proposition symmetric_rel_total_order (R : nat -> nat -> Type) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index eaf4ae354d9..a3b56f590d7 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -832,12 +832,10 @@ Proof. - intros [H' | H'] nq; destruct nq; exact (not_lt_n_n _ H'). Defined. -(** TODO: rename mixed_trans1 -> leq_lt_trans *) -Definition mixed_trans1 {n m k} : n <= m -> m < k -> n < k +Definition leq_lt_trans {n m k} : n <= m -> m < k -> n < k := fun leq lt => leq_trans (leq_succ leq) lt. -(** TODO: rename mixed_trans2 -> lt_leq_trans *) -Definition mixed_trans2 {n m k} : n < m -> m <= k -> n < k +Definition lt_leq_trans {n m k} : n < m -> m <= k -> n < k := fun lt leq => leq_trans lt leq. (** ** Arithmetic relations between [trunc_index] and [nat]. *) From d81de85f360a9e61605336e021a3ccda32bbac36 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 16:06:43 +0200 Subject: [PATCH 116/282] Nat: move and rename mixed transitivity results about < and <= Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 16 ++------------ theories/Spaces/Nat/Core.v | 38 +++++++++++++++++--------------- 2 files changed, 22 insertions(+), 32 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index f4ceb8c6094..33e7a5cfd0b 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -33,18 +33,6 @@ Proof. apply add_n_sub_n_eq. Defined. -Proposition n_lt_m_n_leq_m { n m : nat } : n < m -> n <= m. -Proof. - intro H. apply leq_succ_r, leq_succ' in H; exact H. -Defined. - -Proposition lt_trans {n m k : nat} : n < m -> m < k -> n < k. -Proof. - intros H1 H2. - nrapply n_lt_m_n_leq_m. - exact (leq_lt_trans H1 H2). -Defined. - Proposition not_both_less (n m : nat) : n < m -> ~(m < n). Proof. intros l a; contradiction (not_lt_n_n _ (lt_trans l a)). @@ -133,7 +121,7 @@ Proof. destruct (@leq_dichot m n) as [l | g]. - destruct (symmetric_paths _ _ (natminuspluseq _ _ l)); constructor. - - apply n_lt_m_n_leq_m in g. + - apply leq_lt in g. now destruct (symmetric_paths _ _ (nat_sub_leq _)). Defined. @@ -315,7 +303,7 @@ Proof. - destruct (symmetric_paths _ _ (nataddsub_comm _ _ m l)). destruct (symmetric_paths _ _ (add_n_sub_n_eq n m)). apply leq_refl; done. - - apply n_lt_m_n_leq_m in gt. + - apply leq_lt in gt. destruct (symmetric_paths _ _ (nat_sub_leq _)). assumption. Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index a3b56f590d7..88c2751cb07 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -523,15 +523,29 @@ Proof. apply equiv_leq_succ. Defined. -(** *** Basic Properties of [lt] *) - -Global Instance transitive_lt : Transitive lt. +(** [n.+1 <= m] implies [n <= m]. *) +Definition leq_succ_l {n m} : n.+1 <= m -> n <= m. Proof. - hnf; unfold lt in *. - intros x y z p q. - rapply leq_trans. + intro l; apply leq_succ'; exact _. Defined. +(** *** Basic Properties of [lt] *) + +(** [<=] and [<] imply [<] *) +Definition leq_lt_trans {n m k} : n <= m -> m < k -> n < k + := fun leq lt => leq_trans (leq_succ leq) lt. + +(** [<=] and [<] imply [<] *) +Definition lt_leq_trans {n m k} : n < m -> m <= k -> n < k + := fun lt leq => leq_trans lt leq. + +Definition leq_lt {n m} : n < m -> n <= m + := leq_succ_l. + +Definition lt_trans {n m k} : n < m -> m < k -> n < k + := fun H1 H2 => leq_lt (leq_lt_trans H1 H2). + +Global Instance transitive_lt : Transitive lt := @lt_trans. Global Instance ishprop_lt n m : IsHProp (n < m) := _. Global Instance decidable_lt n m : Decidable (lt n m) := _. @@ -714,12 +728,6 @@ Defined. (** ** More Theory of Comparison Predicates *) -(** [n.+1 <= m] implies [n <= m]. *) -Definition leq_succ_l {n m} : n.+1 <= m -> n <= m. -Proof. - intro l; apply leq_succ'; exact _. -Defined. - (** TODO: rename to [leq_add_r'] *) Fixpoint leq_add n m : n <= (m + n). Proof. @@ -832,12 +840,6 @@ Proof. - intros [H' | H'] nq; destruct nq; exact (not_lt_n_n _ H'). Defined. -Definition leq_lt_trans {n m k} : n <= m -> m < k -> n < k - := fun leq lt => leq_trans (leq_succ leq) lt. - -Definition lt_leq_trans {n m k} : n < m -> m <= k -> n < k - := fun lt leq => leq_trans lt leq. - (** ** Arithmetic relations between [trunc_index] and [nat]. *) Lemma trunc_index_add_nat_add (n : nat) From 379c5a07f1df9ec81d045006b6d86af5af2b08a7 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 16:17:46 +0200 Subject: [PATCH 117/282] Nat: update comments with todos Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 15 ++++++++++++--- 1 file changed, 12 insertions(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 33e7a5cfd0b..eda7084e186 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,8 +8,7 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) - - +(** TODO: move and rename *) Proposition add_n_sub_n_eq (m n : nat) : m + n - n = m. Proof. destruct m. @@ -21,11 +20,13 @@ Proof. + simpl. destruct (nat_add_succ_r m n)^. assumption. Defined. +(** TODO: move and rename *) Proposition add_n_sub_n_eq' (m n : nat) : n + m - n = m. Proof. destruct (nat_add_comm m n). exact (add_n_sub_n_eq m n). Defined. +(** TODO: move and rename *) Lemma summand_is_sub k m n (p : k + n = m) : k = m - n. Proof. destruct p. @@ -33,11 +34,13 @@ Proof. apply add_n_sub_n_eq. Defined. +(** TODO: unused, remove *) Proposition not_both_less (n m : nat) : n < m -> ~(m < n). Proof. intros l a; contradiction (not_lt_n_n _ (lt_trans l a)). Defined. +(** TODO: move, rename *) Proposition n_leq_add_n_k (n m : nat) : n <= n + m. Proof. simple_induction n n IHn. @@ -45,6 +48,7 @@ Proof. - simpl; apply leq_succ, IHn. Defined. +(** TODO: move, rename *) Proposition n_leq_add_n_k' (n m : nat) : n <= m + n. Proof. simple_induction' m. @@ -52,6 +56,7 @@ Proof. - simpl. apply leq_succ_r. assumption. Defined. +(** TODO: move, rename *) Proposition natineq0eq0 {n : nat} : n <= 0 -> n = 0. Proof. destruct n. @@ -59,6 +64,7 @@ Proof. - intro. contradiction (not_leq_Sn_0 n). Defined. +(** TODO: unused, remove *) Definition nleqSm_dichot {n m : nat} : (n <= m.+1) -> (n <= m) + (n = m.+1). Proof. @@ -72,7 +78,7 @@ Proof. * right. apply ap; exact b. Defined. - +(** TODO: move, merge with converse as equiv *) Proposition sub_leq_0_converse (n m : nat) : n - m = 0 -> n <= m. Proof. revert m; simple_induction n n IHn. @@ -84,6 +90,7 @@ Proof. + simpl in eq. apply leq_succ, IHn, eq. Defined. +(** TODO: move, rename *) Proposition sub_gt_0_lt (n m : nat) : n - m > 0 -> m < n. Proof. intro ineq. @@ -92,6 +99,7 @@ Proof. contradiction (not_lt_n_n 0). now rewrite n_leq_m in ineq. Defined. +(** TODO: move, rename *) Proposition lt_sub_gt_0 (n m : nat) : m < n -> 0 < n - m. Proof. revert m; simple_induction n n IHn. @@ -102,6 +110,7 @@ Proof. now apply IHn in ineq. Defined. +(** TODO: move, rename *) Proposition natminuspluseq (n m : nat) : n <= m -> (m - n) + n = m. Proof. From fc0a270159c2bb65d948e53494de22dfccdba259 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 17:09:47 +0200 Subject: [PATCH 118/282] Nat: move add_n_sub_n_eq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 12 ------------ theories/Spaces/Nat/Core.v | 12 ++++++++++++ 2 files changed, 12 insertions(+), 12 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index eda7084e186..cb3614f8b19 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,18 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move and rename *) -Proposition add_n_sub_n_eq (m n : nat) : m + n - n = m. -Proof. - destruct m. - - simple_induction' n. - + reflexivity. - + assumption. - - simple_induction' n. - + simpl. destruct (nat_add_zero_r m)^; reflexivity. - + simpl. destruct (nat_add_succ_r m n)^. assumption. -Defined. - (** TODO: move and rename *) Proposition add_n_sub_n_eq' (m n : nat) : n + m - n = m. Proof. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 88c2751cb07..6a28176b5d6 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -611,6 +611,18 @@ Proof. by destruct IHl^. Defined. +(** TODO: rename add_n_sub_n_eq -> [nat_add_sub_cancel_r] *) +Definition add_n_sub_n_eq m n : m + n - n = m. +Proof. + destruct m. + - simple_induction' n. + + reflexivity. + + assumption. + - simple_induction' n. + + simpl. destruct (nat_add_zero_r m)^; reflexivity. + + simpl. destruct (nat_add_succ_r m n)^. assumption. +Defined. + (** ** Properties of Maximum and Minimum *) (** *** Properties of Maxima *) From e132c3888145d8213d1512a2f90bdc83b1239805 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 17:10:33 +0200 Subject: [PATCH 119/282] Nat: rename add_n_sub_n_eq -> nat_add_sub_cancel_r and revise proof Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 10 +++++----- theories/Spaces/Nat/Core.v | 19 ++++++++++--------- 2 files changed, 15 insertions(+), 14 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index cb3614f8b19..b4a6ebb47c4 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -11,7 +11,7 @@ Local Open Scope nat_scope. (** TODO: move and rename *) Proposition add_n_sub_n_eq' (m n : nat) : n + m - n = m. Proof. - destruct (nat_add_comm m n). exact (add_n_sub_n_eq m n). + destruct (nat_add_comm m n). exact (nat_add_sub_cancel_r m n). Defined. (** TODO: move and rename *) @@ -19,7 +19,7 @@ Lemma summand_is_sub k m n (p : k + n = m) : k = m - n. Proof. destruct p. symmetry. - apply add_n_sub_n_eq. + apply nat_add_sub_cancel_r. Defined. (** TODO: unused, remove *) @@ -297,8 +297,8 @@ Defined. Proposition nat_sub_add_ineq (n m : nat) : n <= n - m + m. Proof. destruct (@leq_dichot m n) as [l | gt]. - - destruct (symmetric_paths _ _ (nataddsub_comm _ _ m l)). - destruct (symmetric_paths _ _ (add_n_sub_n_eq n m)). + - destruct (nataddsub_comm _ _ m l)^. + destruct (nat_add_sub_cancel_r n m)^. apply leq_refl; done. - apply leq_lt in gt. destruct (symmetric_paths _ _ (nat_sub_leq _)). @@ -327,7 +327,7 @@ Proposition nat_add_sub_eq (n : nat) {k: nat} Proof. intro H. destruct (symmetric_paths _ _ (nataddsub_assoc k H)); - destruct (nat_add_comm n k); exact (add_n_sub_n_eq _ _). + destruct (nat_add_comm n k); exact (nat_add_sub_cancel_r _ _). Defined. Proposition predeqminus1 { n : nat } : n - 1 = nat_pred n. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 6a28176b5d6..d68dbfc99e2 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -611,16 +611,17 @@ Proof. by destruct IHl^. Defined. -(** TODO: rename add_n_sub_n_eq -> [nat_add_sub_cancel_r] *) -Definition add_n_sub_n_eq m n : m + n - n = m. +(** We can cancel a right summand when subtracting it from a sum. *) +Definition nat_add_sub_cancel_r m n : m + n - n = m. Proof. - destruct m. - - simple_induction' n. - + reflexivity. - + assumption. - - simple_induction' n. - + simpl. destruct (nat_add_zero_r m)^; reflexivity. - + simpl. destruct (nat_add_succ_r m n)^. assumption. + induction n as [|n IHn]. + - lhs nrapply nat_sub_zero_r. + nrapply nat_add_zero_r. + - destruct m. + + apply nat_sub_cancel. + + rhs_V nrapply IHn. + nrapply (ap (fun x => x - n)). + nrapply nat_add_succ_r. Defined. (** ** Properties of Maximum and Minimum *) From 8b9d28d15294ac7506f0b1fe7e313a67114eaaaa Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 17:18:23 +0200 Subject: [PATCH 120/282] Nat: move add_n_sub_n_eq' Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 6 ------ theories/Spaces/Nat/Core.v | 7 +++++++ 2 files changed, 7 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index b4a6ebb47c4..ead09c9340d 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,12 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move and rename *) -Proposition add_n_sub_n_eq' (m n : nat) : n + m - n = m. -Proof. - destruct (nat_add_comm m n). exact (nat_add_sub_cancel_r m n). -Defined. - (** TODO: move and rename *) Lemma summand_is_sub k m n (p : k + n = m) : k = m - n. Proof. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index d68dbfc99e2..59866e702b3 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -624,6 +624,13 @@ Proof. nrapply nat_add_succ_r. Defined. +(** TODO: rename add_n_sub_n_eq' -> nat_add_sub_cancel_l *) +Definition add_n_sub_n_eq' m n : n + m - n = m. +Proof. + destruct (nat_add_comm m n). + exact (nat_add_sub_cancel_r m n). +Defined. + (** ** Properties of Maximum and Minimum *) (** *** Properties of Maxima *) From f493806d8163a549721df80632b3b1a32f490ebb Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 17:18:56 +0200 Subject: [PATCH 121/282] Nat: rename add_n_sub_n_eq' -> nat_add_sub_cancel_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 8 ++++---- theories/Spaces/Nat/Core.v | 4 ++-- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index ead09c9340d..520cbbc2e1d 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -459,8 +459,8 @@ Proposition natpmswap2 (k m n : nat) Proof. intros l q. apply (@nataddpreservesleq' (k - n) m n) in q. - destruct (symmetric_paths _ _ (nataddsub_assoc n l)) in q. - destruct (symmetric_paths _ _ (add_n_sub_n_eq' k n)) in q; + destruct (nataddsub_assoc n l)^ in q. + destruct (nat_add_sub_cancel_l k n)^ in q; assumption. Defined. @@ -469,8 +469,8 @@ Proposition natpmswap3 (k m n : nat) Proof. intros ineq qe. apply (@nataddpreservesleq' m (n - k) k) in qe. - destruct (symmetric_paths _ _ (nataddsub_assoc k ineq)) in qe. - destruct (symmetric_paths _ _ (add_n_sub_n_eq' n k)) in qe; + destruct (nataddsub_assoc k ineq)^ in qe. + destruct (nat_add_sub_cancel_l n k)^ in qe; assumption. Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 59866e702b3..4880b411d61 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -624,8 +624,8 @@ Proof. nrapply nat_add_succ_r. Defined. -(** TODO: rename add_n_sub_n_eq' -> nat_add_sub_cancel_l *) -Definition add_n_sub_n_eq' m n : n + m - n = m. +(** We can cancel a left summand when subtracting it from a sum. *) +Definition nat_add_sub_cancel_l m n : n + m - n = m. Proof. destruct (nat_add_comm m n). exact (nat_add_sub_cancel_r m n). From c0b2a73fc37cb56e07b34d5ce991447033957528 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 17:22:41 +0200 Subject: [PATCH 122/282] Nat: reprove nat_add_sub_cancel_* Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 26 +++++++++++--------------- 1 file changed, 11 insertions(+), 15 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 4880b411d61..53f498b40ac 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -611,24 +611,20 @@ Proof. by destruct IHl^. Defined. -(** We can cancel a right summand when subtracting it from a sum. *) -Definition nat_add_sub_cancel_r m n : m + n - n = m. -Proof. - induction n as [|n IHn]. - - lhs nrapply nat_sub_zero_r. - nrapply nat_add_zero_r. - - destruct m. - + apply nat_sub_cancel. - + rhs_V nrapply IHn. - nrapply (ap (fun x => x - n)). - nrapply nat_add_succ_r. -Defined. - (** We can cancel a left summand when subtracting it from a sum. *) Definition nat_add_sub_cancel_l m n : n + m - n = m. Proof. - destruct (nat_add_comm m n). - exact (nat_add_sub_cancel_r m n). + induction n as [|n IHn]. + - nrapply nat_sub_zero_r. + - exact IHn. +Defined. + +(** We can cancel a right summand when subtracting it from a sum. *) +Definition nat_add_sub_cancel_r m n : m + n - n = m. +Proof. + rhs_V nrapply (nat_add_sub_cancel_l m n). + nrapply (ap (fun x => x - n)). + nrapply nat_add_comm. Defined. (** ** Properties of Maximum and Minimum *) From 7dbb15131c2c9f619d5735421af85532c0344a35 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 17:25:23 +0200 Subject: [PATCH 123/282] Nat: move summand_is_sub Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 8 -------- theories/Spaces/Nat/Core.v | 8 ++++++++ 2 files changed, 8 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 520cbbc2e1d..cbc5a803a3b 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,14 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move and rename *) -Lemma summand_is_sub k m n (p : k + n = m) : k = m - n. -Proof. - destruct p. - symmetry. - apply nat_add_sub_cancel_r. -Defined. - (** TODO: unused, remove *) Proposition not_both_less (n m : nat) : n < m -> ~(m < n). Proof. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 53f498b40ac..ca96d87481c 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -627,6 +627,14 @@ Proof. nrapply nat_add_comm. Defined. +(** TODO: rename summand_is_sub -> nat_moveR_nV *) +Definition summand_is_sub k m n (p : k + n = m) : k = m - n. +Proof. + destruct p. + symmetry. + apply nat_add_sub_cancel_r. +Defined. + (** ** Properties of Maximum and Minimum *) (** *** Properties of Maxima *) From 3e16f81217316932a73156050444b708433615a6 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 17:26:06 +0200 Subject: [PATCH 124/282] Nat: rename summand_is_sub -> nat_moveL_nV Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index cbc5a803a3b..d1c12fbedf9 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -121,7 +121,7 @@ Proof. srapply equiv_iff_hprop. - apply hprop_allpath. intros [x p] [y q]. - pose (r := summand_is_sub x _ _ p @ (summand_is_sub y _ _ q)^). + pose (r := nat_moveR_nV x _ _ p @ (nat_moveR_nV y _ _ q)^). destruct r. apply ap. apply path_ishprop. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ca96d87481c..522b57e0470 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -627,9 +627,9 @@ Proof. nrapply nat_add_comm. Defined. -(** TODO: rename summand_is_sub -> nat_moveR_nV *) -Definition summand_is_sub k m n (p : k + n = m) : k = m - n. +Definition nat_moveL_nV k m n : k + n = m -> k = m - n. Proof. + intros p. destruct p. symmetry. apply nat_add_sub_cancel_r. From 1534dd48d35100dc312409aa8fd14de65c77e508 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 17:32:59 +0200 Subject: [PATCH 125/282] Nat: complete movement lemmas for add and sub Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 7 ++++++- 2 files changed, 7 insertions(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index d1c12fbedf9..48af50d045d 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -121,7 +121,7 @@ Proof. srapply equiv_iff_hprop. - apply hprop_allpath. intros [x p] [y q]. - pose (r := nat_moveR_nV x _ _ p @ (nat_moveR_nV y _ _ q)^). + pose (r := nat_moveL_nV _ p @ (nat_moveL_nV _ q)^). destruct r. apply ap. apply path_ishprop. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 522b57e0470..e53379a3873 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -627,7 +627,8 @@ Proof. nrapply nat_add_comm. Defined. -Definition nat_moveL_nV k m n : k + n = m -> k = m - n. +(** We can move a subtracted number to the left-hand side of an equation. *) +Definition nat_moveL_nV {k m} n : k + n = m -> k = m - n. Proof. intros p. destruct p. @@ -635,6 +636,10 @@ Proof. apply nat_add_sub_cancel_r. Defined. +(** We can move a subtracted number to the right-hand side of an equation. *) +Definition nat_moveR_nV {k m} n : k = n + m -> k - m = n + := fun p => (nat_moveL_nV _ p^)^. + (** ** Properties of Maximum and Minimum *) (** *** Properties of Maxima *) From 81b386b9789d3cdfa8254fdb99af70beec5e4206 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 17:34:12 +0200 Subject: [PATCH 126/282] Nat: remove not_both_less Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 6 ------ 1 file changed, 6 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 48af50d045d..09e6f36ec41 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,12 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: unused, remove *) -Proposition not_both_less (n m : nat) : n < m -> ~(m < n). -Proof. - intros l a; contradiction (not_lt_n_n _ (lt_trans l a)). -Defined. - (** TODO: move, rename *) Proposition n_leq_add_n_k (n m : nat) : n <= n + m. Proof. From 6daee5780e23fdb6a579aea6786cd906e40412a1 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 19:17:31 +0200 Subject: [PATCH 127/282] Nat: move n_leq_add_n_k, n_leq_add_n_k' to Core.v Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 16 ---------------- theories/Spaces/Nat/Core.v | 16 ++++++++++++++++ 2 files changed, 16 insertions(+), 16 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 09e6f36ec41..ae3ab4c43fd 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,22 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move, rename *) -Proposition n_leq_add_n_k (n m : nat) : n <= n + m. -Proof. - simple_induction n n IHn. - - apply leq_zero. - - simpl; apply leq_succ, IHn. -Defined. - -(** TODO: move, rename *) -Proposition n_leq_add_n_k' (n m : nat) : n <= m + n. -Proof. - simple_induction' m. - - exact(leq_refl n). - - simpl. apply leq_succ_r. assumption. -Defined. - (** TODO: move, rename *) Proposition natineq0eq0 {n : nat} : n <= 0 -> n = 0. Proof. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index e53379a3873..4a005f8d2d8 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -757,6 +757,22 @@ Defined. (** ** More Theory of Comparison Predicates *) +(** TODO: move, rename *) +Proposition n_leq_add_n_k (n m : nat) : n <= n + m. +Proof. + simple_induction n n IHn. + - apply leq_zero. + - simpl; apply leq_succ, IHn. +Defined. + +(** TODO: move, rename *) +Proposition n_leq_add_n_k' (n m : nat) : n <= m + n. +Proof. + simple_induction' m. + - exact(leq_refl n). + - simpl. apply leq_succ_r. assumption. +Defined. + (** TODO: rename to [leq_add_r'] *) Fixpoint leq_add n m : n <= (m + n). Proof. From d90fcecfa50339026fba2b6391a1a11fdcbf5471 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 19:21:45 +0200 Subject: [PATCH 128/282] Nat: update comments and proofs about leq_add lemmas Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 16 ++++++++-------- 1 file changed, 8 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 4a005f8d2d8..54e4a18105c 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -757,23 +757,23 @@ Defined. (** ** More Theory of Comparison Predicates *) -(** TODO: move, rename *) -Proposition n_leq_add_n_k (n m : nat) : n <= n + m. +(** TODO: rename n_leq_add_n_k -> leq_add_l *) +Definition n_leq_add_n_k n m : n <= n + m. Proof. simple_induction n n IHn. - apply leq_zero. - simpl; apply leq_succ, IHn. Defined. -(** TODO: move, rename *) -Proposition n_leq_add_n_k' (n m : nat) : n <= m + n. +(** TODO: rename n_leq_add_n_k' -> leq_add_r *) +Definition n_leq_add_n_k' n m : n <= m + n. Proof. - simple_induction' m. - - exact(leq_refl n). - - simpl. apply leq_succ_r. assumption. + simple_induction m m IH. + - exact (leq_refl n). + - exact (leq_succ_r IH). Defined. -(** TODO: rename to [leq_add_r'] *) +(** TODO: remove *) Fixpoint leq_add n m : n <= (m + n). Proof. destruct m. From 62da8e50d9052a9daf755e7cf60b66a27c63ad25 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 19:22:24 +0200 Subject: [PATCH 129/282] Nat: rename n_leq_add_n_k -> leq_add_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index ae3ab4c43fd..48bd037bdb2 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -455,7 +455,7 @@ Defined. Proposition n_leq_m_n_leq_plus_m_k (n m k : nat) : n <= m -> n <= m + k. Proof. - intro l; apply (leq_trans l); exact (n_leq_add_n_k m k). + intro l; apply (leq_trans l); exact (leq_add_l m k). Defined. Proposition nat_add_bifunctor (n n' m m' : nat) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 54e4a18105c..4094346db68 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -757,8 +757,8 @@ Defined. (** ** More Theory of Comparison Predicates *) -(** TODO: rename n_leq_add_n_k -> leq_add_l *) -Definition n_leq_add_n_k n m : n <= n + m. +(** The first summand is less than or equal to the sum. *) +Definition leq_add_l n m : n <= n + m. Proof. simple_induction n n IHn. - apply leq_zero. From a845473234cdc22483173e0b0205249a30abb23d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 19:23:39 +0200 Subject: [PATCH 130/282] Nat: rename n_leq_add_n_k' -> leq_add_r Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 4 ++-- 2 files changed, 3 insertions(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 48bd037bdb2..0e833941934 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -463,7 +463,7 @@ Proposition nat_add_bifunctor (n n' m m' : nat) Proof. revert n' m m'; simple_induction n n IHn. - intros n' m m' l l'. simpl. - apply (leq_trans l'). exact (n_leq_add_n_k' m' m). + apply (leq_trans l'). exact (leq_add_r m' m). - intros n' m; destruct m. + intros. contradiction (not_leq_Sn_0 n). + intros m' l l'. apply leq_succ' in l. simpl. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 4094346db68..2b00769a9f3 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -765,8 +765,8 @@ Proof. - simpl; apply leq_succ, IHn. Defined. -(** TODO: rename n_leq_add_n_k' -> leq_add_r *) -Definition n_leq_add_n_k' n m : n <= m + n. +(** The second summand is less than or equal to the sum. *) +Definition leq_add_r n m : n <= m + n. Proof. simple_induction m m IH. - exact (leq_refl n). From d44b5609913ff13fad3e76618716ce3f6d7dfaad Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 19:24:58 +0200 Subject: [PATCH 131/282] Nat: remove leq_add Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 8 -------- 2 files changed, 1 insertion(+), 9 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 0e833941934..73d985f0ae5 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -108,7 +108,7 @@ Proof. apply natminuspluseq, p. - intros [k p]. destruct p. - apply leq_add. + apply leq_add_r. Defined. Proposition nataddpreservesleq { n m k : nat } diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 2b00769a9f3..3bc87797b49 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -773,14 +773,6 @@ Proof. - exact (leq_succ_r IH). Defined. -(** TODO: remove *) -Fixpoint leq_add n m : n <= (m + n). -Proof. - destruct m. - 1: apply leq_refl. - apply leq_succ_r, leq_add. -Defined. - (** [<=] is an antisymmetric relation. *) Definition leq_antisym {x y} : x <= y -> y <= x -> x = y. Proof. From 3d64ec7c42cdad537dc94abbb904e70c978dd781 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 19:26:44 +0200 Subject: [PATCH 132/282] Nat: improve proof of leq_add_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 3bc87797b49..7e5da710a96 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -761,8 +761,8 @@ Defined. Definition leq_add_l n m : n <= n + m. Proof. simple_induction n n IHn. - - apply leq_zero. - - simpl; apply leq_succ, IHn. + - exact (leq_zero m). + - exact (leq_succ IHn). Defined. (** The second summand is less than or equal to the sum. *) From e6e8f8c8d72ae542dd5e9750fdaf3501beeb7a17 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 19:31:59 +0200 Subject: [PATCH 133/282] Nat: remove unused nleqSm_dichot Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 14 -------------- 1 file changed, 14 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 73d985f0ae5..24bb9f71c7b 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -16,20 +16,6 @@ Proof. - intro. contradiction (not_leq_Sn_0 n). Defined. -(** TODO: unused, remove *) -Definition nleqSm_dichot {n m : nat} - : (n <= m.+1) -> (n <= m) + (n = m.+1). -Proof. - revert m; simple_induction n n IHn. - - intro. left. exact (leq_zero m). - - destruct m. - + intro l. apply leq_succ', natineq0eq0 in l. - right; apply ap; exact l. - + intro l. apply leq_succ', IHn in l; destruct l as [a | b]. - * left. apply leq_succ; exact a. - * right. apply ap; exact b. -Defined. - (** TODO: move, merge with converse as equiv *) Proposition sub_leq_0_converse (n m : nat) : n - m = 0 -> n <= m. Proof. From 1c461efa6014ffb6a5fe68dc4f7361a6002b09ab Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 19:33:11 +0200 Subject: [PATCH 134/282] Nat: remove unused natineq0eq0 Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 8 -------- 1 file changed, 8 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 24bb9f71c7b..70ad3659691 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,14 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move, rename *) -Proposition natineq0eq0 {n : nat} : n <= 0 -> n = 0. -Proof. - destruct n. - - reflexivity. - - intro. contradiction (not_leq_Sn_0 n). -Defined. - (** TODO: move, merge with converse as equiv *) Proposition sub_leq_0_converse (n m : nat) : n - m = 0 -> n <= m. Proof. From 247443fc554c89cfea520ea42283a815f8f7b758 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 19:48:29 +0200 Subject: [PATCH 135/282] Nat: include converse in nat_sub_leq as inverse equiv Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 25 +++++++++++++++++-------- 1 file changed, 17 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 7e5da710a96..390d0954821 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -600,15 +600,24 @@ Proof. nrapply nat_sub_add. Defined. +(** TODO: rename nat_sub_leq -> equiv_nat_sub_leq *) (** Subtracting a larger number from a smaller number is [0]. *) -Definition nat_sub_leq {n m} : n <= m -> n - m = 0. -Proof. - intro l; induction l. - - exact (nat_sub_cancel n). - - change (m.+1) with (1 + m). - lhs nrapply nat_sub_add. - lhs nrapply nat_sub_comm_r. - by destruct IHl^. +Definition nat_sub_leq {n m} : n <= m <~> n - m = 0. +Proof. + srapply equiv_iff_hprop. + - intro l; induction l. + + exact (nat_sub_cancel n). + + change (m.+1) with (1 + m). + lhs nrapply nat_sub_add. + lhs nrapply nat_sub_comm_r. + by destruct IHl^. + - induction n as [|n IHn] in m |- *. + 1: intro; exact _. + destruct m. + + intros p; by destruct p. + + intros p. + apply leq_succ, IHn. + exact p. Defined. (** We can cancel a left summand when subtracting it from a sum. *) From d7d085da9b86c064d0a084c108960192f482793c Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 19:49:01 +0200 Subject: [PATCH 136/282] Nat: rename nat_sub_leq -> equiv_nat_sub_leq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 10 +++++----- theories/Spaces/Nat/Core.v | 3 +-- 2 files changed, 6 insertions(+), 7 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 70ad3659691..356b1723b4c 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -25,7 +25,7 @@ Proposition sub_gt_0_lt (n m : nat) : n - m > 0 -> m < n. Proof. intro ineq. destruct (@leq_dichot n m) as [n_leq_m |]; [ | assumption]. - apply nat_sub_leq in n_leq_m. + apply equiv_nat_sub_leq in n_leq_m. contradiction (not_lt_n_n 0). now rewrite n_leq_m in ineq. Defined. @@ -61,7 +61,7 @@ Proof. - destruct (symmetric_paths _ _ (natminuspluseq _ _ l)); constructor. - apply leq_lt in g. - now destruct (symmetric_paths _ _ (nat_sub_leq _)). + now destruct (equiv_nat_sub_leq _)^. Defined. Proposition natminuspluseq' (n m : nat) @@ -243,7 +243,7 @@ Proof. destruct (nat_add_sub_cancel_r n m)^. apply leq_refl; done. - apply leq_lt in gt. - destruct (symmetric_paths _ _ (nat_sub_leq _)). + destruct (equiv_nat_sub_leq _)^. assumption. Defined. @@ -505,9 +505,9 @@ Proof. simple_induction k k IHk. - destruct (symmetric_paths _ _ (nat_sub_zero_r n)); constructor. - destruct (@leq_dichot n k) as [l | g]. - + destruct (symmetric_paths _ _ (nat_sub_leq _)) in IHk. + + destruct (equiv_nat_sub_leq _)^ in IHk. apply leq_succ_r in l. - destruct (symmetric_paths _ _ (nat_sub_leq _)). exact IHk. + destruct (equiv_nat_sub_leq _)^. exact IHk. + change k.+1 with (1 + k). destruct (nat_add_comm k 1). destruct (symmetric_paths _ _ (nat_sub_add n k 1)). destruct (symmetric_paths _ _ (@predeqminus1 (n - k))). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 390d0954821..aed1ffb29bb 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -600,9 +600,8 @@ Proof. nrapply nat_sub_add. Defined. -(** TODO: rename nat_sub_leq -> equiv_nat_sub_leq *) (** Subtracting a larger number from a smaller number is [0]. *) -Definition nat_sub_leq {n m} : n <= m <~> n - m = 0. +Definition equiv_nat_sub_leq {n m} : n <= m <~> n - m = 0. Proof. srapply equiv_iff_hprop. - intro l; induction l. From 5b8a134dbb5ff2169fe618f4d6f8ed305e77491e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 19:53:28 +0200 Subject: [PATCH 137/282] Nat: remove sub_leq_0_converse It is given by the inverse of equiv_nat_sub_leq. Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 12 ------------ 1 file changed, 12 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 356b1723b4c..9d5bb4d6b72 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,18 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move, merge with converse as equiv *) -Proposition sub_leq_0_converse (n m : nat) : n - m = 0 -> n <= m. -Proof. - revert m; simple_induction n n IHn. - - simpl. intros m eq. - apply leq_zero. - - intros m eq. destruct m. - + simpl in eq. apply symmetric_paths in eq. - contradiction (neq_nat_zero_succ n eq). - + simpl in eq. apply leq_succ, IHn, eq. -Defined. - (** TODO: move, rename *) Proposition sub_gt_0_lt (n m : nat) : n - m > 0 -> m < n. Proof. From cb8b7d3d238fe8ddbd533f36289519bcfb32aada Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 29 Jul 2024 19:57:21 +0200 Subject: [PATCH 138/282] Nat: more todo comments Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 25 +++++++++++++++++++++++++ 1 file changed, 25 insertions(+) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 9d5bb4d6b72..cf00ea116d2 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -52,6 +52,7 @@ Proof. now destruct (equiv_nat_sub_leq _)^. Defined. +(** TODO: move, rename *) Proposition natminuspluseq' (n m : nat) : n <= m -> n + (m - n) = m. Proof. @@ -59,6 +60,7 @@ Proof. apply natminuspluseq. assumption. Defined. +(** TODO: move, rename *) Lemma equiv_leq_add n m : leq n m <~> exists k, k + n = m. Proof. @@ -77,6 +79,7 @@ Proof. apply leq_add_r. Defined. +(** TODO: move, rename *) Proposition nataddpreservesleq { n m k : nat } : n <= m -> n + k <= m + k. Proof. @@ -87,6 +90,7 @@ Proof. apply leq_succ; exact IHk. Defined. +(** TODO: move, rename *) Proposition nataddpreservesleq' { n m k : nat } : n <= m -> k + n <= k + m. Proof. @@ -95,6 +99,7 @@ Proof. exact nataddpreservesleq. Defined. +(** TODO: move, rename *) Proposition nataddpreserveslt { n m k : nat } : n < m -> n + k < m + k. Proof. @@ -107,6 +112,7 @@ Proof. apply leq_succ; exact IHk. Defined. +(** TODO: move, rename *) Proposition nataddpreserveslt' { n m k : nat } : n < m -> k + n < k + m. Proof. @@ -115,6 +121,7 @@ Proof. exact nataddpreserveslt. Defined. +(** TODO: move, rename *) Proposition nataddreflectslt { n m k : nat } : n + k < m + k -> n < m. Proof. @@ -124,6 +131,7 @@ Proof. apply leq_succ', IHk in l; exact l. Defined. +(** TODO: move, rename *) Proposition nataddreflectsleq { n m k : nat } : n + k <= m + k -> n <= m. Proof. @@ -133,6 +141,7 @@ Proof. now apply (@nataddreflectslt n m k). Defined. +(** TODO: move, rename *) Proposition nataddreflectslt' { n m k : nat } : k + n < k + m -> n < m. Proof. @@ -141,6 +150,7 @@ Proof. exact nataddreflectslt. Defined. +(** TODO: move, rename *) Proposition nataddreflectsleq' { n m k : nat } : k + n <= k + m -> n <= m. Proof. @@ -149,6 +159,7 @@ Proof. exact nataddreflectsleq. Defined. +(** TODO: move, rename *) Proposition natsubreflectsleq { n m k : nat } : k <= m -> n - k <= m - k -> n <= m. Proof. @@ -307,6 +318,7 @@ Proof. - assumption. Defined. +(** TODO: move, rename *) Proposition lt_implies_pred_geq (i j : nat) : i < j -> i <= nat_pred j. Proof. intro l; apply leq_pred in l; assumption. @@ -321,6 +333,7 @@ Proof. - now simpl; apply leq_succ'. Defined. +(** TODO: move, rename *) Proposition pred_gt_implies_lt (i j : nat) : i < nat_pred j -> i.+1 < j. Proof. @@ -333,6 +346,7 @@ Proof. assumption. Defined. +(** TODO: move, rename *) Proposition pred_preserves_lt {i n: nat} (p : i < n) m : (n < m) -> (nat_pred n < nat_pred m). Proof. @@ -343,6 +357,7 @@ Proof. assumption. Defined. +(** TODO: move, rename *) Proposition natsubpreservesleq { n m k : nat } : n <= m -> n - k <= m - k. Proof. @@ -358,6 +373,7 @@ Proof. apply leq_pred, IHk. exact l. Defined. +(** TODO: move, rename *) Proposition sub_less { n k : nat } : n - k <= n. Proof. revert k. @@ -369,6 +385,7 @@ Proof. [ apply IHn | apply leq_succ_r, leq_refl]. Defined. +(** TODO: move, rename *) Proposition sub_less_strict { n k : nat } : 0 < n -> 0 < k -> n - k < n. Proof. @@ -379,6 +396,7 @@ Proof. simpl; apply leq_succ, sub_less. Defined. +(** TODO: move, rename *) Proposition natpmswap1 (k m n : nat) : n <= k -> k < n + m -> k - n < m. Proof. @@ -390,6 +408,7 @@ Proof. exact (nataddreflectslt q'). Defined. +(** TODO: move, rename *) Proposition natpmswap2 (k m n : nat) : n <= k -> k - n <= m -> k <= n + m. Proof. @@ -400,6 +419,7 @@ Proof. assumption. Defined. +(** TODO: move, rename *) Proposition natpmswap3 (k m n : nat) : k <= n -> m <= n - k -> k + m <= n. Proof. @@ -410,6 +430,7 @@ Proof. assumption. Defined. +(** TODO: move, rename *) Proposition natpmswap4 (k m n : nat) : k - n < m -> k < n + m. Proof. @@ -418,12 +439,14 @@ Proof. now rapply (leq_lt_trans (nat_sub_add_ineq _ _)). Defined. +(** TODO: move, rename *) Proposition n_leq_m_n_leq_plus_m_k (n m k : nat) : n <= m -> n <= m + k. Proof. intro l; apply (leq_trans l); exact (leq_add_l m k). Defined. +(** TODO: move, rename *) Proposition nat_add_bifunctor (n n' m m' : nat) : n <= m -> n' <= m' -> n + n' <= m + m'. Proof. @@ -438,6 +461,7 @@ Proof. * exact l'. Defined. +(** TODO: move, rename *) Proposition strong_induction (P : nat -> Type) : (forall n : nat, (forall m : nat, (m < n) -> P m) -> P n) -> forall n : nat, P n. @@ -548,6 +572,7 @@ Proof. + exact (leq_succ_l _). Defined. +(** TODO: remove *) (** This tautology accepts a (potentially opaqued or QED'ed) proof of [n <= m], and returns a transparent proof which can be computed with (i.e., one can loop from n to m) *) Definition leq_wrapper {n m : nat} : n <= m -> n <= m. Proof. From befa2cc3f30d4ac8bd4630de303a7008b2fae198 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 12:36:52 +0200 Subject: [PATCH 139/282] Nat: reorganize comparison lemmas, remove inferrable leq_1_Sn Signed-off-by: Ali Caglayan --- theories/Spaces/Finite/FinInduction.v | 2 +- theories/Spaces/Finite/FinNat.v | 4 +- theories/Spaces/Nat/Core.v | 94 ++++++++++++++++----------- 3 files changed, 58 insertions(+), 42 deletions(-) diff --git a/theories/Spaces/Finite/FinInduction.v b/theories/Spaces/Finite/FinInduction.v index ab3fefe470a..ab31e8af8d0 100644 --- a/theories/Spaces/Finite/FinInduction.v +++ b/theories/Spaces/Finite/FinInduction.v @@ -33,7 +33,7 @@ Proof. intro p. destruct (hset_path2 1 p). cbn. - by destruct (hset_path2 1 (path_zero_finnat n leq_1_Sn)). + by destruct (hset_path2 1 (path_zero_finnat n _)). Defined. Lemma compute_fin_ind_fsucc (P : forall n : nat, Fin n -> Type) diff --git a/theories/Spaces/Finite/FinNat.v b/theories/Spaces/Finite/FinNat.v index 18e8434e664..1882d88573b 100644 --- a/theories/Spaces/Finite/FinNat.v +++ b/theories/Spaces/Finite/FinNat.v @@ -10,7 +10,7 @@ Local Open Scope nat_scope. Definition FinNat (n : nat) : Type0 := {x : nat | x < n}. Definition zero_finnat (n : nat) : FinNat n.+1 - := (0; leq_1_Sn). + := (0; _ : 0 < n.+1). Lemma path_zero_finnat (n : nat) (h : 0 < n.+1) : zero_finnat n = (0; h). Proof. @@ -65,7 +65,7 @@ Lemma compute_finnat_ind_zero (P : forall n : nat, FinNat n -> Type) (n : nat) : finnat_ind P z s (zero_finnat n) = z n. Proof. - cbn. by induction (hset_path2 1 (path_zero_finnat n leq_1_Sn)). + cbn. by induction (hset_path2 1 (path_zero_finnat n _)). Defined. Lemma compute_finnat_ind_succ (P : forall n : nat, FinNat n -> Type) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index aed1ffb29bb..d9828c175ed 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -765,21 +765,7 @@ Defined. (** ** More Theory of Comparison Predicates *) -(** The first summand is less than or equal to the sum. *) -Definition leq_add_l n m : n <= n + m. -Proof. - simple_induction n n IHn. - - exact (leq_zero m). - - exact (leq_succ IHn). -Defined. - -(** The second summand is less than or equal to the sum. *) -Definition leq_add_r n m : n <= m + n. -Proof. - simple_induction m m IH. - - exact (leq_refl n). - - exact (leq_succ_r IH). -Defined. +(** *** Antisymmetry of [<=] and [>=] *) (** [<=] is an antisymmetric relation. *) Definition leq_antisym {x y} : x <= y -> y <= x -> x = y. @@ -793,15 +779,59 @@ Proof. Defined. Global Instance antisymmetric_leq : AntiSymmetric leq := @leq_antisym. +Global Instance antisymemtric_geq : AntiSymmetric geq + := fun _ _ p q => leq_antisym q p. + +(** *** Irreflexivity of [<] and [>] *) (** TODO: rename to [lt_irrefl] *) (** [<] is an irreflexive relation. *) Definition not_lt_n_n n : ~ (n < n) := not_leq_Sn_n n. + Global Instance irreflexive_lt : Irreflexive lt := not_lt_n_n. +Global Instance irreflexive_gt : Irreflexive gt := not_lt_n_n. -(** TODO: rename *) -(** [1] is less than or equal to the successor of any number. *) -Definition leq_1_Sn {n} : 1 <= n.+1 := leq_succ (leq_zero _). +(** *** Addition Lemmas *) + +(** The first summand is less than or equal to the sum. *) +Definition leq_add_l n m : n <= n + m. +Proof. + simple_induction n n IHn. + - exact (leq_zero m). + - exact (leq_succ IHn). +Defined. + +(** The second summand is less than or equal to the sum. *) +Definition leq_add_r n m : n <= m + n. +Proof. + simple_induction m m IH. + - exact (leq_refl n). + - exact (leq_succ_r IH). +Defined. + +(** Characterization of [<=] *) + +(** [n <= m] is equivalent to [(n < m) + (n = m)]. Note that it is not immediately obvious that the latter type is a hprop, hence we have to explicitly show the back and forth maps are inverses of eachother. This is possible and justifies the name "less than or equal to". *) +Definition equiv_leq_lt_or_eq {n m : nat} : (n <= m) <~> (n < m) + (n = m). +Proof. + snrapply equiv_adjointify. + - intro l; induction l. + + now right. + + left; exact (leq_succ l). + - intros [l|p]. + + exact (leq_succ_l l). + + destruct p. + exact (leq_refl _). + - intros [l|p]. + + induction l. + 1: reflexivity. + snrapply (ap (inl)). + rapply path_ishprop. + + by destruct p. + - intro; rapply path_ishprop. +Defined. + +(** *** Dichotomy of [<=] *) (** TODO: rename *) Fixpoint leq_dichot {m} {n} : (m <= n) + (m > n). @@ -809,13 +839,17 @@ Proof. simple_induction' m; simple_induction' n. - left; reflexivity. - left; apply leq_zero. - - right; unfold lt; apply leq_1_Sn. + - right; unfold lt; exact _. - assert ((m <= n) + (n < m)) as X by apply leq_dichot. destruct X as [leqmn|ltnm]. + left; apply leq_succ; assumption. + right; apply leq_succ; assumption. Defined. +(** TODO: Trichotomy *) + +(** *** Negation Lemmas *) + (** TODO: rename *) Lemma not_lt_n_0 n : ~ (n < 0). Proof. @@ -850,27 +884,9 @@ Proof. contradiction (not_lt_n_n n); by refine (leq_trans _ ineq2). Defined. -(** [n <= m] is equivalent to [(n < m) + (n = m)]. Note that it is not immediately obvious that the latter type is a hprop, hence we have to explicitly show the back and forth maps are inverses of eachother. This is possible and justifies the name "less than or equal to". *) -Definition equiv_leq_lt_or_eq {n m : nat} : (n <= m) <~> (n < m) + (n = m). -Proof. - snrapply equiv_adjointify. - - intro l; induction l. - + now right. - + left; exact (leq_succ l). - - intros [l|p]. - + exact (leq_succ_l l). - + destruct p. - exact (leq_refl _). - - intros [l|p]. - + induction l. - 1: reflexivity. - snrapply (ap (inl)). - rapply path_ishprop. - + by destruct p. - - intro; rapply path_ishprop. -Defined. +(** *** Dichotomy of [<>] *) -(** The inequality of natural numbers is equivalent to [n < m] or [n > m]. This could be an equivalence however one of the sections requires funext since we are comparing two inequality proofs. It is therefore more useful to keep it as a biimplication. *) +(** The inequality of natural numbers is equivalent to [n < m] or [n > m]. This could be an equivalence however one of the sections requires funext since we are comparing two inequality proofs. It is therefore more useful to keep it as a biimplication. Note that this is a negated version of antisymmetry of [<=]. *) Definition neq_iff_lt_or_gt {n m} : n <> m <-> (n < m) + (n > m). Proof. split. From eebed4212174cbc3f2580f58d62528ca2ffbd3fc Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 12:55:08 +0200 Subject: [PATCH 140/282] Nat: move sub_gt_0_lt Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 9 --------- theories/Spaces/Nat/Core.v | 11 +++++++++++ 2 files changed, 11 insertions(+), 9 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index cf00ea116d2..1d7c3e839c0 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,15 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move, rename *) -Proposition sub_gt_0_lt (n m : nat) : n - m > 0 -> m < n. -Proof. - intro ineq. - destruct (@leq_dichot n m) as [n_leq_m |]; [ | assumption]. - apply equiv_nat_sub_leq in n_leq_m. - contradiction (not_lt_n_n 0). now rewrite n_leq_m in ineq. -Defined. - (** TODO: move, rename *) Proposition lt_sub_gt_0 (n m : nat) : m < n -> 0 < n - m. Proof. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index d9828c175ed..8209a684ab4 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -917,6 +917,17 @@ Proof. reflexivity. Defined. +(** *** Subtraction *) + +(** TODO: rename sub_gt_0_lt to lt_moveL_m (?) , reprove *) +Definition sub_gt_0_lt n m : 0 < n - m -> m < n. +Proof. + intro ineq. + destruct (@leq_dichot n m) as [n_leq_m |]; [ | assumption]. + apply equiv_nat_sub_leq in n_leq_m. + contradiction (not_lt_n_n 0). now rewrite n_leq_m in ineq. +Defined. + (** ** Properties of Powers *) (** [0] to any power is [0] unless that power is [0] in which case it is [1]. *) From 019afe7969f441d5d577e3163788cf5c7a29b904 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 12:57:56 +0200 Subject: [PATCH 141/282] Nat: move lt_sub_gt_0 Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 11 ----------- theories/Spaces/Nat/Core.v | 11 +++++++++++ 2 files changed, 11 insertions(+), 11 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 1d7c3e839c0..01c27ffbe90 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,17 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move, rename *) -Proposition lt_sub_gt_0 (n m : nat) : m < n -> 0 < n - m. -Proof. - revert m; simple_induction n n IHn. - - intros m ineq. contradiction (not_lt_n_0 m). - - destruct m. - + simpl. easy. - + simpl. intro ineq. apply leq_succ' in ineq. - now apply IHn in ineq. -Defined. - (** TODO: move, rename *) Proposition natminuspluseq (n m : nat) : n <= m -> (m - n) + n = m. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 8209a684ab4..270d6c6592f 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -927,6 +927,17 @@ Proof. apply equiv_nat_sub_leq in n_leq_m. contradiction (not_lt_n_n 0). now rewrite n_leq_m in ineq. Defined. + +(** TODO: merge with above, reprove *) +Definition lt_sub_gt_0 n m : m < n -> 0 < n - m. +Proof. + revert m; simple_induction n n IHn. + - intros m ineq. contradiction (not_lt_n_0 m). + - destruct m. + + simpl. easy. + + simpl. intro ineq. apply leq_succ' in ineq. + now apply IHn in ineq. +Defined. (** ** Properties of Powers *) From 0bff0e980689b115f8a2e855fe072a1142472fd6 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 13:03:11 +0200 Subject: [PATCH 142/282] Nat: move natminuspluseq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 21 +++------------------ theories/Spaces/Nat/Core.v | 16 +++++++++++++++- 2 files changed, 18 insertions(+), 19 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 01c27ffbe90..c91d14757c9 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,25 +8,10 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move, rename *) -Proposition natminuspluseq (n m : nat) - : n <= m -> (m - n) + n = m. -Proof. - revert m; simple_induction n n IHn. - - intros. destruct m; [reflexivity |]. simpl. - apply (ap S), symmetric_paths, (nat_add_zero_r _)^. - - intros m l. destruct m. - + contradiction (not_leq_Sn_0 n). - + simpl. apply leq_succ', IHn in l. - destruct (nat_add_succ_r (m - n) n)^. - destruct (symmetric_paths _ _ l). - reflexivity. -Defined. - Proposition natminusplusineq (n m : nat) : n <= n - m + m. Proof. destruct (@leq_dichot m n) as [l | g]. - - destruct (symmetric_paths _ _ (natminuspluseq _ _ l)); + - destruct (natminuspluseq l)^; constructor. - apply leq_lt in g. now destruct (equiv_nat_sub_leq _)^. @@ -147,7 +132,7 @@ Proof. apply (@nataddpreservesleq _ _ k) in ineq2. apply (@leq_trans _ (n - k + k) _ (natminusplusineq _ _)). apply (@leq_trans _ (m - k + k) _ _). - destruct (symmetric_paths _ _ (natminuspluseq k m ineq1)); easy. + destruct (symmetric_paths _ _ (natminuspluseq ineq1)); easy. Defined. Proposition nataddsub_assoc_lemma {k m : nat} @@ -382,7 +367,7 @@ Proposition natpmswap1 (k m n : nat) Proof. intros l q. assert (q' : k - n + n < m + n) by - (destruct (symmetric_paths _ _ (natminuspluseq n k l)); + (destruct (symmetric_paths _ _ (natminuspluseq l)); destruct (nat_add_comm n m); assumption). exact (nataddreflectslt q'). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 270d6c6592f..ddc3b7bb476 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -635,6 +635,20 @@ Proof. nrapply nat_add_comm. Defined. +(** TODO: rename natminuspluseq -> nat_sub_add_cancel_r, reprove *) +Definition natminuspluseq {n m : nat} : n <= m -> (m - n) + n = m. +Proof. + revert m; simple_induction n n IHn. + - intros. destruct m; [reflexivity |]. simpl. + apply (ap S), symmetric_paths, (nat_add_zero_r _)^. + - intros m l. destruct m. + + contradiction (not_leq_Sn_0 n). + + simpl. apply leq_succ', IHn in l. + destruct (nat_add_succ_r (m - n) n)^. + destruct (symmetric_paths _ _ l). + reflexivity. +Defined. + (** We can move a subtracted number to the left-hand side of an equation. *) Definition nat_moveL_nV {k m} n : k + n = m -> k = m - n. Proof. @@ -917,7 +931,7 @@ Proof. reflexivity. Defined. -(** *** Subtraction *) +(** *** Movement Lemmas *) (** TODO: rename sub_gt_0_lt to lt_moveL_m (?) , reprove *) Definition sub_gt_0_lt n m : 0 < n - m -> m < n. From 4d311a5635bc7c3751099aac644d1d3d98d58f2d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 13:06:23 +0200 Subject: [PATCH 143/282] Nat: rename natminuspluseq -> nat_sub_add_cancel Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 10 +++++----- theories/Spaces/Nat/Core.v | 5 +++-- 2 files changed, 8 insertions(+), 7 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index c91d14757c9..04ef54a11f0 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -11,7 +11,7 @@ Local Open Scope nat_scope. Proposition natminusplusineq (n m : nat) : n <= n - m + m. Proof. destruct (@leq_dichot m n) as [l | g]. - - destruct (natminuspluseq l)^; + - destruct (nat_sub_add_cancel l)^; constructor. - apply leq_lt in g. now destruct (equiv_nat_sub_leq _)^. @@ -22,7 +22,7 @@ Proposition natminuspluseq' (n m : nat) : n <= m -> n + (m - n) = m. Proof. intros. destruct (symmetric_paths _ _ (nat_add_comm n (m - n))). - apply natminuspluseq. assumption. + apply nat_sub_add_cancel. assumption. Defined. (** TODO: move, rename *) @@ -38,7 +38,7 @@ Proof. apply path_ishprop. - intros p. exists (m - n). - apply natminuspluseq, p. + apply nat_sub_add_cancel, p. - intros [k p]. destruct p. apply leq_add_r. @@ -132,7 +132,7 @@ Proof. apply (@nataddpreservesleq _ _ k) in ineq2. apply (@leq_trans _ (n - k + k) _ (natminusplusineq _ _)). apply (@leq_trans _ (m - k + k) _ _). - destruct (symmetric_paths _ _ (natminuspluseq ineq1)); easy. + destruct (nat_sub_add_cancel ineq1)^; easy. Defined. Proposition nataddsub_assoc_lemma {k m : nat} @@ -367,7 +367,7 @@ Proposition natpmswap1 (k m n : nat) Proof. intros l q. assert (q' : k - n + n < m + n) by - (destruct (symmetric_paths _ _ (natminuspluseq l)); + (destruct (nat_sub_add_cancel l)^; destruct (nat_add_comm n m); assumption). exact (nataddreflectslt q'). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ddc3b7bb476..e9d9ef19e57 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -635,8 +635,9 @@ Proof. nrapply nat_add_comm. Defined. -(** TODO: rename natminuspluseq -> nat_sub_add_cancel_r, reprove *) -Definition natminuspluseq {n m : nat} : n <= m -> (m - n) + n = m. +(** TODO: reprove *) +(** We can cancel a right subtrahend when adding it to a subtraction if the subtrahend is less than the number being subtracted from. *) +Definition nat_sub_add_cancel {n m} : n <= m -> (m - n) + n = m. Proof. revert m; simple_induction n n IHn. - intros. destruct m; [reflexivity |]. simpl. From f9151053bfca8e194e2039d640bac82faa6e2867 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 13:12:24 +0200 Subject: [PATCH 144/282] Nat: update comments Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 1 + theories/Spaces/Nat/Core.v | 22 ++++++++++++---------- 2 files changed, 13 insertions(+), 10 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 04ef54a11f0..fa47f7c9238 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,6 +8,7 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) +(** TODO: move, rename *) Proposition natminusplusineq (n m : nat) : n <= n - m + m. Proof. destruct (@leq_dichot m n) as [l | g]. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index e9d9ef19e57..ece57c9f87f 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -130,7 +130,7 @@ Fixpoint factorial (n : nat) : nat (** ** Comparison Predicates *) -(** *** Less Than or Equal To *) +(** *** Less Than or Equal To [<=] *) Inductive leq (n : nat) : nat -> Type0 := | leq_refl : leq n n @@ -147,7 +147,7 @@ Notation "n <= m" := (leq n m) : nat_scope. Existing Class leq. Global Existing Instances leq_refl leq_succ_r. -(** *** Less Than *) +(** *** Less Than [<] *) (** We define the less-than relation [lt] in terms of [leq] *) Definition lt n m : Type0 := leq (S n) m. @@ -158,7 +158,7 @@ Existing Class lt. Infix "<" := lt : nat_scope. Global Instance lt_is_leq n m : leq n.+1 m -> lt n m | 100 := idmap. -(** *** Greater Than or Equal To *) +(** *** Greater Than or Equal To [>=] *) Definition geq n m := leq m n. Existing Class geq. @@ -166,7 +166,7 @@ Existing Class geq. Infix ">=" := geq : nat_scope. Global Instance geq_is_leq n m : leq m n -> geq n m | 100 := idmap. -(*** Greater Than *) +(*** Greater Than [>] *) Definition gt n m := lt m n. Existing Class gt. @@ -397,9 +397,9 @@ Definition nat_mul_one_r@{} n : n * 1 = n (** ** Basic Properties of Comparison Predicates *) -(** *** Basic Properties of [leq] *) +(** *** Basic Properties of [<=] *) -(** [leq] is reflexive by definition. *) +(** [<=] is reflexive by definition. *) Global Instance reflexive_leq : Reflexive leq := leq_refl. (** Being less than or equal to is a transitive relation. *) @@ -408,7 +408,7 @@ Proof. intros H1 H2; induction H2; exact _. Defined. -(** [leq] is transtiive. *) +(** [<=] is transtiive. *) Global Instance transitive_leq : Transitive leq := @leq_trans. (** [0] is less than or equal to any natural number. *) @@ -529,7 +529,7 @@ Proof. intro l; apply leq_succ'; exact _. Defined. -(** *** Basic Properties of [lt] *) +(** *** Basic Properties of [<] *) (** [<=] and [<] imply [<] *) Definition leq_lt_trans {n m k} : n <= m -> m < k -> n < k @@ -549,14 +549,14 @@ Global Instance transitive_lt : Transitive lt := @lt_trans. Global Instance ishprop_lt n m : IsHProp (n < m) := _. Global Instance decidable_lt n m : Decidable (lt n m) := _. -(** *** Basic Properties of [geq] *) +(** *** Basic Properties of [>=] *) Global Instance reflexive_geq : Reflexive geq := leq_refl. Global Instance transitive_geq : Transitive geq := fun x y z p q => leq_trans q p. Global Instance ishprop_geq n m : IsHProp (geq n m) := _. Global Instance decidable_geq n m : Decidable (geq n m) := _. -(** *** Basic Properties of [gt] *) +(** *** Basic Properties of [>] *) Global Instance transitive_gt : Transitive gt := fun x y z p q => transitive_lt _ _ _ q p. @@ -824,6 +824,8 @@ Proof. - exact (leq_succ_r IH). Defined. +(** TODO: monotonicity of addition *) + (** Characterization of [<=] *) (** [n <= m] is equivalent to [(n < m) + (n = m)]. Note that it is not immediately obvious that the latter type is a hprop, hence we have to explicitly show the back and forth maps are inverses of eachother. This is possible and justifies the name "less than or equal to". *) From 47625aa19d994285b93511e80a714833073e56dd Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 13:14:30 +0200 Subject: [PATCH 145/282] Nat: move equiv_leq_add Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 19 ------------------- theories/Spaces/Nat/Core.v | 20 +++++++++++++++++++- 2 files changed, 19 insertions(+), 20 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index fa47f7c9238..0822fba62b9 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -26,25 +26,6 @@ Proof. apply nat_sub_add_cancel. assumption. Defined. -(** TODO: move, rename *) -Lemma equiv_leq_add n m - : leq n m <~> exists k, k + n = m. -Proof. - srapply equiv_iff_hprop. - - apply hprop_allpath. - intros [x p] [y q]. - pose (r := nat_moveL_nV _ p @ (nat_moveL_nV _ q)^). - destruct r. - apply ap. - apply path_ishprop. - - intros p. - exists (m - n). - apply nat_sub_add_cancel, p. - - intros [k p]. - destruct p. - apply leq_add_r. -Defined. - (** TODO: move, rename *) Proposition nataddpreservesleq { n m k : nat } : n <= m -> n + k <= m + k. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ece57c9f87f..74f3a18cc36 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -826,7 +826,7 @@ Defined. (** TODO: monotonicity of addition *) -(** Characterization of [<=] *) +(** Characterizations of [<=] *) (** [n <= m] is equivalent to [(n < m) + (n = m)]. Note that it is not immediately obvious that the latter type is a hprop, hence we have to explicitly show the back and forth maps are inverses of eachother. This is possible and justifies the name "less than or equal to". *) Definition equiv_leq_lt_or_eq {n m : nat} : (n <= m) <~> (n < m) + (n = m). @@ -848,6 +848,24 @@ Proof. - intro; rapply path_ishprop. Defined. +(** Here is an alternative characterization of [<=] in terms of an existence predicate and addition. *) +Definition equiv_leq_add n m : n <= m <~> exists k, k + n = m. +Proof. + srapply equiv_iff_hprop. + - apply hprop_allpath. + intros [x p] [y q]. + pose (r := nat_moveL_nV _ p @ (nat_moveL_nV _ q)^). + destruct r. + apply ap. + apply path_ishprop. + - intros p. + exists (m - n). + apply nat_sub_add_cancel, p. + - intros [k p]. + destruct p. + apply leq_add_r. +Defined. + (** *** Dichotomy of [<=] *) (** TODO: rename *) From 9353f0c6891eef4515dc5e733b5627a33d8dc511 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 13:38:55 +0200 Subject: [PATCH 146/282] Nat: move natminusplusineq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 10 ---------- theories/Spaces/Nat/Core.v | 12 ++++++++++++ 2 files changed, 12 insertions(+), 10 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 0822fba62b9..acfd9ff9002 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,16 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move, rename *) -Proposition natminusplusineq (n m : nat) : n <= n - m + m. -Proof. - destruct (@leq_dichot m n) as [l | g]. - - destruct (nat_sub_add_cancel l)^; - constructor. - - apply leq_lt in g. - now destruct (equiv_nat_sub_leq _)^. -Defined. - (** TODO: move, rename *) Proposition natminuspluseq' (n m : nat) : n <= m -> n + (m - n) = m. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 74f3a18cc36..12deb5aa0b1 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -952,6 +952,18 @@ Proof. reflexivity. Defined. +(** *** Subtraction *) + +(** TODO: rename natminusplusineq -> leq_sub_add *) +Definition natminusplusineq n m : n <= n - m + m. +Proof. + destruct (@leq_dichot m n) as [l | g]. + - destruct (nat_sub_add_cancel l)^; + constructor. + - apply leq_lt in g. + now destruct (equiv_nat_sub_leq _)^. +Defined. + (** *** Movement Lemmas *) (** TODO: rename sub_gt_0_lt to lt_moveL_m (?) , reprove *) From 680b4a68375bf2e836544482a0e16ee0a0ec9f8f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 13:41:23 +0200 Subject: [PATCH 147/282] Nat: rename natminusplusineq -> leq_sub_add Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 3 +-- 2 files changed, 2 insertions(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index acfd9ff9002..8bff7e98c58 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -102,7 +102,7 @@ Proposition natsubreflectsleq { n m k : nat } Proof. intros ineq1 ineq2. apply (@nataddpreservesleq _ _ k) in ineq2. - apply (@leq_trans _ (n - k + k) _ (natminusplusineq _ _)). + apply (@leq_trans _ (n - k + k) _ (leq_sub_add _ _)). apply (@leq_trans _ (m - k + k) _ _). destruct (nat_sub_add_cancel ineq1)^; easy. Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 12deb5aa0b1..dfeac9c64f2 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -954,8 +954,7 @@ Defined. (** *** Subtraction *) -(** TODO: rename natminusplusineq -> leq_sub_add *) -Definition natminusplusineq n m : n <= n - m + m. +Definition leq_sub_add n m : n <= n - m + m. Proof. destruct (@leq_dichot m n) as [l | g]. - destruct (nat_sub_add_cancel l)^; From 0b3cabb087f2326957230e95758533536a18ff13 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 13:47:59 +0200 Subject: [PATCH 148/282] Nat: rename nat_sub_add_cancel -> nat_add_sub_l_cancel Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 6 +++--- theories/Spaces/Nat/Core.v | 6 +++--- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 8bff7e98c58..ce4e9b144ba 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -13,7 +13,7 @@ Proposition natminuspluseq' (n m : nat) : n <= m -> n + (m - n) = m. Proof. intros. destruct (symmetric_paths _ _ (nat_add_comm n (m - n))). - apply nat_sub_add_cancel. assumption. + apply nat_add_sub_l_cancel. assumption. Defined. (** TODO: move, rename *) @@ -104,7 +104,7 @@ Proof. apply (@nataddpreservesleq _ _ k) in ineq2. apply (@leq_trans _ (n - k + k) _ (leq_sub_add _ _)). apply (@leq_trans _ (m - k + k) _ _). - destruct (nat_sub_add_cancel ineq1)^; easy. + destruct (nat_add_sub_l_cancel ineq1)^; easy. Defined. Proposition nataddsub_assoc_lemma {k m : nat} @@ -339,7 +339,7 @@ Proposition natpmswap1 (k m n : nat) Proof. intros l q. assert (q' : k - n + n < m + n) by - (destruct (nat_sub_add_cancel l)^; + (destruct (nat_add_sub_l_cancel l)^; destruct (nat_add_comm n m); assumption). exact (nataddreflectslt q'). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index dfeac9c64f2..a87e56740a5 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -637,7 +637,7 @@ Defined. (** TODO: reprove *) (** We can cancel a right subtrahend when adding it to a subtraction if the subtrahend is less than the number being subtracted from. *) -Definition nat_sub_add_cancel {n m} : n <= m -> (m - n) + n = m. +Definition nat_add_sub_l_cancel {n m} : n <= m -> (m - n) + n = m. Proof. revert m; simple_induction n n IHn. - intros. destruct m; [reflexivity |]. simpl. @@ -860,7 +860,7 @@ Proof. apply path_ishprop. - intros p. exists (m - n). - apply nat_sub_add_cancel, p. + apply nat_add_sub_l_cancel, p. - intros [k p]. destruct p. apply leq_add_r. @@ -957,7 +957,7 @@ Defined. Definition leq_sub_add n m : n <= n - m + m. Proof. destruct (@leq_dichot m n) as [l | g]. - - destruct (nat_sub_add_cancel l)^; + - destruct (nat_add_sub_l_cancel l)^; constructor. - apply leq_lt in g. now destruct (equiv_nat_sub_leq _)^. From 6decdc73aa8c3d1935b96ba3dbc9e85188261c85 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 13:50:24 +0200 Subject: [PATCH 149/282] Nat: move antminuspluseq' Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 8 -------- theories/Spaces/Nat/Core.v | 7 +++++++ 2 files changed, 7 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index ce4e9b144ba..e440eb17481 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,14 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move, rename *) -Proposition natminuspluseq' (n m : nat) - : n <= m -> n + (m - n) = m. -Proof. - intros. destruct (symmetric_paths _ _ (nat_add_comm n (m - n))). - apply nat_add_sub_l_cancel. assumption. -Defined. - (** TODO: move, rename *) Proposition nataddpreservesleq { n m k : nat } : n <= m -> n + k <= m + k. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index a87e56740a5..3c1b7612a04 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -650,6 +650,13 @@ Proof. reflexivity. Defined. +(** TODO: rename natminuspluseq' -> nat_add_sub_r_cancel *) +Definition natminuspluseq' {n m} : n <= m -> n + (m - n) = m. +Proof. + intros. destruct (symmetric_paths _ _ (nat_add_comm n (m - n))). + apply nat_add_sub_l_cancel. assumption. +Defined. + (** We can move a subtracted number to the left-hand side of an equation. *) Definition nat_moveL_nV {k m} n : k + n = m -> k = m - n. Proof. From 8f0fc37741f7ef884a1dc66743e6ee7bd2ddc5d2 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 13:52:02 +0200 Subject: [PATCH 150/282] Nat: rename natminuspluseq' -> nat_add_sub_r_cancel Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 31 ++++++++++++++++--------------- 1 file changed, 16 insertions(+), 15 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 3c1b7612a04..3307cb2107c 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -635,26 +635,27 @@ Proof. nrapply nat_add_comm. Defined. -(** TODO: reprove *) -(** We can cancel a right subtrahend when adding it to a subtraction if the subtrahend is less than the number being subtracted from. *) +(** We can cancel a right subtrahend when adding it on the right to a subtraction if the subtrahend is less than the number being subtracted from. *) Definition nat_add_sub_l_cancel {n m} : n <= m -> (m - n) + n = m. Proof. - revert m; simple_induction n n IHn. - - intros. destruct m; [reflexivity |]. simpl. - apply (ap S), symmetric_paths, (nat_add_zero_r _)^. - - intros m l. destruct m. - + contradiction (not_leq_Sn_0 n). - + simpl. apply leq_succ', IHn in l. - destruct (nat_add_succ_r (m - n) n)^. - destruct (symmetric_paths _ _ l). - reflexivity. + intros H. + induction n as [|n IHn] in m, H |- *. + - lhs nrapply nat_add_zero_r. + nrapply nat_sub_zero_r. + - destruct m. + 1: contradiction (not_leq_Sn_0 n). + lhs nrapply nat_add_succ_r. + nrapply (ap nat_succ). + nrapply IHn. + exact (leq_succ' H). Defined. -(** TODO: rename natminuspluseq' -> nat_add_sub_r_cancel *) -Definition natminuspluseq' {n m} : n <= m -> n + (m - n) = m. +(** We can cancel a right subtrahend when adding it on the left to a subtraction if the subtrahend is less than the nubmer being subtracted from. *) +Definition nat_add_sub_r_cancel {n m} : n <= m -> n + (m - n) = m. Proof. - intros. destruct (symmetric_paths _ _ (nat_add_comm n (m - n))). - apply nat_add_sub_l_cancel. assumption. + intros H. + rhs_V nrapply (nat_add_sub_l_cancel H). + apply nat_add_comm. Defined. (** We can move a subtracted number to the left-hand side of an equation. *) From d93b6de513e0b768f0dad2ad6251700c8861ecde Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 15:47:11 +0200 Subject: [PATCH 151/282] Nat: move addition monoticity lemmas Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 42 ---------------------------- theories/Spaces/Nat/Core.v | 48 ++++++++++++++++++++++++++++++-- 2 files changed, 46 insertions(+), 44 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index e440eb17481..d833a51d8fb 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,48 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move, rename *) -Proposition nataddpreservesleq { n m k : nat } - : n <= m -> n + k <= m + k. -Proof. - intro l. - simple_induction k k IHk. - - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; exact l. - - destruct (nat_add_succ_r n k)^, (nat_add_succ_r m k)^; - apply leq_succ; exact IHk. -Defined. - -(** TODO: move, rename *) -Proposition nataddpreservesleq' { n m k : nat } - : n <= m -> k + n <= k + m. -Proof. - destruct (symmetric_paths _ _ (nat_add_comm k m)), - (symmetric_paths _ _ (nat_add_comm k n)). - exact nataddpreservesleq. -Defined. - -(** TODO: move, rename *) -Proposition nataddpreserveslt { n m k : nat } - : n < m -> n + k < m + k. -Proof. - unfold "<". - change (n + k).+1 with (n.+1 + k). - generalize (n.+1). intros n' l. - simple_induction k k IHk. - - destruct (nat_add_zero_r n')^, (nat_add_zero_r m)^; exact l. - - destruct (nat_add_succ_r n' k)^, (nat_add_succ_r m k)^; - apply leq_succ; exact IHk. -Defined. - -(** TODO: move, rename *) -Proposition nataddpreserveslt' { n m k : nat } - : n < m -> k + n < k + m. -Proof. - destruct (symmetric_paths _ _ (nat_add_comm k n)), - (symmetric_paths _ _ (nat_add_comm k m)); - exact nataddpreserveslt. -Defined. - (** TODO: move, rename *) Proposition nataddreflectslt { n m k : nat } : n + k < m + k -> n < m. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 3307cb2107c..453f2ae2ef9 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -832,8 +832,6 @@ Proof. - exact (leq_succ_r IH). Defined. -(** TODO: monotonicity of addition *) - (** Characterizations of [<=] *) (** [n <= m] is equivalent to [(n < m) + (n = m)]. Note that it is not immediately obvious that the latter type is a hprop, hence we have to explicitly show the back and forth maps are inverses of eachother. This is possible and justifies the name "less than or equal to". *) @@ -993,6 +991,52 @@ Proof. now apply IHn in ineq. Defined. +(** *** Monotonicity of Addition *) + +(** TODO: use OrderPreserving from canonical_names *) + +(** TODO: rename nataddpreservesleq *) +Proposition nataddpreservesleq { n m k : nat } + : n <= m -> n + k <= m + k. +Proof. + intro l. + simple_induction k k IHk. + - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; exact l. + - destruct (nat_add_succ_r n k)^, (nat_add_succ_r m k)^; + apply leq_succ; exact IHk. +Defined. + +(** TODO: move, rename *) +Proposition nataddpreservesleq' { n m k : nat } + : n <= m -> k + n <= k + m. +Proof. + destruct (symmetric_paths _ _ (nat_add_comm k m)), + (symmetric_paths _ _ (nat_add_comm k n)). + exact nataddpreservesleq. +Defined. + +(** TODO: move, rename *) +Proposition nataddpreserveslt { n m k : nat } + : n < m -> n + k < m + k. +Proof. + unfold "<". + change (n + k).+1 with (n.+1 + k). + generalize (n.+1). intros n' l. + simple_induction k k IHk. + - destruct (nat_add_zero_r n')^, (nat_add_zero_r m)^; exact l. + - destruct (nat_add_succ_r n' k)^, (nat_add_succ_r m k)^; + apply leq_succ; exact IHk. +Defined. + +(** TODO: move, rename *) +Proposition nataddpreserveslt' { n m k : nat } + : n < m -> k + n < k + m. +Proof. + destruct (symmetric_paths _ _ (nat_add_comm k n)), + (symmetric_paths _ _ (nat_add_comm k m)); + exact nataddpreserveslt. +Defined. + (** ** Properties of Powers *) (** [0] to any power is [0] unless that power is [0] in which case it is [1]. *) From cbf5a28c6bdf7987be9595e3388ba9dc03832792 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 15:48:56 +0200 Subject: [PATCH 152/282] Nat: move order reflection of addition lemmas Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 38 ----------------------------- theories/Spaces/Nat/Core.v | 42 +++++++++++++++++++++++++++++++- 2 files changed, 41 insertions(+), 39 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index d833a51d8fb..4cf1219e7e0 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,44 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move, rename *) -Proposition nataddreflectslt { n m k : nat } - : n + k < m + k -> n < m. -Proof. - simple_induction k k IHk. - - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; trivial. - - intro l. destruct (nat_add_succ_r n k)^, (nat_add_succ_r m k)^ in l. - apply leq_succ', IHk in l; exact l. -Defined. - -(** TODO: move, rename *) -Proposition nataddreflectsleq { n m k : nat } - : n + k <= m + k -> n <= m. -Proof. - destruct n. - - intros ?; apply leq_zero. - - intro a. change (n.+1 + k) with (n + k).+1 in a. - now apply (@nataddreflectslt n m k). -Defined. - -(** TODO: move, rename *) -Proposition nataddreflectslt' { n m k : nat } - : k + n < k + m -> n < m. -Proof. - destruct (symmetric_paths _ _ (nat_add_comm k n)), - (symmetric_paths _ _ (nat_add_comm k m)); - exact nataddreflectslt. -Defined. - -(** TODO: move, rename *) -Proposition nataddreflectsleq' { n m k : nat } - : k + n <= k + m -> n <= m. -Proof. - destruct (symmetric_paths _ _ (nat_add_comm k n)), - (symmetric_paths _ _ (nat_add_comm k m)); - exact nataddreflectsleq. -Defined. - (** TODO: move, rename *) Proposition natsubreflectsleq { n m k : nat } : k <= m -> n - k <= m - k -> n <= m. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 453f2ae2ef9..cfd7f6d44c3 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -995,7 +995,7 @@ Defined. (** TODO: use OrderPreserving from canonical_names *) -(** TODO: rename nataddpreservesleq *) +(** TODO: rename nataddpreservesleq -> *) Proposition nataddpreservesleq { n m k : nat } : n <= m -> n + k <= m + k. Proof. @@ -1037,6 +1037,46 @@ Proof. exact nataddpreserveslt. Defined. +(** *** Order-reflection of Addition *) + +(** TODO: move, rename *) +Proposition nataddreflectslt { n m k : nat } + : n + k < m + k -> n < m. +Proof. + simple_induction k k IHk. + - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; trivial. + - intro l. destruct (nat_add_succ_r n k)^, (nat_add_succ_r m k)^ in l. + apply leq_succ', IHk in l; exact l. +Defined. + +(** TODO: move, rename *) +Proposition nataddreflectsleq { n m k : nat } + : n + k <= m + k -> n <= m. +Proof. + destruct n. + - intros ?; apply leq_zero. + - intro a. change (n.+1 + k) with (n + k).+1 in a. + now apply (@nataddreflectslt n m k). +Defined. + +(** TODO: move, rename *) +Proposition nataddreflectslt' { n m k : nat } + : k + n < k + m -> n < m. +Proof. + destruct (symmetric_paths _ _ (nat_add_comm k n)), + (symmetric_paths _ _ (nat_add_comm k m)); + exact nataddreflectslt. +Defined. + +(** TODO: move, rename *) +Proposition nataddreflectsleq' { n m k : nat } + : k + n <= k + m -> n <= m. +Proof. + destruct (symmetric_paths _ _ (nat_add_comm k n)), + (symmetric_paths _ _ (nat_add_comm k m)); + exact nataddreflectsleq. +Defined. + (** ** Properties of Powers *) (** [0] to any power is [0] unless that power is [0] in which case it is [1]. *) From 8e66a32818854e01870fe1a670be396b6b443b99 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 15:51:59 +0200 Subject: [PATCH 153/282] Nat: comments and move order-reflection of subtraction Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 11 ----------- theories/Spaces/Nat/Core.v | 18 ++++++++++++++++-- 2 files changed, 16 insertions(+), 13 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 4cf1219e7e0..55ee3225b81 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,17 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -(** TODO: move, rename *) -Proposition natsubreflectsleq { n m k : nat } - : k <= m -> n - k <= m - k -> n <= m. -Proof. - intros ineq1 ineq2. - apply (@nataddpreservesleq _ _ k) in ineq2. - apply (@leq_trans _ (n - k + k) _ (leq_sub_add _ _)). - apply (@leq_trans _ (m - k + k) _ _). - destruct (nat_add_sub_l_cancel ineq1)^; easy. -Defined. - Proposition nataddsub_assoc_lemma {k m : nat} : (k <= m) -> m.+1 - k = (m - k).+1. Proof. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index cfd7f6d44c3..c200f737af2 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -991,7 +991,7 @@ Proof. now apply IHn in ineq. Defined. -(** *** Monotonicity of Addition *) +(** *** Monotonicity *) (** TODO: use OrderPreserving from canonical_names *) @@ -1037,7 +1037,10 @@ Proof. exact nataddpreserveslt. Defined. -(** *** Order-reflection of Addition *) +(** TODO: monotonicity of subtraction *) +(** TODO: monotonicity of multiplication *) + +(** *** Order-reflection *) (** TODO: move, rename *) Proposition nataddreflectslt { n m k : nat } @@ -1077,6 +1080,17 @@ Proof. exact nataddreflectsleq. Defined. +(** TODO: move, rename *) +Proposition natsubreflectsleq { n m k : nat } + : k <= m -> n - k <= m - k -> n <= m. +Proof. + intros ineq1 ineq2. + apply (@nataddpreservesleq _ _ k) in ineq2. + apply (@leq_trans _ (n - k + k) _ (leq_sub_add _ _)). + apply (@leq_trans _ (m - k + k) _ _). + destruct (nat_add_sub_l_cancel ineq1)^; easy. +Defined. + (** ** Properties of Powers *) (** [0] to any power is [0] unless that power is [0] in which case it is [1]. *) From f6bd3d72cad3279d1cefc99bf334aee27fbb6e2e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 15:55:19 +0200 Subject: [PATCH 154/282] Nat: move nataddsub_assoc and nataddsub_comm Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 43 ---------------------------- theories/Spaces/Nat/Core.v | 48 ++++++++++++++++++++++++++++++++ 2 files changed, 48 insertions(+), 43 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 55ee3225b81..af71dc7306a 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,49 +8,6 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) -Proposition nataddsub_assoc_lemma {k m : nat} - : (k <= m) -> m.+1 - k = (m - k).+1. -Proof. - revert m; simple_induction k k IHk. - - intros m l; simpl. destruct m; reflexivity. - - destruct m. - + simpl; intro g; contradiction (not_leq_Sn_0 _ g). - + intro l; apply leq_succ' in l. - change (m.+2 - k.+1) with (m.+1 - k). - change (m.+1 - k.+1) with (m - k). - exact (IHk _ l). -Defined. - -Proposition nataddsub_assoc (n : nat) {m k : nat} - : (k <= m) -> n + (m - k) = n + m - k. -Proof. - revert m k. simple_induction n n IHn. - - reflexivity. - - intros m k l. - change (n.+1 + (m - k)) with (n + (m - k)).+1; - change (n.+1 + m) with (n +m).+1. - destruct k, m; - [ reflexivity - | reflexivity - | contradiction (not_lt_n_0 k _) - | ]. - simpl "-". apply leq_succ' in l. - destruct (nat_add_succ_r n (m - k)). - destruct (nataddsub_assoc_lemma l). - apply (IHn m.+1 k). - apply leq_succ_r. - assumption. -Defined. - -Proposition nataddsub_comm (n m k : nat) - : m <= n -> (n - m) + k = (n + k) - m. -Proof. - intro l. - destruct (nat_add_comm k n). - destruct (nataddsub_assoc k l). - apply nat_add_comm. -Defined. - Proposition nataddsub_comm_ineq_lemma (n m : nat) : n.+1 - m <= (n - m).+1. Proof. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index c200f737af2..7ee549949b8 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1091,6 +1091,54 @@ Proof. destruct (nat_add_sub_l_cancel ineq1)^; easy. Defined. +(** ** Further Properties of Subtraction *) + +(** TODO: rename *) +Proposition nataddsub_assoc_lemma {k m : nat} + : (k <= m) -> m.+1 - k = (m - k).+1. +Proof. + revert m; simple_induction k k IHk. + - intros m l; simpl. destruct m; reflexivity. + - destruct m. + + simpl; intro g; contradiction (not_leq_Sn_0 _ g). + + intro l; apply leq_succ' in l. + change (m.+2 - k.+1) with (m.+1 - k). + change (m.+1 - k.+1) with (m - k). + exact (IHk _ l). +Defined. + +(** TODO: rename *) +Proposition nataddsub_assoc(n : nat) {m k : nat} + : (k <= m) -> n + (m - k) = n + m - k. +Proof. + revert m k. simple_induction n n IHn. + - reflexivity. + - intros m k l. + change (n.+1 + (m - k)) with (n + (m - k)).+1; + change (n.+1 + m) with (n +m).+1. + destruct k, m; + [ reflexivity + | reflexivity + | contradiction (not_lt_n_0 k _) + | ]. + simpl "-". apply leq_succ' in l. + destruct (nat_add_succ_r n (m - k)). + destruct (nataddsub_assoc_lemma l). + apply (IHn m.+1 k). + apply leq_succ_r. + assumption. +Defined. + +(** TODO: rename *) +Proposition nataddsub_comm (n m k : nat) + : m <= n -> (n - m) + k = (n + k) - m. +Proof. + intro l. + destruct (nat_add_comm k n). + destruct (nataddsub_assoc k l). + apply nat_add_comm. +Defined. + (** ** Properties of Powers *) (** [0] to any power is [0] unless that power is [0] in which case it is [1]. *) From 9c37bed648b71260daff05b0c646cc9ec8e51445 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 16:01:07 +0200 Subject: [PATCH 155/282] Nat: show trichotomy of natural numbers Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 14 +++++++++++--- 1 file changed, 11 insertions(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 7ee549949b8..8d7c775acf8 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1,5 +1,5 @@ Require Import Basics.Overture Basics.Tactics Basics.PathGroupoids - Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat Basics.Classes. + Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat Basics.Classes Types.Sum. Export Basics.Nat. Local Set Universe Minimization ToSet. @@ -875,7 +875,7 @@ Defined. (** *** Dichotomy of [<=] *) (** TODO: rename *) -Fixpoint leq_dichot {m} {n} : (m <= n) + (m > n). +Fixpoint leq_dichot m n : (m <= n) + (m > n). Proof. simple_induction' m; simple_induction' n. - left; reflexivity. @@ -887,7 +887,15 @@ Proof. + right; apply leq_succ; assumption. Defined. -(** TODO: Trichotomy *) +(** *** Trichotomy *) + +(** Every two natural numbers are either equal, less than, or greater than each other. *) +Definition nat_trichotomy m n : (m < n) + (m = n) + (m > n). +Proof. + generalize (leq_dichot m n). + snrapply (functor_sum _ idmap). + snrapply equiv_leq_lt_or_eq. +Defined. (** *** Negation Lemmas *) From d2f86aef5dfc15848cb600f78ea4964130d41a23 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 16:01:50 +0200 Subject: [PATCH 156/282] Nat: rename not_lt_n_0 -> not_lt_zero_r Signed-off-by: Ali Caglayan --- theories/Spaces/Finite/FinNat.v | 8 ++++---- theories/Spaces/Finite/FinSeq.v | 2 +- theories/Spaces/Nat/Arithmetic.v | 16 ++++++++-------- theories/Spaces/Nat/Core.v | 7 +++---- 4 files changed, 16 insertions(+), 17 deletions(-) diff --git a/theories/Spaces/Finite/FinNat.v b/theories/Spaces/Finite/FinNat.v index 1882d88573b..e4a239208b5 100644 --- a/theories/Spaces/Finite/FinNat.v +++ b/theories/Spaces/Finite/FinNat.v @@ -51,7 +51,7 @@ Definition finnat_ind (P : forall n : nat, FinNat n -> Type) : P n u. Proof. induction n as [| n IHn]. - - elim (not_lt_n_0 u.1 u.2). + - elim (not_lt_zero_r u.1 u.2). - destruct u as [x h]. destruct x as [| x]. + exact (transport (P n.+1) (path_zero_finnat _ h) (z _)). @@ -101,7 +101,7 @@ Monomorphic Definition fin_to_finnat {n} (k : Fin n) : FinNat n Monomorphic Fixpoint finnat_to_fin {n : nat} : FinNat n -> Fin n := match n with - | 0 => fun u => Empty_rec (not_lt_n_0 _ u.2) + | 0 => fun u => Empty_rec (not_lt_zero_r _ u.2) | n.+1 => fun u => match u with | (0; _) => fin_zero @@ -151,7 +151,7 @@ Lemma path_finnat_to_fin_incl {n : nat} (u : FinNat n) : finnat_to_fin (incl_finnat u) = fin_incl (finnat_to_fin u). Proof. induction n as [| n IHn]. - - elim (not_lt_n_0 _ u.2). + - elim (not_lt_zero_r _ u.2). - destruct u as [x h]. destruct x as [| x]; [reflexivity|]. refine ((ap _ (ap _ (path_succ_finnat (x; leq_succ' h) h)))^ @ _). @@ -171,7 +171,7 @@ Lemma path_finnat_to_fin_to_finnat {n : nat} (u : FinNat n) : fin_to_finnat (finnat_to_fin u) = u. Proof. induction n as [| n IHn]. - - elim (not_lt_n_0 _ u.2). + - elim (not_lt_zero_r _ u.2). - destruct u as [x h]. apply path_sigma_hprop. destruct x as [| x]. diff --git a/theories/Spaces/Finite/FinSeq.v b/theories/Spaces/Finite/FinSeq.v index eed66355d50..18e5f935ab0 100644 --- a/theories/Spaces/Finite/FinSeq.v +++ b/theories/Spaces/Finite/FinSeq.v @@ -40,7 +40,7 @@ Definition fscons {A : Type} {n : nat} : A -> FinSeq n A -> FinSeq n.+1 A Definition fshead' {A} (n : nat) : 0 < n -> FinSeq n A -> A := match n with - | 0 => fun N _ => Empty_rec (not_lt_n_0 _ N) + | 0 => fun N _ => Empty_rec (not_lt_zero_r _ N) | n'.+1 => fun _ v => v fin_zero end. diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index af71dc7306a..07126d22175 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -45,7 +45,7 @@ Proposition i_lt_n_sum_m (n m i : nat) : i < n - m -> m <= n. Proof. revert m i; simple_induction n n IHn. - - intros m i l. simpl in l. contradiction (not_lt_n_0 _ _). + - intros m i l. simpl in l. contradiction (not_lt_zero_r _ _). - intros m i l. destruct m. + apply leq_zero. + apply leq_succ. simpl in l. apply (IHn m i l). @@ -81,14 +81,14 @@ Defined. Proposition S_predn (n i: nat) : (i < n) -> S(nat_pred n) = n. Proof. simple_induction' n; intros l. - - contradiction (not_lt_n_0 i). + - contradiction (not_lt_zero_r i). - reflexivity. Defined. Proposition pred_equiv (k n : nat) : k < n -> k < S (nat_pred n). Proof. intro ineq; destruct n. - - contradiction (not_lt_n_0 _ _). + - contradiction (not_lt_zero_r _ _). - assumption. Defined. @@ -101,7 +101,7 @@ Proposition leq_implies_pred_lt (i n k : nat) : (n > i) -> n <= k -> nat_pred n < k. Proof. intro ineq; destruct n. - - contradiction (not_lt_n_0 i). + - contradiction (not_lt_zero_r i). - intro; assumption. Defined. @@ -124,7 +124,7 @@ Proposition j_geq_0_lt_implies_pred_geq (i j k : nat) Proof. intros l ineq. destruct j. - - contradiction (not_lt_n_0 i). + - contradiction (not_lt_zero_r i). - now simpl; apply leq_succ'. Defined. @@ -187,7 +187,7 @@ Proof. intros l l'. unfold "<". destruct k, n; - try (contradiction (not_lt_n_0 _ _)). + try (contradiction (not_lt_zero_r _ _)). simpl; apply leq_succ, sub_less. Defined. @@ -264,7 +264,7 @@ Proof. intro a. assert (forall n m: nat, m < n -> P m) as X. { simple_induction n n IHn. - - intros m l. contradiction (not_lt_n_0 m). + - intros m l. contradiction (not_lt_zero_r m). - intros m l. apply leq_succ' in l. destruct l as [ | n]. + apply a; intros ? ?; now apply IHn. @@ -328,7 +328,7 @@ Lemma ineq_sub' (n k : nat) : k < n -> n - k = (n - k.+1).+1. Proof. intro ineq. destruct n. - - contradiction (not_lt_n_0 k). + - contradiction (not_lt_zero_r k). - change (n.+1 - k.+1) with (n - k). apply leq_succ' in ineq. apply (nataddsub_assoc_lemma _). Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 8d7c775acf8..7ee228152b1 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -899,8 +899,7 @@ Defined. (** *** Negation Lemmas *) -(** TODO: rename *) -Lemma not_lt_n_0 n : ~ (n < 0). +Definition not_lt_zero_r n : ~ (n < 0). Proof. apply not_leq_Sn_0. Defined. @@ -992,7 +991,7 @@ Defined. Definition lt_sub_gt_0 n m : m < n -> 0 < n - m. Proof. revert m; simple_induction n n IHn. - - intros m ineq. contradiction (not_lt_n_0 m). + - intros m ineq. contradiction (not_lt_zero_r m). - destruct m. + simpl. easy. + simpl. intro ineq. apply leq_succ' in ineq. @@ -1127,7 +1126,7 @@ Proof. destruct k, m; [ reflexivity | reflexivity - | contradiction (not_lt_n_0 k _) + | contradiction (not_lt_zero_r k _) | ]. simpl "-". apply leq_succ' in l. destruct (nat_add_succ_r n (m - k)). From 2e99a3e26d92662d6184610eb0e66f4c4894ff42 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 16:10:57 +0200 Subject: [PATCH 157/282] Nat: add comments and adjustments to negation lemmas Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 18 ++++++++++-------- 1 file changed, 10 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 7ee228152b1..88b96cafd2e 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -904,29 +904,31 @@ Proof. apply not_leq_Sn_0. Defined. -(** TODO: rename *) -Definition not_lt_implies_geq {n m : nat} : ~(n < m) -> m <= n. +(** There are various lemmas we can state about negating the comparison operators on [nat]. To aid readability, we opt to keep the order of the variables in each statement consistent. *) + +(** TODO: rename not_lt_implies_geq -> geq_not_lt *) +Definition not_lt_implies_geq {n m} : ~(n < m) -> n >= m. Proof. intros not_lt. destruct (@leq_dichot m n); [ assumption | contradiction]. Defined. -(** TODO: rename *) -Definition not_leq_implies_gt {n m : nat} : ~(n <= m) -> m < n. +(** TODO: rename not_leq_implies_gt -> gt_not_leq *) +Definition not_leq_implies_gt {n m} : ~(n <= m) -> n > m. Proof. intros not_leq. destruct (@leq_dichot n m); [ contradiction | assumption]. Defined. -(** TODO: rename *) -Definition lt_implies_not_geq {n m : nat} : (n < m) -> ~(m <= n). +(** TODO: rename lt_implies_not_geq -> not_geq_lt *) +Definition lt_implies_not_geq {n m} : (n < m) -> ~(n >= m). Proof. intros ineq1 ineq2. contradiction (not_lt_n_n n). by apply (leq_trans ineq1). Defined. -(** TODO: rename *) -Definition leq_implies_not_gt {n m : nat} : (n <= m) -> ~(m < n). +(** TODO: rename leq_implies_not_gt -> not_gt_leq *) +Definition leq_implies_not_gt {n m} : (n <= m) -> ~(n > m). Proof. intros ineq1 ineq2. contradiction (not_lt_n_n n); by refine (leq_trans _ ineq2). From db7682b59162686ceafa75c17f79358e17fe6ce5 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 16:13:09 +0200 Subject: [PATCH 158/282] Nat: rename not_lt_implies_geq -> geq_not_lt Signed-off-by: Ali Caglayan --- theories/Spaces/List/Theory.v | 2 +- theories/Spaces/Nat/Core.v | 5 ++--- 2 files changed, 3 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index e9ee1c6bf37..64a442dc557 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -838,7 +838,7 @@ Proof. - lhs nrapply nth_app. 1: by rewrite length_seq. by apply IHn. - - apply not_lt_implies_geq in H'. + - apply geq_not_lt in H'. apply leq_succ' in H. destruct (leq_antisym H H'). lhs nrapply nth_last. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 88b96cafd2e..4f163af20a1 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -906,8 +906,7 @@ Defined. (** There are various lemmas we can state about negating the comparison operators on [nat]. To aid readability, we opt to keep the order of the variables in each statement consistent. *) -(** TODO: rename not_lt_implies_geq -> geq_not_lt *) -Definition not_lt_implies_geq {n m} : ~(n < m) -> n >= m. +Definition geq_not_lt {n m} : ~(n < m) -> n >= m. Proof. intros not_lt. destruct (@leq_dichot m n); [ assumption | contradiction]. @@ -942,7 +941,7 @@ Proof. split. - intros diseq. destruct (dec (n < m)) as [| a]; [ now left |]. - apply not_lt_implies_geq in a. + apply geq_not_lt in a. apply equiv_leq_lt_or_eq in a. destruct a as [lt | eq]. 1: by right. From c33c8ddccd99ec1d55147751d5b76c0ba29a8b4e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 16:14:44 +0200 Subject: [PATCH 159/282] Nat: rename not_leq_implies_gt -> gt_not_leq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 4f163af20a1..66d167c5f82 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -912,8 +912,7 @@ Proof. destruct (@leq_dichot m n); [ assumption | contradiction]. Defined. -(** TODO: rename not_leq_implies_gt -> gt_not_leq *) -Definition not_leq_implies_gt {n m} : ~(n <= m) -> n > m. +Definition gt_not_leq {n m} : ~(n <= m) -> n > m. Proof. intros not_leq. destruct (@leq_dichot n m); [ contradiction | assumption]. From d30985841c49159ac5c85786c2fdf6e7db792244 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 16:16:01 +0200 Subject: [PATCH 160/282] Nat: rename lt_implies_not_geq -> not_geq_lt Signed-off-by: Ali Caglayan --- theories/Algebra/Rings/KroneckerDelta.v | 2 +- theories/Spaces/Nat/Core.v | 3 +-- 2 files changed, 2 insertions(+), 3 deletions(-) diff --git a/theories/Algebra/Rings/KroneckerDelta.v b/theories/Algebra/Rings/KroneckerDelta.v index b78e6822136..58e0297c2b8 100644 --- a/theories/Algebra/Rings/KroneckerDelta.v +++ b/theories/Algebra/Rings/KroneckerDelta.v @@ -112,7 +112,7 @@ Proof. + nrapply IHn. apply neq_iff_lt_or_gt in p. destruct p; [assumption|]. - contradiction (lt_implies_not_geq Hi). + contradiction (not_geq_lt Hi). + rewrite (kronecker_delta_neq p). rewrite rng_mult_zero_l. rewrite grp_unit_l. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 66d167c5f82..5b498062097 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -918,8 +918,7 @@ Proof. destruct (@leq_dichot n m); [ contradiction | assumption]. Defined. -(** TODO: rename lt_implies_not_geq -> not_geq_lt *) -Definition lt_implies_not_geq {n m} : (n < m) -> ~(n >= m). +Definition not_geq_lt {n m} : (n < m) -> ~(n >= m). Proof. intros ineq1 ineq2. contradiction (not_lt_n_n n). by apply (leq_trans ineq1). From a9cdd3e70d1e56dbc93b063e2d0e74bb4c5111f6 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 16:17:14 +0200 Subject: [PATCH 161/282] Nat: rename leq_implies_not_gt -> not_gt_leq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 5b498062097..7f257809f18 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -924,8 +924,7 @@ Proof. contradiction (not_lt_n_n n). by apply (leq_trans ineq1). Defined. -(** TODO: rename leq_implies_not_gt -> not_gt_leq *) -Definition leq_implies_not_gt {n m} : (n <= m) -> ~(n > m). +Definition not_gt_leq {n m} : (n <= m) -> ~(n > m). Proof. intros ineq1 ineq2. contradiction (not_lt_n_n n); by refine (leq_trans _ ineq2). From 0e51bb1bc1530454535f9a00ecca640080cb7783 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 16:21:32 +0200 Subject: [PATCH 162/282] Nat: rename gt_not_leq -> gt_iff_not_leq merge with not_leq_gt Signed-off-by: Ali Caglayan --- theories/Algebra/Rings/KroneckerDelta.v | 3 ++- theories/Spaces/Nat/Core.v | 15 ++++++--------- 2 files changed, 8 insertions(+), 10 deletions(-) diff --git a/theories/Algebra/Rings/KroneckerDelta.v b/theories/Algebra/Rings/KroneckerDelta.v index 58e0297c2b8..75e754294fd 100644 --- a/theories/Algebra/Rings/KroneckerDelta.v +++ b/theories/Algebra/Rings/KroneckerDelta.v @@ -112,7 +112,8 @@ Proof. + nrapply IHn. apply neq_iff_lt_or_gt in p. destruct p; [assumption|]. - contradiction (not_geq_lt Hi). + apply gt_iff_not_leq in Hi. + contradiction Hi. + rewrite (kronecker_delta_neq p). rewrite rng_mult_zero_l. rewrite grp_unit_l. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 7f257809f18..532a4c840c4 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -912,16 +912,13 @@ Proof. destruct (@leq_dichot m n); [ assumption | contradiction]. Defined. -Definition gt_not_leq {n m} : ~(n <= m) -> n > m. +Definition gt_iff_not_leq {n m} : ~(n <= m) <-> n > m. Proof. - intros not_leq. - destruct (@leq_dichot n m); [ contradiction | assumption]. -Defined. - -Definition not_geq_lt {n m} : (n < m) -> ~(n >= m). -Proof. - intros ineq1 ineq2. - contradiction (not_lt_n_n n). by apply (leq_trans ineq1). + split. + - intros not_leq. + destruct (@leq_dichot n m); [ contradiction | assumption]. + - intros ineq1 ineq2. + contradiction (not_lt_n_n m). by apply (leq_trans ineq1). Defined. Definition not_gt_leq {n m} : (n <= m) -> ~(n > m). From 9af5a4cfca46f1b67f100ce141b8a46254eb6317 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 16:27:17 +0200 Subject: [PATCH 163/282] Nat: rename gt_not_leq -> gt_iff_not_leq and merge with not_leq_gt Signed-off-by: Ali Caglayan --- theories/Spaces/List/Theory.v | 2 +- theories/Spaces/Nat/Core.v | 17 +++++++---------- 2 files changed, 8 insertions(+), 11 deletions(-) diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index 64a442dc557..074ce8bf9a8 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -838,7 +838,7 @@ Proof. - lhs nrapply nth_app. 1: by rewrite length_seq. by apply IHn. - - apply geq_not_lt in H'. + - apply geq_iff_not_lt in H'. apply leq_succ' in H. destruct (leq_antisym H H'). lhs nrapply nth_last. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 532a4c840c4..6313a2d2d47 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -906,10 +906,13 @@ Defined. (** There are various lemmas we can state about negating the comparison operators on [nat]. To aid readability, we opt to keep the order of the variables in each statement consistent. *) -Definition geq_not_lt {n m} : ~(n < m) -> n >= m. +Definition geq_iff_not_lt {n m} : ~(n < m) <-> n >= m. Proof. - intros not_lt. - destruct (@leq_dichot m n); [ assumption | contradiction]. + split. + - intros not_lt. + destruct (@leq_dichot m n); [ assumption | contradiction]. + - intros ineq1 ineq2. + contradiction (not_lt_n_n n); by refine (leq_trans _ ineq1). Defined. Definition gt_iff_not_leq {n m} : ~(n <= m) <-> n > m. @@ -921,12 +924,6 @@ Proof. contradiction (not_lt_n_n m). by apply (leq_trans ineq1). Defined. -Definition not_gt_leq {n m} : (n <= m) -> ~(n > m). -Proof. - intros ineq1 ineq2. - contradiction (not_lt_n_n n); by refine (leq_trans _ ineq2). -Defined. - (** *** Dichotomy of [<>] *) (** The inequality of natural numbers is equivalent to [n < m] or [n > m]. This could be an equivalence however one of the sections requires funext since we are comparing two inequality proofs. It is therefore more useful to keep it as a biimplication. Note that this is a negated version of antisymmetry of [<=]. *) @@ -935,7 +932,7 @@ Proof. split. - intros diseq. destruct (dec (n < m)) as [| a]; [ now left |]. - apply geq_not_lt in a. + apply geq_iff_not_lt in a. apply equiv_leq_lt_or_eq in a. destruct a as [lt | eq]. 1: by right. From dc89c8043e061fab5b22b0a3df62efad8cb89178 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 16:35:39 +0200 Subject: [PATCH 164/282] Nat: reprove and complete comparison negation lemmas Signed-off-by: Ali Caglayan --- theories/Algebra/AbGroups/AbelianGroup.v | 4 +--- theories/Spaces/Nat/Core.v | 19 +++++++++++-------- 2 files changed, 12 insertions(+), 11 deletions(-) diff --git a/theories/Algebra/AbGroups/AbelianGroup.v b/theories/Algebra/AbGroups/AbelianGroup.v index 4df356aabb3..e6f2d4fb734 100644 --- a/theories/Algebra/AbGroups/AbelianGroup.v +++ b/theories/Algebra/AbGroups/AbelianGroup.v @@ -302,9 +302,7 @@ Proof. - exact zero. - refine (f n _ + IHn _). intros k Hk. - refine (f k _). - apply leq_succ_r. - exact Hk. + exact (f k _). Defined. (** If the function is constant in the range of a finite sum then the sum is equal to the constant times [n]. This is a group power in the underlying group. *) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 6313a2d2d47..1aac78bb56f 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -407,6 +407,7 @@ Definition leq_trans {x y z} : x <= y -> y <= z -> x <= z. Proof. intros H1 H2; induction H2; exact _. Defined. +Hint Immediate leq_trans : typeclass_instances. (** [<=] is transtiive. *) Global Instance transitive_leq : Transitive leq := @leq_trans. @@ -909,21 +910,23 @@ Defined. Definition geq_iff_not_lt {n m} : ~(n < m) <-> n >= m. Proof. split. - - intros not_lt. - destruct (@leq_dichot m n); [ assumption | contradiction]. - - intros ineq1 ineq2. - contradiction (not_lt_n_n n); by refine (leq_trans _ ineq1). + - intro; by destruct (leq_dichot m n). + - intros ? ?; contradiction (not_lt_n_n n); exact _. Defined. Definition gt_iff_not_leq {n m} : ~(n <= m) <-> n > m. Proof. split. - - intros not_leq. - destruct (@leq_dichot n m); [ contradiction | assumption]. - - intros ineq1 ineq2. - contradiction (not_lt_n_n m). by apply (leq_trans ineq1). + - intro; by destruct (leq_dichot n m). + - intros ? ?; contradiction (not_lt_n_n m); exact _. Defined. +Definition leq_iff_not_gt {n m} : ~(n > m) <-> n <= m + := geq_iff_not_lt. + +Definition lt_iff_not_geq {n m} : ~(n >= m) <-> n < m + := gt_iff_not_leq. + (** *** Dichotomy of [<>] *) (** The inequality of natural numbers is equivalent to [n < m] or [n > m]. This could be an equivalence however one of the sections requires funext since we are comparing two inequality proofs. It is therefore more useful to keep it as a biimplication. Note that this is a negated version of antisymmetry of [<=]. *) From dbf8cc8df457b76d7ac3d50259ff1d14e7229ced Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 16:45:30 +0200 Subject: [PATCH 165/282] Nat: update comment on alternative chracterizations of <= Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 1aac78bb56f..6850a12c021 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -833,10 +833,10 @@ Proof. - exact (leq_succ_r IH). Defined. -(** Characterizations of [<=] *) +(** Alternative Characterizations of [<=] *) (** [n <= m] is equivalent to [(n < m) + (n = m)]. Note that it is not immediately obvious that the latter type is a hprop, hence we have to explicitly show the back and forth maps are inverses of eachother. This is possible and justifies the name "less than or equal to". *) -Definition equiv_leq_lt_or_eq {n m : nat} : (n <= m) <~> (n < m) + (n = m). +Definition equiv_leq_lt_or_eq {n m} : (n <= m) <~> (n < m) + (n = m). Proof. snrapply equiv_adjointify. - intro l; induction l. From 1dbb72b1f7a0af690f525c085e57b143323c3c80 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 16:57:01 +0200 Subject: [PATCH 166/282] Nat: rename nataddpreservesleq -> nat_add_r_monotone Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 8 +++----- 1 file changed, 3 insertions(+), 5 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 6850a12c021..aa0bd0d6cc9 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -997,9 +997,7 @@ Defined. (** TODO: use OrderPreserving from canonical_names *) -(** TODO: rename nataddpreservesleq -> *) -Proposition nataddpreservesleq { n m k : nat } - : n <= m -> n + k <= m + k. +Definition nat_add_r_monotone {n m} k : n <= m -> n + k <= m + k. Proof. intro l. simple_induction k k IHk. @@ -1014,7 +1012,7 @@ Proposition nataddpreservesleq' { n m k : nat } Proof. destruct (symmetric_paths _ _ (nat_add_comm k m)), (symmetric_paths _ _ (nat_add_comm k n)). - exact nataddpreservesleq. + exact (nat_add_r_monotone _). Defined. (** TODO: move, rename *) @@ -1087,7 +1085,7 @@ Proposition natsubreflectsleq { n m k : nat } : k <= m -> n - k <= m - k -> n <= m. Proof. intros ineq1 ineq2. - apply (@nataddpreservesleq _ _ k) in ineq2. + apply (nat_add_r_monotone k) in ineq2. apply (@leq_trans _ (n - k + k) _ (leq_sub_add _ _)). apply (@leq_trans _ (m - k + k) _ _). destruct (nat_add_sub_l_cancel ineq1)^; easy. From f19e0ee81617148ebf1eaafe7a147b8245c4fe55 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 17:00:48 +0200 Subject: [PATCH 167/282] Nat: update comment on monotonicity Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 +++- 1 file changed, 3 insertions(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index aa0bd0d6cc9..e7d6ea79019 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -993,7 +993,7 @@ Proof. now apply IHn in ineq. Defined. -(** *** Monotonicity *) +(** *** Monotonicity of Addition *) (** TODO: use OrderPreserving from canonical_names *) @@ -1015,6 +1015,8 @@ Proof. exact (nat_add_r_monotone _). Defined. +(** *** Strict Monotonicity of Addition *) + (** TODO: move, rename *) Proposition nataddpreserveslt { n m k : nat } : n < m -> n + k < m + k. From ef63e9b071c4345d037709be4c66c8116d30b018 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 17:02:05 +0200 Subject: [PATCH 168/282] Nat: rename nataddpreservesleq' -> nat_add_l_monotone Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 4 ++-- theories/Spaces/Nat/Core.v | 4 +--- 2 files changed, 3 insertions(+), 5 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 07126d22175..2bc9617d8ff 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -208,7 +208,7 @@ Proposition natpmswap2 (k m n : nat) : n <= k -> k - n <= m -> k <= n + m. Proof. intros l q. - apply (@nataddpreservesleq' (k - n) m n) in q. + apply (nat_add_l_monotone n) in q. destruct (nataddsub_assoc n l)^ in q. destruct (nat_add_sub_cancel_l k n)^ in q; assumption. @@ -219,7 +219,7 @@ Proposition natpmswap3 (k m n : nat) : k <= n -> m <= n - k -> k + m <= n. Proof. intros ineq qe. - apply (@nataddpreservesleq' m (n - k) k) in qe. + apply (nat_add_l_monotone k) in qe. destruct (nataddsub_assoc k ineq)^ in qe. destruct (nat_add_sub_cancel_l n k)^ in qe; assumption. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index e7d6ea79019..4d81e42c3fc 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1006,9 +1006,7 @@ Proof. apply leq_succ; exact IHk. Defined. -(** TODO: move, rename *) -Proposition nataddpreservesleq' { n m k : nat } - : n <= m -> k + n <= k + m. +Definition nat_add_l_monotone {n m} k : n <= m -> k + n <= k + m. Proof. destruct (symmetric_paths _ _ (nat_add_comm k m)), (symmetric_paths _ _ (nat_add_comm k n)). From d16b980969d9a51f5df9522034de6550320a83bd Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 17:07:10 +0200 Subject: [PATCH 169/282] Nat: reprove monotonicity lemmas Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 16 ++++++---------- 1 file changed, 6 insertions(+), 10 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 4d81e42c3fc..db8332ccf7d 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -997,21 +997,17 @@ Defined. (** TODO: use OrderPreserving from canonical_names *) -Definition nat_add_r_monotone {n m} k : n <= m -> n + k <= m + k. +Definition nat_add_l_monotone {n m} k : n <= m -> k + n <= k + m. Proof. - intro l. - simple_induction k k IHk. - - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; exact l. - - destruct (nat_add_succ_r n k)^, (nat_add_succ_r m k)^; - apply leq_succ; exact IHk. + intros H; induction k as [|k IHk] in n, m, H |- *; exact _. Defined. +Hint Immediate nat_add_l_monotone : typeclass_instances. -Definition nat_add_l_monotone {n m} k : n <= m -> k + n <= k + m. +Definition nat_add_r_monotone {n m} k : n <= m -> n + k <= m + k. Proof. - destruct (symmetric_paths _ _ (nat_add_comm k m)), - (symmetric_paths _ _ (nat_add_comm k n)). - exact (nat_add_r_monotone _). + intros H; rewrite 2 (nat_add_comm _ k); exact _. Defined. +Hint Immediate nat_add_r_monotone : typeclass_instances. (** *** Strict Monotonicity of Addition *) From 883721d05cd999aae1148110185fa23322c94b25 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 17:08:13 +0200 Subject: [PATCH 170/282] Nat: move nat_add_bifunctor Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 15 --------------- theories/Spaces/Nat/Core.v | 15 +++++++++++++++ 2 files changed, 15 insertions(+), 15 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 2bc9617d8ff..e7c5fa7efd2 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -241,21 +241,6 @@ Proof. intro l; apply (leq_trans l); exact (leq_add_l m k). Defined. -(** TODO: move, rename *) -Proposition nat_add_bifunctor (n n' m m' : nat) - : n <= m -> n' <= m' -> n + n' <= m + m'. -Proof. - revert n' m m'; simple_induction n n IHn. - - intros n' m m' l l'. simpl. - apply (leq_trans l'). exact (leq_add_r m' m). - - intros n' m; destruct m. - + intros. contradiction (not_leq_Sn_0 n). - + intros m' l l'. apply leq_succ' in l. simpl. - apply leq_succ, IHn. - * exact l. - * exact l'. -Defined. - (** TODO: move, rename *) Proposition strong_induction (P : nat -> Type) : (forall n : nat, (forall m : nat, (m < n) -> P m) -> P n) -> diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index db8332ccf7d..108846e8347 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1009,6 +1009,21 @@ Proof. Defined. Hint Immediate nat_add_r_monotone : typeclass_instances. +(** TODO: rename nat_add_bifunctor -> nat_add_monotone *) +Definition nat_add_bifunctor {n n' m m'} + : n <= m -> n' <= m' -> n + n' <= m + m'. +Proof. + revert n' m m'; simple_induction n n IHn. + - intros n' m m' l l'. simpl. + apply (leq_trans l'). exact (leq_add_r m' m). + - intros n' m; destruct m. + + intros. contradiction (not_leq_Sn_0 n). + + intros m' l l'. apply leq_succ' in l. simpl. + apply leq_succ, IHn. + * exact l. + * exact l'. +Defined. + (** *** Strict Monotonicity of Addition *) (** TODO: move, rename *) From 18f8757b76bfe59e322a4ba60b204e0db03392cd Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 17:08:58 +0200 Subject: [PATCH 171/282] Nat: rename nat_add_bifunctor -> nat_add_monotone Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 108846e8347..b7dd98b3b3d 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1009,8 +1009,7 @@ Proof. Defined. Hint Immediate nat_add_r_monotone : typeclass_instances. -(** TODO: rename nat_add_bifunctor -> nat_add_monotone *) -Definition nat_add_bifunctor {n n' m m'} +Definition nat_add_monotone {n n' m m'} : n <= m -> n' <= m' -> n + n' <= m + m'. Proof. revert n' m m'; simple_induction n n IHn. From b569a2541050a7f1773239cb40e3c19543d1f109 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 17:11:10 +0200 Subject: [PATCH 172/282] Nat: add comments and clean up monotonicity of addition proof Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 14 +++++--------- 1 file changed, 5 insertions(+), 9 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index b7dd98b3b3d..09c2e1df109 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -997,31 +997,27 @@ Defined. (** TODO: use OrderPreserving from canonical_names *) +(** Addition on the left is monotone. *) Definition nat_add_l_monotone {n m} k : n <= m -> k + n <= k + m. Proof. intros H; induction k as [|k IHk] in n, m, H |- *; exact _. Defined. Hint Immediate nat_add_l_monotone : typeclass_instances. +(** Addition on the right is monotone. *) Definition nat_add_r_monotone {n m} k : n <= m -> n + k <= m + k. Proof. intros H; rewrite 2 (nat_add_comm _ k); exact _. Defined. Hint Immediate nat_add_r_monotone : typeclass_instances. +(** Addition is monotone in both arguments. (This makes [+] a bifunctor when treating [nat] as a category (as a preorder)). *) Definition nat_add_monotone {n n' m m'} : n <= m -> n' <= m' -> n + n' <= m + m'. Proof. - revert n' m m'; simple_induction n n IHn. - - intros n' m m' l l'. simpl. - apply (leq_trans l'). exact (leq_add_r m' m). - - intros n' m; destruct m. - + intros. contradiction (not_leq_Sn_0 n). - + intros m' l l'. apply leq_succ' in l. simpl. - apply leq_succ, IHn. - * exact l. - * exact l'. + intros H1 H2; induction H1; exact _. Defined. +Hint Immediate nat_add_monotone : typeclass_isntances. (** *** Strict Monotonicity of Addition *) From c60fbf709141d74c5ac60e5e055dcf8264a8d46b Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 17:13:38 +0200 Subject: [PATCH 173/282] Nat: rename nataddpreserveslt -> nat_add_r_strictly_monotone Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 5 ++--- 2 files changed, 3 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index e7c5fa7efd2..8185ce0ebc3 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -229,7 +229,7 @@ Defined. Proposition natpmswap4 (k m n : nat) : k - n < m -> k < n + m. Proof. - intro l; apply (@nataddpreserveslt (k - n) m n) in l. + intro l; apply (nat_add_r_strictly_monotone n) in l. destruct (nat_add_comm m n). now rapply (leq_lt_trans (nat_sub_add_ineq _ _)). Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 09c2e1df109..5b1416b0fc2 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1021,8 +1021,7 @@ Hint Immediate nat_add_monotone : typeclass_isntances. (** *** Strict Monotonicity of Addition *) -(** TODO: move, rename *) -Proposition nataddpreserveslt { n m k : nat } +Definition nat_add_r_strictly_monotone {n m} k : n < m -> n + k < m + k. Proof. unfold "<". @@ -1040,7 +1039,7 @@ Proposition nataddpreserveslt' { n m k : nat } Proof. destruct (symmetric_paths _ _ (nat_add_comm k n)), (symmetric_paths _ _ (nat_add_comm k m)); - exact nataddpreserveslt. + exact (nat_add_r_strictly_monotone k). Defined. (** TODO: monotonicity of subtraction *) From e020c8b0e5f970de320f1ffd49b55e0111aac9f4 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 17:16:14 +0200 Subject: [PATCH 174/282] Nat: rename nataddpreserveslt' -> nat_add_l_strictly_monotone Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 5b1416b0fc2..d5d6e83bfdd 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1033,8 +1033,7 @@ Proof. apply leq_succ; exact IHk. Defined. -(** TODO: move, rename *) -Proposition nataddpreserveslt' { n m k : nat } +Definition nat_add_l_strictly_monotone {n m} k : n < m -> k + n < k + m. Proof. destruct (symmetric_paths _ _ (nat_add_comm k n)), From fa504f964cf1da0cf45276ce598352fc25f1cfcd Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 17:23:23 +0200 Subject: [PATCH 175/282] Nat: reprove and reorganise strict monotinicity lemmas for add Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 33 +++++++++++++++++++++------------ 1 file changed, 21 insertions(+), 12 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index d5d6e83bfdd..c33e941dd97 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1021,25 +1021,34 @@ Hint Immediate nat_add_monotone : typeclass_isntances. (** *** Strict Monotonicity of Addition *) +(** [nat_succ] is strictly monotone. *) +Global Instance lt_succ {n m} : n < m -> n.+1 < m.+1 := _. + +Global Instance lt_succ_r {n m} : n < m -> n < m.+1 := _. + +(** Addition on the left is strictly monotone. *) +Definition nat_add_l_strictly_monotone {n m} k + : n < m -> k + n < k + m. +Proof. + intros H; induction k as [|k IHk] in n, m, H |- *; exact _. +Defined. +Hint Immediate nat_add_l_strictly_monotone : typeclass_instances. + +(** Addition on the right is strictly monotone. *) Definition nat_add_r_strictly_monotone {n m} k : n < m -> n + k < m + k. Proof. - unfold "<". - change (n + k).+1 with (n.+1 + k). - generalize (n.+1). intros n' l. - simple_induction k k IHk. - - destruct (nat_add_zero_r n')^, (nat_add_zero_r m)^; exact l. - - destruct (nat_add_succ_r n' k)^, (nat_add_succ_r m k)^; - apply leq_succ; exact IHk. + intros H; rewrite 2 (nat_add_comm _ k); exact _. Defined. +Hint Immediate nat_add_r_strictly_monotone : typeclass_instances. -Definition nat_add_l_strictly_monotone {n m} k - : n < m -> k + n < k + m. +(** Addition is strictly monotone in both arguments. *) +Definition nat_add_strictly_monotone {n n' m m'} + : n < m -> n' < m' -> n + n' < m + m'. Proof. - destruct (symmetric_paths _ _ (nat_add_comm k n)), - (symmetric_paths _ _ (nat_add_comm k m)); - exact (nat_add_r_strictly_monotone k). + intros H1 H2; induction H1; exact _. Defined. +Hint Immediate nat_add_strictly_monotone : typeclass_instances. (** TODO: monotonicity of subtraction *) (** TODO: monotonicity of multiplication *) From 1a46f80243f2bef92e865774d9481c5a4a5e3c08 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 17:32:09 +0200 Subject: [PATCH 176/282] Nat: monotonicity of multiplication Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 46 +++++++++++++++++++++++++++----------- 1 file changed, 33 insertions(+), 13 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index c33e941dd97..70978dc6f9f 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -998,26 +998,25 @@ Defined. (** TODO: use OrderPreserving from canonical_names *) (** Addition on the left is monotone. *) -Definition nat_add_l_monotone {n m} k : n <= m -> k + n <= k + m. +Global Instance nat_add_l_monotone {n m} k + : n <= m -> k + n <= k + m. Proof. intros H; induction k as [|k IHk] in n, m, H |- *; exact _. Defined. -Hint Immediate nat_add_l_monotone : typeclass_instances. (** Addition on the right is monotone. *) -Definition nat_add_r_monotone {n m} k : n <= m -> n + k <= m + k. +Global Instance nat_add_r_monotone {n m} k + : n <= m -> n + k <= m + k. Proof. intros H; rewrite 2 (nat_add_comm _ k); exact _. Defined. -Hint Immediate nat_add_r_monotone : typeclass_instances. (** Addition is monotone in both arguments. (This makes [+] a bifunctor when treating [nat] as a category (as a preorder)). *) -Definition nat_add_monotone {n n' m m'} +Global Instance nat_add_monotone {n n' m m'} : n <= m -> n' <= m' -> n + n' <= m + m'. Proof. intros H1 H2; induction H1; exact _. Defined. -Hint Immediate nat_add_monotone : typeclass_isntances. (** *** Strict Monotonicity of Addition *) @@ -1027,31 +1026,52 @@ Global Instance lt_succ {n m} : n < m -> n.+1 < m.+1 := _. Global Instance lt_succ_r {n m} : n < m -> n < m.+1 := _. (** Addition on the left is strictly monotone. *) -Definition nat_add_l_strictly_monotone {n m} k +Global Instance nat_add_l_strictly_monotone {n m} k : n < m -> k + n < k + m. Proof. intros H; induction k as [|k IHk] in n, m, H |- *; exact _. Defined. -Hint Immediate nat_add_l_strictly_monotone : typeclass_instances. (** Addition on the right is strictly monotone. *) -Definition nat_add_r_strictly_monotone {n m} k +Global Instance nat_add_r_strictly_monotone {n m} k : n < m -> n + k < m + k. Proof. intros H; rewrite 2 (nat_add_comm _ k); exact _. Defined. -Hint Immediate nat_add_r_strictly_monotone : typeclass_instances. (** Addition is strictly monotone in both arguments. *) -Definition nat_add_strictly_monotone {n n' m m'} +Global Instance nat_add_strictly_monotone {n n' m m'} : n < m -> n' < m' -> n + n' < m + m'. Proof. intros H1 H2; induction H1; exact _. Defined. -Hint Immediate nat_add_strictly_monotone : typeclass_instances. + +(** *** Monotonicity of Multiplication *) + +(** Multiplication on the left is monotone. *) +Global Instance nat_mul_l_monotone {n m} k + : n <= m -> k * n <= k * m. +Proof. + intros H; induction k as [|k IHk] in |- *; exact _. +Defined. + +(** Multiplication on the right is monotone. *) +Global Instance nat_mul_r_monotone {n m} k + : n <= m -> n * k <= m * k. +Proof. + intros H; rewrite 2 (nat_mul_comm _ k); exact _. +Defined. + +(** Multiplication is monotone in both arguments. *) +Global Instance nat_mul_monotone {n n' m m'} + : n <= m -> n' <= m' -> n * n' <= m * m'. +Proof. + intros H1 H2; induction H1. + - exact _. + - rapply leq_trans. +Defined. (** TODO: monotonicity of subtraction *) -(** TODO: monotonicity of multiplication *) (** *** Order-reflection *) From 23a2dd6d0511ba160e6cc0282c0aef60b1f428de Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 30 Jul 2024 17:36:48 +0200 Subject: [PATCH 177/282] Nat: add conditional strict monotonicity of multiplication Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 18 +++++++++++++++++- 1 file changed, 17 insertions(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 70978dc6f9f..daea1973c5d 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -420,7 +420,7 @@ Defined. Global Existing Instance leq_zero | 10. (** A predecessor is less than or equal to a predecessor if the original number is less than or equal. *) -Definition leq_pred {n m} : leq n m -> leq (nat_pred n) (nat_pred m). +Definition leq_pred {n m} : n <= m -> nat_pred n <= nat_pred m. Proof. intros H; induction H. 1: exact _. @@ -1071,6 +1071,22 @@ Proof. - rapply leq_trans. Defined. +(** *** Strict Monotonicity of Multiplication *) + +(** Multiplication on the left by a positive number is strictly monotone. *) +Global Instance nat_mul_l_strictly_monotone {n m} k + : n < m -> k.+1 * n < k.+1 * m. +Proof. + intros H; induction k as [|k IHk] in |- *; exact _. +Defined. + +(** Multiplication on the right by a positive number is strictly monotone. *) +Global Instance nat_mul_r_strictly_monotone {n m} k + : n < m -> n * k.+1 < m * k.+1. +Proof. + intros H; rewrite 2 (nat_mul_comm _ k.+1); exact _. +Defined. + (** TODO: monotonicity of subtraction *) (** *** Order-reflection *) From 08775a6bc809158f3dee9ac9d3bd6a73d609b358 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 12:27:07 +0200 Subject: [PATCH 178/282] Nat: fixups later in library Signed-off-by: Ali Caglayan --- theories/Algebra/Rings/Matrix.v | 13 +++++-------- 1 file changed, 5 insertions(+), 8 deletions(-) diff --git a/theories/Algebra/Rings/Matrix.v b/theories/Algebra/Rings/Matrix.v index b2aa8b6885f..e8c2dea4db9 100644 --- a/theories/Algebra/Rings/Matrix.v +++ b/theories/Algebra/Rings/Matrix.v @@ -596,17 +596,14 @@ Proof. contradiction (H' (leq_trans _ H)). - destruct p. contradiction (H (leq_trans _ H')). - - by apply path_nat_S. + - by apply path_nat_succ. Defined. Local Instance lt_n1_skip k i n (H : (i < n.+1)%nat) (H' : (k < n)%nat) : (skip i k < n.+1)%nat. Proof. unfold skip. - destruct (dec (k < i))%nat as [H''|H'']. - - exact (transitive_lt _ _ _ H'' H) . - - apply leq_S_n'. - exact H'. + destruct (dec (k < i))%nat as [H''|H'']; exact _. Defined. Definition matrix_minor {R : Ring@{i}} {n : nat} (i j : nat) @@ -759,8 +756,8 @@ Proof. destruct (dec (k <= i)%nat) as [leq_k_i|gt_k_i]. { rewrite H2. 1: by rewrite rng_mult_zero_r. - rapply mixed_trans1. } - apply not_leq_implies_gt in gt_k_i. + rapply leq_lt_trans. } + apply gt_iff_not_leq in gt_k_i. rewrite H1. 1: by rewrite rng_mult_zero_l. assumption. @@ -932,7 +929,7 @@ Proof. rewrite kronecker_delta_refl. rewrite rng_mult_one_l. f_ap; apply path_ishprop. } - apply diseq_implies_lt in np. + apply neq_iff_lt_or_gt in np. destruct np as [l | l]. - rewrite (kronecker_delta_lt l). rewrite rng_mult_zero_l. From 6705c0dceae97bec148541d25a7c34c509b29699 Mon Sep 17 00:00:00 2001 From: Dan Christensen Date: Sat, 3 Aug 2024 09:41:23 -0400 Subject: [PATCH 179/282] EMSpace: every n-connected and (n+1)-truncated type is an EM space --- theories/Homotopy/EMSpace.v | 41 ++++++++++++++++++++++++++++++++++--- theories/Homotopy/Hopf.v | 12 +++++++++++ 2 files changed, 50 insertions(+), 3 deletions(-) diff --git a/theories/Homotopy/EMSpace.v b/theories/Homotopy/EMSpace.v index 5f14e8b5aae..b433d8e5306 100644 --- a/theories/Homotopy/EMSpace.v +++ b/theories/Homotopy/EMSpace.v @@ -10,7 +10,7 @@ Require Import Homotopy.Hopf. Require Import Truncations.Core Truncations.Connectedness. Require Import WildCat. -(* Formalisation of Eilenberg-MacLane spaces *) +(** * Eilenberg-Mac Lane spaces *) Local Open Scope pointed_scope. Local Open Scope nat_scope. @@ -57,7 +57,7 @@ Section EilenbergMacLane. Local Open Scope trunc_scope. - (* This is a variant of [pequiv_ptr_loop_psusp] from pSusp.v. All we are really using is that [n.+2 <= n +2+ n], but because of the use of [isconnmap_pred_add], the proof is a bit more specific to this case. *) + (** This is a variant of [pequiv_ptr_loop_psusp] from pSusp.v. All we are really using is that [n.+2 <= n +2+ n], but because of the use of [isconnmap_pred_add], the proof is a bit more specific to this case. *) Local Lemma pequiv_ptr_loop_psusp' (X : pType) (n : nat) `{IsConnected n.+1 X} : pTr n.+2 X <~>* pTr n.+2 (loops (psusp X)). Proof. @@ -91,7 +91,7 @@ Section EilenbergMacLane. exact (emap (iterated_loops n) (pequiv_loops_em_em _ _)). Defined. - (* For positive indices, we in fact get a group isomorphism. *) + (** For positive indices, we in fact get a group isomorphism. *) Definition equiv_g_pi_n_em (G : AbGroup) (n : nat) : GroupIsomorphism G (Pi n.+1 K(G, n.+1)). Proof. @@ -110,4 +110,39 @@ Section EilenbergMacLane. apply iscohhspace_loops. Defined. + (** If [G] and [G'] are isomorphic, then [K(G,n)] and [K(G',n)] are equivalent. TODO: We should show that [K(-,n)] is a functor, which implies this. *) + Definition pequiv_em_group_iso {G G' : Group} (n : nat) + (e : G $<~> G') + : K(G, n) <~>* K(G', n). + Proof. + by destruct (equiv_path_group e). + Defined. + + (** Every pointed (n-1)-connected n-type is an Eilenberg-Mac Lane space. *) + Definition pequiv_em_connected_truncated (X : pType) + (n : nat) `{IsConnected n X} `{IsTrunc n.+1 X} + : K(Pi n.+1 X, n.+1) <~>* X. + Proof. + generalize dependent X; induction n; intros X isC isT. + 1: rapply pequiv_pclassifyingspace_pi1. + (* The equivalence will be the composite +<< + K( (Pi n.+2 X) n.+2) + <~>* K( (Pi n.+1 (loops X)), n.+2) + = pTr n.+2 (psusp K( (Pi n.+1 (loops X)), n.+1)) [by definition] + <~>* pTr n.+2 (psusp (loops X)) + <~>* pTr n.+2 X + <~>* X +>> + and we'll work from right to left. +*) + refine ((pequiv_ptr (n:=n.+2))^-1* o*E _). + refine (pequiv_ptr_psusp_loops X n o*E _). + change (K(?G, n.+2)) with (pTr n.+2 (psusp K( G, n.+1 ))). + refine (emap (pTr n.+2 o psusp) _). + refine ((IHn (loops X) _ _) o*E _). + apply pequiv_em_group_iso. + apply groupiso_pi_loops. + Defined. + End EilenbergMacLane. diff --git a/theories/Homotopy/Hopf.v b/theories/Homotopy/Hopf.v index b01bb22fc25..16da9f43743 100644 --- a/theories/Homotopy/Hopf.v +++ b/theories/Homotopy/Hopf.v @@ -176,3 +176,15 @@ Proof. + nrapply (isconnected_equiv' _ _ (pequiv_pfiber_loops_susp_counit_join X)^-1). nrapply isconnected_join; exact _. Defined. + +(** In particular, we get the following result. All we are really using is that [n.+2 <= n +2+ n], but because of the use of [isconnmap_pred_add], the proof is a bit more specific to this case. *) +Definition pequiv_ptr_psusp_loops `{Univalence} (X : pType) (n : nat) `{IsConnected n.+1 X} + : pTr n.+2 (psusp (loops X)) <~>* pTr n.+2 X. +Proof. + snrapply Build_pEquiv. + 1: rapply (fmap (pTr _) (loop_susp_counit _)). + nrapply O_inverts_conn_map. + nrapply (isconnmap_pred_add n.-2). + rewrite 2 trunc_index_add_succ. + rapply (conn_map_loop_susp_counit X). +Defined. From a1f351dac1fdad6b04f8fdd6005b673988892166 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 12:47:34 +0200 Subject: [PATCH 180/282] Nat: improve proof of trunc_index_add_nat_add Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 17 +++++++---------- 1 file changed, 7 insertions(+), 10 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index daea1973c5d..aca5fb92b55 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -946,18 +946,15 @@ Defined. (** ** Arithmetic relations between [trunc_index] and [nat]. *) -Lemma trunc_index_add_nat_add (n : nat) +Definition trunc_index_add_nat_add {n : nat} : trunc_index_add n n = n.+1 + n.+1. Proof. - induction n as [|n IH]; only 1: reflexivity. - refine (trunc_index_add_succ _ _ @ _). - refine (ap trunc_S _ @ _). - { refine (trunc_index_add_comm _ _ @ _). - refine (trunc_index_add_succ _ _ @ _). - exact (ap trunc_S IH). } - refine (_ @ ap nat_to_trunc_index _). - 2: exact (nat_add_succ_r _ _)^. - reflexivity. + induction n as [|n IHn]. + 1: reflexivity. + lhs nrapply trunc_index_add_succ. + rhs nrapply (ap nat_to_trunc_index). + 2: nrapply nat_add_succ_r. + exact (ap (fun x => x.+2%trunc) IHn). Defined. (** *** Subtraction *) From 42745b3349d19dc549e243875d93fc846fd95bc0 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 12:56:08 +0200 Subject: [PATCH 181/282] Nat: rework proof of equiv_leq_lt_eq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 21 ++++++++++----------- 1 file changed, 10 insertions(+), 11 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index aca5fb92b55..4911449d904 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -835,10 +835,17 @@ Defined. (** Alternative Characterizations of [<=] *) -(** [n <= m] is equivalent to [(n < m) + (n = m)]. Note that it is not immediately obvious that the latter type is a hprop, hence we have to explicitly show the back and forth maps are inverses of eachother. This is possible and justifies the name "less than or equal to". *) +(** [n <= m] is equivalent to [(n < m) + (n = m)]. This justifies the name "less than or equal to". Note that it is not immediately obvious that the latter type is a hprop. *) Definition equiv_leq_lt_or_eq {n m} : (n <= m) <~> (n < m) + (n = m). Proof. - snrapply equiv_adjointify. + srapply equiv_iff_hprop. + - srapply hprop_allpath. + intros x y. + snrapply equiv_path_sum. + destruct x as [l|p], y as [q|r]. + 1,4: rapply path_ishprop. + + destruct r; contradiction (not_lt_n_n _ _). + + destruct p; contradiction (not_lt_n_n _ _). - intro l; induction l. + now right. + left; exact (leq_succ l). @@ -846,13 +853,6 @@ Proof. + exact (leq_succ_l l). + destruct p. exact (leq_refl _). - - intros [l|p]. - + induction l. - 1: reflexivity. - snrapply (ap (inl)). - rapply path_ishprop. - + by destruct p. - - intro; rapply path_ishprop. Defined. (** Here is an alternative characterization of [<=] in terms of an existence predicate and addition. *) @@ -946,8 +946,7 @@ Defined. (** ** Arithmetic relations between [trunc_index] and [nat]. *) -Definition trunc_index_add_nat_add {n : nat} - : trunc_index_add n n = n.+1 + n.+1. +Definition trunc_index_add_nat_add {n : nat}: trunc_index_add n n = n.+1 + n.+1. Proof. induction n as [|n IHn]. 1: reflexivity. From eed0c410e1cc75d9758a9824396b964a2a61f82d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 13:03:40 +0200 Subject: [PATCH 182/282] Nat: revise proof of leq_dichot Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 20 +++++++++----------- 1 file changed, 9 insertions(+), 11 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 4911449d904..d1b4e7025e6 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -875,17 +875,15 @@ Defined. (** *** Dichotomy of [<=] *) -(** TODO: rename *) -Fixpoint leq_dichot m n : (m <= n) + (m > n). -Proof. - simple_induction' m; simple_induction' n. - - left; reflexivity. - - left; apply leq_zero. - - right; unfold lt; exact _. - - assert ((m <= n) + (n < m)) as X by apply leq_dichot. - destruct X as [leqmn|ltnm]. - + left; apply leq_succ; assumption. - + right; apply leq_succ; assumption. +Definition leq_dichot m n : (m <= n) + (m > n). +Proof. + induction m as [|m IHm] in n |- *. + 1: left; exact _. + destruct n. + 1: right; exact _. + destruct (IHm n). + 1: left; exact _. + 1: right; exact _. Defined. (** *** Trichotomy *) From fd10f2a43c43e08ff8d98b88f4a10386a3fd677f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 13:04:05 +0200 Subject: [PATCH 183/282] Nat: rename leq_dichot -> leq_dichotomy Signed-off-by: Ali Caglayan --- theories/BoundedSearch.v | 4 ++-- theories/Spaces/Nat/Arithmetic.v | 8 ++++---- theories/Spaces/Nat/Core.v | 12 ++++++------ 3 files changed, 12 insertions(+), 12 deletions(-) diff --git a/theories/BoundedSearch.v b/theories/BoundedSearch.v index 19f66d29c71..2af95c4891a 100644 --- a/theories/BoundedSearch.v +++ b/theories/BoundedSearch.v @@ -57,14 +57,14 @@ Section bounded_search. * left. refine (n.+1;(h,_,_)). -- intros m pm. - assert ((n.+1 <= m)+(n.+1>m)) as X by apply leq_dichot. + assert ((n.+1 <= m)+(n.+1>m)) as X by apply leq_dichotomy. destruct X as [leqSnm|ltmSn]. ++ assumption. ++ unfold gt, lt in ltmSn. assert (m <= n) as X by rapply leq_succ'. destruct (n0 m X pm). * right. intros l q. - assert ((l <= n) + (l > n)) as X by apply leq_dichot. + assert ((l <= n) + (l > n)) as X by apply leq_dichotomy. destruct X as [h|h]. -- exact (n0 l h). -- unfold lt in h. diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 8185ce0ebc3..70c8dcfd7b6 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -32,7 +32,7 @@ Defined. Proposition nat_sub_add_ineq (n m : nat) : n <= n - m + m. Proof. - destruct (@leq_dichot m n) as [l | gt]. + destruct (@leq_dichotomy m n) as [l | gt]. - destruct (nataddsub_comm _ _ m l)^. destruct (nat_add_sub_cancel_r n m)^. apply leq_refl; done. @@ -296,7 +296,7 @@ Lemma increasing_geq_minus (n k : nat) Proof. simple_induction k k IHk. - destruct (symmetric_paths _ _ (nat_sub_zero_r n)); constructor. - - destruct (@leq_dichot n k) as [l | g]. + - destruct (@leq_dichotomy n k) as [l | g]. + destruct (equiv_nat_sub_leq _)^ in IHk. apply leq_succ_r in l. destruct (equiv_nat_sub_leq _)^. exact IHk. @@ -357,7 +357,7 @@ Defined. Definition leq_wrapper {n m : nat} : n <= m -> n <= m. Proof. intro ineq. - destruct (@leq_dichot n m) as [l | g]. + destruct (@leq_dichotomy n m) as [l | g]. - exact l. - contradiction (not_lt_n_n m (lt_leq_trans g ineq)). Defined. @@ -367,7 +367,7 @@ Proposition symmetric_rel_total_order (R : nat -> nat -> Type) : (forall n m : nat, n < m -> R n m) -> (forall n m : nat, R n m). Proof. intros A n m. - destruct (@leq_dichot m n) as [m_leq_n | m_gt_n]. + destruct (@leq_dichotomy m n) as [m_leq_n | m_gt_n]. - apply symmetry. destruct m_leq_n. + apply reflexivity. + apply A. apply leq_succ. assumption. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index d1b4e7025e6..2188fa8d660 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -875,7 +875,7 @@ Defined. (** *** Dichotomy of [<=] *) -Definition leq_dichot m n : (m <= n) + (m > n). +Definition leq_dichotomy m n : (m <= n) + (m > n). Proof. induction m as [|m IHm] in n |- *. 1: left; exact _. @@ -891,7 +891,7 @@ Defined. (** Every two natural numbers are either equal, less than, or greater than each other. *) Definition nat_trichotomy m n : (m < n) + (m = n) + (m > n). Proof. - generalize (leq_dichot m n). + generalize (leq_dichotomy m n). snrapply (functor_sum _ idmap). snrapply equiv_leq_lt_or_eq. Defined. @@ -908,14 +908,14 @@ Defined. Definition geq_iff_not_lt {n m} : ~(n < m) <-> n >= m. Proof. split. - - intro; by destruct (leq_dichot m n). + - intro; by destruct (leq_dichotomy m n). - intros ? ?; contradiction (not_lt_n_n n); exact _. Defined. Definition gt_iff_not_leq {n m} : ~(n <= m) <-> n > m. Proof. split. - - intro; by destruct (leq_dichot n m). + - intro; by destruct (leq_dichotomy n m). - intros ? ?; contradiction (not_lt_n_n m); exact _. Defined. @@ -958,7 +958,7 @@ Defined. Definition leq_sub_add n m : n <= n - m + m. Proof. - destruct (@leq_dichot m n) as [l | g]. + destruct (@leq_dichotomy m n) as [l | g]. - destruct (nat_add_sub_l_cancel l)^; constructor. - apply leq_lt in g. @@ -971,7 +971,7 @@ Defined. Definition sub_gt_0_lt n m : 0 < n - m -> m < n. Proof. intro ineq. - destruct (@leq_dichot n m) as [n_leq_m |]; [ | assumption]. + destruct (@leq_dichotomy n m) as [n_leq_m |]; [ | assumption]. apply equiv_nat_sub_leq in n_leq_m. contradiction (not_lt_n_n 0). now rewrite n_leq_m in ineq. Defined. From ffaedb886908cfc21c53bb0399783172b98f852a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 13:13:26 +0200 Subject: [PATCH 184/282] Nat: tweak decidability proof Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 2188fa8d660..11c1ba211f4 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -242,8 +242,8 @@ Defined. (** [nat] has decidable paths. *) Global Instance decidable_paths_nat@{} : DecidablePaths nat. Proof. - intros n; simple_induction n n IHn; - intros m; destruct m. + intros n m. + induction n as [|n IHn] in m |- *; destruct m. - exact (inl idpath). - exact (inr (neq_nat_zero_succ m)). - exact (inr (fun p => neq_nat_zero_succ n p^)). From 5c6029a52a2b0ec3b5b771dcd0afff52a52d0bd4 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 13:16:41 +0200 Subject: [PATCH 185/282] Nat: remove Theorem, Lemma, Proposition and replace with Definition Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 28 ++++++++++++++-------------- 1 file changed, 14 insertions(+), 14 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 11c1ba211f4..ba72b29604b 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -183,14 +183,14 @@ Notation "x < y <= z" := (x < y /\ y <= z) : nat_scope. (** ** Properties of [nat_iter]. *) -Lemma nat_iter_succ_r n {A} (f : A -> A) (x : A) +Definition nat_iter_succ_r n {A} (f : A -> A) (x : A) : nat_iter (S n) f x = nat_iter n f (f x). Proof. simple_induction n n IHn; simpl; trivial. exact (ap f IHn). Defined. -Theorem nat_iter_add (n m : nat) {A} (f : A -> A) (x : A) +Definition nat_iter_add (n m : nat) {A} (f : A -> A) (x : A) : nat_iter (n + m) f x = nat_iter n f (nat_iter m f x). Proof. simple_induction n n IHn; simpl; trivial. @@ -198,7 +198,7 @@ Proof. Defined. (** Preservation of invariants : if [f : A -> A] preserves the invariant [P], then the iterates of [f] also preserve it. *) -Theorem nat_iter_invariant (n : nat) {A} (f : A -> A) (P : A -> Type) +Definition nat_iter_invariant (n : nat) {A} (f : A -> A) (P : A -> Type) : (forall x, P x -> P (f x)) -> forall x, P x -> P (nat_iter n f x). Proof. simple_induction n n IHn; simpl; trivial. @@ -230,7 +230,7 @@ Proof. Defined. (** A natural number cannot be equal to its own successor. *) -Theorem neq_nat_succ'@{} n : n <> S n. +Definition neq_nat_succ'@{} n : n <> S n. Proof. simple_induction' n. - apply neq_nat_zero_succ. @@ -439,7 +439,7 @@ Definition leq_succ' {n m} : n.+1 <= m.+1 -> n <= m := leq_pred. (** TODO: rename *) (** TODO: use lemmas about negating predicate *) -Lemma not_leq_Sn_n n : ~ (n.+1 <= n). +Definition not_leq_Sn_n n : ~ (n.+1 <= n). Proof. simple_induction n n IHn. { intro p. @@ -449,7 +449,7 @@ Proof. Defined. (** TODO: rename *) -Lemma not_leq_Sn_0 n : ~ (n.+1 <= 0). +Definition not_leq_Sn_0 n : ~ (n.+1 <= 0). Proof. intros p. apply (fun x => leq_trans x (leq_zero n)) in p. @@ -1086,7 +1086,7 @@ Defined. (** *** Order-reflection *) (** TODO: move, rename *) -Proposition nataddreflectslt { n m k : nat } +Definition nataddreflectslt { n m k : nat } : n + k < m + k -> n < m. Proof. simple_induction k k IHk. @@ -1096,7 +1096,7 @@ Proof. Defined. (** TODO: move, rename *) -Proposition nataddreflectsleq { n m k : nat } +Definition nataddreflectsleq { n m k : nat } : n + k <= m + k -> n <= m. Proof. destruct n. @@ -1106,7 +1106,7 @@ Proof. Defined. (** TODO: move, rename *) -Proposition nataddreflectslt' { n m k : nat } +Definition nataddreflectslt' { n m k : nat } : k + n < k + m -> n < m. Proof. destruct (symmetric_paths _ _ (nat_add_comm k n)), @@ -1115,7 +1115,7 @@ Proof. Defined. (** TODO: move, rename *) -Proposition nataddreflectsleq' { n m k : nat } +Definition nataddreflectsleq' { n m k : nat } : k + n <= k + m -> n <= m. Proof. destruct (symmetric_paths _ _ (nat_add_comm k n)), @@ -1124,7 +1124,7 @@ Proof. Defined. (** TODO: move, rename *) -Proposition natsubreflectsleq { n m k : nat } +Definition natsubreflectsleq { n m k : nat } : k <= m -> n - k <= m - k -> n <= m. Proof. intros ineq1 ineq2. @@ -1137,7 +1137,7 @@ Defined. (** ** Further Properties of Subtraction *) (** TODO: rename *) -Proposition nataddsub_assoc_lemma {k m : nat} +Definition nataddsub_assoc_lemma {k m : nat} : (k <= m) -> m.+1 - k = (m - k).+1. Proof. revert m; simple_induction k k IHk. @@ -1151,7 +1151,7 @@ Proof. Defined. (** TODO: rename *) -Proposition nataddsub_assoc(n : nat) {m k : nat} +Definition nataddsub_assoc(n : nat) {m k : nat} : (k <= m) -> n + (m - k) = n + m - k. Proof. revert m k. simple_induction n n IHn. @@ -1173,7 +1173,7 @@ Proof. Defined. (** TODO: rename *) -Proposition nataddsub_comm (n m k : nat) +Definition nataddsub_comm (n m k : nat) : m <= n -> (n - m) + k = (n + k) - m. Proof. intro l. From 59e15bd2805f99fb860468526043c4229514b484 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 13:30:11 +0200 Subject: [PATCH 186/282] Nat: replace not_leq_Sn_n with renamed not_lt_n_n -> lt_irrefl Signed-off-by: Ali Caglayan --- theories/Algebra/Rings/KroneckerDelta.v | 4 +-- theories/Algebra/Universal/Operation.v | 2 +- theories/Spaces/Finite/FinSeq.v | 2 +- theories/Spaces/Finite/Finite.v | 2 +- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 45 ++++++++++--------------- 6 files changed, 24 insertions(+), 33 deletions(-) diff --git a/theories/Algebra/Rings/KroneckerDelta.v b/theories/Algebra/Rings/KroneckerDelta.v index 75e754294fd..fed6ff0e5b6 100644 --- a/theories/Algebra/Rings/KroneckerDelta.v +++ b/theories/Algebra/Rings/KroneckerDelta.v @@ -78,7 +78,7 @@ Definition kronecker_delta_lt {R : Ring} {i j : nat} (p : (i < j)%nat) Proof. apply kronecker_delta_neq. intros q; destruct q. - by apply not_lt_n_n in p. + by apply lt_irrefl in p. Defined. (** Kronecker delta where the first index is strictly greater than the second is 0. *) @@ -87,7 +87,7 @@ Definition kronecker_delta_gt {R : Ring} {i j : nat} (p : (j < i)%nat) Proof. apply kronecker_delta_neq. intros q; destruct q. - by apply not_lt_n_n in p. + by apply lt_irrefl in p. Defined. (** Kronecker delta can be used to extract a single term from a finite sum. *) diff --git a/theories/Algebra/Universal/Operation.v b/theories/Algebra/Universal/Operation.v index 901ee9092b1..bcdc0c2d876 100644 --- a/theories/Algebra/Universal/Operation.v +++ b/theories/Algebra/Universal/Operation.v @@ -20,7 +20,7 @@ Monomorphic Definition head_dom' {σ} (A : Carriers σ) (n : nat) : forall (N : n > 0) (ss : FinSeq n (Sort σ)) (a : forall i, A (ss i)), A (fshead' n N ss) := match n with - | 0 => fun N ss _ => Empty_rec (not_lt_n_n _ N) + | 0 => fun N ss _ => Empty_rec (lt_irrefl _ N) | n'.+1 => fun N ss a => a fin_zero end. diff --git a/theories/Spaces/Finite/FinSeq.v b/theories/Spaces/Finite/FinSeq.v index 18e5f935ab0..a5841f68c84 100644 --- a/theories/Spaces/Finite/FinSeq.v +++ b/theories/Spaces/Finite/FinSeq.v @@ -49,7 +49,7 @@ Definition fshead {A} {n : nat} : FinSeq n.+1 A -> A := fshead' n.+1 _. Definition compute_fshead' {A} n (N : n > 0) (a : A) (v : FinSeq (nat_pred n) A) : fshead' n N (fscons' n a v) = a. Proof. - destruct n; [elim (not_lt_n_n _ N)|]. + destruct n; [elim (lt_irrefl _ N)|]. exact (apD10 (compute_fin_rec_fin_zero _ _ _ _) v). Defined. diff --git a/theories/Spaces/Finite/Finite.v b/theories/Spaces/Finite/Finite.v index 8f140a9afc2..0c7841e8a12 100644 --- a/theories/Spaces/Finite/Finite.v +++ b/theories/Spaces/Finite/Finite.v @@ -751,7 +751,7 @@ Section Enumeration. Proof. destruct (finite_enumeration_stage (fcard X).+1) as [p|?]. - assert (q := leq_inj_finite (er (fcard X).+1) p); simpl in q. - elim (not_lt_n_n _ q). + elim (lt_irrefl _ q). - assumption. Defined. diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 70c8dcfd7b6..7245df1cca7 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -359,7 +359,7 @@ Proof. intro ineq. destruct (@leq_dichotomy n m) as [l | g]. - exact l. - - contradiction (not_lt_n_n m (lt_leq_trans g ineq)). + - contradiction (lt_irrefl m (lt_leq_trans g ineq)). Defined. Proposition symmetric_rel_total_order (R : nat -> nat -> Type) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ba72b29604b..ce5ef2f2584 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -437,23 +437,23 @@ Global Existing Instance leq_succ | 100. (** The converse to [leq_succ] also holds. *) Definition leq_succ' {n m} : n.+1 <= m.+1 -> n <= m := leq_pred. -(** TODO: rename *) -(** TODO: use lemmas about negating predicate *) -Definition not_leq_Sn_n n : ~ (n.+1 <= n). +(** [<] is an irreflexive relation. *) +Definition lt_irrefl n : ~ (n < n). Proof. - simple_induction n n IHn. - { intro p. - inversion p. } - intros p. - by apply IHn, leq_succ'. + induction n as [|n IHn]. + 1: intro p; inversion p. + intros p; by apply IHn, leq_succ'. Defined. +Global Instance irreflexive_lt : Irreflexive lt := lt_irrefl. +Global Instance irreflexive_gt : Irreflexive gt := lt_irrefl. + (** TODO: rename *) Definition not_leq_Sn_0 n : ~ (n.+1 <= 0). Proof. intros p. apply (fun x => leq_trans x (leq_zero n)) in p. - contradiction (not_leq_Sn_n _ p). + contradiction (lt_irrefl _ p). Defined. (** A general form for injectivity of this constructor *) @@ -464,7 +464,7 @@ Proof. destruct c. reflexivity. + destruct r^. - contradiction (not_leq_Sn_n _ p). + contradiction (lt_irrefl _ p). Defined. (** Which we specialise to this lemma *) @@ -478,7 +478,7 @@ Proof. destruct p. + intros k p r. destruct r. - contradiction (not_leq_Sn_n _ p). + contradiction (lt_irrefl _ p). + intros m' q r. pose (r' := path_nat_succ _ _ r). destruct r'. @@ -799,22 +799,13 @@ Proof. 1: reflexivity. destruct x; [inversion q|]. apply leq_succ' in q. - contradiction (not_leq_Sn_n _ (leq_trans p q)). + contradiction (lt_irrefl _ (leq_trans p q)). Defined. Global Instance antisymmetric_leq : AntiSymmetric leq := @leq_antisym. Global Instance antisymemtric_geq : AntiSymmetric geq := fun _ _ p q => leq_antisym q p. -(** *** Irreflexivity of [<] and [>] *) - -(** TODO: rename to [lt_irrefl] *) -(** [<] is an irreflexive relation. *) -Definition not_lt_n_n n : ~ (n < n) := not_leq_Sn_n n. - -Global Instance irreflexive_lt : Irreflexive lt := not_lt_n_n. -Global Instance irreflexive_gt : Irreflexive gt := not_lt_n_n. - (** *** Addition Lemmas *) (** The first summand is less than or equal to the sum. *) @@ -844,8 +835,8 @@ Proof. snrapply equiv_path_sum. destruct x as [l|p], y as [q|r]. 1,4: rapply path_ishprop. - + destruct r; contradiction (not_lt_n_n _ _). - + destruct p; contradiction (not_lt_n_n _ _). + + destruct r; contradiction (lt_irrefl _ _). + + destruct p; contradiction (lt_irrefl _ _). - intro l; induction l. + now right. + left; exact (leq_succ l). @@ -909,14 +900,14 @@ Definition geq_iff_not_lt {n m} : ~(n < m) <-> n >= m. Proof. split. - intro; by destruct (leq_dichotomy m n). - - intros ? ?; contradiction (not_lt_n_n n); exact _. + - intros ? ?; contradiction (lt_irrefl n); exact _. Defined. Definition gt_iff_not_leq {n m} : ~(n <= m) <-> n > m. Proof. split. - intro; by destruct (leq_dichotomy n m). - - intros ? ?; contradiction (not_lt_n_n m); exact _. + - intros ? ?; contradiction (lt_irrefl m); exact _. Defined. Definition leq_iff_not_gt {n m} : ~(n > m) <-> n <= m @@ -939,7 +930,7 @@ Proof. 1: by right. symmetry in eq. contradiction. - - intros [H' | H'] nq; destruct nq; exact (not_lt_n_n _ H'). + - intros [H' | H'] nq; destruct nq; exact (lt_irrefl _ H'). Defined. (** ** Arithmetic relations between [trunc_index] and [nat]. *) @@ -973,7 +964,7 @@ Proof. intro ineq. destruct (@leq_dichotomy n m) as [n_leq_m |]; [ | assumption]. apply equiv_nat_sub_leq in n_leq_m. - contradiction (not_lt_n_n 0). now rewrite n_leq_m in ineq. + contradiction (lt_irrefl 0). now rewrite n_leq_m in ineq. Defined. (** TODO: merge with above, reprove *) From 9b0d300d0749a1dca4321e23247f067e15737438 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 13:36:00 +0200 Subject: [PATCH 187/282] Nat: replace not_leq_Sn_0 with not_lt_zero_r Signed-off-by: Ali Caglayan --- theories/Algebra/Rings/KroneckerDelta.v | 2 +- theories/Spaces/List/Theory.v | 22 +++++++++++----------- theories/Spaces/Nat/Core.v | 19 +++++++------------ 3 files changed, 19 insertions(+), 24 deletions(-) diff --git a/theories/Algebra/Rings/KroneckerDelta.v b/theories/Algebra/Rings/KroneckerDelta.v index fed6ff0e5b6..0a8dd7154bf 100644 --- a/theories/Algebra/Rings/KroneckerDelta.v +++ b/theories/Algebra/Rings/KroneckerDelta.v @@ -96,7 +96,7 @@ Definition rng_sum_kronecker_delta_l {R : Ring} (n i : nat) (Hi : (i < n)%nat) : ab_sum n (fun j Hj => kronecker_delta i j * f j Hj) = f i Hi. Proof. revert i Hi f; simple_induction n n IHn; intros i Hi f. - 1: destruct (not_leq_Sn_0 _ Hi). + 1: destruct (not_lt_zero_r _ Hi). destruct (dec (i = n)) as [p|p]. - destruct p; simpl. rewrite kronecker_delta_refl. diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index 074ce8bf9a8..2c84d47acd6 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -332,7 +332,7 @@ Definition nth_lt@{i|} {A : Type@{i}} (l : list A) (n : nat) : { x : A & nth l n = Some x }. Proof. induction l as [|a l IHa] in n, H |- * using list_ind@{i i}. - 1: destruct (not_leq_Sn_0 _ H). + 1: destruct (not_lt_zero_r _ H). destruct n. 1: by exists a. apply IHa. @@ -357,7 +357,7 @@ Definition inlist_nth'@{i|} {A : Type@{i}} (l : list A) (n : nat) : InList (nth' l n H) l. Proof. induction l as [|a l IHa] in n, H |- * using list_ind@{i i}. - 1: destruct (not_leq_Sn_0 _ H). + 1: destruct (not_lt_zero_r _ H). destruct n. 1: by left. right. @@ -423,7 +423,7 @@ Definition nth'_list_map@{i j|} {A : Type@{i}} {B : Type@{j}} : nth' (list_map f l) n H' = f (nth' l n H). Proof. induction l as [|a l IHl] in n, H, H' |- * using list_ind@{i j}. - 1: destruct (not_leq_Sn_0 _ H). + 1: destruct (not_lt_zero_r _ H). destruct n. 1: reflexivity. apply IHl. @@ -441,7 +441,7 @@ Proof. simple_list_induction l1 a l1 IHl1; intros l2 n defl defr H H' H'' p. - destruct l2 as [|b l2]. - + destruct (not_leq_Sn_0 _ H). + + destruct (not_lt_zero_r _ H). + inversion p. - destruct l2 as [|b l2]. + inversion p. @@ -463,7 +463,7 @@ Definition nth'_repeat@{i|} {A : Type@{i}} (x : A) (i n : nat) : nth' (repeat x n) i H = x. Proof. induction n as [|n IHn] in i, H |- * using nat_ind@{i}. - 1: destruct (not_leq_Sn_0 _ H). + 1: destruct (not_lt_zero_r _ H). destruct i. 1: reflexivity. apply IHn. @@ -499,7 +499,7 @@ Definition nth_app@{i|} {A : Type@{i}} (l l' : list A) (n : nat) : nth (l ++ l') n = nth l n. Proof. induction l as [|a l IHl] in l', n, H |- * using list_ind@{i i}. - 1: destruct (not_leq_Sn_0 _ H). + 1: destruct (not_lt_zero_r _ H). destruct n. 1: reflexivity. by apply IHl, leq_succ'. @@ -572,7 +572,7 @@ Proof. induction l as [|a l IHl] in H, n |- * using list_ind@{i i}. 1: apply drop_nil. destruct n. - 1: destruct (not_leq_Sn_0 _ H). + 1: destruct (not_lt_zero_r _ H). cbn; apply IHl. apply leq_succ'. exact H. @@ -630,7 +630,7 @@ Proof. induction l as [|a l IHl] in H, n |- * using list_ind@{i i}. 1: apply take_nil. destruct n. - 1: destruct (not_leq_Sn_0 _ H). + 1: destruct (not_lt_zero_r _ H). cbn; f_ap. by apply IHl, leq_succ'. Defined. @@ -820,10 +820,10 @@ Definition nth_seq_rev@{} {n i} (H : (i < n)%nat) Proof. induction i as [|i IHi] in n, H |- *. - induction n. - 1: destruct (not_leq_Sn_0 _ H). + 1: destruct (not_lt_zero_r _ H). cbn; by rewrite nat_sub_zero_r. - induction n as [|n IHn]. - 1: destruct (not_leq_Sn_0 _ H). + 1: destruct (not_lt_zero_r _ H). by apply IHi, leq_succ'. Defined. @@ -832,7 +832,7 @@ Definition nth_seq@{} {n i} (H : (i < n)%nat) : nth (seq n) i = Some i. Proof. induction n. - 1: destruct (not_leq_Sn_0 _ H). + 1: destruct (not_lt_zero_r _ H). rewrite seq_succ. destruct (dec (i < n)%nat) as [H'|H']. - lhs nrapply nth_app. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ce5ef2f2584..5bf9896788b 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -448,12 +448,12 @@ Defined. Global Instance irreflexive_lt : Irreflexive lt := lt_irrefl. Global Instance irreflexive_gt : Irreflexive gt := lt_irrefl. -(** TODO: rename *) -Definition not_leq_Sn_0 n : ~ (n.+1 <= 0). +(** Nothing can be less than [0]. *) +Definition not_lt_zero_r n : ~ (n < 0). Proof. intros p. - apply (fun x => leq_trans x (leq_zero n)) in p. - contradiction (lt_irrefl _ p). + apply (lt_irrefl n), (leq_trans p). + exact _. Defined. (** A general form for injectivity of this constructor *) @@ -516,7 +516,7 @@ Proof. simple_induction' m; intros n. - destruct n. + left; exact _. - + right; apply not_leq_Sn_0. + + right; apply not_lt_zero_r. - destruct n. + left; exact _. + rapply decidable_equiv'. @@ -644,7 +644,7 @@ Proof. - lhs nrapply nat_add_zero_r. nrapply nat_sub_zero_r. - destruct m. - 1: contradiction (not_leq_Sn_0 n). + 1: contradiction (not_lt_zero_r n). lhs nrapply nat_add_succ_r. nrapply (ap nat_succ). nrapply IHn. @@ -889,11 +889,6 @@ Defined. (** *** Negation Lemmas *) -Definition not_lt_zero_r n : ~ (n < 0). -Proof. - apply not_leq_Sn_0. -Defined. - (** There are various lemmas we can state about negating the comparison operators on [nat]. To aid readability, we opt to keep the order of the variables in each statement consistent. *) Definition geq_iff_not_lt {n m} : ~(n < m) <-> n >= m. @@ -1134,7 +1129,7 @@ Proof. revert m; simple_induction k k IHk. - intros m l; simpl. destruct m; reflexivity. - destruct m. - + simpl; intro g; contradiction (not_leq_Sn_0 _ g). + + simpl; intro g; contradiction (not_lt_zero_r _ g). + intro l; apply leq_succ' in l. change (m.+2 - k.+1) with (m.+1 - k). change (m.+1 - k.+1) with (m - k). From 74beb500a14fdbbbd4cab0d5a0ed3fd104022187 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 13:41:19 +0200 Subject: [PATCH 188/282] Nat: make leq_succ' an immediate hint Signed-off-by: Ali Caglayan --- theories/Spaces/List/Theory.v | 7 +------ theories/Spaces/Nat/Core.v | 2 +- 2 files changed, 2 insertions(+), 7 deletions(-) diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index 2c84d47acd6..ad191e4a727 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -447,12 +447,7 @@ Proof. + inversion p. + destruct n. * reflexivity. - * unshelve erewrite nth'_cons. - 1: apply leq_succ', H. - unshelve erewrite nth'_cons. - 1: apply leq_succ', H'. - unshelve erewrite nth'_cons. - 1: apply leq_succ', H''. + * erewrite 3 nth'_cons. apply IHl1. by apply path_nat_succ. Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 5bf9896788b..f99af2cfbe0 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -436,6 +436,7 @@ Global Existing Instance leq_succ | 100. (** The converse to [leq_succ] also holds. *) Definition leq_succ' {n m} : n.+1 <= m.+1 -> n <= m := leq_pred. +Hint Immediate leq_succ' : typeclass_instances. (** [<] is an irreflexive relation. *) Definition lt_irrefl n : ~ (n < n). @@ -507,7 +508,6 @@ Defined. Definition equiv_leq_succ n m : n.+1 <= m.+1 <~> n <= m. Proof. srapply equiv_iff_hprop. - apply leq_succ'. Defined. Global Instance decidable_leq n m : Decidable (n <= m). From cc5a1c2fd8ad5c20ee649d8af88255b9e8429f09 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 14:25:56 +0200 Subject: [PATCH 189/282] Nat: add comments for renaming reflection lemmas Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index f99af2cfbe0..6294d28abcf 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1071,7 +1071,7 @@ Defined. (** *** Order-reflection *) -(** TODO: move, rename *) +(** TODO: rename nataddreflectslt -> lt_reflects_add_r *) Definition nataddreflectslt { n m k : nat } : n + k < m + k -> n < m. Proof. @@ -1081,7 +1081,7 @@ Proof. apply leq_succ', IHk in l; exact l. Defined. -(** TODO: move, rename *) +(** TODO: rename nataddreflectsleq -> leq_reflects_add_r *) Definition nataddreflectsleq { n m k : nat } : n + k <= m + k -> n <= m. Proof. @@ -1091,7 +1091,7 @@ Proof. now apply (@nataddreflectslt n m k). Defined. -(** TODO: move, rename *) +(** TODO: rename nataddreflectslt' -> lt_reflects_add_l *) Definition nataddreflectslt' { n m k : nat } : k + n < k + m -> n < m. Proof. @@ -1100,7 +1100,7 @@ Proof. exact nataddreflectslt. Defined. -(** TODO: move, rename *) +(** TODO: rename nataddreflectsleq' -> leq_reflects_add_l *) Definition nataddreflectsleq' { n m k : nat } : k + n <= k + m -> n <= m. Proof. From d7ccbd0d50b2aa0e4af27935d52fc863300f4a53 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 14:26:50 +0200 Subject: [PATCH 190/282] nat: rename nataddreflectslt -> lt_reflects_add_r Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 9 ++++----- 2 files changed, 5 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 7245df1cca7..98e4ac75d15 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -200,7 +200,7 @@ Proof. (destruct (nat_add_sub_l_cancel l)^; destruct (nat_add_comm n m); assumption). - exact (nataddreflectslt q'). + exact (lt_reflects_add_r q'). Defined. (** TODO: move, rename *) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 6294d28abcf..fb8035a6fd5 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1071,9 +1071,8 @@ Defined. (** *** Order-reflection *) -(** TODO: rename nataddreflectslt -> lt_reflects_add_r *) -Definition nataddreflectslt { n m k : nat } - : n + k < m + k -> n < m. +(** Addition on the right is strictly order-reflecting. *) +Definition lt_reflects_add_r {n m k} : n + k < m + k -> n < m. Proof. simple_induction k k IHk. - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; trivial. @@ -1088,7 +1087,7 @@ Proof. destruct n. - intros ?; apply leq_zero. - intro a. change (n.+1 + k) with (n + k).+1 in a. - now apply (@nataddreflectslt n m k). + now apply (@lt_reflects_add_r n m k). Defined. (** TODO: rename nataddreflectslt' -> lt_reflects_add_l *) @@ -1097,7 +1096,7 @@ Definition nataddreflectslt' { n m k : nat } Proof. destruct (symmetric_paths _ _ (nat_add_comm k n)), (symmetric_paths _ _ (nat_add_comm k m)); - exact nataddreflectslt. + exact lt_reflects_add_r. Defined. (** TODO: rename nataddreflectsleq' -> leq_reflects_add_l *) From 8ac8d56792ad2d1fe8beac5ace210af4cdaa9ffe Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 14:30:51 +0200 Subject: [PATCH 191/282] Nat: rename nataddreflectsleq -> leq_reflects_add_r Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index fb8035a6fd5..06e2d426477 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1080,9 +1080,7 @@ Proof. apply leq_succ', IHk in l; exact l. Defined. -(** TODO: rename nataddreflectsleq -> leq_reflects_add_r *) -Definition nataddreflectsleq { n m k : nat } - : n + k <= m + k -> n <= m. +Definition leq_reflects_add_r {n m k} : n + k <= m + k -> n <= m. Proof. destruct n. - intros ?; apply leq_zero. @@ -1105,7 +1103,7 @@ Definition nataddreflectsleq' { n m k : nat } Proof. destruct (symmetric_paths _ _ (nat_add_comm k n)), (symmetric_paths _ _ (nat_add_comm k m)); - exact nataddreflectsleq. + exact leq_reflects_add_r. Defined. (** TODO: move, rename *) From b090fd33c95a5c5f8a2a57e1e96f0e539ae38b41 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 14:33:31 +0200 Subject: [PATCH 192/282] Nat: rename lt_reflects_add_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 06e2d426477..33e71981a84 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1088,9 +1088,7 @@ Proof. now apply (@lt_reflects_add_r n m k). Defined. -(** TODO: rename nataddreflectslt' -> lt_reflects_add_l *) -Definition nataddreflectslt' { n m k : nat } - : k + n < k + m -> n < m. +Definition lt_reflects_add_l {n m k} : k + n < k + m -> n < m. Proof. destruct (symmetric_paths _ _ (nat_add_comm k n)), (symmetric_paths _ _ (nat_add_comm k m)); From 6beb01ab6ed43d1cce39ad23f28c3e0adb47f816 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 14:35:04 +0200 Subject: [PATCH 193/282] Nat: rename nataddreflectsleq' -> leq_reflects_add_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 33e71981a84..51dc70d65c3 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1095,9 +1095,7 @@ Proof. exact lt_reflects_add_r. Defined. -(** TODO: rename nataddreflectsleq' -> leq_reflects_add_l *) -Definition nataddreflectsleq' { n m k : nat } - : k + n <= k + m -> n <= m. +Definition leq_reflects_add_l {n m k} : k + n <= k + m -> n <= m. Proof. destruct (symmetric_paths _ _ (nat_add_comm k n)), (symmetric_paths _ _ (nat_add_comm k m)); From 1d4be93103972e994467e79332055bb4ddbdbfcc Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 14:36:46 +0200 Subject: [PATCH 194/282] Nat: add comments to order-reflection lemmas Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 3 +++ 1 file changed, 3 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 51dc70d65c3..7b66cda47f2 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1080,6 +1080,7 @@ Proof. apply leq_succ', IHk in l; exact l. Defined. +(** Addition on the right is order-reflecting. *) Definition leq_reflects_add_r {n m k} : n + k <= m + k -> n <= m. Proof. destruct n. @@ -1088,6 +1089,7 @@ Proof. now apply (@lt_reflects_add_r n m k). Defined. +(** Addition on the lef is strictly order-reflecting. *) Definition lt_reflects_add_l {n m k} : k + n < k + m -> n < m. Proof. destruct (symmetric_paths _ _ (nat_add_comm k n)), @@ -1095,6 +1097,7 @@ Proof. exact lt_reflects_add_r. Defined. +(** Addition on the left is order-reflecting. *) Definition leq_reflects_add_l {n m k} : k + n <= k + m -> n <= m. Proof. destruct (symmetric_paths _ _ (nat_add_comm k n)), From 6e58b1abf7203d4b6f56aec377d4c3f0f919d283 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 14:54:23 +0200 Subject: [PATCH 195/282] Nat: revise hints and proofs of reflection Signed-off-by: Ali Caglayan --- theories/Spaces/Finite/FinInduction.v | 4 +- theories/Spaces/Finite/FinNat.v | 5 +- theories/Spaces/Nat/Core.v | 75 +++++++++++++-------------- 3 files changed, 43 insertions(+), 41 deletions(-) diff --git a/theories/Spaces/Finite/FinInduction.v b/theories/Spaces/Finite/FinInduction.v index ab31e8af8d0..6cea070d52a 100644 --- a/theories/Spaces/Finite/FinInduction.v +++ b/theories/Spaces/Finite/FinInduction.v @@ -32,8 +32,8 @@ Proof. induction (path_fin_to_finnat_fin_zero n)^. intro p. destruct (hset_path2 1 p). - cbn. - by destruct (hset_path2 1 (path_zero_finnat n _)). + lhs nrapply transport_1. + nrapply compute_finnat_ind_zero. Defined. Lemma compute_fin_ind_fsucc (P : forall n : nat, Fin n -> Type) diff --git a/theories/Spaces/Finite/FinNat.v b/theories/Spaces/Finite/FinNat.v index e4a239208b5..af12985ac84 100644 --- a/theories/Spaces/Finite/FinNat.v +++ b/theories/Spaces/Finite/FinNat.v @@ -65,7 +65,10 @@ Lemma compute_finnat_ind_zero (P : forall n : nat, FinNat n -> Type) (n : nat) : finnat_ind P z s (zero_finnat n) = z n. Proof. - cbn. by induction (hset_path2 1 (path_zero_finnat n _)). + unshelve lhs snrapply transport2. + - reflexivity. + - rapply hset_path2. + - reflexivity. Defined. Lemma compute_finnat_ind_succ (P : forall n : nat, FinNat n -> Type) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 7b66cda47f2..44f29fbb958 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -809,7 +809,7 @@ Global Instance antisymemtric_geq : AntiSymmetric geq (** *** Addition Lemmas *) (** The first summand is less than or equal to the sum. *) -Definition leq_add_l n m : n <= n + m. +Global Instance leq_add_l n m : n <= n + m. Proof. simple_induction n n IHn. - exact (leq_zero m). @@ -817,7 +817,7 @@ Proof. Defined. (** The second summand is less than or equal to the sum. *) -Definition leq_add_r n m : n <= m + n. +Global Instance leq_add_r n m : n <= m + n. Proof. simple_induction m m IH. - exact (leq_refl n). @@ -978,25 +978,28 @@ Defined. (** TODO: use OrderPreserving from canonical_names *) (** Addition on the left is monotone. *) -Global Instance nat_add_l_monotone {n m} k +Definition nat_add_l_monotone {n m} k : n <= m -> k + n <= k + m. Proof. - intros H; induction k as [|k IHk] in n, m, H |- *; exact _. + intros H; induction k; exact _. Defined. +Hint Immediate nat_add_l_monotone : typeclass_instances. (** Addition on the right is monotone. *) -Global Instance nat_add_r_monotone {n m} k +Definition nat_add_r_monotone {n m} k : n <= m -> n + k <= m + k. Proof. intros H; rewrite 2 (nat_add_comm _ k); exact _. Defined. +Hint Immediate nat_add_r_monotone : typeclass_instances. (** Addition is monotone in both arguments. (This makes [+] a bifunctor when treating [nat] as a category (as a preorder)). *) -Global Instance nat_add_monotone {n n' m m'} +Definition nat_add_monotone {n n' m m'} : n <= m -> n' <= m' -> n + n' <= m + m'. Proof. intros H1 H2; induction H1; exact _. Defined. +Hint Immediate nat_add_monotone : typeclass_instances. (** *** Strict Monotonicity of Addition *) @@ -1006,103 +1009,99 @@ Global Instance lt_succ {n m} : n < m -> n.+1 < m.+1 := _. Global Instance lt_succ_r {n m} : n < m -> n < m.+1 := _. (** Addition on the left is strictly monotone. *) -Global Instance nat_add_l_strictly_monotone {n m} k +Definition nat_add_l_strictly_monotone {n m} k : n < m -> k + n < k + m. Proof. - intros H; induction k as [|k IHk] in n, m, H |- *; exact _. + intros H; induction k; exact _. Defined. +Hint Immediate nat_add_l_strictly_monotone : typeclass_instances. (** Addition on the right is strictly monotone. *) -Global Instance nat_add_r_strictly_monotone {n m} k +Definition nat_add_r_strictly_monotone {n m} k : n < m -> n + k < m + k. Proof. intros H; rewrite 2 (nat_add_comm _ k); exact _. Defined. +Hint Immediate nat_add_r_strictly_monotone : typeclass_instances. (** Addition is strictly monotone in both arguments. *) -Global Instance nat_add_strictly_monotone {n n' m m'} +Definition nat_add_strictly_monotone {n n' m m'} : n < m -> n' < m' -> n + n' < m + m'. Proof. intros H1 H2; induction H1; exact _. Defined. +Hint Immediate nat_add_strictly_monotone : typeclass_instances. (** *** Monotonicity of Multiplication *) (** Multiplication on the left is monotone. *) -Global Instance nat_mul_l_monotone {n m} k +Definition nat_mul_l_monotone {n m} k : n <= m -> k * n <= k * m. Proof. - intros H; induction k as [|k IHk] in |- *; exact _. + intros H; induction k; exact _. Defined. +Hint Immediate nat_mul_l_monotone : typeclass_instances. (** Multiplication on the right is monotone. *) -Global Instance nat_mul_r_monotone {n m} k +Definition nat_mul_r_monotone {n m} k : n <= m -> n * k <= m * k. Proof. intros H; rewrite 2 (nat_mul_comm _ k); exact _. Defined. +Hint Immediate nat_mul_r_monotone : typeclass_instances. (** Multiplication is monotone in both arguments. *) -Global Instance nat_mul_monotone {n n' m m'} +Definition nat_mul_monotone {n n' m m'} : n <= m -> n' <= m' -> n * n' <= m * m'. Proof. - intros H1 H2; induction H1. - - exact _. - - rapply leq_trans. + intros H1 H2; induction H1; exact _. Defined. +Hint Immediate nat_mul_monotone : typeclass_instances. (** *** Strict Monotonicity of Multiplication *) (** Multiplication on the left by a positive number is strictly monotone. *) -Global Instance nat_mul_l_strictly_monotone {n m} k +Definition nat_mul_l_strictly_monotone {n m} k : n < m -> k.+1 * n < k.+1 * m. Proof. intros H; induction k as [|k IHk] in |- *; exact _. Defined. +Hint Immediate nat_mul_l_strictly_monotone : typeclass_instances. (** Multiplication on the right by a positive number is strictly monotone. *) -Global Instance nat_mul_r_strictly_monotone {n m} k +Definition nat_mul_r_strictly_monotone {n m} k : n < m -> n * k.+1 < m * k.+1. Proof. intros H; rewrite 2 (nat_mul_comm _ k.+1); exact _. Defined. +Hint Immediate nat_mul_r_strictly_monotone : typeclass_instances. (** TODO: monotonicity of subtraction *) (** *** Order-reflection *) -(** Addition on the right is strictly order-reflecting. *) -Definition lt_reflects_add_r {n m k} : n + k < m + k -> n < m. +(** Addition on the left is order-reflecting. *) +Definition leq_reflects_add_l {n m k} : k + n <= k + m -> n <= m. Proof. - simple_induction k k IHk. - - destruct (nat_add_zero_r n)^, (nat_add_zero_r m)^; trivial. - - intro l. destruct (nat_add_succ_r n k)^, (nat_add_succ_r m k)^ in l. - apply leq_succ', IHk in l; exact l. + intros H; induction k; exact _. Defined. (** Addition on the right is order-reflecting. *) Definition leq_reflects_add_r {n m k} : n + k <= m + k -> n <= m. Proof. - destruct n. - - intros ?; apply leq_zero. - - intro a. change (n.+1 + k) with (n + k).+1 in a. - now apply (@lt_reflects_add_r n m k). + rewrite 2 (nat_add_comm _ k); nrapply leq_reflects_add_l. Defined. -(** Addition on the lef is strictly order-reflecting. *) +(** Addition on the left is strictly order-reflecting. *) Definition lt_reflects_add_l {n m k} : k + n < k + m -> n < m. Proof. - destruct (symmetric_paths _ _ (nat_add_comm k n)), - (symmetric_paths _ _ (nat_add_comm k m)); - exact lt_reflects_add_r. + intros H; induction k; exact _. Defined. -(** Addition on the left is order-reflecting. *) -Definition leq_reflects_add_l {n m k} : k + n <= k + m -> n <= m. +(** Addition on the right is strictly order-reflecting. *) +Definition lt_reflects_add_r {n m k} : n + k < m + k -> n < m. Proof. - destruct (symmetric_paths _ _ (nat_add_comm k n)), - (symmetric_paths _ _ (nat_add_comm k m)); - exact leq_reflects_add_r. + rewrite 2 (nat_add_comm _ k); nrapply lt_reflects_add_l. Defined. (** TODO: move, rename *) From 5ee53aeecec9fa7c3202ba697b5a00a4efb1e5da Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 17:17:09 +0200 Subject: [PATCH 196/282] Nat: more properties of subtraction Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 89 ++++++++++++++++++++++++++------ 2 files changed, 74 insertions(+), 17 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 98e4ac75d15..aa507d42e47 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -200,7 +200,7 @@ Proof. (destruct (nat_add_sub_l_cancel l)^; destruct (nat_add_comm n m); assumption). - exact (lt_reflects_add_r q'). + exact (lt_reflects_add_r _ q'). Defined. (** TODO: move, rename *) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 44f29fbb958..5547b506901 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -211,6 +211,11 @@ Defined. Definition nat_pred_succ@{} n : nat_pred (nat_succ n) = n := idpath. +Definition nat_succ_pred@{} n : 0 < n -> nat_succ (nat_pred n) = n. +Proof. + by intros []. +Defined. + (** Injectivity of successor. *) Definition path_nat_succ@{} n m (H : S n = S m) : n = m := ap nat_pred H. Global Instance isinj_succ : IsInjective nat_succ := path_nat_succ. @@ -671,6 +676,16 @@ Defined. (** We can move a subtracted number to the right-hand side of an equation. *) Definition nat_moveR_nV {k m} n : k = n + m -> k - m = n := fun p => (nat_moveL_nV _ p^)^. + +(** Subtracting a successor is the predecessor of subtracting the original number. *) +Definition nat_sub_succ_r n m : n - m.+1 = nat_pred (n - m). +Proof. + induction n as [|n IHn] in m |- *. + 1: reflexivity. + destruct m. + 1: apply nat_sub_zero_r. + apply IHn. +Defined. (** ** Properties of Maximum and Minimum *) @@ -1076,47 +1091,47 @@ Proof. Defined. Hint Immediate nat_mul_r_strictly_monotone : typeclass_instances. -(** TODO: monotonicity of subtraction *) - (** *** Order-reflection *) (** Addition on the left is order-reflecting. *) -Definition leq_reflects_add_l {n m k} : k + n <= k + m -> n <= m. +Definition leq_reflects_add_l {n m} k : k + n <= k + m -> n <= m. Proof. intros H; induction k; exact _. Defined. (** Addition on the right is order-reflecting. *) -Definition leq_reflects_add_r {n m k} : n + k <= m + k -> n <= m. +Definition leq_reflects_add_r {n m} k : n + k <= m + k -> n <= m. Proof. rewrite 2 (nat_add_comm _ k); nrapply leq_reflects_add_l. Defined. (** Addition on the left is strictly order-reflecting. *) -Definition lt_reflects_add_l {n m k} : k + n < k + m -> n < m. +Definition lt_reflects_add_l {n m} k : k + n < k + m -> n < m. Proof. intros H; induction k; exact _. Defined. (** Addition on the right is strictly order-reflecting. *) -Definition lt_reflects_add_r {n m k} : n + k < m + k -> n < m. +Definition lt_reflects_add_r {n m} k : n + k < m + k -> n < m. Proof. rewrite 2 (nat_add_comm _ k); nrapply lt_reflects_add_l. Defined. -(** TODO: move, rename *) -Definition natsubreflectsleq { n m k : nat } - : k <= m -> n - k <= m - k -> n <= m. +(** ** Further Properties of Subtraction *) + +(** Subtracting from a successor is the successor of subtracting from the original number, as long as the amount being subtracted is less than or equal to the original number. *) +Definition nat_sub_succ_l n m : m <= n -> n.+1 - m = (n - m).+1. Proof. - intros ineq1 ineq2. - apply (nat_add_r_monotone k) in ineq2. - apply (@leq_trans _ (n - k + k) _ (leq_sub_add _ _)). - apply (@leq_trans _ (m - k + k) _ _). - destruct (nat_add_sub_l_cancel ineq1)^; easy. + intros H. + induction m as [|m IHm] in n, H |- *. + - by rewrite 2 nat_sub_zero_r. + - simpl. + rewrite nat_sub_succ_r. + symmetry. + apply nat_succ_pred. + by apply lt_sub_gt_0. Defined. -(** ** Further Properties of Subtraction *) - (** TODO: rename *) Definition nataddsub_assoc_lemma {k m : nat} : (k <= m) -> m.+1 - k = (m - k).+1. @@ -1163,6 +1178,48 @@ Proof. apply nat_add_comm. Defined. +(** *** Monotonicity of Subtraction *) + +(** Subtraction is monotone in the left argument. *) +Definition nat_sub_monotone_l {n m} k : n <= m -> n - k <= m - k. +Proof. + intros H. + destruct (leq_dichotomy k n) as [l|r]. + - apply (leq_reflects_add_l k). + rewrite 2 nat_add_sub_r_cancel. + + exact H. + + rapply leq_trans. + + exact l. + - apply leq_succ_l in r. + apply equiv_nat_sub_leq in r. + destruct r^. + exact _. +Defined. + +(** Subtraction is contravariantly monotone in the right argument. *) +Definition nat_sub_monotone_r {n m} k : n <= m -> k - m <= k - n. +Proof. + intros H. + induction k. + - by rewrite nat_sub_zero_l. + - destruct (leq_dichotomy m k) as [l|r]. + + rewrite 2 nat_sub_succ_l; exact _. + + apply equiv_nat_sub_leq in r. + destruct r^. + exact _. +Defined. + +(** TODO: rename natsubreflectsleq -> leq_reflects_sub_l *) +Definition natsubreflectsleq { n m k : nat } + : k <= m -> n - k <= m - k -> n <= m. +Proof. + intros ineq1 ineq2. + apply (nat_add_r_monotone k) in ineq2. + apply (@leq_trans _ (n - k + k) _ (leq_sub_add _ _)). + apply (@leq_trans _ (m - k + k) _ _). + by rewrite nat_add_sub_l_cancel. +Defined. + (** ** Properties of Powers *) (** [0] to any power is [0] unless that power is [0] in which case it is [1]. *) From 9069b02398f559368c6710eb7ead0751108122ae Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 17:18:18 +0200 Subject: [PATCH 197/282] Nat: rename natsubreflectsleq -> leq_reflects_sub_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 5547b506901..0c659cd9d33 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1209,9 +1209,7 @@ Proof. exact _. Defined. -(** TODO: rename natsubreflectsleq -> leq_reflects_sub_l *) -Definition natsubreflectsleq { n m k : nat } - : k <= m -> n - k <= m - k -> n <= m. +Definition leq_reflects_sub_l {n m} k : k <= m -> n - k <= m - k -> n <= m. Proof. intros ineq1 ineq2. apply (nat_add_r_monotone k) in ineq2. From f7593c9fdcb95af3a57e6ce3bfc48715791b2c22 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 18:31:41 +0200 Subject: [PATCH 198/282] Nat: rename leq_zero -> leq_zero_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 6 +++--- theories/Spaces/Nat/Core.v | 6 +++--- 2 files changed, 6 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index aa507d42e47..0c4a811fd98 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -47,7 +47,7 @@ Proof. revert m i; simple_induction n n IHn. - intros m i l. simpl in l. contradiction (not_lt_zero_r _ _). - intros m i l. destruct m. - + apply leq_zero. + + apply leq_zero_l. + apply leq_succ. simpl in l. apply (IHn m i l). Defined. @@ -109,7 +109,7 @@ Proposition pred_lt_implies_leq (n k : nat) : nat_pred n < k -> n <= k. Proof. intro l; destruct n. - - apply leq_zero. + - apply leq_zero_l. - assumption. Defined. @@ -173,7 +173,7 @@ Proposition sub_less { n k : nat } : n - k <= n. Proof. revert k. simple_induction n n IHn. - - intros; apply leq_zero. + - intros; apply leq_zero_l. - destruct k. + apply leq_refl. + simpl; apply (@leq_trans _ n _); diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 0c659cd9d33..52b19f15ddc 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -418,11 +418,11 @@ Hint Immediate leq_trans : typeclass_instances. Global Instance transitive_leq : Transitive leq := @leq_trans. (** [0] is less than or equal to any natural number. *) -Definition leq_zero n : 0 <= n. +Definition leq_zero_l n : 0 <= n. Proof. simple_induction' n; exact _. Defined. -Global Existing Instance leq_zero | 10. +Global Existing Instance leq_zero_l | 10. (** A predecessor is less than or equal to a predecessor if the original number is less than or equal. *) Definition leq_pred {n m} : n <= m -> nat_pred n <= nat_pred m. @@ -827,7 +827,7 @@ Global Instance antisymemtric_geq : AntiSymmetric geq Global Instance leq_add_l n m : n <= n + m. Proof. simple_induction n n IHn. - - exact (leq_zero m). + - exact (leq_zero_l m). - exact (leq_succ IHn). Defined. From 43fbf2ceafb9a654153ebc996e5f1f192ffd4fc8 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 18:34:53 +0200 Subject: [PATCH 199/282] Nat: hint tweaks and comments Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 8 ++++++++ 1 file changed, 8 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 52b19f15ddc..f1b127a16fc 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1195,6 +1195,7 @@ Proof. destruct r^. exact _. Defined. +Hint Immediate nat_sub_monotone_l : typeclass_instances. (** Subtraction is contravariantly monotone in the right argument. *) Definition nat_sub_monotone_r {n m} k : n <= m -> k - m <= k - n. @@ -1208,6 +1209,13 @@ Proof. destruct r^. exact _. Defined. +Hint Immediate nat_sub_monotone_r : typeclass_instances. + +(** *** Movement Lemmas *) + + + +(** *** Order-reflection Lemmas *) Definition leq_reflects_sub_l {n m} k : k <= m -> n - k <= m - k -> n <= m. Proof. From 0f117ee9657f240260cdcb303fb29534015cf2ae Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 18:36:01 +0200 Subject: [PATCH 200/282] Nat: move natpmswap{1,2,3,4}, nat_sub_add_ineq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 43 -------------------------- theories/Spaces/Nat/Core.v | 53 ++++++++++++++++++++++++++++++++ 2 files changed, 53 insertions(+), 43 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 0c4a811fd98..0c7c9f7a502 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -191,49 +191,6 @@ Proof. simpl; apply leq_succ, sub_less. Defined. -(** TODO: move, rename *) -Proposition natpmswap1 (k m n : nat) - : n <= k -> k < n + m -> k - n < m. -Proof. - intros l q. - assert (q' : k - n + n < m + n) by - (destruct (nat_add_sub_l_cancel l)^; - destruct (nat_add_comm n m); - assumption). - exact (lt_reflects_add_r _ q'). -Defined. - -(** TODO: move, rename *) -Proposition natpmswap2 (k m n : nat) - : n <= k -> k - n <= m -> k <= n + m. -Proof. - intros l q. - apply (nat_add_l_monotone n) in q. - destruct (nataddsub_assoc n l)^ in q. - destruct (nat_add_sub_cancel_l k n)^ in q; - assumption. -Defined. - -(** TODO: move, rename *) -Proposition natpmswap3 (k m n : nat) - : k <= n -> m <= n - k -> k + m <= n. -Proof. - intros ineq qe. - apply (nat_add_l_monotone k) in qe. - destruct (nataddsub_assoc k ineq)^ in qe. - destruct (nat_add_sub_cancel_l n k)^ in qe; - assumption. -Defined. - -(** TODO: move, rename *) -Proposition natpmswap4 (k m n : nat) - : k - n < m -> k < n + m. -Proof. - intro l; apply (nat_add_r_strictly_monotone n) in l. - destruct (nat_add_comm m n). - now rapply (leq_lt_trans (nat_sub_add_ineq _ _)). -Defined. - (** TODO: move, rename *) Proposition n_leq_m_n_leq_plus_m_k (n m k : nat) : n <= m -> n <= m + k. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index f1b127a16fc..a6fff9f8cf1 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1213,7 +1213,60 @@ Hint Immediate nat_sub_monotone_r : typeclass_instances. (** *** Movement Lemmas *) +(** TODO: rename *) +Proposition natpmswap1 (k m n : nat) + : n <= k -> k < n + m -> k - n < m. +Proof. + intros l q. + assert (q' : k - n + n < m + n) by + (destruct (nat_add_sub_l_cancel l)^; + destruct (nat_add_comm n m); + assumption). + exact (lt_reflects_add_r _ q'). +Defined. + +(** TODO: rename *) +Proposition natpmswap2 (k m n : nat) + : n <= k -> k - n <= m -> k <= n + m. +Proof. + intros l q. + apply (nat_add_l_monotone n) in q. + destruct (nataddsub_assoc n l)^ in q. + destruct (nat_add_sub_cancel_l k n)^ in q; + assumption. +Defined. + +(** TODO: rename *) +Proposition natpmswap3 (k m n : nat) + : k <= n -> m <= n - k -> k + m <= n. +Proof. + intros ineq qe. + apply (nat_add_l_monotone k) in qe. + destruct (nataddsub_assoc k ineq)^ in qe. + destruct (nat_add_sub_cancel_l n k)^ in qe; + assumption. +Defined. + +(** TODO: move, rename *) +Proposition nat_sub_add_ineq (n m : nat) : n <= n - m + m. +Proof. + destruct (@leq_dichotomy m n) as [l | gt]. + - destruct (nataddsub_comm _ _ m l)^. + destruct (nat_add_sub_cancel_r n m)^. + apply leq_refl; done. + - apply leq_lt in gt. + destruct (equiv_nat_sub_leq _)^. + assumption. +Defined. +(** TODO: rename *) +Proposition natpmswap4 (k m n : nat) + : k - n < m -> k < n + m. +Proof. + intro l; apply (nat_add_r_strictly_monotone n) in l. + destruct (nat_add_comm m n). + now rapply (leq_lt_trans (nat_sub_add_ineq _ _)). +Defined. (** *** Order-reflection Lemmas *) From 9fbd3758dab13c63d3079aabf8267e91a7c703b8 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 31 Jul 2024 18:41:32 +0200 Subject: [PATCH 201/282] Nat: subtraction order reflection lemma Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 50 +++++++++++++++++++++++++------------- 1 file changed, 33 insertions(+), 17 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index a6fff9f8cf1..e4045ee0756 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -454,6 +454,27 @@ Defined. Global Instance irreflexive_lt : Irreflexive lt := lt_irrefl. Global Instance irreflexive_gt : Irreflexive gt := lt_irrefl. +(** [<=] is an antisymmetric relation. *) +Definition leq_antisym {x y} : x <= y -> y <= x -> x = y. +Proof. + intros p q. + destruct p. + 1: reflexivity. + destruct x; [inversion q|]. + apply leq_succ' in q. + contradiction (lt_irrefl _ (leq_trans p q)). +Defined. + +Global Instance antisymmetric_leq : AntiSymmetric leq := @leq_antisym. +Global Instance antisymemtric_geq : AntiSymmetric geq + := fun _ _ p q => leq_antisym q p. + +(** Being less than or equal to [0] implies being [0]. *) +Definition path_zero_leq_zero_r n : n <= 0 -> n = 0. +Proof. + intros H; rapply leq_antisym. +Defined. + (** Nothing can be less than [0]. *) Definition not_lt_zero_r n : ~ (n < 0). Proof. @@ -804,23 +825,6 @@ Defined. (** ** More Theory of Comparison Predicates *) -(** *** Antisymmetry of [<=] and [>=] *) - -(** [<=] is an antisymmetric relation. *) -Definition leq_antisym {x y} : x <= y -> y <= x -> x = y. -Proof. - intros p q. - destruct p. - 1: reflexivity. - destruct x; [inversion q|]. - apply leq_succ' in q. - contradiction (lt_irrefl _ (leq_trans p q)). -Defined. - -Global Instance antisymmetric_leq : AntiSymmetric leq := @leq_antisym. -Global Instance antisymemtric_geq : AntiSymmetric geq - := fun _ _ p q => leq_antisym q p. - (** *** Addition Lemmas *) (** The first summand is less than or equal to the sum. *) @@ -1270,6 +1274,7 @@ Defined. (** *** Order-reflection Lemmas *) +(** Subtraction reflects [<=] in the left argument. *) Definition leq_reflects_sub_l {n m} k : k <= m -> n - k <= m - k -> n <= m. Proof. intros ineq1 ineq2. @@ -1279,6 +1284,17 @@ Proof. by rewrite nat_add_sub_l_cancel. Defined. +(** Subtraction reflects [<=] in the right argument contravariantly. *) +Definition leq_reflects_sub_r {n m} k + : m <= k -> n <= k -> k - n <= k - m -> m <= n. +Proof. + intros H1 H2 H3. + rapply (leq_reflects_add_r k). + rapply natpmswap3. + rewrite <- nataddsub_assoc; only 2: exact _. + rapply natpmswap2. +Defined. + (** ** Properties of Powers *) (** [0] to any power is [0] unless that power is [0] in which case it is [1]. *) From 881453c3a9fbdcc404e6ba87428fd62afc0d3b7d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 14:10:22 +0200 Subject: [PATCH 202/282] Nat: add nat_sub_pred_r Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 15 ++++++++++++++- 1 file changed, 14 insertions(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index e4045ee0756..58aa0dce907 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -425,7 +425,7 @@ Defined. Global Existing Instance leq_zero_l | 10. (** A predecessor is less than or equal to a predecessor if the original number is less than or equal. *) -Definition leq_pred {n m} : n <= m -> nat_pred n <= nat_pred m. +Global Instance leq_pred {n m} : n <= m -> nat_pred n <= nat_pred m. Proof. intros H; induction H. 1: exact _. @@ -1136,6 +1136,19 @@ Proof. by apply lt_sub_gt_0. Defined. +(** Under certain conditions, subtracting a predecessor is the successor of the subtraction. *) +Definition nat_sub_pred_r n m : 0 < m -> m < n -> n - nat_pred m = (n - m).+1. +Proof. + intros H1 H2. + induction m as [|m IHm] in n, H1, H2 |- *. + 1: contradiction (not_lt_zero_r _ H1). + rewrite nat_sub_succ_r. + rewrite nat_succ_pred. + 1: reflexivity. + apply lt_sub_gt_0. + exact (lt_trans _ H2). +Defined. + (** TODO: rename *) Definition nataddsub_assoc_lemma {k m : nat} : (k <= m) -> m.+1 - k = (m - k).+1. From 91754618fe1862a9105271b52c2dd73f73a27ca9 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 14:10:38 +0200 Subject: [PATCH 203/282] Nat: add nat_sub_sub_cancel_r Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 58aa0dce907..b0ee9d30ec4 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1195,6 +1195,15 @@ Proof. apply nat_add_comm. Defined. +(** Subtracting a subtraction is the subtrahend. *) +Definition nat_sub_sub_cancel_r {n m} : n <= m -> m - (m - n) = n. +Proof. + intros H; induction H. + - by rewrite nat_sub_cancel, nat_sub_zero_r. + - rewrite (nat_sub_succ_l m n); only 2: exact _. + exact IHleq. +Defined. + (** *** Monotonicity of Subtraction *) (** Subtraction is monotone in the left argument. *) From 920216f55d869b0ecf0be4da7fc6fe6940ced245 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 14:12:12 +0200 Subject: [PATCH 204/282] Nat: revise proof of leq_reflects_sub_r Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index b0ee9d30ec4..52d82719557 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1311,10 +1311,8 @@ Definition leq_reflects_sub_r {n m} k : m <= k -> n <= k -> k - n <= k - m -> m <= n. Proof. intros H1 H2 H3. - rapply (leq_reflects_add_r k). - rapply natpmswap3. - rewrite <- nataddsub_assoc; only 2: exact _. - rapply natpmswap2. + apply (nat_sub_monotone_r k) in H3. + rewrite 2 nat_sub_sub_cancel_r in H3; exact _. Defined. (** ** Properties of Powers *) From 26d4132b3b5031a9a1a1ba56d725d69b9c0a9478 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 14:12:41 +0200 Subject: [PATCH 205/282] Nat: add leq_moveL_nM Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 8 +++++++- 1 file changed, 7 insertions(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 52d82719557..d124d5776ab 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1262,7 +1262,13 @@ Proof. assumption. Defined. -(** TODO: rename *) +Definition leq_moveL_nM {k m n} + : n <= k -> k - n <= m -> k <= m + n. +Proof. + rewrite nat_add_comm. + apply natpmswap2. +Defined. + Proposition natpmswap3 (k m n : nat) : k <= n -> m <= n - k -> k + m <= n. Proof. From 2c582e01262dae76b86cfb6ad8689d5ef3592e87 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 14:13:21 +0200 Subject: [PATCH 206/282] Nat: add todo comments for movement lemmas Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 9 +++++---- 1 file changed, 5 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index d124d5776ab..ae2c3e0ccef 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1239,7 +1239,7 @@ Hint Immediate nat_sub_monotone_r : typeclass_instances. (** *** Movement Lemmas *) -(** TODO: rename *) +(** TODO: rename natpmswap1 -> lt_moveR_nV *) Proposition natpmswap1 (k m n : nat) : n <= k -> k < n + m -> k - n < m. Proof. @@ -1251,7 +1251,7 @@ Proof. exact (lt_reflects_add_r _ q'). Defined. -(** TODO: rename *) +(** TODO: rename natpmswap2 -> leq_moveL_Mn *) Proposition natpmswap2 (k m n : nat) : n <= k -> k - n <= m -> k <= n + m. Proof. @@ -1269,6 +1269,7 @@ Proof. apply natpmswap2. Defined. +(** TODO: rename natpmswap3 -> leq_moveR_Mn *) Proposition natpmswap3 (k m n : nat) : k <= n -> m <= n - k -> k + m <= n. Proof. @@ -1279,7 +1280,7 @@ Proof. assumption. Defined. -(** TODO: move, rename *) +(** TODO: rename, move *) Proposition nat_sub_add_ineq (n m : nat) : n <= n - m + m. Proof. destruct (@leq_dichotomy m n) as [l | gt]. @@ -1291,7 +1292,7 @@ Proof. assumption. Defined. -(** TODO: rename *) +(** TODO: rename natpmswap4 -> lt_moveL_Mm *) Proposition natpmswap4 (k m n : nat) : k - n < m -> k < n + m. Proof. From 0d4bec3230935a76d6f1200eee5ced240301cf0d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 14:26:59 +0200 Subject: [PATCH 207/282] Nat: merge sub_gt_0_lt and lt_sub_gt_0 into equiv_lt_lt_sub Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 6 +++--- theories/Spaces/Nat/Core.v | 36 +++++++++++++------------------- 2 files changed, 18 insertions(+), 24 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 0c7c9f7a502..3d18ce1eaa3 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -262,7 +262,7 @@ Proof. destruct (symmetric_paths _ _ (@predeqminus1 (n - k))). apply increasing_geq_S. unfold ">", "<" in *. - apply lt_sub_gt_0 in g. + apply equiv_lt_lt_sub in g. now (destruct (symmetric_paths _ _ (S_predn (n - k) 0 _))). Defined. @@ -284,9 +284,9 @@ Proof. - intros m ineq. change (m - n.+1) with (m - (1 + n)). (destruct (nat_add_comm n 1)). destruct (symmetric_paths _ _ (nat_sub_add m n 1)). - destruct (S_predn (m - n) 0 (lt_sub_gt_0 _ _ ineq)); simpl; + destruct (S_predn (m - n) 0 (equiv_lt_lt_sub _ _ ineq)); simpl; destruct (symmetric_paths _ _ (nat_sub_zero_r (nat_pred (m - n)))). - assert (0 < m - n) as dp by exact (lt_sub_gt_0 _ _ ineq). + assert (0 < m - n) as dp by exact (equiv_lt_lt_sub _ _ ineq). assert (nat_pred (m - n) < m) as sh by ( unfold "<"; destruct (symmetric_paths _ _ (S_predn _ 0 _)); diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ae2c3e0ccef..32e51912e44 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -970,26 +970,20 @@ Proof. now destruct (equiv_nat_sub_leq _)^. Defined. -(** *** Movement Lemmas *) - -(** TODO: rename sub_gt_0_lt to lt_moveL_m (?) , reprove *) -Definition sub_gt_0_lt n m : 0 < n - m -> m < n. -Proof. - intro ineq. - destruct (@leq_dichotomy n m) as [n_leq_m |]; [ | assumption]. - apply equiv_nat_sub_leq in n_leq_m. - contradiction (lt_irrefl 0). now rewrite n_leq_m in ineq. -Defined. - -(** TODO: merge with above, reprove *) -Definition lt_sub_gt_0 n m : m < n -> 0 < n - m. +(** A number being less than another is equivalent to their difference being greater than zero. *) +Definition equiv_lt_lt_sub n m : m < n <~> 0 < n - m. Proof. - revert m; simple_induction n n IHn. - - intros m ineq. contradiction (not_lt_zero_r m). - - destruct m. - + simpl. easy. - + simpl. intro ineq. apply leq_succ' in ineq. - now apply IHn in ineq. + srapply equiv_iff_hprop. + - revert m; simple_induction n n IHn. + + intros m ineq. contradiction (not_lt_zero_r m). + + destruct m. + * simpl. easy. + * simpl. intro ineq. apply leq_succ' in ineq. + now apply IHn in ineq. + - intro ineq. + destruct (@leq_dichotomy n m) as [n_leq_m |]; [ | assumption]. + apply equiv_nat_sub_leq in n_leq_m. + contradiction (lt_irrefl 0). now rewrite n_leq_m in ineq. Defined. (** *** Monotonicity of Addition *) @@ -1133,7 +1127,7 @@ Proof. rewrite nat_sub_succ_r. symmetry. apply nat_succ_pred. - by apply lt_sub_gt_0. + by apply equiv_lt_lt_sub. Defined. (** Under certain conditions, subtracting a predecessor is the successor of the subtraction. *) @@ -1145,7 +1139,7 @@ Proof. rewrite nat_sub_succ_r. rewrite nat_succ_pred. 1: reflexivity. - apply lt_sub_gt_0. + apply equiv_lt_lt_sub. exact (lt_trans _ H2). Defined. From 2fd203472d7294efecf416c174ee67cf298eed4e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 14:37:36 +0200 Subject: [PATCH 208/282] Nat: revise proof of nataddsub_assoc Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 24 +++++++----------------- 1 file changed, 7 insertions(+), 17 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 32e51912e44..842dd44b6dc 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1158,25 +1158,15 @@ Proof. Defined. (** TODO: rename *) -Definition nataddsub_assoc(n : nat) {m k : nat} - : (k <= m) -> n + (m - k) = n + m - k. +Definition nataddsub_assoc {m k} n + : k <= m -> n + (m - k) = n + m - k. Proof. - revert m k. simple_induction n n IHn. + intros H; induction n as [|n IHn] in |- *. - reflexivity. - - intros m k l. - change (n.+1 + (m - k)) with (n + (m - k)).+1; - change (n.+1 + m) with (n +m).+1. - destruct k, m; - [ reflexivity - | reflexivity - | contradiction (not_lt_zero_r k _) - | ]. - simpl "-". apply leq_succ' in l. - destruct (nat_add_succ_r n (m - k)). - destruct (nataddsub_assoc_lemma l). - apply (IHn m.+1 k). - apply leq_succ_r. - assumption. + - change (?n.+1 + ?m) with (n + m).+1. + rhs nrapply nat_sub_succ_l. + 1: exact (ap nat_succ IHn). + exact _. Defined. (** TODO: rename *) From 379b34eb07b8e7482e6aacadc5e7b03fcb897c99 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 14:38:34 +0200 Subject: [PATCH 209/282] Nat: replace use of nataddsub_lemma with nat_sub_succ_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 14 -------------- 2 files changed, 1 insertion(+), 15 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 3d18ce1eaa3..be2c65125d8 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -272,7 +272,7 @@ Proof. destruct n. - contradiction (not_lt_zero_r k). - change (n.+1 - k.+1) with (n - k). apply leq_succ' in ineq. - apply (nataddsub_assoc_lemma _). + by apply nat_sub_succ_l. Defined. Lemma ineq_sub (n m : nat) : n <= m -> m - (m - n) = n. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 842dd44b6dc..2e94eec69c6 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1143,20 +1143,6 @@ Proof. exact (lt_trans _ H2). Defined. -(** TODO: rename *) -Definition nataddsub_assoc_lemma {k m : nat} - : (k <= m) -> m.+1 - k = (m - k).+1. -Proof. - revert m; simple_induction k k IHk. - - intros m l; simpl. destruct m; reflexivity. - - destruct m. - + simpl; intro g; contradiction (not_lt_zero_r _ g). - + intro l; apply leq_succ' in l. - change (m.+2 - k.+1) with (m.+1 - k). - change (m.+1 - k.+1) with (m - k). - exact (IHk _ l). -Defined. - (** TODO: rename *) Definition nataddsub_assoc {m k} n : k <= m -> n + (m - k) = n + m - k. From e92918ca8d5c102543495fa6ffdeb4db3d287a52 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 14:39:35 +0200 Subject: [PATCH 210/282] Nat: remove comment about renaming nat lemmas since they have been Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 2 -- 1 file changed, 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 2e94eec69c6..40f476579e9 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -11,8 +11,6 @@ Local Unset Elimination Schemes. Local Close Scope trunc_scope. Local Open Scope nat_scope. -(** TODO: Some results are prefixed with [nat_] and some are not. Should we be more consistent? *) - (** ** Basic operations on naturals *) (** *** Iteration *) From 8dc1ab9057b0d5624705677b35040cdb17c63320 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 14:41:51 +0200 Subject: [PATCH 211/282] Nat: rename nataddsub_assoc -> nat_sub_l_add_r Signed-off-by: Ali Caglayan --- theories/Spaces/List/Theory.v | 13 ++++++------- theories/Spaces/Nat/Arithmetic.v | 4 ++-- theories/Spaces/Nat/Core.v | 16 ++++++++-------- 3 files changed, 16 insertions(+), 17 deletions(-) diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index ad191e4a727..7bfa375e329 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -694,13 +694,12 @@ Proof. rewrite length_take. rewrite length_drop. rewrite nat_min_l. - - rewrite nataddsub_assoc. - 2: exact H. - rewrite <- predeqminus1. - induction n in |- *. - 1: reflexivity. - cbn; assumption. - - exact (leq_trans _ H). + 2: exact (leq_trans _ H). + rewrite <- nat_sub_l_add_r. + 2: exact _. + lhs nrapply nat_sub_succ_r. + apply ap. + apply nat_add_sub_cancel_l. Defined. (** An element of a [remove] is an element of the original list. *) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index be2c65125d8..26b39802dec 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -55,14 +55,14 @@ Proposition nataddsub_assoc_implication (n : nat) {m k z : nat} : (k <= m) -> n + (m - k) = z -> n + m - k = z. Proof. intro H. - destruct (symmetric_paths _ _ (nataddsub_assoc n H)); done. + destruct (symmetric_paths _ _ (nat_sub_l_add_r n H)); done. Defined. Proposition nat_add_sub_eq (n : nat) {k: nat} : (k <= n) -> k + (n - k) = n. Proof. intro H. - destruct (symmetric_paths _ _ (nataddsub_assoc k H)); + destruct (nat_sub_l_add_r k H); destruct (nat_add_comm n k); exact (nat_add_sub_cancel_r _ _). Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 40f476579e9..a2c09fade83 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1141,15 +1141,15 @@ Proof. exact (lt_trans _ H2). Defined. -(** TODO: rename *) -Definition nataddsub_assoc {m k} n - : k <= m -> n + (m - k) = n + m - k. +(** Subtraction and addition satisfy an associativity law. *) +Definition nat_sub_l_add_r {m k} n + : k <= m -> (n + m) - k = n + (m - k). Proof. intros H; induction n as [|n IHn] in |- *. - reflexivity. - change (?n.+1 + ?m) with (n + m).+1. - rhs nrapply nat_sub_succ_l. - 1: exact (ap nat_succ IHn). + lhs nrapply nat_sub_succ_l. + 2: exact (ap nat_succ IHn). exact _. Defined. @@ -1159,7 +1159,7 @@ Definition nataddsub_comm (n m k : nat) Proof. intro l. destruct (nat_add_comm k n). - destruct (nataddsub_assoc k l). + rewrite (nat_sub_l_add_r k l). apply nat_add_comm. Defined. @@ -1225,7 +1225,7 @@ Proposition natpmswap2 (k m n : nat) Proof. intros l q. apply (nat_add_l_monotone n) in q. - destruct (nataddsub_assoc n l)^ in q. + destruct (nat_sub_l_add_r n l) in q. destruct (nat_add_sub_cancel_l k n)^ in q; assumption. Defined. @@ -1243,7 +1243,7 @@ Proposition natpmswap3 (k m n : nat) Proof. intros ineq qe. apply (nat_add_l_monotone k) in qe. - destruct (nataddsub_assoc k ineq)^ in qe. + destruct (nat_sub_l_add_r k ineq) in qe. destruct (nat_add_sub_cancel_l n k)^ in qe; assumption. Defined. From 456ebc939d046a78c4386af47647e9722f0d48ce Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:13:58 +0200 Subject: [PATCH 212/282] Nat: rename nat_sub_add -> nat_sub_r_add, revise proofs, update comments Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 14 +++++++------- theories/Spaces/Nat/Core.v | 27 ++++++++++++++------------- 2 files changed, 21 insertions(+), 20 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 26b39802dec..d19a7dc2f2b 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -33,7 +33,7 @@ Defined. Proposition nat_sub_add_ineq (n m : nat) : n <= n - m + m. Proof. destruct (@leq_dichotomy m n) as [l | gt]. - - destruct (nataddsub_comm _ _ m l)^. + - rewrite <- nataddsub_comm; trivial. destruct (nat_add_sub_cancel_r n m)^. apply leq_refl; done. - apply leq_lt in gt. @@ -55,14 +55,14 @@ Proposition nataddsub_assoc_implication (n : nat) {m k z : nat} : (k <= m) -> n + (m - k) = z -> n + m - k = z. Proof. intro H. - destruct (symmetric_paths _ _ (nat_sub_l_add_r n H)); done. + by rewrite nat_sub_l_add_r. Defined. Proposition nat_add_sub_eq (n : nat) {k: nat} : (k <= n) -> k + (n - k) = n. Proof. intro H. - destruct (nat_sub_l_add_r k H); + destruct (nat_sub_l_add_r _ k _ H); destruct (nat_add_comm n k); exact (nat_add_sub_cancel_r _ _). Defined. @@ -161,8 +161,8 @@ Proof. (symmetric_paths _ _ (nat_sub_zero_r m)); done. - intro l. change (k.+1) with (1 + k). destruct (nat_add_comm k 1). - destruct (symmetric_paths _ _ (nat_sub_add n k 1)). - destruct (symmetric_paths _ _ (nat_sub_add m k 1)). + destruct (symmetric_paths _ _ (nat_sub_r_add n k 1)). + destruct (symmetric_paths _ _ (nat_sub_r_add m k 1)). destruct (symmetric_paths _ _ (@predeqminus1 (n -k))). destruct (symmetric_paths _ _ (@predeqminus1 (m -k))). apply leq_pred, IHk. exact l. @@ -258,7 +258,7 @@ Proof. apply leq_succ_r in l. destruct (equiv_nat_sub_leq _)^. exact IHk. + change k.+1 with (1 + k). destruct (nat_add_comm k 1). - destruct (symmetric_paths _ _ (nat_sub_add n k 1)). + destruct (symmetric_paths _ _ (nat_sub_r_add n k 1)). destruct (symmetric_paths _ _ (@predeqminus1 (n - k))). apply increasing_geq_S. unfold ">", "<" in *. @@ -283,7 +283,7 @@ Proof. reflexivity. - intros m ineq. change (m - n.+1) with (m - (1 + n)). (destruct (nat_add_comm n 1)). - destruct (symmetric_paths _ _ (nat_sub_add m n 1)). + destruct (symmetric_paths _ _ (nat_sub_r_add m n 1)). destruct (S_predn (m - n) 0 (equiv_lt_lt_sub _ _ ineq)); simpl; destruct (symmetric_paths _ _ (nat_sub_zero_r (nat_pred (m - n)))). assert (0 < m - n) as dp by exact (equiv_lt_lt_sub _ _ ineq). diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index a2c09fade83..27066aa7232 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -608,7 +608,7 @@ Proof. Defined. (** Subtracting an addition is the same as subtracting the two numbers separately. *) -Definition nat_sub_add@{} n m k : n - (m + k) = n - m - k. +Definition nat_sub_r_add@{} n m k : n - (m + k) = n - m - k. Proof. induction n as [|n IHn] in m, k |- *. - reflexivity. @@ -620,9 +620,9 @@ Defined. (** The order in which two numbers are subtracted does not matter. *) Definition nat_sub_comm_r@{} n m k : n - m - k = n - k - m. Proof. - lhs_V nrapply nat_sub_add. + lhs_V nrapply nat_sub_r_add. rewrite nat_add_comm. - nrapply nat_sub_add. + nrapply nat_sub_r_add. Defined. (** Subtracting a larger number from a smaller number is [0]. *) @@ -632,7 +632,7 @@ Proof. - intro l; induction l. + exact (nat_sub_cancel n). + change (m.+1) with (1 + m). - lhs nrapply nat_sub_add. + lhs nrapply nat_sub_r_add. lhs nrapply nat_sub_comm_r. by destruct IHl^. - induction n as [|n IHn] in m |- *. @@ -1142,7 +1142,7 @@ Proof. Defined. (** Subtraction and addition satisfy an associativity law. *) -Definition nat_sub_l_add_r {m k} n +Definition nat_sub_l_add_r m n k : k <= m -> (n + m) - k = n + (m - k). Proof. intros H; induction n as [|n IHn] in |- *. @@ -1153,13 +1153,13 @@ Proof. exact _. Defined. -(** TODO: rename *) -Definition nataddsub_comm (n m k : nat) - : m <= n -> (n - m) + k = (n + k) - m. +(** TODO: rename nataddsub_comm -> nat_sub_l_add_r *) +Definition nataddsub_comm n m k + : k <= n -> (n + m) - k = (n - k) + m. Proof. intro l. - destruct (nat_add_comm k n). - rewrite (nat_sub_l_add_r k l). + rewrite nat_add_comm. + lhs rapply nat_sub_l_add_r. apply nat_add_comm. Defined. @@ -1225,7 +1225,7 @@ Proposition natpmswap2 (k m n : nat) Proof. intros l q. apply (nat_add_l_monotone n) in q. - destruct (nat_sub_l_add_r n l) in q. + rewrite <- nat_sub_l_add_r in q; trivial. destruct (nat_add_sub_cancel_l k n)^ in q; assumption. Defined. @@ -1243,7 +1243,7 @@ Proposition natpmswap3 (k m n : nat) Proof. intros ineq qe. apply (nat_add_l_monotone k) in qe. - destruct (nat_sub_l_add_r k ineq) in qe. + rewrite <- nat_sub_l_add_r in qe; trivial. destruct (nat_add_sub_cancel_l n k)^ in qe; assumption. Defined. @@ -1252,7 +1252,8 @@ Defined. Proposition nat_sub_add_ineq (n m : nat) : n <= n - m + m. Proof. destruct (@leq_dichotomy m n) as [l | gt]. - - destruct (nataddsub_comm _ _ m l)^. + - rewrite <- nataddsub_comm; trivial. + rewrite nat_add_sub_cancel_r. destruct (nat_add_sub_cancel_r n m)^. apply leq_refl; done. - apply leq_lt in gt. From 169fbca3230a092219cd04fa34d39da0c10c4927 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:17:20 +0200 Subject: [PATCH 213/282] Nat: rename natsubadd_comm -> nat_sub_l_add_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 6 +++--- 2 files changed, 4 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index d19a7dc2f2b..4bf1c6f337d 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -33,7 +33,7 @@ Defined. Proposition nat_sub_add_ineq (n m : nat) : n <= n - m + m. Proof. destruct (@leq_dichotomy m n) as [l | gt]. - - rewrite <- nataddsub_comm; trivial. + - rewrite <- nat_sub_l_add_l; trivial. destruct (nat_add_sub_cancel_r n m)^. apply leq_refl; done. - apply leq_lt in gt. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 27066aa7232..92c083bb0a5 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1153,8 +1153,8 @@ Proof. exact _. Defined. -(** TODO: rename nataddsub_comm -> nat_sub_l_add_r *) -Definition nataddsub_comm n m k +(** Subtraction and addition satisfy a special commutativity law. *) +Definition nat_sub_l_add_l n m k : k <= n -> (n + m) - k = (n - k) + m. Proof. intro l. @@ -1252,7 +1252,7 @@ Defined. Proposition nat_sub_add_ineq (n m : nat) : n <= n - m + m. Proof. destruct (@leq_dichotomy m n) as [l | gt]. - - rewrite <- nataddsub_comm; trivial. + - rewrite <- nat_sub_l_add_l; trivial. rewrite nat_add_sub_cancel_r. destruct (nat_add_sub_cancel_r n m)^. apply leq_refl; done. From b05c46e5c267e07f74e5bd8a6f8903da80eec7b3 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:18:23 +0200 Subject: [PATCH 214/282] Nat: better comments for nat_sub_l_add_{l,r} Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 92c083bb0a5..5d993ad8aea 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1141,7 +1141,7 @@ Proof. exact (lt_trans _ H2). Defined. -(** Subtraction and addition satisfy an associativity law. *) +(** Subtracting from a sum is the sum of subtracting from the second summand. *) Definition nat_sub_l_add_r m n k : k <= m -> (n + m) - k = n + (m - k). Proof. @@ -1153,7 +1153,7 @@ Proof. exact _. Defined. -(** Subtraction and addition satisfy a special commutativity law. *) +(** Subtracting from a sum is the sum of subtracting from the first summand. *) Definition nat_sub_l_add_l n m k : k <= n -> (n + m) - k = (n - k) + m. Proof. From b2bc8a9d9bea896fb1d08840247b5030ce7a5b69 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:35:23 +0200 Subject: [PATCH 215/282] Nat: rename natpmswap1 -> lt_moveR_nV Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 5d993ad8aea..a6e61bcbaa7 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1207,9 +1207,7 @@ Hint Immediate nat_sub_monotone_r : typeclass_instances. (** *** Movement Lemmas *) -(** TODO: rename natpmswap1 -> lt_moveR_nV *) -Proposition natpmswap1 (k m n : nat) - : n <= k -> k < n + m -> k - n < m. +Definition lt_moveR_nV {k m} n : n <= k -> k < n + m -> k - n < m. Proof. intros l q. assert (q' : k - n + n < m + n) by From 1188a2eeab8d5befae0e75836c287d60e6ff18ae Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:44:36 +0200 Subject: [PATCH 216/282] Nat: revise proof of lt_moveR_nV Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 11 +++++------ 1 file changed, 5 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index a6e61bcbaa7..49cf7e628b5 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1209,12 +1209,11 @@ Hint Immediate nat_sub_monotone_r : typeclass_instances. Definition lt_moveR_nV {k m} n : n <= k -> k < n + m -> k - n < m. Proof. - intros l q. - assert (q' : k - n + n < m + n) by - (destruct (nat_add_sub_l_cancel l)^; - destruct (nat_add_comm n m); - assumption). - exact (lt_reflects_add_r _ q'). + intros H1 H2; unfold lt. + rewrite <- nat_sub_succ_l; only 2: exact _. + rewrite <- (nat_add_sub_cancel_l m n). + apply nat_sub_monotone_l. + exact H2. Defined. (** TODO: rename natpmswap2 -> leq_moveL_Mn *) From b7ebeae2d7d97940d3662031e92baf3a95062523 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:45:18 +0200 Subject: [PATCH 217/282] Nat: rename natpmswap2 -> leq_moveL_Mn Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 6 ++---- 1 file changed, 2 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 49cf7e628b5..56a71ddb22f 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1216,9 +1216,7 @@ Proof. exact H2. Defined. -(** TODO: rename natpmswap2 -> leq_moveL_Mn *) -Proposition natpmswap2 (k m n : nat) - : n <= k -> k - n <= m -> k <= n + m. +Definition leq_moveL_Mn {k m} n : n <= k -> k - n <= m -> k <= n + m. Proof. intros l q. apply (nat_add_l_monotone n) in q. @@ -1231,7 +1229,7 @@ Definition leq_moveL_nM {k m n} : n <= k -> k - n <= m -> k <= m + n. Proof. rewrite nat_add_comm. - apply natpmswap2. + apply leq_moveL_Mn. Defined. (** TODO: rename natpmswap3 -> leq_moveR_Mn *) From 61d726875322b7fb28aee8c66c77c825ada59c5c Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:47:21 +0200 Subject: [PATCH 218/282] Nat: rename natpmswap3 -> leq_moveR_Mn Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 56a71ddb22f..8833657417a 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1233,8 +1233,7 @@ Proof. Defined. (** TODO: rename natpmswap3 -> leq_moveR_Mn *) -Proposition natpmswap3 (k m n : nat) - : k <= n -> m <= n - k -> k + m <= n. +Definition leq_moveR_Mn {m n} k : k <= n -> m <= n - k -> k + m <= n. Proof. intros ineq qe. apply (nat_add_l_monotone k) in qe. From 1395d2821510dfae3d0795773b4fe4791a75d52f Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:49:43 +0200 Subject: [PATCH 219/282] Nat: revise proof of leq_moveL_Mn Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 10 ++++------ 1 file changed, 4 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 8833657417a..3910f14e967 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1218,11 +1218,10 @@ Defined. Definition leq_moveL_Mn {k m} n : n <= k -> k - n <= m -> k <= n + m. Proof. - intros l q. - apply (nat_add_l_monotone n) in q. - rewrite <- nat_sub_l_add_r in q; trivial. - destruct (nat_add_sub_cancel_l k n)^ in q; - assumption. + intros H1 H2. + apply (nat_add_l_monotone n) in H2. + rewrite <- nat_sub_l_add_r in H2; only 2: exact _. + by rewrite nat_add_sub_cancel_l in H2. Defined. Definition leq_moveL_nM {k m n} @@ -1232,7 +1231,6 @@ Proof. apply leq_moveL_Mn. Defined. -(** TODO: rename natpmswap3 -> leq_moveR_Mn *) Definition leq_moveR_Mn {m n} k : k <= n -> m <= n - k -> k + m <= n. Proof. intros ineq qe. From fe9996bfe4ceda9c7ec98d31b2a094023a424455 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:50:42 +0200 Subject: [PATCH 220/282] Nat: revise proof of leq_moveR_Mn Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 9 ++++----- 1 file changed, 4 insertions(+), 5 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 3910f14e967..de43d07c2e3 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1233,11 +1233,10 @@ Defined. Definition leq_moveR_Mn {m n} k : k <= n -> m <= n - k -> k + m <= n. Proof. - intros ineq qe. - apply (nat_add_l_monotone k) in qe. - rewrite <- nat_sub_l_add_r in qe; trivial. - destruct (nat_add_sub_cancel_l n k)^ in qe; - assumption. + intros H1 H2. + apply (nat_add_l_monotone k) in H2. + rewrite <- nat_sub_l_add_r in H2; trivial. + by rewrite nat_add_sub_cancel_l in H2. Defined. (** TODO: rename, move *) From 4cd5e59b20e3468bd35ca5f055002acec2027c9c Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:51:38 +0200 Subject: [PATCH 221/282] Nat: adjust style of leq_moveL_nM Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 3 +-- 1 file changed, 1 insertion(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index de43d07c2e3..857ad84be5d 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1224,8 +1224,7 @@ Proof. by rewrite nat_add_sub_cancel_l in H2. Defined. -Definition leq_moveL_nM {k m n} - : n <= k -> k - n <= m -> k <= m + n. +Definition leq_moveL_nM {k m} n : n <= k -> k - n <= m -> k <= m + n. Proof. rewrite nat_add_comm. apply leq_moveL_Mn. From 84613feec72a932f23fb3aab92e76f7358fb118e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 15:52:35 +0200 Subject: [PATCH 222/282] Nat: rename natpmswap4 -> lt_moveL_Mn Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 +--- 1 file changed, 1 insertion(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 857ad84be5d..11cae571ed7 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1251,9 +1251,7 @@ Proof. assumption. Defined. -(** TODO: rename natpmswap4 -> lt_moveL_Mm *) -Proposition natpmswap4 (k m n : nat) - : k - n < m -> k < n + m. +Definition lt_moveL_Mn {k m} n : k - n < m -> k < n + m. Proof. intro l; apply (nat_add_r_strictly_monotone n) in l. destruct (nat_add_comm m n). From acc576c210135e54ae3c4ceb90d9d97c304e982a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 16:01:34 +0200 Subject: [PATCH 223/282] Nat: revise proof of lt_moveL_Mn Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 28 ++++++++-------------------- 1 file changed, 8 insertions(+), 20 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 11cae571ed7..e8a32f780d1 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -959,13 +959,12 @@ Defined. (** *** Subtraction *) -Definition leq_sub_add n m : n <= n - m + m. +Global Instance leq_sub_add n m : n <= n - m + m. Proof. destruct (@leq_dichotomy m n) as [l | g]. - - destruct (nat_add_sub_l_cancel l)^; - constructor. + - by rewrite nat_add_sub_l_cancel. - apply leq_lt in g. - now destruct (equiv_nat_sub_leq _)^. + by destruct (equiv_nat_sub_leq _)^. Defined. (** A number being less than another is equivalent to their difference being greater than zero. *) @@ -1216,6 +1215,7 @@ Proof. exact H2. Defined. +(** TODO: [n <= k] can be dropped *) Definition leq_moveL_Mn {k m} n : n <= k -> k - n <= m -> k <= n + m. Proof. intros H1 H2. @@ -1238,24 +1238,12 @@ Proof. by rewrite nat_add_sub_cancel_l in H2. Defined. -(** TODO: rename, move *) -Proposition nat_sub_add_ineq (n m : nat) : n <= n - m + m. -Proof. - destruct (@leq_dichotomy m n) as [l | gt]. - - rewrite <- nat_sub_l_add_l; trivial. - rewrite nat_add_sub_cancel_r. - destruct (nat_add_sub_cancel_r n m)^. - apply leq_refl; done. - - apply leq_lt in gt. - destruct (equiv_nat_sub_leq _)^. - assumption. -Defined. - Definition lt_moveL_Mn {k m} n : k - n < m -> k < n + m. Proof. - intro l; apply (nat_add_r_strictly_monotone n) in l. - destruct (nat_add_comm m n). - now rapply (leq_lt_trans (nat_sub_add_ineq _ _)). + intros H. + rewrite nat_add_comm. + apply (nat_add_r_strictly_monotone n) in H. + rapply leq_lt_trans. Defined. (** *** Order-reflection Lemmas *) From a2ca46eac14a1c4faec390eeb51ebb381718c8d9 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 16:02:22 +0200 Subject: [PATCH 224/282] Nat: move lt_moveR_nV somewhere more appropriate Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 18 +++++++++--------- 1 file changed, 9 insertions(+), 9 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index e8a32f780d1..71e42d7d5ff 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1206,15 +1206,6 @@ Hint Immediate nat_sub_monotone_r : typeclass_instances. (** *** Movement Lemmas *) -Definition lt_moveR_nV {k m} n : n <= k -> k < n + m -> k - n < m. -Proof. - intros H1 H2; unfold lt. - rewrite <- nat_sub_succ_l; only 2: exact _. - rewrite <- (nat_add_sub_cancel_l m n). - apply nat_sub_monotone_l. - exact H2. -Defined. - (** TODO: [n <= k] can be dropped *) Definition leq_moveL_Mn {k m} n : n <= k -> k - n <= m -> k <= n + m. Proof. @@ -1246,6 +1237,15 @@ Proof. rapply leq_lt_trans. Defined. +Definition lt_moveR_nV {k m} n : n <= k -> k < n + m -> k - n < m. +Proof. + intros H1 H2; unfold lt. + rewrite <- nat_sub_succ_l; only 2: exact _. + rewrite <- (nat_add_sub_cancel_l m n). + apply nat_sub_monotone_l. + exact H2. +Defined. + (** *** Order-reflection Lemmas *) (** Subtraction reflects [<=] in the left argument. *) From 2f81a1c31580f81ba598a82e66cbd1cc3b95fa68 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 16:05:00 +0200 Subject: [PATCH 225/282] Nat: weaken hyps of leq_moveL_Mn and leq_moveL_nM Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 16 +++++++--------- 1 file changed, 7 insertions(+), 9 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 71e42d7d5ff..ef2eb940850 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1206,16 +1206,15 @@ Hint Immediate nat_sub_monotone_r : typeclass_instances. (** *** Movement Lemmas *) -(** TODO: [n <= k] can be dropped *) -Definition leq_moveL_Mn {k m} n : n <= k -> k - n <= m -> k <= n + m. +Definition leq_moveL_Mn {k m} n : k - n <= m -> k <= n + m. Proof. - intros H1 H2. - apply (nat_add_l_monotone n) in H2. - rewrite <- nat_sub_l_add_r in H2; only 2: exact _. - by rewrite nat_add_sub_cancel_l in H2. + intros H. + rewrite nat_add_comm. + apply (nat_add_r_monotone n) in H. + rapply leq_trans. Defined. -Definition leq_moveL_nM {k m} n : n <= k -> k - n <= m -> k <= m + n. +Definition leq_moveL_nM {k m} n : k - n <= m -> k <= m + n. Proof. rewrite nat_add_comm. apply leq_moveL_Mn. @@ -1242,8 +1241,7 @@ Proof. intros H1 H2; unfold lt. rewrite <- nat_sub_succ_l; only 2: exact _. rewrite <- (nat_add_sub_cancel_l m n). - apply nat_sub_monotone_l. - exact H2. + by apply nat_sub_monotone_l. Defined. (** *** Order-reflection Lemmas *) From 54fd06301751bbf7074542753194b173506382e7 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 16:13:38 +0200 Subject: [PATCH 226/282] Nat: add todos for missing movement lemmas Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 9 +++++++++ 1 file changed, 9 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ef2eb940850..e53647ddb1d 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1228,6 +1228,10 @@ Proof. by rewrite nat_add_sub_cancel_l in H2. Defined. +(** TODO: leq_moveR_nM *) +(** TODO: leq_moveL_nV *) +(** TODO: leq_moveR_nV *) + Definition lt_moveL_Mn {k m} n : k - n < m -> k < n + m. Proof. intros H. @@ -1236,6 +1240,11 @@ Proof. rapply leq_lt_trans. Defined. +(** TODO: lt_moveL_nM *) +(** TODO: lt_moveR_Mn *) +(** TODO: lt_moveR_nM *) +(** TODO: lt_moveL_nV *) + Definition lt_moveR_nV {k m} n : n <= k -> k < n + m -> k - n < m. Proof. intros H1 H2; unfold lt. From f5fc787b1d17d5933c46d80ed76aef4eda804a22 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 17:44:39 +0200 Subject: [PATCH 227/282] Nat: remove natsubpreservesleq (replaced by nat_sub_monotone_l) Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 16 ---------------- 1 file changed, 16 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 4bf1c6f337d..2ae2aacb8ad 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -152,22 +152,6 @@ Proof. assumption. Defined. -(** TODO: move, rename *) -Proposition natsubpreservesleq { n m k : nat } - : n <= m -> n - k <= m - k. -Proof. - simple_induction k k IHk. - - destruct (symmetric_paths _ _ (nat_sub_zero_r n)), - (symmetric_paths _ _ (nat_sub_zero_r m)); done. - - intro l. change (k.+1) with (1 + k). - destruct (nat_add_comm k 1). - destruct (symmetric_paths _ _ (nat_sub_r_add n k 1)). - destruct (symmetric_paths _ _ (nat_sub_r_add m k 1)). - destruct (symmetric_paths _ _ (@predeqminus1 (n -k))). - destruct (symmetric_paths _ _ (@predeqminus1 (m -k))). - apply leq_pred, IHk. exact l. -Defined. - (** TODO: move, rename *) Proposition sub_less { n k : nat } : n - k <= n. Proof. From 5203d49f8e4bf0618fedd49301488abad5adad48 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 18:15:35 +0200 Subject: [PATCH 228/282] Nat: add monotonicity of powers Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 14 ++++++++++++++ 2 files changed, 15 insertions(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 2ae2aacb8ad..a19596e61c9 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -21,7 +21,7 @@ Defined. Proposition nataddsub_comm_ineq (n m k : nat) : (n + k) - m <= (n - m) + k. -Proof. +Proof. simple_induction k k IHk. - destruct (nat_add_zero_r n)^, (nat_add_zero_r (n - m))^; constructor. - destruct (nat_add_succ_r n k)^. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index e53647ddb1d..84ecf658610 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1337,3 +1337,17 @@ Proof. nrapply ap. exact IHm. Defined. + +(** *** Monotonicity of powers *) + +Definition nat_pow_l_monotone {n m} k + : n <= m -> nat_pow k.+1 n <= nat_pow k.+1 m. +Proof. + intros H; induction H; exact _. +Defined. + +Definition nat_pow_r_monotone {n m} k + : n <= m -> nat_pow n k <= nat_pow m k. +Proof. + intros H; induction k; exact _. +Defined. From d0dc5f4846b356b4f2316b94e6b20fc08da85e1a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 18:38:43 +0200 Subject: [PATCH 229/282] Nat: give shorter proof of equiv_lt_lt_sub Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 14 ++++---------- 1 file changed, 4 insertions(+), 10 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 84ecf658610..4ea4390e007 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -970,17 +970,11 @@ Defined. (** A number being less than another is equivalent to their difference being greater than zero. *) Definition equiv_lt_lt_sub n m : m < n <~> 0 < n - m. Proof. + induction n as [|n IHn] in m |- *. + 1: srapply equiv_iff_hprop; intro; contradiction (not_lt_zero_r _ _). + destruct m; only 1: reflexivity. + nrefine (IHn m oE _). srapply equiv_iff_hprop. - - revert m; simple_induction n n IHn. - + intros m ineq. contradiction (not_lt_zero_r m). - + destruct m. - * simpl. easy. - * simpl. intro ineq. apply leq_succ' in ineq. - now apply IHn in ineq. - - intro ineq. - destruct (@leq_dichotomy n m) as [n_leq_m |]; [ | assumption]. - apply equiv_nat_sub_leq in n_leq_m. - contradiction (lt_irrefl 0). now rewrite n_leq_m in ineq. Defined. (** *** Monotonicity of Addition *) From ebdcf385eeb2e1fa73f2873cf9e28b00e82f37f1 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 19:42:13 +0200 Subject: [PATCH 230/282] Nat: add all leq/lt movement lemmas and homegenize variables Signed-off-by: Ali Caglayan --- theories/Algebra/Rings/Matrix.v | 2 +- theories/Spaces/Nat/Arithmetic.v | 4 +- theories/Spaces/Nat/Core.v | 126 ++++++++++++++++++++----------- 3 files changed, 87 insertions(+), 45 deletions(-) diff --git a/theories/Algebra/Rings/Matrix.v b/theories/Algebra/Rings/Matrix.v index e8c2dea4db9..422e24e15ef 100644 --- a/theories/Algebra/Rings/Matrix.v +++ b/theories/Algebra/Rings/Matrix.v @@ -756,7 +756,7 @@ Proof. destruct (dec (k <= i)%nat) as [leq_k_i|gt_k_i]. { rewrite H2. 1: by rewrite rng_mult_zero_r. - rapply leq_lt_trans. } + rapply lt_leq_lt_trans. } apply gt_iff_not_leq in gt_k_i. rewrite H1. 1: by rewrite rng_mult_zero_l. diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index a19596e61c9..ffa4f549c62 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -134,7 +134,7 @@ Proposition pred_gt_implies_lt (i j : nat) Proof. intros ineq. assert (H := leq_succ ineq). assert (i < j) as X. { - apply (@lt_leq_trans _ (nat_pred j) _); + apply (@lt_lt_leq_trans _ (nat_pred j) _); [assumption | apply predn_leq_n]. } destruct (symmetric_paths _ _ (S_predn _ _ X)) in H. @@ -300,7 +300,7 @@ Proof. intro ineq. destruct (@leq_dichotomy n m) as [l | g]. - exact l. - - contradiction (lt_irrefl m (lt_leq_trans g ineq)). + - contradiction (lt_irrefl m (lt_lt_leq_trans g ineq)). Defined. Proposition symmetric_rel_total_order (R : nat -> nat -> Type) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 4ea4390e007..b41a770d672 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -557,18 +557,19 @@ Defined. (** *** Basic Properties of [<] *) (** [<=] and [<] imply [<] *) -Definition leq_lt_trans {n m k} : n <= m -> m < k -> n < k +Definition lt_leq_lt_trans {n m k} : n <= m -> m < k -> n < k := fun leq lt => leq_trans (leq_succ leq) lt. (** [<=] and [<] imply [<] *) -Definition lt_leq_trans {n m k} : n < m -> m <= k -> n < k +Definition lt_lt_leq_trans {n m k} : n < m -> m <= k -> n < k := fun lt leq => leq_trans lt leq. Definition leq_lt {n m} : n < m -> n <= m := leq_succ_l. +Hint Immediate leq_lt : typeclass_instances. Definition lt_trans {n m k} : n < m -> m < k -> n < k - := fun H1 H2 => leq_lt (leq_lt_trans H1 H2). + := fun H1 H2 => leq_lt (lt_leq_lt_trans H1 H2). Global Instance transitive_lt : Transitive lt := @lt_trans. Global Instance ishprop_lt n m : IsHProp (n < m) := _. @@ -959,7 +960,7 @@ Defined. (** *** Subtraction *) -Global Instance leq_sub_add n m : n <= n - m + m. +Global Instance leq_sub_add_l n m : n <= n - m + m. Proof. destruct (@leq_dichotomy m n) as [l | g]. - by rewrite nat_add_sub_l_cancel. @@ -967,6 +968,11 @@ Proof. by destruct (equiv_nat_sub_leq _)^. Defined. +Global Instance leq_sub_add_r n m : n <= m + (n - m). +Proof. + rewrite nat_add_comm; exact _. +Defined. + (** A number being less than another is equivalent to their difference being greater than zero. *) Definition equiv_lt_lt_sub n m : m < n <~> 0 < n - m. Proof. @@ -1198,74 +1204,110 @@ Proof. Defined. Hint Immediate nat_sub_monotone_r : typeclass_instances. -(** *** Movement Lemmas *) +(** *** Order-reflection Lemmas *) -Definition leq_moveL_Mn {k m} n : k - n <= m -> k <= n + m. +(** Subtraction reflects [<=] in the left argument. *) +Definition leq_reflects_sub_l {n m} k : k <= m -> n - k <= m - k -> n <= m. +Proof. + intros ineq1 ineq2. + apply (nat_add_r_monotone k) in ineq2. + apply (@leq_trans _ (n - k + k) _ (leq_sub_add_l _ _)). + apply (@leq_trans _ (m - k + k) _ _). + by rewrite nat_add_sub_l_cancel. +Defined. + +(** Subtraction reflects [<=] in the right argument contravariantly. *) +Definition leq_reflects_sub_r {n m} k + : m <= k -> n <= k -> k - n <= k - m -> m <= n. +Proof. + intros H1 H2 H3. + apply (nat_sub_monotone_r k) in H3. + rewrite 2 nat_sub_sub_cancel_r in H3; exact _. +Defined. + +(** *** Movement lemmas *) + +(** Given an inequality [n < m] we can move around a summand or subtrahend [k] from either side. *) + +Definition leq_moveL_Mn {n m} k : n - k <= m -> n <= k + m. Proof. intros H. rewrite nat_add_comm. - apply (nat_add_r_monotone n) in H. + apply (nat_add_r_monotone k) in H. rapply leq_trans. Defined. -Definition leq_moveL_nM {k m} n : k - n <= m -> k <= m + n. +Definition leq_moveL_nM {n m} k : n - k <= m -> n <= m + k. Proof. rewrite nat_add_comm. apply leq_moveL_Mn. Defined. -Definition leq_moveR_Mn {m n} k : k <= n -> m <= n - k -> k + m <= n. +Definition leq_moveR_Mn {n m} k : k <= m -> n <= m - k -> k + n <= m. Proof. intros H1 H2. - apply (nat_add_l_monotone k) in H2. - rewrite <- nat_sub_l_add_r in H2; trivial. - by rewrite nat_add_sub_cancel_l in H2. + rapply (leq_reflects_sub_l k). + by rewrite nat_add_sub_cancel_l. Defined. -(** TODO: leq_moveR_nM *) -(** TODO: leq_moveL_nV *) -(** TODO: leq_moveR_nV *) +Definition leq_moveR_nM {n m} k : k <= m -> n <= m - k -> n + k <= m. +Proof. + rewrite nat_add_comm. + apply leq_moveR_Mn. +Defined. -Definition lt_moveL_Mn {k m} n : k - n < m -> k < n + m. +Definition leq_moveL_nV {n m} k : n + k <= m -> n <= m - k. Proof. intros H. - rewrite nat_add_comm. - apply (nat_add_r_strictly_monotone n) in H. - rapply leq_lt_trans. + apply (leq_reflects_add_r k). + rapply leq_trans. Defined. -(** TODO: lt_moveL_nM *) -(** TODO: lt_moveR_Mn *) -(** TODO: lt_moveR_nM *) -(** TODO: lt_moveL_nV *) +Definition leq_moveR_nV {n m} k : n <= m + k -> n - k <= m. +Proof. + intros H. + apply (nat_sub_monotone_l k) in H. + by rewrite nat_add_sub_cancel_r in H. +Defined. -Definition lt_moveR_nV {k m} n : n <= k -> k < n + m -> k - n < m. +Definition lt_moveL_Mn {n m} k : n - k < m -> n < k + m. Proof. - intros H1 H2; unfold lt. - rewrite <- nat_sub_succ_l; only 2: exact _. - rewrite <- (nat_add_sub_cancel_l m n). - by apply nat_sub_monotone_l. + intros H. + apply (nat_add_l_strictly_monotone k) in H. + rapply lt_leq_lt_trans. Defined. -(** *** Order-reflection Lemmas *) +Definition lt_moveL_nM {n m} k : n - k < m -> n < m + k. +Proof. + rewrite nat_add_comm. + apply lt_moveL_Mn. +Defined. -(** Subtraction reflects [<=] in the left argument. *) -Definition leq_reflects_sub_l {n m} k : k <= m -> n - k <= m - k -> n <= m. +Definition lt_moveR_Mn {n m} k : k < m -> n < m - k -> k + n < m. Proof. - intros ineq1 ineq2. - apply (nat_add_r_monotone k) in ineq2. - apply (@leq_trans _ (n - k + k) _ (leq_sub_add _ _)). - apply (@leq_trans _ (m - k + k) _ _). - by rewrite nat_add_sub_l_cancel. + intros H1 H2. + rewrite nat_add_comm. + rapply (leq_moveR_nM (n:=n.+1) k). Defined. -(** Subtraction reflects [<=] in the right argument contravariantly. *) -Definition leq_reflects_sub_r {n m} k - : m <= k -> n <= k -> k - n <= k - m -> m <= n. +Definition lt_moveR_nM {n m} k : k < m -> n < m - k -> n + k < m. Proof. - intros H1 H2 H3. - apply (nat_sub_monotone_r k) in H3. - rewrite 2 nat_sub_sub_cancel_r in H3; exact _. + rewrite nat_add_comm. + apply lt_moveR_Mn. +Defined. + +Definition lt_moveL_nV {n m} k : n + k < m -> n < m - k. +Proof. + intros H. + rapply leq_moveL_nV. +Defined. + +Definition lt_moveR_nV {n m} k : k <= n -> n < k + m -> n - k < m. +Proof. + intros H1 H2; unfold lt. + rewrite <- nat_sub_succ_l; only 2: exact _. + rewrite <- (nat_add_sub_cancel_l m k). + by apply nat_sub_monotone_l. Defined. (** ** Properties of Powers *) From a18a3cb0ebde3f523c9672986355a671481c9408 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 19:53:39 +0200 Subject: [PATCH 231/282] Nat: adjust style of factorial Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 10 +++++----- 1 file changed, 5 insertions(+), 5 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index b41a770d672..b535998bd9c 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -120,11 +120,11 @@ Fixpoint gcd a b := (** *** Factorial *) -Fixpoint factorial (n : nat) : nat - := match n with - | 0 => 1 - | S n => S n * factorial n - end. +Fixpoint factorial n := + match n with + | 0 => 1 + | S n => S n * factorial n + end. (** ** Comparison Predicates *) From b9e8d70ea246745a2b325607a2c75612af4a80a2 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 1 Aug 2024 19:54:21 +0200 Subject: [PATCH 232/282] Nat: adjust style of div and modulo Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 8 ++++---- 1 file changed, 4 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index b535998bd9c..9c9b23d4d72 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -95,14 +95,14 @@ Fixpoint divmod x y q u : nat * nat := Definition div x y : nat := match y with - | 0 => y - | S y' => fst (divmod x y' 0 y') + | 0 => y + | S y' => fst (divmod x y' 0 y') end. Definition modulo x y : nat := match y with - | 0 => y - | S y' => y' - snd (divmod x y' 0 y') + | 0 => y + | S y' => y' - snd (divmod x y' 0 y') end. Infix "/" := div : nat_scope. From 7a1e159508fb15f0c1088be5d7770a7587db78cc Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Fri, 2 Aug 2024 14:13:07 +0200 Subject: [PATCH 233/282] Nat: move strong_induction Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 17 ----------------- theories/Spaces/Nat/Core.v | 19 +++++++++++++++++++ 2 files changed, 19 insertions(+), 17 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index ffa4f549c62..68ec8e2798c 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -182,23 +182,6 @@ Proof. intro l; apply (leq_trans l); exact (leq_add_l m k). Defined. -(** TODO: move, rename *) -Proposition strong_induction (P : nat -> Type) - : (forall n : nat, (forall m : nat, (m < n) -> P m) -> P n) -> - forall n : nat, P n. -Proof. - intro a. - assert (forall n m: nat, m < n -> P m) as X. { - simple_induction n n IHn. - - intros m l. contradiction (not_lt_zero_r m). - - intros m l. apply leq_succ' in l. - destruct l as [ | n]. - + apply a; intros ? ?; now apply IHn. - + now apply (IHn m), leq_succ. - } - intro n. apply (X (n.+1) n), (leq_refl n.+1). -Defined. - (** This inductive type is defined because it lets you loop from [i = 0] up to [i = n] by structural induction on a proof of [increasing_geq n 0]. With the existing [leq] type and the inductive structure of [n], it is easier and more natural to loop downwards from [i = n] to [i = 0], but harder to find the least natural number in the interval $[0,n]$ satisfying a given property. *) Local Unset Elimination Schemes. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 9c9b23d4d72..4e476333a2b 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1387,3 +1387,22 @@ Definition nat_pow_r_monotone {n m} k Proof. intros H; induction k; exact _. Defined. + +(** *** Strong induction *) + +(** TODO: rename nat_ind_strong *) +Proposition strong_induction (P : nat -> Type) + : (forall n : nat, (forall m : nat, (m < n) -> P m) -> P n) -> + forall n : nat, P n. +Proof. + intro a. + assert (forall n m: nat, m < n -> P m) as X. { + simple_induction n n IHn. + - intros m l. contradiction (not_lt_zero_r m). + - intros m l. apply leq_succ' in l. + destruct l as [ | n]. + + apply a; intros ? ?; now apply IHn. + + now apply (IHn m), leq_succ. + } + intro n. apply (X (n.+1) n), (leq_refl n.+1). +Defined. From 9f7df7dd81bde384be37c08f4e4c184060aa09ae Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Fri, 2 Aug 2024 14:14:07 +0200 Subject: [PATCH 234/282] Nat: rename strong_induction -> nat_ind_strong Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 4e476333a2b..61c99c61de5 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1391,7 +1391,7 @@ Defined. (** *** Strong induction *) (** TODO: rename nat_ind_strong *) -Proposition strong_induction (P : nat -> Type) +Proposition nat_ind_strong (P : nat -> Type) : (forall n : nat, (forall m : nat, (m < n) -> P m) -> P n) -> forall n : nat, P n. Proof. From ff72e992ee2fd3427849f16c08cdc881c1514751 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Fri, 2 Aug 2024 14:26:34 +0200 Subject: [PATCH 235/282] Nat: reprove strong induction Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 32 +++++++++++++++++--------------- 1 file changed, 17 insertions(+), 15 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 61c99c61de5..ee462ed9046 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1390,19 +1390,21 @@ Defined. (** *** Strong induction *) -(** TODO: rename nat_ind_strong *) -Proposition nat_ind_strong (P : nat -> Type) - : (forall n : nat, (forall m : nat, (m < n) -> P m) -> P n) -> - forall n : nat, P n. -Proof. - intro a. - assert (forall n m: nat, m < n -> P m) as X. { - simple_induction n n IHn. - - intros m l. contradiction (not_lt_zero_r m). - - intros m l. apply leq_succ' in l. - destruct l as [ | n]. - + apply a; intros ? ?; now apply IHn. - + now apply (IHn m), leq_succ. - } - intro n. apply (X (n.+1) n), (leq_refl n.+1). +Definition nat_ind_strong (P : nat -> Type) + (IH_strong : forall n, (forall m, m < n -> P m) -> P n) + : forall n, P n. +Proof. + intros n. + apply IH_strong. + intros m H. + induction n as [|n IHn] in m, H |- *. + 1: contradiction (not_lt_zero_r m). + apply leq_succ' in H. + apply equiv_leq_lt_or_eq in H. + destruct H as [H|p]. + - apply IHn. + exact H. + - destruct p. + apply IH_strong. + exact IHn. Defined. From 2de76a71c80bbdce39200e40190095dc121a98e4 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Fri, 2 Aug 2024 15:22:08 +0200 Subject: [PATCH 236/282] Nat: improve universe for nat_ind_strong and add comment Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 15 +++++++-------- 1 file changed, 7 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ee462ed9046..9ff01c763b5 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1388,23 +1388,22 @@ Proof. intros H; induction k; exact _. Defined. -(** *** Strong induction *) +(** ** Strong induction *) -Definition nat_ind_strong (P : nat -> Type) +(** Sometimes using [nat_ind] is not sufficient to prove a statement as it may be difficult to prove [P n -> P n.+1]. We can strengthen the induction hypothesis by assuming that [P m] holds for all [m] less than [n]. This is known as strong induction. *) +Definition nat_ind_strong@{u} (P : nat -> Type@{u}) (IH_strong : forall n, (forall m, m < n -> P m) -> P n) : forall n, P n. Proof. intros n. apply IH_strong. - intros m H. - induction n as [|n IHn] in m, H |- *. + simple_induction n n IHn; intros m H. 1: contradiction (not_lt_zero_r m). apply leq_succ' in H. apply equiv_leq_lt_or_eq in H. destruct H as [H|p]. - - apply IHn. - exact H. + - by apply IHn. - destruct p. - apply IH_strong. - exact IHn. + by apply IH_strong. Defined. + From 3c162fb825bde2474ab9fafba34d2fa387ed7bf1 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Fri, 2 Aug 2024 20:08:16 +0200 Subject: [PATCH 237/282] Nat: add more todo comments Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 14 ++++++++++++++ 1 file changed, 14 insertions(+) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 68ec8e2798c..3494acba199 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -8,6 +8,7 @@ Local Open Scope nat_scope. (** TODO: The results in this file are in the process of being moved over to Core.v *) +(** TODO: move, rename *) Proposition nataddsub_comm_ineq_lemma (n m : nat) : n.+1 - m <= (n - m).+1. Proof. @@ -19,6 +20,7 @@ Proof. + apply IHn. Defined. +(** TODO: move, rename *) Proposition nataddsub_comm_ineq (n m k : nat) : (n + k) - m <= (n - m) + k. Proof. @@ -30,6 +32,7 @@ Proof. now apply leq_succ. Defined. +(** TODO: move, rename *) Proposition nat_sub_add_ineq (n m : nat) : n <= n - m + m. Proof. destruct (@leq_dichotomy m n) as [l | gt]. @@ -41,6 +44,7 @@ Proof. assumption. Defined. +(** TODO: move, rename *) Proposition i_lt_n_sum_m (n m i : nat) : i < n - m -> m <= n. Proof. @@ -51,6 +55,7 @@ Proof. + apply leq_succ. simpl in l. apply (IHn m i l). Defined. +(** TODO: remove *) Proposition nataddsub_assoc_implication (n : nat) {m k z : nat} : (k <= m) -> n + (m - k) = z -> n + m - k = z. Proof. @@ -58,6 +63,7 @@ Proof. by rewrite nat_sub_l_add_r. Defined. +(** TODO: remove *) Proposition nat_add_sub_eq (n : nat) {k: nat} : (k <= n) -> k + (n - k) = n. Proof. @@ -66,6 +72,7 @@ Proof. destruct (nat_add_comm n k); exact (nat_add_sub_cancel_r _ _). Defined. +(** TODO: move, rename *) Proposition predeqminus1 { n : nat } : n - 1 = nat_pred n. Proof. simple_induction' n. @@ -73,11 +80,13 @@ Proof. - apply nat_sub_zero_r. Defined. +(** TODO: move, rename *) Proposition predn_leq_n (n : nat) : nat_pred n <= n. Proof. destruct n; exact _. Defined. +(** TODO: remove *) Proposition S_predn (n i: nat) : (i < n) -> S(nat_pred n) = n. Proof. simple_induction' n; intros l. @@ -85,6 +94,7 @@ Proof. - reflexivity. Defined. +(** TODO: move, rename *) Proposition pred_equiv (k n : nat) : k < n -> k < S (nat_pred n). Proof. intro ineq; destruct n. @@ -92,11 +102,13 @@ Proof. - assumption. Defined. +(** TODO: move, rename *) Proposition n_leq_pred_Sn (n : nat) : n <= S (nat_pred n). Proof. destruct n; exact _. Defined. +(** TODO: move, rename *) Proposition leq_implies_pred_lt (i n k : nat) : (n > i) -> n <= k -> nat_pred n < k. Proof. @@ -105,6 +117,7 @@ Proof. - intro; assumption. Defined. +(** TODO: move, rename *) Proposition pred_lt_implies_leq (n k : nat) : nat_pred n < k -> n <= k. Proof. @@ -119,6 +132,7 @@ Proof. intro l; apply leq_pred in l; assumption. Defined. +(** TODO: move, rename *) Proposition j_geq_0_lt_implies_pred_geq (i j k : nat) : i < j -> k.+1 <= j -> k <= nat_pred j. Proof. From 864e778bb6696b73214e215dcd2ce4404be16623 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Fri, 2 Aug 2024 20:09:12 +0200 Subject: [PATCH 238/282] Nat: remove nataddsub_assoc_implication (replaced by nat_sub_l_add_r) Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 8 -------- 1 file changed, 8 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 3494acba199..5b3d1840966 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -55,14 +55,6 @@ Proof. + apply leq_succ. simpl in l. apply (IHn m i l). Defined. -(** TODO: remove *) -Proposition nataddsub_assoc_implication (n : nat) {m k z : nat} - : (k <= m) -> n + (m - k) = z -> n + m - k = z. -Proof. - intro H. - by rewrite nat_sub_l_add_r. -Defined. - (** TODO: remove *) Proposition nat_add_sub_eq (n : nat) {k: nat} : (k <= n) -> k + (n - k) = n. From 9a01df8b1847dbb4d4d2494a32ae651028f1a039 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Fri, 2 Aug 2024 20:11:03 +0200 Subject: [PATCH 239/282] Nat: remove nat_add_sub_eq (replaced by nat_add_sub_r_cancel) Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 9 --------- 1 file changed, 9 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 5b3d1840966..08036691fd2 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -54,15 +54,6 @@ Proof. + apply leq_zero_l. + apply leq_succ. simpl in l. apply (IHn m i l). Defined. - -(** TODO: remove *) -Proposition nat_add_sub_eq (n : nat) {k: nat} - : (k <= n) -> k + (n - k) = n. -Proof. - intro H. - destruct (nat_sub_l_add_r _ k _ H); - destruct (nat_add_comm n k); exact (nat_add_sub_cancel_r _ _). -Defined. (** TODO: move, rename *) Proposition predeqminus1 { n : nat } : n - 1 = nat_pred n. From 5a6c40cd7a9be50cf3254d50fa4173d90dd3db1b Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Fri, 2 Aug 2024 20:23:44 +0200 Subject: [PATCH 240/282] Nat: merge S_predn into nat_succ_pred and generalise slightly Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 23 +++++++---------------- theories/Spaces/Nat/Core.v | 8 +++++++- 2 files changed, 14 insertions(+), 17 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 08036691fd2..4974cd9c10e 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -69,14 +69,6 @@ Proof. destruct n; exact _. Defined. -(** TODO: remove *) -Proposition S_predn (n i: nat) : (i < n) -> S(nat_pred n) = n. -Proof. - simple_induction' n; intros l. - - contradiction (not_lt_zero_r i). - - reflexivity. -Defined. - (** TODO: move, rename *) Proposition pred_equiv (k n : nat) : k < n -> k < S (nat_pred n). Proof. @@ -134,8 +126,7 @@ Proof. apply (@lt_lt_leq_trans _ (nat_pred j) _); [assumption | apply predn_leq_n]. } - destruct (symmetric_paths _ _ (S_predn _ _ X)) in H. - assumption. + by rewrite <- (nat_succ_pred' j i). Defined. (** TODO: move, rename *) @@ -143,9 +134,9 @@ Proposition pred_preserves_lt {i n: nat} (p : i < n) m : (n < m) -> (nat_pred n < nat_pred m). Proof. intro l. - apply leq_succ'. destruct (symmetric_paths _ _ (S_predn n i _)). + apply leq_succ'. destruct (symmetric_paths _ _ (nat_succ_pred' n i _)). set (k := transitive_lt i n m p l). - destruct (symmetric_paths _ _ (S_predn m i _)). + destruct (symmetric_paths _ _ (nat_succ_pred' m i _)). assumption. Defined. @@ -227,7 +218,7 @@ Proof. apply increasing_geq_S. unfold ">", "<" in *. apply equiv_lt_lt_sub in g. - now (destruct (symmetric_paths _ _ (S_predn (n - k) 0 _))). + now (destruct (symmetric_paths _ _ (nat_succ_pred (n - k) _))). Defined. Lemma ineq_sub' (n k : nat) : k < n -> n - k = (n - k.+1).+1. @@ -248,15 +239,15 @@ Proof. - intros m ineq. change (m - n.+1) with (m - (1 + n)). (destruct (nat_add_comm n 1)). destruct (symmetric_paths _ _ (nat_sub_r_add m n 1)). - destruct (S_predn (m - n) 0 (equiv_lt_lt_sub _ _ ineq)); simpl; + destruct (nat_succ_pred (m - n) (equiv_lt_lt_sub _ _ ineq)); simpl; destruct (symmetric_paths _ _ (nat_sub_zero_r (nat_pred (m - n)))). assert (0 < m - n) as dp by exact (equiv_lt_lt_sub _ _ ineq). assert (nat_pred (m - n) < m) as sh by ( unfold "<"; - destruct (symmetric_paths _ _ (S_predn _ 0 _)); + destruct (symmetric_paths _ _ (nat_succ_pred _ _)); exact sub_less). destruct (symmetric_paths _ _ (ineq_sub' _ _ _)). - destruct (symmetric_paths _ _ (S_predn _ 0 _)). + destruct (symmetric_paths _ _ (nat_succ_pred _ _)). apply (ap S), IHn, leq_succ_l, ineq. Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 9ff01c763b5..7ab1aa55731 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -206,14 +206,20 @@ Defined. (** ** Properties of successors *) +(** The predecessor of a successor is the original number. *) Definition nat_pred_succ@{} n : nat_pred (nat_succ n) = n := idpath. -Definition nat_succ_pred@{} n : 0 < n -> nat_succ (nat_pred n) = n. +(** The successor of a predecessor is the original as long as there is a strict lower bound. *) +Definition nat_succ_pred'@{} n i : i < n -> nat_succ (nat_pred n) = n. Proof. by intros []. Defined. +(** The most common lower bound is to take [0]. *) +Definition nat_succ_pred@{} n : 0 < n -> nat_succ (nat_pred n) = n + := nat_succ_pred' n 0. + (** Injectivity of successor. *) Definition path_nat_succ@{} n m (H : S n = S m) : n = m := ap nat_pred H. Global Instance isinj_succ : IsInjective nat_succ := path_nat_succ. From 1c4d208f82e7427984412c8b32dbfb693a982c2e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 4 Aug 2024 14:40:28 +0100 Subject: [PATCH 241/282] Nat: fix capitalisation in headers Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 48 +++++++++++++++++++------------------- 1 file changed, 24 insertions(+), 24 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 7ab1aa55731..0aa2061308c 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -126,9 +126,9 @@ Fixpoint factorial n := | S n => S n * factorial n end. -(** ** Comparison Predicates *) +(** ** Comparison predicates *) -(** *** Less Than or Equal To [<=] *) +(** *** Less than or equal To [<=] *) Inductive leq (n : nat) : nat -> Type0 := | leq_refl : leq n n @@ -145,7 +145,7 @@ Notation "n <= m" := (leq n m) : nat_scope. Existing Class leq. Global Existing Instances leq_refl leq_succ_r. -(** *** Less Than [<] *) +(** *** Less than [<] *) (** We define the less-than relation [lt] in terms of [leq] *) Definition lt n m : Type0 := leq (S n) m. @@ -156,7 +156,7 @@ Existing Class lt. Infix "<" := lt : nat_scope. Global Instance lt_is_leq n m : leq n.+1 m -> lt n m | 100 := idmap. -(** *** Greater Than or Equal To [>=] *) +(** *** Greater than or equal To [>=] *) Definition geq n m := leq m n. Existing Class geq. @@ -172,7 +172,7 @@ Existing Class gt. Infix ">" := gt : nat_scope. Global Instance gt_is_leq n m : leq m.+1 n -> gt n m | 100 := idmap. -(** *** Combined Comparison Predicates *) +(** *** Combined comparison predicates *) Notation "x <= y <= z" := (x <= y /\ y <= z) : nat_scope. Notation "x <= y < z" := (x <= y /\ y < z) : nat_scope. @@ -404,9 +404,9 @@ Definition nat_mul_one_l@{} n : 1 * n = n Definition nat_mul_one_r@{} n : n * 1 = n := nat_mul_comm _ _ @ nat_mul_one_l _. -(** ** Basic Properties of Comparison Predicates *) +(** ** Basic properties of comparison predicates *) -(** *** Basic Properties of [<=] *) +(** *** Basic properties of [<=] *) (** [<=] is reflexive by definition. *) Global Instance reflexive_leq : Reflexive leq := leq_refl. @@ -560,7 +560,7 @@ Proof. intro l; apply leq_succ'; exact _. Defined. -(** *** Basic Properties of [<] *) +(** *** Basic properties of [<] *) (** [<=] and [<] imply [<] *) Definition lt_leq_lt_trans {n m k} : n <= m -> m < k -> n < k @@ -581,21 +581,21 @@ Global Instance transitive_lt : Transitive lt := @lt_trans. Global Instance ishprop_lt n m : IsHProp (n < m) := _. Global Instance decidable_lt n m : Decidable (lt n m) := _. -(** *** Basic Properties of [>=] *) +(** *** Basic properties of [>=] *) Global Instance reflexive_geq : Reflexive geq := leq_refl. Global Instance transitive_geq : Transitive geq := fun x y z p q => leq_trans q p. Global Instance ishprop_geq n m : IsHProp (geq n m) := _. Global Instance decidable_geq n m : Decidable (geq n m) := _. -(** *** Basic Properties of [>] *) +(** *** Basic properties of [>] *) Global Instance transitive_gt : Transitive gt := fun x y z p q => transitive_lt _ _ _ q p. Global Instance ishprop_gt n m : IsHProp (gt n m) := _. Global Instance decidable_gt n m : Decidable (gt n m) := _. -(** ** Properties of Subtraction *) +(** ** Properties of subtraction *) (** Subtracting a number from [0] is [0]. *) Definition nat_sub_zero_l@{} n : 0 - n = 0 := idpath. @@ -713,9 +713,9 @@ Proof. apply IHn. Defined. -(** ** Properties of Maximum and Minimum *) +(** ** Properties of maximum and minimum *) -(** *** Properties of Maxima *) +(** *** Properties of maximum *) (** [nat_max] is idempotent. *) Definition nat_max_idem@{} n : nat_max n n = n. @@ -828,9 +828,9 @@ Proof. by apply (ap S), IHn. Defined. -(** ** More Theory of Comparison Predicates *) +(** ** More theory of comparison predicates *) -(** *** Addition Lemmas *) +(** *** Addition lemmas *) (** The first summand is less than or equal to the sum. *) Global Instance leq_add_l n m : n <= n + m. @@ -911,7 +911,7 @@ Proof. snrapply equiv_leq_lt_or_eq. Defined. -(** *** Negation Lemmas *) +(** *** Negation lemmas *) (** There are various lemmas we can state about negating the comparison operators on [nat]. To aid readability, we opt to keep the order of the variables in each statement consistent. *) @@ -989,7 +989,7 @@ Proof. srapply equiv_iff_hprop. Defined. -(** *** Monotonicity of Addition *) +(** *** Monotonicity of addition *) (** TODO: use OrderPreserving from canonical_names *) @@ -1017,7 +1017,7 @@ Proof. Defined. Hint Immediate nat_add_monotone : typeclass_instances. -(** *** Strict Monotonicity of Addition *) +(** *** Strict monotonicity of addition *) (** [nat_succ] is strictly monotone. *) Global Instance lt_succ {n m} : n < m -> n.+1 < m.+1 := _. @@ -1048,7 +1048,7 @@ Proof. Defined. Hint Immediate nat_add_strictly_monotone : typeclass_instances. -(** *** Monotonicity of Multiplication *) +(** *** Monotonicity of multiplication *) (** Multiplication on the left is monotone. *) Definition nat_mul_l_monotone {n m} k @@ -1074,7 +1074,7 @@ Proof. Defined. Hint Immediate nat_mul_monotone : typeclass_instances. -(** *** Strict Monotonicity of Multiplication *) +(** *** Strict monotonicity of multiplication *) (** Multiplication on the left by a positive number is strictly monotone. *) Definition nat_mul_l_strictly_monotone {n m} k @@ -1118,7 +1118,7 @@ Proof. rewrite 2 (nat_add_comm _ k); nrapply lt_reflects_add_l. Defined. -(** ** Further Properties of Subtraction *) +(** ** Further properties of subtraction *) (** Subtracting from a successor is the successor of subtracting from the original number, as long as the amount being subtracted is less than or equal to the original number. *) Definition nat_sub_succ_l n m : m <= n -> n.+1 - m = (n - m).+1. @@ -1177,7 +1177,7 @@ Proof. exact IHleq. Defined. -(** *** Monotonicity of Subtraction *) +(** *** Monotonicity of subtraction *) (** Subtraction is monotone in the left argument. *) Definition nat_sub_monotone_l {n m} k : n <= m -> n - k <= m - k. @@ -1210,7 +1210,7 @@ Proof. Defined. Hint Immediate nat_sub_monotone_r : typeclass_instances. -(** *** Order-reflection Lemmas *) +(** *** Order-reflection lemmas *) (** Subtraction reflects [<=] in the left argument. *) Definition leq_reflects_sub_l {n m} k : k <= m -> n - k <= m - k -> n <= m. @@ -1316,7 +1316,7 @@ Proof. by apply nat_sub_monotone_l. Defined. -(** ** Properties of Powers *) +(** ** Properties of powers *) (** [0] to any power is [0] unless that power is [0] in which case it is [1]. *) Definition nat_pow_zero_l@{} n : nat_pow 0 n = if dec (n = 0) then 1 else 0. From 2a1d6870ea26847192ac0b649480ed1401b784ab Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 4 Aug 2024 14:49:47 +0100 Subject: [PATCH 242/282] Nat: fix comment Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 0aa2061308c..ed272dc97c5 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -903,7 +903,7 @@ Defined. (** *** Trichotomy *) -(** Every two natural numbers are either equal, less than, or greater than each other. *) +(** Any two natural numbers are either equal, less than, or greater than each other. *) Definition nat_trichotomy m n : (m < n) + (m = n) + (m > n). Proof. generalize (leq_dichotomy m n). From 88efc92feafaa22b62b2ed0924e3549385dcc8aa Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 4 Aug 2024 21:03:42 +0100 Subject: [PATCH 243/282] Nat: rename leq_succ' -> leq_pred' Signed-off-by: Ali Caglayan --- theories/BoundedSearch.v | 2 +- theories/Spaces/Finite/FinNat.v | 14 +++++++------- theories/Spaces/List/Theory.v | 12 ++++++------ theories/Spaces/Nat/Arithmetic.v | 6 +++--- theories/Spaces/Nat/Core.v | 18 +++++++++--------- 5 files changed, 26 insertions(+), 26 deletions(-) diff --git a/theories/BoundedSearch.v b/theories/BoundedSearch.v index 2af95c4891a..01e30fb5335 100644 --- a/theories/BoundedSearch.v +++ b/theories/BoundedSearch.v @@ -61,7 +61,7 @@ Section bounded_search. destruct X as [leqSnm|ltmSn]. ++ assumption. ++ unfold gt, lt in ltmSn. - assert (m <= n) as X by rapply leq_succ'. + assert (m <= n) as X by rapply leq_pred'. destruct (n0 m X pm). * right. intros l q. assert ((l <= n) + (l > n)) as X by apply leq_dichotomy. diff --git a/theories/Spaces/Finite/FinNat.v b/theories/Spaces/Finite/FinNat.v index af12985ac84..9d7132b5312 100644 --- a/theories/Spaces/Finite/FinNat.v +++ b/theories/Spaces/Finite/FinNat.v @@ -55,7 +55,7 @@ Proof. - destruct u as [x h]. destruct x as [| x]. + exact (transport (P n.+1) (path_zero_finnat _ h) (z _)). - + refine (transport (P n.+1) (path_succ_finnat (x; leq_succ' h) _) _). + + refine (transport (P n.+1) (path_succ_finnat (x; leq_pred' h) _) _). apply s. apply IHn. Defined. @@ -83,7 +83,7 @@ Proof. (fun p => transport _ p (s n u _) = s n u (finnat_ind P z s u)) (hset_path2 1 (path_succ_finnat u (leq_succ u.2))) 1). destruct u as [u1 u2]. - assert (u2 = leq_succ' (leq_succ u2)) as p by apply path_ishprop. + assert (u2 = leq_pred' (leq_succ u2)) as p by apply path_ishprop. simpl; by induction p. Defined. @@ -108,7 +108,7 @@ Monomorphic Fixpoint finnat_to_fin {n : nat} : FinNat n -> Fin n | n.+1 => fun u => match u with | (0; _) => fin_zero - | (x.+1; h) => fsucc (finnat_to_fin (x; leq_succ' h)) + | (x.+1; h) => fsucc (finnat_to_fin (x; leq_pred' h)) end end. @@ -157,9 +157,9 @@ Proof. - elim (not_lt_zero_r _ u.2). - destruct u as [x h]. destruct x as [| x]; [reflexivity|]. - refine ((ap _ (ap _ (path_succ_finnat (x; leq_succ' h) h)))^ @ _). - refine (_ @ ap fsucc (IHn (x; leq_succ' h))). - by induction (path_finnat_to_fin_succ (incl_finnat (x; leq_succ' h))). + refine ((ap _ (ap _ (path_succ_finnat (x; leq_pred' h) h)))^ @ _). + refine (_ @ ap fsucc (IHn (x; leq_pred' h))). + by induction (path_finnat_to_fin_succ (incl_finnat (x; leq_pred' h))). Defined. Lemma path_finnat_to_fin_last (n : nat) @@ -180,7 +180,7 @@ Proof. destruct x as [| x]. + exact (ap pr1 (path_fin_to_finnat_fin_zero n)). + refine ((path_fin_to_finnat_fsucc _)..1 @ _). - exact (ap S (IHn (x; leq_succ' h))..1). + exact (ap S (IHn (x; leq_pred' h))..1). Defined. Lemma path_fin_to_finnat_to_fin {n : nat} (k : Fin n) diff --git a/theories/Spaces/List/Theory.v b/theories/Spaces/List/Theory.v index 7bfa375e329..cf000251841 100644 --- a/theories/Spaces/List/Theory.v +++ b/theories/Spaces/List/Theory.v @@ -336,7 +336,7 @@ Proof. destruct n. 1: by exists a. apply IHa. - apply leq_succ'. + apply leq_pred'. exact H. Defined. @@ -497,7 +497,7 @@ Proof. 1: destruct (not_lt_zero_r _ H). destruct n. 1: reflexivity. - by apply IHl, leq_succ'. + by apply IHl, leq_pred'. Defined. (** The [nth i] element where [pred (length l) = i] is the last element of the list. *) @@ -569,7 +569,7 @@ Proof. destruct n. 1: destruct (not_lt_zero_r _ H). cbn; apply IHl. - apply leq_succ'. + apply leq_pred'. exact H. Defined. @@ -627,7 +627,7 @@ Proof. destruct n. 1: destruct (not_lt_zero_r _ H). cbn; f_ap. - by apply IHl, leq_succ'. + by apply IHl, leq_pred'. Defined. (** The length of a [take n] is the minimum of [n] and the length of the original list. *) @@ -818,7 +818,7 @@ Proof. cbn; by rewrite nat_sub_zero_r. - induction n as [|n IHn]. 1: destruct (not_lt_zero_r _ H). - by apply IHi, leq_succ'. + by apply IHi, leq_pred'. Defined. (** The [nth] element of a [seq] is [i]. *) @@ -833,7 +833,7 @@ Proof. 1: by rewrite length_seq. by apply IHn. - apply geq_iff_not_lt in H'. - apply leq_succ' in H. + apply leq_pred' in H. destruct (leq_antisym H H'). lhs nrapply nth_last. { rewrite length_app. diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index 4974cd9c10e..a49a34447a3 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -114,7 +114,7 @@ Proof. intros l ineq. destruct j. - contradiction (not_lt_zero_r i). - - now simpl; apply leq_succ'. + - now simpl; apply leq_pred'. Defined. (** TODO: move, rename *) @@ -134,7 +134,7 @@ Proposition pred_preserves_lt {i n: nat} (p : i < n) m : (n < m) -> (nat_pred n < nat_pred m). Proof. intro l. - apply leq_succ'. destruct (symmetric_paths _ _ (nat_succ_pred' n i _)). + apply leq_pred'. destruct (symmetric_paths _ _ (nat_succ_pred' n i _)). set (k := transitive_lt i n m p l). destruct (symmetric_paths _ _ (nat_succ_pred' m i _)). assumption. @@ -226,7 +226,7 @@ Proof. intro ineq. destruct n. - contradiction (not_lt_zero_r k). - - change (n.+1 - k.+1) with (n - k). apply leq_succ' in ineq. + - change (n.+1 - k.+1) with (n - k). apply leq_pred' in ineq. by apply nat_sub_succ_l. Defined. diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index ed272dc97c5..e61676c8a81 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -444,15 +444,15 @@ Defined. Global Existing Instance leq_succ | 100. (** The converse to [leq_succ] also holds. *) -Definition leq_succ' {n m} : n.+1 <= m.+1 -> n <= m := leq_pred. -Hint Immediate leq_succ' : typeclass_instances. +Definition leq_pred' {n m} : n.+1 <= m.+1 -> n <= m := leq_pred. +Hint Immediate leq_pred' : typeclass_instances. (** [<] is an irreflexive relation. *) Definition lt_irrefl n : ~ (n < n). Proof. induction n as [|n IHn]. 1: intro p; inversion p. - intros p; by apply IHn, leq_succ'. + intros p; by apply IHn, leq_pred'. Defined. Global Instance irreflexive_lt : Irreflexive lt := lt_irrefl. @@ -465,7 +465,7 @@ Proof. destruct p. 1: reflexivity. destruct x; [inversion q|]. - apply leq_succ' in q. + apply leq_pred' in q. contradiction (lt_irrefl _ (leq_trans p q)). Defined. @@ -557,7 +557,7 @@ Defined. (** [n.+1 <= m] implies [n <= m]. *) Definition leq_succ_l {n m} : n.+1 <= m -> n <= m. Proof. - intro l; apply leq_succ'; exact _. + intro l; apply leq_pred'; exact _. Defined. (** *** Basic properties of [<] *) @@ -679,7 +679,7 @@ Proof. lhs nrapply nat_add_succ_r. nrapply (ap nat_succ). nrapply IHn. - exact (leq_succ' H). + exact (leq_pred' H). Defined. (** We can cancel a right subtrahend when adding it on the left to a subtraction if the subtrahend is less than the nubmer being subtracted from. *) @@ -760,7 +760,7 @@ Proof. 1: nrapply nat_max_zero_r. destruct n. 1: inversion H. - cbn; by apply (ap S), IHm, leq_succ'. + cbn; by apply (ap S), IHm, leq_pred'. Defined. (** [nat_max n m] is [m] if [n <= m]. *) @@ -810,7 +810,7 @@ Proof. simple_induction n n IHn; auto. intros [] p. 1: inversion p. - cbn; by apply (ap S), IHn, leq_succ'. + cbn; by apply (ap S), IHn, leq_pred'. Defined. (** [nat_min n m] is [m] if [m <= n]. *) @@ -1405,7 +1405,7 @@ Proof. apply IH_strong. simple_induction n n IHn; intros m H. 1: contradiction (not_lt_zero_r m). - apply leq_succ' in H. + apply leq_pred' in H. apply equiv_leq_lt_or_eq in H. destruct H as [H|p]. - by apply IHn. From 76885e7e45803b9a4b8abbb365fedf4c462c3b01 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 4 Aug 2024 21:19:55 +0100 Subject: [PATCH 244/282] Nat: improve proof of equiv_leq_lt_or_eq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 10 +++------- 1 file changed, 3 insertions(+), 7 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index e61676c8a81..fd4436cc8ad 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -854,13 +854,9 @@ Defined. Definition equiv_leq_lt_or_eq {n m} : (n <= m) <~> (n < m) + (n = m). Proof. srapply equiv_iff_hprop. - - srapply hprop_allpath. - intros x y. - snrapply equiv_path_sum. - destruct x as [l|p], y as [q|r]. - 1,4: rapply path_ishprop. - + destruct r; contradiction (lt_irrefl _ _). - + destruct p; contradiction (lt_irrefl _ _). + - rapply ishprop_sum. + intros H1 H2. + contradiction (lt_irrefl _ _). - intro l; induction l. + now right. + left; exact (leq_succ l). From 8a0d5a4df443cf52ed5472b6d9231e2accfc0010 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Sun, 4 Aug 2024 21:22:16 +0100 Subject: [PATCH 245/282] Fix: proof of equiv_leq_lt_or_eq Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 5 +++-- 1 file changed, 3 insertions(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index fd4436cc8ad..d96d70f12ce 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -854,8 +854,9 @@ Defined. Definition equiv_leq_lt_or_eq {n m} : (n <= m) <~> (n < m) + (n = m). Proof. srapply equiv_iff_hprop. - - rapply ishprop_sum. - intros H1 H2. + - nrapply ishprop_sum. + 1,2: exact _. + intros H1 p; destruct p. contradiction (lt_irrefl _ _). - intro l; induction l. + now right. From 5cad8d0d44e4e776d0ba311ef8676b5a09a60759 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Mon, 5 Aug 2024 11:13:19 +0100 Subject: [PATCH 246/282] Nat: distributivity of multiplication over subtraction Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 27 +++++++++++++++++++++++++++ 1 file changed, 27 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index d96d70f12ce..c43c89378f3 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1174,6 +1174,33 @@ Proof. exact IHleq. Defined. +(** Multiplication on the left distributes over subtraction. *) +Definition nat_dist_sub_l n m k + : n * (m - k) = n * m - n * k. +Proof. + induction n as [|n IHn] in m, k |- *. + 1: reflexivity. + destruct (leq_dichotomy k m) as [l|r]. + - simpl; rewrite IHn, <- nat_sub_l_add_r, <- nat_sub_l_add_l, + nat_sub_r_add; trivial; exact _. + - apply leq_lt in r. + apply equiv_nat_sub_leq in r. + rewrite r. + rewrite nat_mul_zero_r. + symmetry. + apply equiv_nat_sub_leq. + apply nat_mul_l_monotone. + by apply equiv_nat_sub_leq. +Defined. + +(** Multiplication on the right distributes over subtraction. *) +Definition nat_dist_sub_r n m k + : (n - m) * k = n * k - m * k. +Proof. + rewrite 3 (nat_mul_comm _ k). + apply nat_dist_sub_l. +Defined. + (** *** Monotonicity of subtraction *) (** Subtraction is monotone in the left argument. *) From 55adc5b09c32c70ed9b5871455f3346f1ae7b508 Mon Sep 17 00:00:00 2001 From: D Date: Tue, 6 Aug 2024 17:51:09 +0200 Subject: [PATCH 247/282] Added definition for cyclic' to AbGroups/Cyclic.v --- theories/Algebra/AbGroups/Cyclic.v | 4 ++++ 1 file changed, 4 insertions(+) diff --git a/theories/Algebra/AbGroups/Cyclic.v b/theories/Algebra/AbGroups/Cyclic.v index 172984a1f52..7759a72a042 100644 --- a/theories/Algebra/AbGroups/Cyclic.v +++ b/theories/Algebra/AbGroups/Cyclic.v @@ -78,3 +78,7 @@ Definition Z1_to_Z `{Funext} : ab_Z1 $-> abgroup_Z (** The [n]-th cyclic group is the cokernel of [Z1_mul_nat n]. *) Definition cyclic@{u v | u < v} `{Funext} (n : nat) : AbGroup@{u} := ab_cokernel@{u v} (Z1_mul_nat n). + +(** The [n]-th cyclic group is the cokernel of [ab_mul n]. *) +Definition cyclic'@{u v | u < v} `{Funext} (n : nat) : AbGroup@{u} + := ab_cokernel@{u v} (ab_mul (A:=abgroup_Z) n). \ No newline at end of file From 5d59e74d5eb5a1f409c8b5231f4257ba93a0afd7 Mon Sep 17 00:00:00 2001 From: D Date: Tue, 6 Aug 2024 18:55:49 +0200 Subject: [PATCH 248/282] Comment, newline and grp_image change --- theories/Algebra/AbGroups/Cyclic.v | 8 +++++--- theories/Algebra/AbSES/SixTerm.v | 2 +- theories/Algebra/Groups/Image.v | 2 +- 3 files changed, 7 insertions(+), 5 deletions(-) diff --git a/theories/Algebra/AbGroups/Cyclic.v b/theories/Algebra/AbGroups/Cyclic.v index 7759a72a042..0896b9fbacb 100644 --- a/theories/Algebra/AbGroups/Cyclic.v +++ b/theories/Algebra/AbGroups/Cyclic.v @@ -77,8 +77,10 @@ Definition Z1_to_Z `{Funext} : ab_Z1 $-> abgroup_Z (** The [n]-th cyclic group is the cokernel of [Z1_mul_nat n]. *) Definition cyclic@{u v | u < v} `{Funext} (n : nat) : AbGroup@{u} - := ab_cokernel@{u v} (Z1_mul_nat n). + := ab_cokernel@{u} (Z1_mul_nat n). + +(** ** Alternative definition of cyclic group *) (** The [n]-th cyclic group is the cokernel of [ab_mul n]. *) -Definition cyclic'@{u v | u < v} `{Funext} (n : nat) : AbGroup@{u} - := ab_cokernel@{u v} (ab_mul (A:=abgroup_Z) n). \ No newline at end of file +Definition cyclic'@{u} (n : nat) : AbGroup@{u} + := ab_cokernel@{u} (ab_mul (A:=abgroup_Z) n). diff --git a/theories/Algebra/AbSES/SixTerm.v b/theories/Algebra/AbSES/SixTerm.v index 560effb9eae..79df757a9fe 100644 --- a/theories/Algebra/AbSES/SixTerm.v +++ b/theories/Algebra/AbSES/SixTerm.v @@ -220,7 +220,7 @@ Defined. (** The main result of this section. *) Theorem ext_cyclic_ab@{u v w | u < v, v < w} `{Univalence} (n : nat) `{emb : IsEmbedding (Z1_mul_nat n)} {A : AbGroup@{u}} - : ab_cokernel@{v w} (ab_mul (A:=A) n) + : ab_cokernel@{v} (ab_mul (A:=A) n) $<~> ab_ext@{u v} (cyclic'@{u v} n) A. (* We take a large cokernel in order to apply [abses_cokernel_iso]. *) Proof. diff --git a/theories/Algebra/Groups/Image.v b/theories/Algebra/Groups/Image.v index 6ce4afc4f3e..0fe6589ed93 100644 --- a/theories/Algebra/Groups/Image.v +++ b/theories/Algebra/Groups/Image.v @@ -11,7 +11,7 @@ Local Open Scope mc_scope. Local Open Scope mc_mult_scope. (** The image of a group homomorphism between groups is a subgroup *) -Definition grp_image {A B : Group} (f : A $-> B) : Subgroup B. +Definition grp_image {A B : Group} (f : GroupHomomorphism A B) : Subgroup B. Proof. snrapply (Build_Subgroup' (fun b => hexists (fun a => f a = b))). - exact _. From a6c4ec7b5517e21b714e8906a33c40d584adfe03 Mon Sep 17 00:00:00 2001 From: D Date: Tue, 6 Aug 2024 19:28:45 +0200 Subject: [PATCH 249/282] Removing universe variables --- theories/Algebra/AbGroups/Cyclic.v | 8 +++++--- theories/Algebra/AbSES/SixTerm.v | 2 +- 2 files changed, 6 insertions(+), 4 deletions(-) diff --git a/theories/Algebra/AbGroups/Cyclic.v b/theories/Algebra/AbGroups/Cyclic.v index 0896b9fbacb..0b5e194a2c5 100644 --- a/theories/Algebra/AbGroups/Cyclic.v +++ b/theories/Algebra/AbGroups/Cyclic.v @@ -2,6 +2,8 @@ Require Import Basics Types WildCat.Core Truncations.Core AbelianGroup AbHom Centralizer AbProjective Groups.FreeGroup AbGroups.Z Spaces.Int. +Local Set Universe Minimization ToSet. + (** * Cyclic groups *) (** ** The free group on one generator *) @@ -13,7 +15,7 @@ Definition Z1 := FreeGroup Unit. Definition Z1_gen : Z1 := freegroup_in tt. (* The generator *) (** The recursion principle of [Z1] and its computation rule. *) -Definition Z1_rec {G : Group@{u}} (g : G) : Z1 $-> G +Definition Z1_rec {G : Group} (g : G) : Z1 $-> G := FreeGroup_rec Unit G (unit_name g). Definition Z1_rec_beta {G : Group} (g : G) : Z1_rec g Z1_gen = g @@ -82,5 +84,5 @@ Definition cyclic@{u v | u < v} `{Funext} (n : nat) : AbGroup@{u} (** ** Alternative definition of cyclic group *) (** The [n]-th cyclic group is the cokernel of [ab_mul n]. *) -Definition cyclic'@{u} (n : nat) : AbGroup@{u} - := ab_cokernel@{u} (ab_mul (A:=abgroup_Z) n). +Definition cyclic'@{} (n : nat) : AbGroup + := ab_cokernel (ab_mul (A:=abgroup_Z) n). diff --git a/theories/Algebra/AbSES/SixTerm.v b/theories/Algebra/AbSES/SixTerm.v index 79df757a9fe..8efe576ca4b 100644 --- a/theories/Algebra/AbSES/SixTerm.v +++ b/theories/Algebra/AbSES/SixTerm.v @@ -218,7 +218,7 @@ Proof. Defined. (** The main result of this section. *) -Theorem ext_cyclic_ab@{u v w | u < v, v < w} `{Univalence} +Theorem ext_cyclic_ab@{u v| u < v} `{Univalence} (n : nat) `{emb : IsEmbedding (Z1_mul_nat n)} {A : AbGroup@{u}} : ab_cokernel@{v} (ab_mul (A:=A) n) $<~> ab_ext@{u v} (cyclic'@{u v} n) A. From 35da1bca3f734118887a70969b298e288b648a81 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 13:11:59 +0100 Subject: [PATCH 250/282] Nat: improve and fix nat_move lemmas Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 12 +++++++----- 1 file changed, 7 insertions(+), 5 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index c43c89378f3..93b482c8ff9 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -691,7 +691,7 @@ Proof. Defined. (** We can move a subtracted number to the left-hand side of an equation. *) -Definition nat_moveL_nV {k m} n : k + n = m -> k = m - n. +Definition nat_moveL_nV {n m} k : n + k = m -> n = m - k. Proof. intros p. destruct p. @@ -700,9 +700,12 @@ Proof. Defined. (** We can move a subtracted number to the right-hand side of an equation. *) -Definition nat_moveR_nV {k m} n : k = n + m -> k - m = n - := fun p => (nat_moveL_nV _ p^)^. - +Definition nat_moveR_nV {n m} k : n = k + m -> n - k = m. +Proof. + rewrite nat_add_comm. + exact (fun p => (nat_moveL_nV _ p^)^). +Defined. + (** Subtracting a successor is the predecessor of subtracting the original number. *) Definition nat_sub_succ_r n m : n - m.+1 = nat_pred (n - m). Proof. @@ -1436,4 +1439,3 @@ Proof. - destruct p. by apply IH_strong. Defined. - From aa3b300efad0b77a0babb046d51ea864050112d3 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 14:46:43 +0100 Subject: [PATCH 251/282] Nat: divisibility Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 93 ++++++++++++++++++++++++++++++++++++++ 1 file changed, 93 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 93b482c8ff9..0a1266d8869 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1,5 +1,10 @@ Require Import Basics.Overture Basics.Tactics Basics.PathGroupoids +<<<<<<< HEAD Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat Basics.Classes Types.Sum. +======= + Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat Basics.Classes + Types.Sum Types.Sigma Types.Prod Types.Paths. +>>>>>>> 740b9439f (fixup divides) Export Basics.Nat. Local Set Universe Minimization ToSet. @@ -1439,3 +1444,91 @@ Proof. - destruct p. by apply IH_strong. Defined. + +(** ** Divisibility *) + +(** We define divisibility as a relation between natural numbers. *) +Class NatDivides (n m : nat) : Type0 := nat_divides : {k : nat & k * n = m}. + +Notation "( n | m )" := (NatDivides n m) (at level 0) : nat_scope. + +(** Any number divides [0]. *) +Global Instance nat_divides_zero_r n : (n | 0) + := (0; idpath). + +(** [1] divides any number. *) +Global Instance nat_divides_one_l n : (1 | n) + := (n; nat_mul_one_r _). + +(** Any number divides itself. Divisibility is a reflexive relation. *) +Global Instance nat_divides_refl n : (n | n) + := (1; nat_mul_one_l _). + +Global Instance reflexive_nat_divides : Reflexive NatDivides := nat_divides_refl. + +(** Divisibility is transitive. *) +Definition nat_divides_trans n m l : (n | m) -> (m | l) -> (n | l). +Proof. + intros [x p] [y q]. + exists (y * x). + lhs_V nrapply nat_mul_assoc. + lhs nrapply (ap _ p). + exact q. +Defined. + +Global Instance transitive_nat_divides : Transitive NatDivides := nat_divides_trans. + +(** A left factor divides a product. *) +Global Instance nat_divides_mul_l' n m : (n | n * m) + := (m; nat_mul_comm _ _). + +(** A right factor divides a product. *) +Global Instance nat_divides_mul_r' n m : (m | n * m) + := (n; idpath). + +(** Divisibility of the product is implied by divisibility of the left factor. *) +Global Instance nat_divides_mul_l n m l : (n | m) -> (n | m * l) + := fun H => nat_divides_trans _ _ _ _ _. + +(** Divisibility of the product is implied by divisibility of the right factor. *) +Global Instance nat_divides_mul_r n m l : (n | m) -> (n | l * m) + := fun H => nat_divides_trans _ _ _ _ _. + +(** Divisibility of the sum is implied by divisibility of the summands. *) +Global Instance nat_divides_add n m l : (n | m) -> (n | l) -> (n | m + l). +Proof. + intros [x p] [y q]. + exists (x + y). + destruct p, q. + nrapply nat_dist_r. +Defined. + +(** Divisibility of the difference is implied by divisibility of the minuend and subtrahend. *) +Global Instance nat_divides_sub n m l : (n | m) -> (n | l) -> (n | m - l). +Proof. + intros [x p] [y q]. + exists (x - y). + destruct p, q. + nrapply nat_dist_sub_r. +Defined. + +(** Divisibility implies that the divisor is less than or equal to the dividend. *) +Definition leq_divides n m : 0 < m -> (n | m) -> n <= m. +Proof. + intros H1 [x p]. + destruct p, x. + 1: contradiction (not_lt_zero_r _ H1). + rapply (leq_mul_l _ _ 0). +Defined. + +(** Divisibility is antisymmetric *) +Definition nat_divides_antisym n m : (n | m) -> (m | n) -> n = m. +Proof. + intros H1 H2. + destruct m; only 1: exact (H2.2^ @ nat_mul_zero_r _). + destruct n; only 1: exact ((nat_mul_zero_r _)^ @ H1.2). + snrapply leq_antisym; nrapply leq_divides; exact _. +Defined. + +Global Instance antisymmetric_divides : AntiSymmetric NatDivides + := nat_divides_antisym. From 261444d65ca1d2825b2889a6d70c4f7a595c5d1a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 13:15:39 +0100 Subject: [PATCH 252/282] Nat: fix summand less than sum lemmas, add multiplicand less than prod Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Arithmetic.v | 2 +- theories/Spaces/Nat/Core.v | 28 +++++++++++++++++++++------- 2 files changed, 22 insertions(+), 8 deletions(-) diff --git a/theories/Spaces/Nat/Arithmetic.v b/theories/Spaces/Nat/Arithmetic.v index a49a34447a3..a117f5a388a 100644 --- a/theories/Spaces/Nat/Arithmetic.v +++ b/theories/Spaces/Nat/Arithmetic.v @@ -167,7 +167,7 @@ Defined. Proposition n_leq_m_n_leq_plus_m_k (n m k : nat) : n <= m -> n <= m + k. Proof. - intro l; apply (leq_trans l); exact (leq_add_l m k). + intro l; apply (leq_trans l); exact (leq_add_r m k). Defined. (** This inductive type is defined because it lets you loop from [i = 0] up to [i = n] by structural induction on a proof of [increasing_geq n 0]. With the existing [leq] type and the inductive structure of [n], it is easier and more natural to loop downwards from [i = n] to [i = 0], but harder to find the least natural number in the interval $[0,n]$ satisfying a given property. *) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 0a1266d8869..0e3c397462e 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -840,22 +840,36 @@ Defined. (** *** Addition lemmas *) +(** The second summand is less than or equal to the sum. *) +Global Instance leq_add_l n m : n <= m + n. +Proof. + simple_induction m m IH. + - exact (leq_refl n). + - exact (leq_succ_r IH). +Defined. + (** The first summand is less than or equal to the sum. *) -Global Instance leq_add_l n m : n <= n + m. +Global Instance leq_add_r n m : n <= n + m. Proof. simple_induction n n IHn. - exact (leq_zero_l m). - exact (leq_succ IHn). Defined. -(** The second summand is less than or equal to the sum. *) -Global Instance leq_add_r n m : n <= m + n. +(** *** Multiplication lemmas *) + +(** The second multiplicand is less than or equal to the product. *) +Global Instance leq_mul_l n m l : l < m -> n <= m * n. Proof. - simple_induction m m IH. - - exact (leq_refl n). - - exact (leq_succ_r IH). + intros H; induction H; exact _. Defined. +(** The first multiplicand is less than or equal to the product. *) +Global Instance leq_mul_r n m l : l < m -> n <= n * m. +Proof. + rewrite nat_mul_comm; exact _. +Defined. + (** Alternative Characterizations of [<=] *) (** [n <= m] is equivalent to [(n < m) + (n = m)]. This justifies the name "less than or equal to". Note that it is not immediately obvious that the latter type is a hprop. *) @@ -890,7 +904,7 @@ Proof. apply nat_add_sub_l_cancel, p. - intros [k p]. destruct p. - apply leq_add_r. + apply leq_add_l. Defined. (** *** Dichotomy of [<=] *) From 346925a27d48db29a23b742d83c2cee6ec244eda Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 13:17:15 +0100 Subject: [PATCH 253/282] Nat: strengthen left/right strict monotonicity of multiplication Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 14 ++++++++------ 1 file changed, 8 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 0e3c397462e..df31d6508c3 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1096,18 +1096,20 @@ Hint Immediate nat_mul_monotone : typeclass_instances. (** *** Strict monotonicity of multiplication *) (** Multiplication on the left by a positive number is strictly monotone. *) -Definition nat_mul_l_strictly_monotone {n m} k - : n < m -> k.+1 * n < k.+1 * m. +Definition nat_mul_l_strictly_monotone {n m l} k + : l < k -> n < m -> k * n < k * m. Proof. - intros H; induction k as [|k IHk] in |- *; exact _. + destruct k. + 1: intro; contradiction (not_lt_zero_r _ H). + intros _ H; induction k as [|k IHk] in |- *; exact _. Defined. Hint Immediate nat_mul_l_strictly_monotone : typeclass_instances. (** Multiplication on the right by a positive number is strictly monotone. *) -Definition nat_mul_r_strictly_monotone {n m} k - : n < m -> n * k.+1 < m * k.+1. +Definition nat_mul_r_strictly_monotone {n m l} k + : l < k -> n < m -> n * k < m * k. Proof. - intros H; rewrite 2 (nat_mul_comm _ k.+1); exact _. + intros ? H; rewrite 2 (nat_mul_comm _ k); exact _. Defined. Hint Immediate nat_mul_r_strictly_monotone : typeclass_instances. From 7709112d976c73ff0542f67a65dd79304c48534e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 13:19:50 +0100 Subject: [PATCH 254/282] Nat: add uniqueness of quotient remainder form Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 69 +++++++++++++++++++++++++++++++++++--- 1 file changed, 64 insertions(+), 5 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index df31d6508c3..6874cc7cf78 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1,10 +1,6 @@ Require Import Basics.Overture Basics.Tactics Basics.PathGroupoids -<<<<<<< HEAD - Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat Basics.Classes Types.Sum. -======= Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat Basics.Classes - Types.Sum Types.Sigma Types.Prod Types.Paths. ->>>>>>> 740b9439f (fixup divides) + Types.Sum Types.Prod Types.Paths. Export Basics.Nat. Local Set Universe Minimization ToSet. @@ -1548,3 +1544,66 @@ Defined. Global Instance antisymmetric_divides : AntiSymmetric NatDivides := nat_divides_antisym. + +(** ** Properties of division *) + +Local Definition divmod_unique_helper b q1 q2 r1 r2 + : r1 < b -> r2 < b -> q1 < q2 -> r2 < r1 -> b * q1 + r1 = b * q2 + r2 + -> q1 = q2 /\ r1 = r2. +Proof. + intros H1 H2 H3 H4 p. + rewrite 2 (nat_add_comm (b * _)) in p. + apply nat_moveL_nV in p. + rewrite nat_sub_l_add_r in p. + - rewrite <- nat_dist_sub_l in p. + apply nat_moveR_nV in p. + symmetry in p. + assert (H5 : b <= r1 - r2). + { rewrite <- p. + snrapply (leq_mul_r _ _ 0). + by apply equiv_lt_lt_sub. } + apply leq_iff_not_gt in H5. + contradiction H5. + rapply lt_moveR_nV. + - rapply nat_mul_l_monotone. +Defined. + +(** Quotients and remainders are uniquely determined. *) +Definition divmod_unique b q1 q2 r1 r2 + : r1 < b -> r2 < b -> b * q1 + r1 = b * q2 + r2 + -> q1 = q2 /\ r1 = r2. +Proof. + intros H1 H2 p. + destruct (nat_trichotomy q1 q2) as [[q | q] | q]; + destruct (nat_trichotomy r1 r2) as [[r | r] | r]. + - apply (nat_mul_l_strictly_monotone b H1) in q. + pose (nat_add_strictly_monotone q r). + rewrite p in l. + contradiction (lt_irrefl _ l). + - destruct r. + split; trivial. + apply isinj_nat_add_r in p. + apply (nat_mul_l_strictly_monotone b H1) in q. + rewrite p in q. + contradiction (lt_irrefl _ q). + - by apply (divmod_unique_helper b). + - destruct q. + split; trivial. + by apply isinj_nat_add_l in p. + - by split. + - split; trivial. + destruct q. + by apply isinj_nat_add_l in p. + - rapply (equiv_functor_prod (f:=inverse) (g:=inverse)). + by apply (divmod_unique_helper b). + - destruct r. + split; trivial. + apply isinj_nat_add_r in p. + apply (nat_mul_l_strictly_monotone b H2) in q. + rewrite p in q. + contradiction (lt_irrefl _ q). + - apply (nat_mul_l_strictly_monotone b H2) in q. + pose (nat_add_strictly_monotone q r). + rewrite p in l. + contradiction (lt_irrefl _ l). +Defined. From 89f4408e2ad0bfcf5a1a2629323bdbd1f9cd3630 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 14:46:15 +0100 Subject: [PATCH 255/282] Nat: basic properties of div and mod Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 35 +++++++++++++++++++++++++++++++++++ 1 file changed, 35 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 6874cc7cf78..4cf6b5a6464 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1607,3 +1607,38 @@ Proof. rewrite p in l. contradiction (lt_irrefl _ l). Defined. + +(** [0] divided by any number is [0]. *) +Definition nat_div_zero_l n : 0 / n = 0. +Proof. + by induction n. +Defined. + +(** [n] divided by [0] is [0] by convention. *) +Definition nat_div_zero_r n : n / 0 = 0 := idpath. + +(** [n] divided by [1] is [n]. *) +Definition nat_div_one_r n : n / 1 = n. +Proof. + destruct n; trivial. + cbn; change n.+1 with (1 + n). + generalize 1; intros k. + induction n in k |- *. + 1: exact (nat_add_zero_r k)^. + rhs nrapply nat_add_succ_r. + apply IHn. +Defined. + +(** [0] modulo any positive number is [0]. *) +Definition nat_mod_zero_l n : 0 < n -> 0 mod n = 0. +Proof. + induction n; trivial. + intro; apply nat_sub_cancel. +Defined. + +(** [n] modulo [1] is [0]. *) +Definition nat_mod_one_r n : n mod 1 = 0. +Proof. + by induction n. +Defined. + From 0d06bfddec01d11a613bacdfce197031c51bd35d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 18:51:01 +0100 Subject: [PATCH 256/282] Nat: rename divmod -> nat_divmod Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 12 ++++++------ 1 file changed, 6 insertions(+), 6 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 4cf6b5a6464..83f48caf7be 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -82,28 +82,28 @@ Fixpoint nat_min n m := (** *** Euclidean division *) -(** This division is linear and tail-recursive. In [divmod], [y] is the predecessor of the actual divisor, and [u] is [y] sub the real remainder. *) +(** This division is linear and tail-recursive. In [nat_divmod], [y] is the predecessor of the actual divisor, and [u] is [y] sub the real remainder. *) -Fixpoint divmod x y q u : nat * nat := +Fixpoint nat_divmod x y q u : nat * nat := match x with | 0 => (q , u) | S x' => match u with - | 0 => divmod x' y (S q) y - | S u' => divmod x' y q u' + | 0 => nat_divmod x' y (S q) y + | S u' => nat_divmod x' y q u' end end. Definition div x y : nat := match y with | 0 => y - | S y' => fst (divmod x y' 0 y') + | S y' => fst (nat_divmod x y' 0 y') end. Definition modulo x y : nat := match y with | 0 => y - | S y' => y' - snd (divmod x y' 0 y') + | S y' => y' - snd (nat_divmod x y' 0 y') end. Infix "/" := div : nat_scope. From 077fbf9712289c49fa048bfbc1a5b34f0312d564 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 18:51:33 +0100 Subject: [PATCH 257/282] Nat: rename div -> nat_div Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 5 ++--- 1 file changed, 2 insertions(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 83f48caf7be..f9be5172b23 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -94,19 +94,18 @@ Fixpoint nat_divmod x y q u : nat * nat := end end. -Definition div x y : nat := +Definition nat_div x y : nat := match y with | 0 => y | S y' => fst (nat_divmod x y' 0 y') end. +Infix "/" := nat_div : nat_scope. Definition modulo x y : nat := match y with | 0 => y | S y' => y' - snd (nat_divmod x y' 0 y') end. - -Infix "/" := div : nat_scope. Infix "mod" := modulo : nat_scope. (** *** Greatest common divisor *) From fe86f3f52804d51f975afc29a0abf60940db4543 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 18:52:05 +0100 Subject: [PATCH 258/282] Nat: rename modulo -> nat_mod Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index f9be5172b23..323fc77f95f 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -101,12 +101,12 @@ Definition nat_div x y : nat := end. Infix "/" := nat_div : nat_scope. -Definition modulo x y : nat := +Definition nat_mod x y : nat := match y with | 0 => y | S y' => y' - snd (nat_divmod x y' 0 y') end. -Infix "mod" := modulo : nat_scope. +Infix "mod" := nat_mod : nat_scope. (** *** Greatest common divisor *) From b8630f0774522d8c8bf7d7f725157f1f5f74e228 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 19:02:29 +0100 Subject: [PATCH 259/282] remove blank line Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 1 - 1 file changed, 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 323fc77f95f..93254521807 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1640,4 +1640,3 @@ Definition nat_mod_one_r n : n mod 1 = 0. Proof. by induction n. Defined. - From d02fe7a06eb9b1def011f184b14570620a2777a1 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 19:16:32 +0100 Subject: [PATCH 260/282] Nat: add nat_div_cancel showing n / n = 1 Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 13 +++++++++++++ 1 file changed, 13 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 93254521807..3b7d45bdd69 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1628,6 +1628,19 @@ Proof. apply IHn. Defined. +(** [n] divided by [n] is [1]. *) +Definition nat_div_cancel n : 0 < n -> n / n = 1. +Proof. + intros H; destruct n; trivial. + 1: contradiction (not_lt_zero_r _ H). + clear H. + destruct n; trivial; cbn. + pose (k := 1); change (n.+1) with (k + n); clearbody k. + induction n in k |- *; trivial. + rewrite nat_add_succ_r. + exact (IHn k.+1). +Defined. + (** [0] modulo any positive number is [0]. *) Definition nat_mod_zero_l n : 0 < n -> 0 mod n = 0. Proof. From db5f035a4cfeb05ff2a9ef429e3bd295ab278282 Mon Sep 17 00:00:00 2001 From: D Date: Tue, 6 Aug 2024 20:23:34 +0200 Subject: [PATCH 261/282] Removing universe variables, now builds --- theories/Algebra/AbGroups/Cyclic.v | 8 +++----- theories/Algebra/AbSES/SixTerm.v | 4 ++-- theories/Algebra/Groups/Image.v | 2 +- 3 files changed, 6 insertions(+), 8 deletions(-) diff --git a/theories/Algebra/AbGroups/Cyclic.v b/theories/Algebra/AbGroups/Cyclic.v index 0b5e194a2c5..72130ccbc4a 100644 --- a/theories/Algebra/AbGroups/Cyclic.v +++ b/theories/Algebra/AbGroups/Cyclic.v @@ -2,8 +2,6 @@ Require Import Basics Types WildCat.Core Truncations.Core AbelianGroup AbHom Centralizer AbProjective Groups.FreeGroup AbGroups.Z Spaces.Int. -Local Set Universe Minimization ToSet. - (** * Cyclic groups *) (** ** The free group on one generator *) @@ -78,11 +76,11 @@ Definition Z1_to_Z `{Funext} : ab_Z1 $-> abgroup_Z (** * Finite cyclic groups *) (** The [n]-th cyclic group is the cokernel of [Z1_mul_nat n]. *) -Definition cyclic@{u v | u < v} `{Funext} (n : nat) : AbGroup@{u} - := ab_cokernel@{u} (Z1_mul_nat n). +Definition cyclic `{Funext} (n : nat) : AbGroup + := ab_cokernel (Z1_mul_nat n). (** ** Alternative definition of cyclic group *) (** The [n]-th cyclic group is the cokernel of [ab_mul n]. *) -Definition cyclic'@{} (n : nat) : AbGroup +Definition cyclic' (n : nat) : AbGroup := ab_cokernel (ab_mul (A:=abgroup_Z) n). diff --git a/theories/Algebra/AbSES/SixTerm.v b/theories/Algebra/AbSES/SixTerm.v index 8efe576ca4b..560effb9eae 100644 --- a/theories/Algebra/AbSES/SixTerm.v +++ b/theories/Algebra/AbSES/SixTerm.v @@ -218,9 +218,9 @@ Proof. Defined. (** The main result of this section. *) -Theorem ext_cyclic_ab@{u v| u < v} `{Univalence} +Theorem ext_cyclic_ab@{u v w | u < v, v < w} `{Univalence} (n : nat) `{emb : IsEmbedding (Z1_mul_nat n)} {A : AbGroup@{u}} - : ab_cokernel@{v} (ab_mul (A:=A) n) + : ab_cokernel@{v w} (ab_mul (A:=A) n) $<~> ab_ext@{u v} (cyclic'@{u v} n) A. (* We take a large cokernel in order to apply [abses_cokernel_iso]. *) Proof. diff --git a/theories/Algebra/Groups/Image.v b/theories/Algebra/Groups/Image.v index 0fe6589ed93..6ce4afc4f3e 100644 --- a/theories/Algebra/Groups/Image.v +++ b/theories/Algebra/Groups/Image.v @@ -11,7 +11,7 @@ Local Open Scope mc_scope. Local Open Scope mc_mult_scope. (** The image of a group homomorphism between groups is a subgroup *) -Definition grp_image {A B : Group} (f : GroupHomomorphism A B) : Subgroup B. +Definition grp_image {A B : Group} (f : A $-> B) : Subgroup B. Proof. snrapply (Build_Subgroup' (fun b => hexists (fun a => f a = b))). - exact _. From 4e342b005b722477144f2a64810c20153ac97ba0 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 19:38:13 +0100 Subject: [PATCH 262/282] Nat: add todo for n mod n Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 2 ++ 1 file changed, 2 insertions(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 3b7d45bdd69..70b668e4c09 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1653,3 +1653,5 @@ Definition nat_mod_one_r n : n mod 1 = 0. Proof. by induction n. Defined. + +(** TODO: [n] modulo [n] is [0]. *) From 3e4792b84531516a2efec7535d8a746487e3fca5 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 20:11:39 +0100 Subject: [PATCH 263/282] remove level on notation Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 70b668e4c09..af425cc6995 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1461,7 +1461,7 @@ Defined. (** We define divisibility as a relation between natural numbers. *) Class NatDivides (n m : nat) : Type0 := nat_divides : {k : nat & k * n = m}. -Notation "( n | m )" := (NatDivides n m) (at level 0) : nat_scope. +Notation "( n | m )" := (NatDivides n m) : nat_scope. (** Any number divides [0]. *) Global Instance nat_divides_zero_r n : (n | 0) From c9b3ab6e2f1132170f2d082ae515b6b22934ad4e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 20:14:54 +0100 Subject: [PATCH 264/282] Nat: strengthen nat_mod_zero_l Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 6 +++--- 1 file changed, 3 insertions(+), 3 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index af425cc6995..f57eb3d46b7 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1641,11 +1641,11 @@ Proof. exact (IHn k.+1). Defined. -(** [0] modulo any positive number is [0]. *) -Definition nat_mod_zero_l n : 0 < n -> 0 mod n = 0. +(** [0] modulo any number is [0]. *) +Definition nat_mod_zero_l n : 0 mod n = 0. Proof. induction n; trivial. - intro; apply nat_sub_cancel. + apply nat_sub_cancel. Defined. (** [n] modulo [1] is [0]. *) From 64b6836457d323ade5707f3f8b9e139b43384fdc Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 20:30:44 +0100 Subject: [PATCH 265/282] replace cyclic with cyclic' Signed-off-by: Ali Caglayan --- theories/Algebra/AbGroups/Cyclic.v | 8 +------- 1 file changed, 1 insertion(+), 7 deletions(-) diff --git a/theories/Algebra/AbGroups/Cyclic.v b/theories/Algebra/AbGroups/Cyclic.v index 72130ccbc4a..517129dfa6a 100644 --- a/theories/Algebra/AbGroups/Cyclic.v +++ b/theories/Algebra/AbGroups/Cyclic.v @@ -75,12 +75,6 @@ Definition Z1_to_Z `{Funext} : ab_Z1 $-> abgroup_Z (** * Finite cyclic groups *) -(** The [n]-th cyclic group is the cokernel of [Z1_mul_nat n]. *) -Definition cyclic `{Funext} (n : nat) : AbGroup - := ab_cokernel (Z1_mul_nat n). - -(** ** Alternative definition of cyclic group *) - (** The [n]-th cyclic group is the cokernel of [ab_mul n]. *) -Definition cyclic' (n : nat) : AbGroup +Definition cyclic (n : nat) : AbGroup := ab_cokernel (ab_mul (A:=abgroup_Z) n). From 76986e028feb49f281af898b58f011838cf5c046 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Tue, 6 Aug 2024 22:41:44 +0100 Subject: [PATCH 266/282] Nat: add nat_mod_cancel for n mod n Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 16 +++++++++++++++- 1 file changed, 15 insertions(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index f57eb3d46b7..c15d6729088 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1654,4 +1654,18 @@ Proof. by induction n. Defined. -(** TODO: [n] modulo [n] is [0]. *) +(** [n] modulo [n] is [0]. *) +Definition nat_mod_cancel n : n mod n = 0. +Proof. + destruct n; trivial. + apply nat_moveR_nV. + rhs nrapply nat_add_zero_r. + destruct n; trivial. + simpl; change n.+1 with (1 + n). + generalize 1; intros k. + induction n in k |- *. + 1: by destruct k. + rewrite nat_add_succ_r. + destruct n; trivial. + exact (IHn k.+1). +Defined. From c78e25f1cfd8f4cf87bfb9a5818d6026cebcaa8d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 7 Aug 2024 12:17:49 +0100 Subject: [PATCH 267/282] Nat: complete specification of div/mod, basics on gcd Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 232 +++++++++++++++++++++++++++++++------ 1 file changed, 196 insertions(+), 36 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index c15d6729088..7d8cf4dc733 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -103,7 +103,7 @@ Infix "/" := nat_div : nat_scope. Definition nat_mod x y : nat := match y with - | 0 => y + | 0 => x | S y' => y' - snd (nat_divmod x y' 0 y') end. Infix "mod" := nat_mod : nat_scope. @@ -112,10 +112,10 @@ Infix "mod" := nat_mod : nat_scope. (** We use Euclid algorithm, which is normally not structural, but Coq is now clever enough to accept this (behind modulo there is a subtraction, which now preserves being a subterm) *) -Fixpoint gcd a b := +Fixpoint nat_gcd a b := match a with | O => b - | S a' => gcd (b mod a'.+1) a'.+1 + | S a' => nat_gcd (b mod a'.+1) a'.+1 end. (** *** Factorial *) @@ -1136,6 +1136,14 @@ Defined. (** ** Further properties of subtraction *) +Global Instance leq_sub_l n m : n - m <= n. +Proof. + apply equiv_nat_sub_leq. + rewrite nat_sub_comm_r. + rewrite nat_sub_cancel. + apply nat_sub_zero_l. +Defined. + (** Subtracting from a successor is the successor of subtracting from the original number, as long as the amount being subtracted is less than or equal to the original number. *) Definition nat_sub_succ_l n m : m <= n -> n.+1 - m = (n - m).+1. Proof. @@ -1514,6 +1522,23 @@ Proof. nrapply nat_dist_r. Defined. +(** If [n] divides a sum and the left summand, then [n] divides the right summand. *) +Definition nat_divides_add_r n m l : (n | m) -> (n | m + l) -> (n | l). +Proof. + intros [x p] [y q]. + exists (y - x). + lhs nrapply nat_dist_sub_r. + apply nat_moveR_nV. + lhs nrapply q. + exact (ap _ p^). +Defined. + +(** If [n] divides a sum and the right summand, then [n] divides the left summand. *) +Definition nat_divides_add_l n m l : (n | l) -> (n | m + l) -> (n | m). +Proof. + rewrite nat_add_comm; apply nat_divides_add_r. +Defined. + (** Divisibility of the difference is implied by divisibility of the minuend and subtrahend. *) Global Instance nat_divides_sub n m l : (n | m) -> (n | l) -> (n | m - l). Proof. @@ -1546,7 +1571,7 @@ Global Instance antisymmetric_divides : AntiSymmetric NatDivides (** ** Properties of division *) -Local Definition divmod_unique_helper b q1 q2 r1 r2 +Local Definition nat_divmod_unique_helper b q1 q2 r1 r2 : r1 < b -> r2 < b -> q1 < q2 -> r2 < r1 -> b * q1 + r1 = b * q2 + r2 -> q1 = q2 /\ r1 = r2. Proof. @@ -1568,24 +1593,24 @@ Proof. Defined. (** Quotients and remainders are uniquely determined. *) -Definition divmod_unique b q1 q2 r1 r2 - : r1 < b -> r2 < b -> b * q1 + r1 = b * q2 + r2 +Definition nat_divmod_unique d q1 q2 r1 r2 + : r1 < d -> r2 < d -> d * q1 + r1 = d * q2 + r2 -> q1 = q2 /\ r1 = r2. Proof. intros H1 H2 p. destruct (nat_trichotomy q1 q2) as [[q | q] | q]; destruct (nat_trichotomy r1 r2) as [[r | r] | r]. - - apply (nat_mul_l_strictly_monotone b H1) in q. + - apply (nat_mul_l_strictly_monotone d H1) in q. pose (nat_add_strictly_monotone q r). rewrite p in l. contradiction (lt_irrefl _ l). - destruct r. split; trivial. apply isinj_nat_add_r in p. - apply (nat_mul_l_strictly_monotone b H1) in q. + apply (nat_mul_l_strictly_monotone d H1) in q. rewrite p in q. contradiction (lt_irrefl _ q). - - by apply (divmod_unique_helper b). + - by apply (nat_divmod_unique_helper d). - destruct q. split; trivial. by apply isinj_nat_add_l in p. @@ -1594,19 +1619,72 @@ Proof. destruct q. by apply isinj_nat_add_l in p. - rapply (equiv_functor_prod (f:=inverse) (g:=inverse)). - by apply (divmod_unique_helper b). + by apply (nat_divmod_unique_helper d). - destruct r. split; trivial. apply isinj_nat_add_r in p. - apply (nat_mul_l_strictly_monotone b H2) in q. + apply (nat_mul_l_strictly_monotone d H2) in q. rewrite p in q. contradiction (lt_irrefl _ q). - - apply (nat_mul_l_strictly_monotone b H2) in q. + - apply (nat_mul_l_strictly_monotone d H2) in q. pose (nat_add_strictly_monotone q r). rewrite p in l. contradiction (lt_irrefl _ l). Defined. +(** The quotient-remainder form is invariant under [nat_divmod]. *) +Definition nat_divmod_spec_helper x y q u (H : u <= y) + : let (q', u') := nat_divmod x y q u in + x + y.+1 * q + (y - u) = y.+1 * q' + (y - u') /\ u' <= y. +Proof. + intros d r. + induction x as [|x IHx] in y, q, u, H, d, r |- *; only 1: by split. + destruct u as [|u]. + - destruct (IHx y q.+1 y _) as [p H']. + split; trivial. + rewrite <- p, nat_sub_zero_r, nat_sub_cancel, nat_add_zero_r. + simpl. + by rewrite nat_add_succ_r, <- 2 nat_add_assoc, nat_mul_succ_r. + - destruct (IHx y q u _) as [p H']. + split; trivial. + rewrite <- p, 2 nat_add_succ_l, <- nat_add_succ_r. + snrapply ap. + rewrite nat_sub_succ_r. + apply nat_succ_pred. + rapply lt_moveL_nV. +Defined. + +(** Division and modulo can be put in quotient-remainder form. *) +Definition nat_divmod_spec x y : x = y * (x / y) + x mod y. +Proof. + destruct y as [|y]; only 1: reflexivity. + pose proof (p := fst (nat_divmod_spec_helper x y 0 y _)). + by rewrite nat_mul_zero_r, nat_sub_cancel, 2 nat_add_zero_r in p. +Defined. + +Definition nat_divmod_spec' x y : x - y * (x / y) = x mod y. +Proof. + apply nat_moveR_nV, nat_divmod_spec. +Defined. + +Global Instance nat_mod_lt' n m r : r < m -> n mod m < m. +Proof. + intros H; destruct H; only 1: exact _. + rapply (lt_leq_lt_trans (m:=m)). +Defined. + +(** [n] modulo [m] is less than [m]. *) +Global Instance nat_mod_lt n m : 0 < m -> n mod m < m + := nat_mod_lt' n m 0. + +(** Division is unique. *) +Definition nat_div_unique x y q r (H : r < y) (p : y * q + r = x) : q = x / y + := fst (nat_divmod_unique y q (x / y) r (x mod y) _ _ (p @ nat_divmod_spec x y)). + +(** Modulo is unique. *) +Definition nat_mod_unique x y q r (H : r < y) (p : y * q + r = x) : r = x mod y + := snd (nat_divmod_unique y q (x / y) r (x mod y) _ _ (p @ nat_divmod_spec x y)). + (** [0] divided by any number is [0]. *) Definition nat_div_zero_l n : 0 / n = 0. Proof. @@ -1619,35 +1697,41 @@ Definition nat_div_zero_r n : n / 0 = 0 := idpath. (** [n] divided by [1] is [n]. *) Definition nat_div_one_r n : n / 1 = n. Proof. - destruct n; trivial. - cbn; change n.+1 with (1 + n). - generalize 1; intros k. - induction n in k |- *. - 1: exact (nat_add_zero_r k)^. - rhs nrapply nat_add_succ_r. - apply IHn. + lhs_V nrapply nat_mul_one_l. + lhs_V nrapply nat_add_zero_r. + symmetry; apply nat_divmod_spec. Defined. (** [n] divided by [n] is [1]. *) Definition nat_div_cancel n : 0 < n -> n / n = 1. Proof. - intros H; destruct n; trivial. - 1: contradiction (not_lt_zero_r _ H). - clear H. - destruct n; trivial; cbn. - pose (k := 1); change (n.+1) with (k + n); clearbody k. - induction n in k |- *; trivial. - rewrite nat_add_succ_r. - exact (IHn k.+1). + intros [|m _]; trivial; symmetry. + nrapply (nat_div_unique _ _ _ 0); only 1: exact _. + lhs nrapply nat_add_zero_r. + nrapply nat_mul_one_r. Defined. -(** [0] modulo any number is [0]. *) +(** [0] modulo [n] is [0]. *) Definition nat_mod_zero_l n : 0 mod n = 0. Proof. induction n; trivial. apply nat_sub_cancel. Defined. +(** [n] modulo [0] is [n]. *) +Definition nat_mod_zero_r n : n mod 0 = n := idpath. + +(** TODO: generalise for all small n *) +Definition nat_mod_one_l n : 1 < n -> 1 mod n = 1. +Proof. + intros H. + destruct H; trivial. + destruct m. + 1: contradiction (not_lt_zero_r _ H). + cbn; clear H. + by induction m. +Defined. + (** [n] modulo [1] is [0]. *) Definition nat_mod_one_r n : n mod 1 = 0. Proof. @@ -1658,14 +1742,90 @@ Defined. Definition nat_mod_cancel n : n mod n = 0. Proof. destruct n; trivial. - apply nat_moveR_nV. - rhs nrapply nat_add_zero_r. + symmetry. + snrapply (nat_mod_unique _ _ 1); only 1: exact _. + lhs nrapply nat_add_zero_r. + nrapply nat_mul_one_r. +Defined. + +(** ** Greatest Common Divisor *) + +(** The greatest common divisor of [0] and a number is the number itself. *) +Definition nat_gcd_zero_l n : nat_gcd 0 n = n := idpath. + +(** The greatest common divisor of a number and [0] is the number itself. *) +Definition nat_gcd_zero_r n : nat_gcd n 0 = n. +Proof. + induction n; simpl; only 1: reflexivity. + by rewrite nat_sub_cancel. +Defined. + +(** The greatest common divisor of [1] and any number is [1]. *) +Definition nat_gcd_one_l n : nat_gcd 1 n = 1 := idpath. + +(** The greatest common divisor of any number and [1] is [1]. *) +Definition nat_gcd_one_r n : nat_gcd n 1 = 1. +Proof. destruct n; trivial. - simpl; change n.+1 with (1 + n). - generalize 1; intros k. - induction n in k |- *. - 1: by destruct k. - rewrite nat_add_succ_r. + simpl. destruct n; trivial. - exact (IHn k.+1). + rewrite nat_sub_succ_l; only 2: exact _. + by rewrite nat_sub_cancel. +Defined. + +(** Idempotency. *) +Definition nat_gcd_idem n : nat_gcd n n = n. +Proof. + induction n. + 1: reflexivity. + unfold nat_gcd; fold nat_gcd. + by rewrite nat_mod_cancel. +Defined. + +(** We can prove that the greatest common divisor of [n] and [m] divides both [n] and [m]. This proof requires strong induction. *) +Definition divides_l_nat_gcd n m : (nat_gcd n m | n) /\ (nat_gcd n m | m). +Proof. + revert n m; snrapply nat_ind_strong; intros n IHn m. + destruct n. + 1: split; exact _. + destruct (IHn (m mod n.+1) _ n.+1) as [H1 H2]. + unfold nat_gcd; fold nat_gcd. + set (n' := n.+1) in *. + split; only 1: exact H2. + set (r := m mod n'); rewrite (nat_divmod_spec m n'); unfold r; clear r. + exact _. +Defined. + +(** The greatest common divisor of [n] and [m] divides [n]. *) +Global Instance divides_l_nat_gcd_l n m : (nat_gcd n m | n) + := fst (divides_l_nat_gcd n m). + +(** The greatest common divisor of [n] and [m] divides [m]. *) +Global Instance divides_l_nat_gcd_r n m : (nat_gcd n m | m) + := snd (divides_l_nat_gcd n m). + +(** We can prove that any common divisor of [n] and [m] divides the greatest common divisor of [n] and [m]. It is in that sense the greatest. *) +Global Instance divides_r_nat_gcd n m p : (p | n) -> (p | m) -> (p | nat_gcd n m). +Proof. + revert n m p; snrapply nat_ind_strong; intros n IHn m p H1 H2. + destruct n; only 1: exact _. + unfold nat_gcd; fold nat_gcd. + apply IHn; only 1,3: exact _. + rewrite (nat_divmod_spec m n.+1) in H2. + apply nat_divides_add_r in H2; exact _. +Defined. + +(** If [p] is divisible by all common divisors of [n] and [m], and [p] is also a common divisor, then it must necesserily be equal to the greatest common divisor. *) +Definition nat_gcd_unique n m p + (H : forall q, (q | n) -> (q | m) -> (q | p)) + : (p | n) -> (p | m) -> nat_gcd n m = p. +Proof. + intros H1 H2. + rapply nat_divides_antisym. +Defined. + +(** As a corollary of uniquness, we get that the greatest common divisor operation is commutative. *) +Definition nat_gcd_comm n m : nat_gcd n m = nat_gcd m n. +Proof. + rapply nat_gcd_unique. Defined. From f05263ddf3a85031e78721ffd40fedcfaeaa0649 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 7 Aug 2024 12:25:39 +0100 Subject: [PATCH 268/282] Nat: fix typeclass loop Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 3 ++- 1 file changed, 2 insertions(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 7d8cf4dc733..d283c9b54f2 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1667,11 +1667,12 @@ Proof. apply nat_moveR_nV, nat_divmod_spec. Defined. -Global Instance nat_mod_lt' n m r : r < m -> n mod m < m. +Definition nat_mod_lt' n m r : r < m -> n mod m < m. Proof. intros H; destruct H; only 1: exact _. rapply (lt_leq_lt_trans (m:=m)). Defined. +Hint Immediate nat_mod_lt' : typeclass_instances. (** [n] modulo [m] is less than [m]. *) Global Instance nat_mod_lt n m : 0 < m -> n mod m < m From 34d2fafa4fba3de8d85fe6ced7ac582c7d136226 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 7 Aug 2024 14:00:13 +0100 Subject: [PATCH 269/282] Nat: naming improvements and associativity of gcd Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 57 ++++++++++++++++++++++++++++++-------- 1 file changed, 45 insertions(+), 12 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index d283c9b54f2..614887c62a7 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1,6 +1,6 @@ Require Import Basics.Overture Basics.Tactics Basics.PathGroupoids Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat Basics.Classes - Types.Sum Types.Prod Types.Paths. + Types.Sum Types.Prod Types.Paths Types.Sigma. Export Basics.Nat. Local Set Universe Minimization ToSet. @@ -1486,7 +1486,7 @@ Global Instance nat_divides_refl n : (n | n) Global Instance reflexive_nat_divides : Reflexive NatDivides := nat_divides_refl. (** Divisibility is transitive. *) -Definition nat_divides_trans n m l : (n | m) -> (m | l) -> (n | l). +Definition nat_divides_trans {n m l} : (n | m) -> (m | l) -> (n | l). Proof. intros [x p] [y q]. exists (y * x). @@ -1494,8 +1494,9 @@ Proof. lhs nrapply (ap _ p). exact q. Defined. +Hint Immediate nat_divides_trans : typeclass_instances. -Global Instance transitive_nat_divides : Transitive NatDivides := nat_divides_trans. +Global Instance transitive_nat_divides : Transitive NatDivides := @nat_divides_trans. (** A left factor divides a product. *) Global Instance nat_divides_mul_l' n m : (n | n * m) @@ -1506,12 +1507,12 @@ Global Instance nat_divides_mul_r' n m : (m | n * m) := (n; idpath). (** Divisibility of the product is implied by divisibility of the left factor. *) -Global Instance nat_divides_mul_l n m l : (n | m) -> (n | m * l) - := fun H => nat_divides_trans _ _ _ _ _. +Global Instance nat_divides_mul_l {n m} l : (n | m) -> (n | m * l) + := fun H => nat_divides_trans _ _. (** Divisibility of the product is implied by divisibility of the right factor. *) -Global Instance nat_divides_mul_r n m l : (n | m) -> (n | l * m) - := fun H => nat_divides_trans _ _ _ _ _. +Global Instance nat_divides_mul_r {n m} l : (n | m) -> (n | l * m) + := fun H => nat_divides_trans _ _. (** Divisibility of the sum is implied by divisibility of the summands. *) Global Instance nat_divides_add n m l : (n | m) -> (n | l) -> (n | m + l). @@ -1632,6 +1633,22 @@ Proof. contradiction (lt_irrefl _ l). Defined. +(** Divisibility by a positive natural number is a hprop. *) +Global Instance ishprop_nat_divides n m : 0 < n -> IsHProp (n | m). +Proof. + intros H. + apply hprop_allpath. + intros [x p] [y q]. + rapply path_sigma_hprop. + destruct H as [|n]; simpl. + 1: exact ((nat_mul_one_r _)^ @ p @ q^ @ nat_mul_one_r _). + refine (fst (nat_divmod_unique n.+1 x y 0 0 _ _ _)). + lhs nrapply nat_add_zero_r. + rhs nrapply nat_add_zero_r. + rewrite 2 (nat_mul_comm n.+1). + exact (p @ q^). +Defined. + (** The quotient-remainder form is invariant under [nat_divmod]. *) Definition nat_divmod_spec_helper x y q u (H : u <= y) : let (q', u') := nat_divmod x y q u in @@ -1784,7 +1801,7 @@ Proof. Defined. (** We can prove that the greatest common divisor of [n] and [m] divides both [n] and [m]. This proof requires strong induction. *) -Definition divides_l_nat_gcd n m : (nat_gcd n m | n) /\ (nat_gcd n m | m). +Definition nat_divides_l_gcd n m : (nat_gcd n m | n) /\ (nat_gcd n m | m). Proof. revert n m; snrapply nat_ind_strong; intros n IHn m. destruct n. @@ -1798,15 +1815,15 @@ Proof. Defined. (** The greatest common divisor of [n] and [m] divides [n]. *) -Global Instance divides_l_nat_gcd_l n m : (nat_gcd n m | n) - := fst (divides_l_nat_gcd n m). +Global Instance nat_divides_l_gcd_l n m : (nat_gcd n m | n) + := fst (nat_divides_l_gcd n m). (** The greatest common divisor of [n] and [m] divides [m]. *) Global Instance divides_l_nat_gcd_r n m : (nat_gcd n m | m) - := snd (divides_l_nat_gcd n m). + := snd (nat_divides_l_gcd n m). (** We can prove that any common divisor of [n] and [m] divides the greatest common divisor of [n] and [m]. It is in that sense the greatest. *) -Global Instance divides_r_nat_gcd n m p : (p | n) -> (p | m) -> (p | nat_gcd n m). +Global Instance nat_divides_r_gcd n m p : (p | n) -> (p | m) -> (p | nat_gcd n m). Proof. revert n m p; snrapply nat_ind_strong; intros n IHn m p H1 H2. destruct n; only 1: exact _. @@ -1816,6 +1833,12 @@ Proof. apply nat_divides_add_r in H2; exact _. Defined. +Definition nat_divides_r_iff_divides_r_gcd n m p + : (p | n) * (p | m) <-> (p | nat_gcd n m). +Proof. + split; [intros [H1 H2] | intros H; split]; exact _. +Defined. + (** If [p] is divisible by all common divisors of [n] and [m], and [p] is also a common divisor, then it must necesserily be equal to the greatest common divisor. *) Definition nat_gcd_unique n m p (H : forall q, (q | n) -> (q | m) -> (q | p)) @@ -1830,3 +1853,13 @@ Definition nat_gcd_comm n m : nat_gcd n m = nat_gcd m n. Proof. rapply nat_gcd_unique. Defined. + +(** [nat_gcd] is associative. *) +Definition nat_gcd_assoc n m k : nat_gcd n (nat_gcd m k) = nat_gcd (nat_gcd n m) k. +Proof. + nrapply nat_gcd_unique. + - intros q H1 H2. + rapply nat_divides_r_gcd. + - rapply (nat_divides_trans (nat_divides_l_gcd_l _ _)). + - apply nat_divides_r_gcd; rapply nat_divides_trans. +Defined. From b14fc55b90b50a089c06e8a1a17e1aad27e94223 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Wed, 7 Aug 2024 18:22:35 +0100 Subject: [PATCH 270/282] Nat: more theory on gcd Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 94 ++++++++++++++++++++++++++++++++++++-- 1 file changed, 90 insertions(+), 4 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 614887c62a7..cb7cc561357 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1684,16 +1684,23 @@ Proof. apply nat_moveR_nV, nat_divmod_spec. Defined. -Definition nat_mod_lt' n m r : r < m -> n mod m < m. +Definition nat_mod_lt_r' n m r : r < m -> n mod m < m. Proof. intros H; destruct H; only 1: exact _. rapply (lt_leq_lt_trans (m:=m)). Defined. -Hint Immediate nat_mod_lt' : typeclass_instances. +Hint Immediate nat_mod_lt_r' : typeclass_instances. (** [n] modulo [m] is less than [m]. *) -Global Instance nat_mod_lt n m : 0 < m -> n mod m < m - := nat_mod_lt' n m 0. +Global Instance nat_mod_lt_r n m : 0 < m -> n mod m < m + := nat_mod_lt_r' n m 0. + +(** [n] modulo [m] is less than or equal to [m]. *) +Global Instance nat_mod_leq_l n m : n mod m <= n. +Proof. + rewrite <- nat_divmod_spec'. + rapply leq_moveR_nV. +Defined. (** Division is unique. *) Definition nat_div_unique x y q r (H : r < y) (p : y * q + r = x) : q = x / y @@ -1729,6 +1736,32 @@ Proof. nrapply nat_mul_one_r. Defined. +(** [n * m] divided by [n] is [m]. *) +Definition nat_div_mul_cancel_l n m : 0 < n -> (n * m) / n = m. +Proof. + intros H. + symmetry. + nrapply (nat_div_unique _ _ _ _ H). + apply nat_add_zero_r. +Defined. + +(** [n * m] divided by [n] is [m]. *) +Definition nat_div_mul_cancel_r n m : 0 < m -> (n * m) / m = n. +Proof. + rewrite nat_mul_comm. + apply nat_div_mul_cancel_l. +Defined. + +(** When [d] divides [n] and [m], division distributes over addition. *) +Definition nat_div_dist n m d + : 0 < d -> (d | n) -> (d | m) -> (n + m) / d = n / d + m / d. +Proof. + intros H1 [x p] [y q]. + destruct p, q. + rewrite <- nat_dist_r. + by rewrite 3 nat_div_mul_cancel_r. +Defined. + (** [0] modulo [n] is [0]. *) Definition nat_mod_zero_l n : 0 mod n = 0. Proof. @@ -1756,6 +1789,19 @@ Proof. by induction n. Defined. +(** If [m] divides [n], then [n mod m = 0]. *) +Definition nat_mod_divides n m : 0 < m -> (m | n) -> n mod m = 0. +Proof. + intros H [x p]. + destruct p. + lhs_V nrapply nat_divmod_spec'. + pose (nat_div_cancel m H). + rewrite nat_div_mul_cancel_r; only 2: exact _. + apply nat_moveR_nV. + rhs nrapply nat_add_zero_r. + apply nat_mul_comm. +Defined. + (** [n] modulo [n] is [0]. *) Definition nat_mod_cancel n : n mod n = 0. Proof. @@ -1863,3 +1909,43 @@ Proof. - rapply (nat_divides_trans (nat_divides_l_gcd_l _ _)). - apply nat_divides_r_gcd; rapply nat_divides_trans. Defined. + +(** If [nat_gcd n m] is [0], then [n] must also be [0]. *) +Definition nat_gcd_is_zero_l n m : nat_gcd n m = 0 -> n = 0. +Proof. + intros H. + generalize (nat_divides_l_gcd_l n m). + rewrite H. + intros [x p]. + exact (p^ @ nat_mul_zero_r _). +Defined. + +(** If [nat_gcd n m] is [0], then [m] must also be [0]. *) +Definition nat_gcd_is_zero_r n m : nat_gcd n m = 0 -> m = 0. +Proof. + rewrite nat_gcd_comm. + apply nat_gcd_is_zero_l. +Defined. + +(** [nat_gcd n m] is [0] if and only if both [n] and [m] are [0]. *) +Definition nat_gcd_zero_iff_zero n m : nat_gcd n m = 0 <-> n = 0 /\ m = 0. +Proof. + split. + - split. + + by apply (nat_gcd_is_zero_l _ m). + + by apply (nat_gcd_is_zero_r n). + - intros [-> ->]. + reflexivity. +Defined. + +(** [nat_gcd] is positive for positive inputs. *) +Global Instance nat_gcd_pos n m : 0 < n -> 0 < m -> 0 < nat_gcd n m. +Proof. + intros H1 H2. + apply lt_iff_not_geq. + intros H3; hnf in H3. + apply path_zero_leq_zero_r in H3. + apply nat_gcd_zero_iff_zero in H3. + destruct H3 as [->]. + contradiction (not_lt_zero_r _ H1). +Defined. From 78ff1883e86d260030b9df497675d082c58b1f64 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 8 Aug 2024 12:52:47 +0100 Subject: [PATCH 271/282] Update theories/Spaces/Nat/Core.v Co-authored-by: Dan Christensen --- theories/Spaces/Nat/Core.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index cb7cc561357..bc7b02b7fa0 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -82,7 +82,7 @@ Fixpoint nat_min n m := (** *** Euclidean division *) -(** This division is linear and tail-recursive. In [nat_divmod], [y] is the predecessor of the actual divisor, and [u] is [y] sub the real remainder. *) +(** This division takes time linear in `x` and is tail-recursive. In [nat_divmod x y q u], [x + y.+1 * q + (y - u)] is the quantity being divided and [y] is the predecessor of the divisor. It will be called with [q] zero and [u] equal to [y], so that [x] is the quantity being divided. The return value is a pair [(q', u')] with [x + y.+1 * q + (y - u) = y.+1 * q' + (y - u')], as least when [u <= y], as shown in [nat_divmod_spec_helper] below. *) Fixpoint nat_divmod x y q u : nat * nat := match x with From def67e195531f3f03d7f2b51e4028d08471d6e2a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 8 Aug 2024 13:14:51 +0100 Subject: [PATCH 272/282] Update theories/Spaces/Nat/Core.v Co-authored-by: Dan Christensen --- theories/Spaces/Nat/Core.v | 1 + 1 file changed, 1 insertion(+) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index bc7b02b7fa0..0519e6b489d 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -101,6 +101,7 @@ Definition nat_div x y : nat := end. Infix "/" := nat_div : nat_scope. +(** [nat_mod x y] is the remainder when [x] is divided by [y]. When [y] is zero, it is defined to be [x]. See [nat_div_mod_spec] and related results below. *) Definition nat_mod x y : nat := match y with | 0 => x From d3a9cdf208ebea3d61f247067367b0fa22774624 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 8 Aug 2024 13:50:27 +0100 Subject: [PATCH 273/282] review suggestions Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 122 ++++++++++++++----------------------- 1 file changed, 45 insertions(+), 77 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index 0519e6b489d..bbfc257f5f7 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1,6 +1,6 @@ Require Import Basics.Overture Basics.Tactics Basics.PathGroupoids Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat Basics.Classes - Types.Sum Types.Prod Types.Paths Types.Sigma. + Types.Sum Types.Sigma. Export Basics.Nat. Local Set Universe Minimization ToSet. @@ -82,22 +82,22 @@ Fixpoint nat_min n m := (** *** Euclidean division *) -(** This division takes time linear in `x` and is tail-recursive. In [nat_divmod x y q u], [x + y.+1 * q + (y - u)] is the quantity being divided and [y] is the predecessor of the divisor. It will be called with [q] zero and [u] equal to [y], so that [x] is the quantity being divided. The return value is a pair [(q', u')] with [x + y.+1 * q + (y - u) = y.+1 * q' + (y - u')], as least when [u <= y], as shown in [nat_divmod_spec_helper] below. *) +(** This division takes time linear in `x` and is tail-recursive. In [nat_div_mod x y q u], [x + y.+1 * q + (y - u)] is the quantity being divided and [y] is the predecessor of the divisor. It will be called with [q] zero and [u] equal to [y], so that [x] is the quantity being divided. The return value is a pair [(q', u')] with [x + y.+1 * q + (y - u) = y.+1 * q' + (y - u')], as least when [u <= y], as shown in [nat_div_mod_spec_helper] below. *) -Fixpoint nat_divmod x y q u : nat * nat := +Fixpoint nat_div_mod x y q u : nat * nat := match x with | 0 => (q , u) | S x' => match u with - | 0 => nat_divmod x' y (S q) y - | S u' => nat_divmod x' y q u' + | 0 => nat_div_mod x' y (S q) y + | S u' => nat_div_mod x' y q u' end end. Definition nat_div x y : nat := match y with | 0 => y - | S y' => fst (nat_divmod x y' 0 y') + | S y' => fst (nat_div_mod x y' 0 y') end. Infix "/" := nat_div : nat_scope. @@ -105,7 +105,7 @@ Infix "/" := nat_div : nat_scope. Definition nat_mod x y : nat := match y with | 0 => x - | S y' => y' - snd (nat_divmod x y' 0 y') + | S y' => y' - snd (nat_div_mod x y' 0 y') end. Infix "mod" := nat_mod : nat_scope. @@ -701,11 +701,8 @@ Proof. Defined. (** We can move a subtracted number to the right-hand side of an equation. *) -Definition nat_moveR_nV {n m} k : n = k + m -> n - k = m. -Proof. - rewrite nat_add_comm. - exact (fun p => (nat_moveL_nV _ p^)^). -Defined. +Definition nat_moveR_nV {n m} k : n = m + k -> n - k = m + := fun p => (nat_moveL_nV _ p^)^. (** Subtracting a successor is the predecessor of subtracting the original number. *) Definition nat_sub_succ_r n m : n - m.+1 = nat_pred (n - m). @@ -1532,6 +1529,7 @@ Proof. lhs nrapply nat_dist_sub_r. apply nat_moveR_nV. lhs nrapply q. + lhs nrapply nat_add_comm. exact (ap _ p^). Defined. @@ -1573,65 +1571,35 @@ Global Instance antisymmetric_divides : AntiSymmetric NatDivides (** ** Properties of division *) -Local Definition nat_divmod_unique_helper b q1 q2 r1 r2 - : r1 < b -> r2 < b -> q1 < q2 -> r2 < r1 -> b * q1 + r1 = b * q2 + r2 - -> q1 = q2 /\ r1 = r2. +Local Definition nat_div_mod_unique_helper b q1 q2 r1 r2 + : r1 < b -> r2 < b -> q1 < q2 -> b * q1 + r1 <> b * q2 + r2. Proof. - intros H1 H2 H3 H4 p. + intros H1 H2 H3 p. rewrite 2 (nat_add_comm (b * _)) in p. apply nat_moveL_nV in p. - rewrite nat_sub_l_add_r in p. - - rewrite <- nat_dist_sub_l in p. - apply nat_moveR_nV in p. - symmetry in p. - assert (H5 : b <= r1 - r2). - { rewrite <- p. - snrapply (leq_mul_r _ _ 0). - by apply equiv_lt_lt_sub. } - apply leq_iff_not_gt in H5. - contradiction H5. - rapply lt_moveR_nV. - - rapply nat_mul_l_monotone. + rewrite nat_sub_l_add_r in p; only 2: rapply nat_mul_l_monotone. + rewrite <- nat_dist_sub_l in p. + rewrite nat_add_comm in p. + apply nat_moveR_nV in p. + nrapply (snd (@leq_iff_not_gt b (r1 - r2))). + 2: exact (lt_leq_lt_trans _ H1). + rewrite p. + snrapply (leq_mul_r _ _ 0). + by apply equiv_lt_lt_sub. Defined. (** Quotients and remainders are uniquely determined. *) -Definition nat_divmod_unique d q1 q2 r1 r2 +Definition nat_div_mod_unique d q1 q2 r1 r2 : r1 < d -> r2 < d -> d * q1 + r1 = d * q2 + r2 -> q1 = q2 /\ r1 = r2. Proof. intros H1 H2 p. - destruct (nat_trichotomy q1 q2) as [[q | q] | q]; - destruct (nat_trichotomy r1 r2) as [[r | r] | r]. - - apply (nat_mul_l_strictly_monotone d H1) in q. - pose (nat_add_strictly_monotone q r). - rewrite p in l. - contradiction (lt_irrefl _ l). - - destruct r. - split; trivial. - apply isinj_nat_add_r in p. - apply (nat_mul_l_strictly_monotone d H1) in q. - rewrite p in q. - contradiction (lt_irrefl _ q). - - by apply (nat_divmod_unique_helper d). - - destruct q. - split; trivial. - by apply isinj_nat_add_l in p. - - by split. + destruct (nat_trichotomy q1 q2) as [[q | q] | q]. + - contradiction (nat_div_mod_unique_helper d q1 q2 r1 r2). - split; trivial. - destruct q. + destruct q. by apply isinj_nat_add_l in p. - - rapply (equiv_functor_prod (f:=inverse) (g:=inverse)). - by apply (nat_divmod_unique_helper d). - - destruct r. - split; trivial. - apply isinj_nat_add_r in p. - apply (nat_mul_l_strictly_monotone d H2) in q. - rewrite p in q. - contradiction (lt_irrefl _ q). - - apply (nat_mul_l_strictly_monotone d H2) in q. - pose (nat_add_strictly_monotone q r). - rewrite p in l. - contradiction (lt_irrefl _ l). + - by contradiction (nat_div_mod_unique_helper d q2 q1 r2 r1). Defined. (** Divisibility by a positive natural number is a hprop. *) @@ -1643,16 +1611,16 @@ Proof. rapply path_sigma_hprop. destruct H as [|n]; simpl. 1: exact ((nat_mul_one_r _)^ @ p @ q^ @ nat_mul_one_r _). - refine (fst (nat_divmod_unique n.+1 x y 0 0 _ _ _)). + refine (fst (nat_div_mod_unique n.+1 x y 0 0 _ _ _)). lhs nrapply nat_add_zero_r. rhs nrapply nat_add_zero_r. rewrite 2 (nat_mul_comm n.+1). exact (p @ q^). Defined. -(** The quotient-remainder form is invariant under [nat_divmod]. *) -Definition nat_divmod_spec_helper x y q u (H : u <= y) - : let (q', u') := nat_divmod x y q u in +(** This specifies the behaviour of [nat_div_mod_helper] when [u <= y]. *) +Definition nat_div_mod_helper_spec x y q u (H : u <= y) + : let (q', u') := nat_div_mod x y q u in x + y.+1 * q + (y - u) = y.+1 * q' + (y - u') /\ u' <= y. Proof. intros d r. @@ -1673,16 +1641,18 @@ Proof. Defined. (** Division and modulo can be put in quotient-remainder form. *) -Definition nat_divmod_spec x y : x = y * (x / y) + x mod y. +Definition nat_div_mod_spec x y : x = y * (x / y) + x mod y. Proof. destruct y as [|y]; only 1: reflexivity. - pose proof (p := fst (nat_divmod_spec_helper x y 0 y _)). + pose proof (p := fst (nat_div_mod_helper_spec x y 0 y _)). by rewrite nat_mul_zero_r, nat_sub_cancel, 2 nat_add_zero_r in p. Defined. -Definition nat_divmod_spec' x y : x - y * (x / y) = x mod y. +Definition nat_div_mod_spec' x y : x - y * (x / y) = x mod y. Proof. - apply nat_moveR_nV, nat_divmod_spec. + apply nat_moveR_nV. + rhs nrapply nat_add_comm. + apply nat_div_mod_spec. Defined. Definition nat_mod_lt_r' n m r : r < m -> n mod m < m. @@ -1699,17 +1669,17 @@ Global Instance nat_mod_lt_r n m : 0 < m -> n mod m < m (** [n] modulo [m] is less than or equal to [m]. *) Global Instance nat_mod_leq_l n m : n mod m <= n. Proof. - rewrite <- nat_divmod_spec'. + rewrite <- nat_div_mod_spec'. rapply leq_moveR_nV. Defined. (** Division is unique. *) Definition nat_div_unique x y q r (H : r < y) (p : y * q + r = x) : q = x / y - := fst (nat_divmod_unique y q (x / y) r (x mod y) _ _ (p @ nat_divmod_spec x y)). + := fst (nat_div_mod_unique y q (x / y) r (x mod y) _ _ (p @ nat_div_mod_spec x y)). (** Modulo is unique. *) Definition nat_mod_unique x y q r (H : r < y) (p : y * q + r = x) : r = x mod y - := snd (nat_divmod_unique y q (x / y) r (x mod y) _ _ (p @ nat_divmod_spec x y)). + := snd (nat_div_mod_unique y q (x / y) r (x mod y) _ _ (p @ nat_div_mod_spec x y)). (** [0] divided by any number is [0]. *) Definition nat_div_zero_l n : 0 / n = 0. @@ -1725,7 +1695,7 @@ Definition nat_div_one_r n : n / 1 = n. Proof. lhs_V nrapply nat_mul_one_l. lhs_V nrapply nat_add_zero_r. - symmetry; apply nat_divmod_spec. + symmetry; apply nat_div_mod_spec. Defined. (** [n] divided by [n] is [1]. *) @@ -1795,12 +1765,10 @@ Definition nat_mod_divides n m : 0 < m -> (m | n) -> n mod m = 0. Proof. intros H [x p]. destruct p. - lhs_V nrapply nat_divmod_spec'. + lhs_V nrapply nat_div_mod_spec'. pose (nat_div_cancel m H). rewrite nat_div_mul_cancel_r; only 2: exact _. - apply nat_moveR_nV. - rhs nrapply nat_add_zero_r. - apply nat_mul_comm. + apply nat_moveR_nV, nat_mul_comm. Defined. (** [n] modulo [n] is [0]. *) @@ -1857,7 +1825,7 @@ Proof. unfold nat_gcd; fold nat_gcd. set (n' := n.+1) in *. split; only 1: exact H2. - set (r := m mod n'); rewrite (nat_divmod_spec m n'); unfold r; clear r. + set (r := m mod n'); rewrite (nat_div_mod_spec m n'); unfold r; clear r. exact _. Defined. @@ -1876,7 +1844,7 @@ Proof. destruct n; only 1: exact _. unfold nat_gcd; fold nat_gcd. apply IHn; only 1,3: exact _. - rewrite (nat_divmod_spec m n.+1) in H2. + rewrite (nat_div_mod_spec m n.+1) in H2. apply nat_divides_add_r in H2; exact _. Defined. From 192a569d07022ca0827e51c328699ac2bdfb1d5a Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 8 Aug 2024 14:20:01 +0100 Subject: [PATCH 274/282] Nat: move division to own file Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 461 +------------------------------- theories/Spaces/Nat/Division.v | 462 +++++++++++++++++++++++++++++++++ 2 files changed, 465 insertions(+), 458 deletions(-) create mode 100644 theories/Spaces/Nat/Division.v diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index bbfc257f5f7..d8b308abc9d 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -82,7 +82,7 @@ Fixpoint nat_min n m := (** *** Euclidean division *) -(** This division takes time linear in `x` and is tail-recursive. In [nat_div_mod x y q u], [x + y.+1 * q + (y - u)] is the quantity being divided and [y] is the predecessor of the divisor. It will be called with [q] zero and [u] equal to [y], so that [x] is the quantity being divided. The return value is a pair [(q', u')] with [x + y.+1 * q + (y - u) = y.+1 * q' + (y - u')], as least when [u <= y], as shown in [nat_div_mod_spec_helper] below. *) +(** This division takes time linear in `x` and is tail-recursive. In [nat_div_mod x y q u], [x + y.+1 * q + (y - u)] is the quantity being divided and [y] is the predecessor of the divisor. It will be called with [q] zero and [u] equal to [y], so that [x] is the quantity being divided. The return value is a pair [(q', u')] with [x + y.+1 * q + (y - u) = y.+1 * q' + (y - u')], at least when [u <= y], as shown in [nat_div_mod_spec_helper] in Nat/Divison.v. *) Fixpoint nat_div_mod x y q u : nat * nat := match x with @@ -109,6 +109,8 @@ Definition nat_mod x y : nat := end. Infix "mod" := nat_mod : nat_scope. +(** For results about division and modulo, see Nat/Division.v. *) + (** *** Greatest common divisor *) (** We use Euclid algorithm, which is normally not structural, but Coq is now clever enough to accept this (behind modulo there is a subtraction, which now preserves being a subterm) *) @@ -1461,460 +1463,3 @@ Proof. - destruct p. by apply IH_strong. Defined. - -(** ** Divisibility *) - -(** We define divisibility as a relation between natural numbers. *) -Class NatDivides (n m : nat) : Type0 := nat_divides : {k : nat & k * n = m}. - -Notation "( n | m )" := (NatDivides n m) : nat_scope. - -(** Any number divides [0]. *) -Global Instance nat_divides_zero_r n : (n | 0) - := (0; idpath). - -(** [1] divides any number. *) -Global Instance nat_divides_one_l n : (1 | n) - := (n; nat_mul_one_r _). - -(** Any number divides itself. Divisibility is a reflexive relation. *) -Global Instance nat_divides_refl n : (n | n) - := (1; nat_mul_one_l _). - -Global Instance reflexive_nat_divides : Reflexive NatDivides := nat_divides_refl. - -(** Divisibility is transitive. *) -Definition nat_divides_trans {n m l} : (n | m) -> (m | l) -> (n | l). -Proof. - intros [x p] [y q]. - exists (y * x). - lhs_V nrapply nat_mul_assoc. - lhs nrapply (ap _ p). - exact q. -Defined. -Hint Immediate nat_divides_trans : typeclass_instances. - -Global Instance transitive_nat_divides : Transitive NatDivides := @nat_divides_trans. - -(** A left factor divides a product. *) -Global Instance nat_divides_mul_l' n m : (n | n * m) - := (m; nat_mul_comm _ _). - -(** A right factor divides a product. *) -Global Instance nat_divides_mul_r' n m : (m | n * m) - := (n; idpath). - -(** Divisibility of the product is implied by divisibility of the left factor. *) -Global Instance nat_divides_mul_l {n m} l : (n | m) -> (n | m * l) - := fun H => nat_divides_trans _ _. - -(** Divisibility of the product is implied by divisibility of the right factor. *) -Global Instance nat_divides_mul_r {n m} l : (n | m) -> (n | l * m) - := fun H => nat_divides_trans _ _. - -(** Divisibility of the sum is implied by divisibility of the summands. *) -Global Instance nat_divides_add n m l : (n | m) -> (n | l) -> (n | m + l). -Proof. - intros [x p] [y q]. - exists (x + y). - destruct p, q. - nrapply nat_dist_r. -Defined. - -(** If [n] divides a sum and the left summand, then [n] divides the right summand. *) -Definition nat_divides_add_r n m l : (n | m) -> (n | m + l) -> (n | l). -Proof. - intros [x p] [y q]. - exists (y - x). - lhs nrapply nat_dist_sub_r. - apply nat_moveR_nV. - lhs nrapply q. - lhs nrapply nat_add_comm. - exact (ap _ p^). -Defined. - -(** If [n] divides a sum and the right summand, then [n] divides the left summand. *) -Definition nat_divides_add_l n m l : (n | l) -> (n | m + l) -> (n | m). -Proof. - rewrite nat_add_comm; apply nat_divides_add_r. -Defined. - -(** Divisibility of the difference is implied by divisibility of the minuend and subtrahend. *) -Global Instance nat_divides_sub n m l : (n | m) -> (n | l) -> (n | m - l). -Proof. - intros [x p] [y q]. - exists (x - y). - destruct p, q. - nrapply nat_dist_sub_r. -Defined. - -(** Divisibility implies that the divisor is less than or equal to the dividend. *) -Definition leq_divides n m : 0 < m -> (n | m) -> n <= m. -Proof. - intros H1 [x p]. - destruct p, x. - 1: contradiction (not_lt_zero_r _ H1). - rapply (leq_mul_l _ _ 0). -Defined. - -(** Divisibility is antisymmetric *) -Definition nat_divides_antisym n m : (n | m) -> (m | n) -> n = m. -Proof. - intros H1 H2. - destruct m; only 1: exact (H2.2^ @ nat_mul_zero_r _). - destruct n; only 1: exact ((nat_mul_zero_r _)^ @ H1.2). - snrapply leq_antisym; nrapply leq_divides; exact _. -Defined. - -Global Instance antisymmetric_divides : AntiSymmetric NatDivides - := nat_divides_antisym. - -(** ** Properties of division *) - -Local Definition nat_div_mod_unique_helper b q1 q2 r1 r2 - : r1 < b -> r2 < b -> q1 < q2 -> b * q1 + r1 <> b * q2 + r2. -Proof. - intros H1 H2 H3 p. - rewrite 2 (nat_add_comm (b * _)) in p. - apply nat_moveL_nV in p. - rewrite nat_sub_l_add_r in p; only 2: rapply nat_mul_l_monotone. - rewrite <- nat_dist_sub_l in p. - rewrite nat_add_comm in p. - apply nat_moveR_nV in p. - nrapply (snd (@leq_iff_not_gt b (r1 - r2))). - 2: exact (lt_leq_lt_trans _ H1). - rewrite p. - snrapply (leq_mul_r _ _ 0). - by apply equiv_lt_lt_sub. -Defined. - -(** Quotients and remainders are uniquely determined. *) -Definition nat_div_mod_unique d q1 q2 r1 r2 - : r1 < d -> r2 < d -> d * q1 + r1 = d * q2 + r2 - -> q1 = q2 /\ r1 = r2. -Proof. - intros H1 H2 p. - destruct (nat_trichotomy q1 q2) as [[q | q] | q]. - - contradiction (nat_div_mod_unique_helper d q1 q2 r1 r2). - - split; trivial. - destruct q. - by apply isinj_nat_add_l in p. - - by contradiction (nat_div_mod_unique_helper d q2 q1 r2 r1). -Defined. - -(** Divisibility by a positive natural number is a hprop. *) -Global Instance ishprop_nat_divides n m : 0 < n -> IsHProp (n | m). -Proof. - intros H. - apply hprop_allpath. - intros [x p] [y q]. - rapply path_sigma_hprop. - destruct H as [|n]; simpl. - 1: exact ((nat_mul_one_r _)^ @ p @ q^ @ nat_mul_one_r _). - refine (fst (nat_div_mod_unique n.+1 x y 0 0 _ _ _)). - lhs nrapply nat_add_zero_r. - rhs nrapply nat_add_zero_r. - rewrite 2 (nat_mul_comm n.+1). - exact (p @ q^). -Defined. - -(** This specifies the behaviour of [nat_div_mod_helper] when [u <= y]. *) -Definition nat_div_mod_helper_spec x y q u (H : u <= y) - : let (q', u') := nat_div_mod x y q u in - x + y.+1 * q + (y - u) = y.+1 * q' + (y - u') /\ u' <= y. -Proof. - intros d r. - induction x as [|x IHx] in y, q, u, H, d, r |- *; only 1: by split. - destruct u as [|u]. - - destruct (IHx y q.+1 y _) as [p H']. - split; trivial. - rewrite <- p, nat_sub_zero_r, nat_sub_cancel, nat_add_zero_r. - simpl. - by rewrite nat_add_succ_r, <- 2 nat_add_assoc, nat_mul_succ_r. - - destruct (IHx y q u _) as [p H']. - split; trivial. - rewrite <- p, 2 nat_add_succ_l, <- nat_add_succ_r. - snrapply ap. - rewrite nat_sub_succ_r. - apply nat_succ_pred. - rapply lt_moveL_nV. -Defined. - -(** Division and modulo can be put in quotient-remainder form. *) -Definition nat_div_mod_spec x y : x = y * (x / y) + x mod y. -Proof. - destruct y as [|y]; only 1: reflexivity. - pose proof (p := fst (nat_div_mod_helper_spec x y 0 y _)). - by rewrite nat_mul_zero_r, nat_sub_cancel, 2 nat_add_zero_r in p. -Defined. - -Definition nat_div_mod_spec' x y : x - y * (x / y) = x mod y. -Proof. - apply nat_moveR_nV. - rhs nrapply nat_add_comm. - apply nat_div_mod_spec. -Defined. - -Definition nat_mod_lt_r' n m r : r < m -> n mod m < m. -Proof. - intros H; destruct H; only 1: exact _. - rapply (lt_leq_lt_trans (m:=m)). -Defined. -Hint Immediate nat_mod_lt_r' : typeclass_instances. - -(** [n] modulo [m] is less than [m]. *) -Global Instance nat_mod_lt_r n m : 0 < m -> n mod m < m - := nat_mod_lt_r' n m 0. - -(** [n] modulo [m] is less than or equal to [m]. *) -Global Instance nat_mod_leq_l n m : n mod m <= n. -Proof. - rewrite <- nat_div_mod_spec'. - rapply leq_moveR_nV. -Defined. - -(** Division is unique. *) -Definition nat_div_unique x y q r (H : r < y) (p : y * q + r = x) : q = x / y - := fst (nat_div_mod_unique y q (x / y) r (x mod y) _ _ (p @ nat_div_mod_spec x y)). - -(** Modulo is unique. *) -Definition nat_mod_unique x y q r (H : r < y) (p : y * q + r = x) : r = x mod y - := snd (nat_div_mod_unique y q (x / y) r (x mod y) _ _ (p @ nat_div_mod_spec x y)). - -(** [0] divided by any number is [0]. *) -Definition nat_div_zero_l n : 0 / n = 0. -Proof. - by induction n. -Defined. - -(** [n] divided by [0] is [0] by convention. *) -Definition nat_div_zero_r n : n / 0 = 0 := idpath. - -(** [n] divided by [1] is [n]. *) -Definition nat_div_one_r n : n / 1 = n. -Proof. - lhs_V nrapply nat_mul_one_l. - lhs_V nrapply nat_add_zero_r. - symmetry; apply nat_div_mod_spec. -Defined. - -(** [n] divided by [n] is [1]. *) -Definition nat_div_cancel n : 0 < n -> n / n = 1. -Proof. - intros [|m _]; trivial; symmetry. - nrapply (nat_div_unique _ _ _ 0); only 1: exact _. - lhs nrapply nat_add_zero_r. - nrapply nat_mul_one_r. -Defined. - -(** [n * m] divided by [n] is [m]. *) -Definition nat_div_mul_cancel_l n m : 0 < n -> (n * m) / n = m. -Proof. - intros H. - symmetry. - nrapply (nat_div_unique _ _ _ _ H). - apply nat_add_zero_r. -Defined. - -(** [n * m] divided by [n] is [m]. *) -Definition nat_div_mul_cancel_r n m : 0 < m -> (n * m) / m = n. -Proof. - rewrite nat_mul_comm. - apply nat_div_mul_cancel_l. -Defined. - -(** When [d] divides [n] and [m], division distributes over addition. *) -Definition nat_div_dist n m d - : 0 < d -> (d | n) -> (d | m) -> (n + m) / d = n / d + m / d. -Proof. - intros H1 [x p] [y q]. - destruct p, q. - rewrite <- nat_dist_r. - by rewrite 3 nat_div_mul_cancel_r. -Defined. - -(** [0] modulo [n] is [0]. *) -Definition nat_mod_zero_l n : 0 mod n = 0. -Proof. - induction n; trivial. - apply nat_sub_cancel. -Defined. - -(** [n] modulo [0] is [n]. *) -Definition nat_mod_zero_r n : n mod 0 = n := idpath. - -(** TODO: generalise for all small n *) -Definition nat_mod_one_l n : 1 < n -> 1 mod n = 1. -Proof. - intros H. - destruct H; trivial. - destruct m. - 1: contradiction (not_lt_zero_r _ H). - cbn; clear H. - by induction m. -Defined. - -(** [n] modulo [1] is [0]. *) -Definition nat_mod_one_r n : n mod 1 = 0. -Proof. - by induction n. -Defined. - -(** If [m] divides [n], then [n mod m = 0]. *) -Definition nat_mod_divides n m : 0 < m -> (m | n) -> n mod m = 0. -Proof. - intros H [x p]. - destruct p. - lhs_V nrapply nat_div_mod_spec'. - pose (nat_div_cancel m H). - rewrite nat_div_mul_cancel_r; only 2: exact _. - apply nat_moveR_nV, nat_mul_comm. -Defined. - -(** [n] modulo [n] is [0]. *) -Definition nat_mod_cancel n : n mod n = 0. -Proof. - destruct n; trivial. - symmetry. - snrapply (nat_mod_unique _ _ 1); only 1: exact _. - lhs nrapply nat_add_zero_r. - nrapply nat_mul_one_r. -Defined. - -(** ** Greatest Common Divisor *) - -(** The greatest common divisor of [0] and a number is the number itself. *) -Definition nat_gcd_zero_l n : nat_gcd 0 n = n := idpath. - -(** The greatest common divisor of a number and [0] is the number itself. *) -Definition nat_gcd_zero_r n : nat_gcd n 0 = n. -Proof. - induction n; simpl; only 1: reflexivity. - by rewrite nat_sub_cancel. -Defined. - -(** The greatest common divisor of [1] and any number is [1]. *) -Definition nat_gcd_one_l n : nat_gcd 1 n = 1 := idpath. - -(** The greatest common divisor of any number and [1] is [1]. *) -Definition nat_gcd_one_r n : nat_gcd n 1 = 1. -Proof. - destruct n; trivial. - simpl. - destruct n; trivial. - rewrite nat_sub_succ_l; only 2: exact _. - by rewrite nat_sub_cancel. -Defined. - -(** Idempotency. *) -Definition nat_gcd_idem n : nat_gcd n n = n. -Proof. - induction n. - 1: reflexivity. - unfold nat_gcd; fold nat_gcd. - by rewrite nat_mod_cancel. -Defined. - -(** We can prove that the greatest common divisor of [n] and [m] divides both [n] and [m]. This proof requires strong induction. *) -Definition nat_divides_l_gcd n m : (nat_gcd n m | n) /\ (nat_gcd n m | m). -Proof. - revert n m; snrapply nat_ind_strong; intros n IHn m. - destruct n. - 1: split; exact _. - destruct (IHn (m mod n.+1) _ n.+1) as [H1 H2]. - unfold nat_gcd; fold nat_gcd. - set (n' := n.+1) in *. - split; only 1: exact H2. - set (r := m mod n'); rewrite (nat_div_mod_spec m n'); unfold r; clear r. - exact _. -Defined. - -(** The greatest common divisor of [n] and [m] divides [n]. *) -Global Instance nat_divides_l_gcd_l n m : (nat_gcd n m | n) - := fst (nat_divides_l_gcd n m). - -(** The greatest common divisor of [n] and [m] divides [m]. *) -Global Instance divides_l_nat_gcd_r n m : (nat_gcd n m | m) - := snd (nat_divides_l_gcd n m). - -(** We can prove that any common divisor of [n] and [m] divides the greatest common divisor of [n] and [m]. It is in that sense the greatest. *) -Global Instance nat_divides_r_gcd n m p : (p | n) -> (p | m) -> (p | nat_gcd n m). -Proof. - revert n m p; snrapply nat_ind_strong; intros n IHn m p H1 H2. - destruct n; only 1: exact _. - unfold nat_gcd; fold nat_gcd. - apply IHn; only 1,3: exact _. - rewrite (nat_div_mod_spec m n.+1) in H2. - apply nat_divides_add_r in H2; exact _. -Defined. - -Definition nat_divides_r_iff_divides_r_gcd n m p - : (p | n) * (p | m) <-> (p | nat_gcd n m). -Proof. - split; [intros [H1 H2] | intros H; split]; exact _. -Defined. - -(** If [p] is divisible by all common divisors of [n] and [m], and [p] is also a common divisor, then it must necesserily be equal to the greatest common divisor. *) -Definition nat_gcd_unique n m p - (H : forall q, (q | n) -> (q | m) -> (q | p)) - : (p | n) -> (p | m) -> nat_gcd n m = p. -Proof. - intros H1 H2. - rapply nat_divides_antisym. -Defined. - -(** As a corollary of uniquness, we get that the greatest common divisor operation is commutative. *) -Definition nat_gcd_comm n m : nat_gcd n m = nat_gcd m n. -Proof. - rapply nat_gcd_unique. -Defined. - -(** [nat_gcd] is associative. *) -Definition nat_gcd_assoc n m k : nat_gcd n (nat_gcd m k) = nat_gcd (nat_gcd n m) k. -Proof. - nrapply nat_gcd_unique. - - intros q H1 H2. - rapply nat_divides_r_gcd. - - rapply (nat_divides_trans (nat_divides_l_gcd_l _ _)). - - apply nat_divides_r_gcd; rapply nat_divides_trans. -Defined. - -(** If [nat_gcd n m] is [0], then [n] must also be [0]. *) -Definition nat_gcd_is_zero_l n m : nat_gcd n m = 0 -> n = 0. -Proof. - intros H. - generalize (nat_divides_l_gcd_l n m). - rewrite H. - intros [x p]. - exact (p^ @ nat_mul_zero_r _). -Defined. - -(** If [nat_gcd n m] is [0], then [m] must also be [0]. *) -Definition nat_gcd_is_zero_r n m : nat_gcd n m = 0 -> m = 0. -Proof. - rewrite nat_gcd_comm. - apply nat_gcd_is_zero_l. -Defined. - -(** [nat_gcd n m] is [0] if and only if both [n] and [m] are [0]. *) -Definition nat_gcd_zero_iff_zero n m : nat_gcd n m = 0 <-> n = 0 /\ m = 0. -Proof. - split. - - split. - + by apply (nat_gcd_is_zero_l _ m). - + by apply (nat_gcd_is_zero_r n). - - intros [-> ->]. - reflexivity. -Defined. - -(** [nat_gcd] is positive for positive inputs. *) -Global Instance nat_gcd_pos n m : 0 < n -> 0 < m -> 0 < nat_gcd n m. -Proof. - intros H1 H2. - apply lt_iff_not_geq. - intros H3; hnf in H3. - apply path_zero_leq_zero_r in H3. - apply nat_gcd_zero_iff_zero in H3. - destruct H3 as [->]. - contradiction (not_lt_zero_r _ H1). -Defined. diff --git a/theories/Spaces/Nat/Division.v b/theories/Spaces/Nat/Division.v new file mode 100644 index 00000000000..22a8f6f24f2 --- /dev/null +++ b/theories/Spaces/Nat/Division.v @@ -0,0 +1,462 @@ +Require Import Basics.Overture Basics.Tactics Basics.PathGroupoids + Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat Basics.Classes + Types.Sum Types.Sigma Spaces.Nat.Core. + +Local Open Scope nat_scope. + +(** ** Divisibility *) + +(** We define divisibility as a relation between natural numbers. *) +Class NatDivides (n m : nat) : Type0 := nat_divides : {k : nat & k * n = m}. + +Notation "( n | m )" := (NatDivides n m) : nat_scope. + +(** Any number divides [0]. *) +Global Instance nat_divides_zero_r n : (n | 0) + := (0; idpath). + +(** [1] divides any number. *) +Global Instance nat_divides_one_l n : (1 | n) + := (n; nat_mul_one_r _). + +(** Any number divides itself. Divisibility is a reflexive relation. *) +Global Instance nat_divides_refl n : (n | n) + := (1; nat_mul_one_l _). + +Global Instance reflexive_nat_divides : Reflexive NatDivides := nat_divides_refl. + +(** Divisibility is transitive. *) +Definition nat_divides_trans {n m l} : (n | m) -> (m | l) -> (n | l). +Proof. + intros [x p] [y q]. + exists (y * x). + lhs_V nrapply nat_mul_assoc. + lhs nrapply (ap _ p). + exact q. +Defined. +Hint Immediate nat_divides_trans : typeclass_instances. + +Global Instance transitive_nat_divides : Transitive NatDivides := @nat_divides_trans. + +(** A left factor divides a product. *) +Global Instance nat_divides_mul_l' n m : (n | n * m) + := (m; nat_mul_comm _ _). + +(** A right factor divides a product. *) +Global Instance nat_divides_mul_r' n m : (m | n * m) + := (n; idpath). + +(** Divisibility of the product is implied by divisibility of the left factor. *) +Global Instance nat_divides_mul_l {n m} l : (n | m) -> (n | m * l) + := fun H => nat_divides_trans _ _. + +(** Divisibility of the product is implied by divisibility of the right factor. *) +Global Instance nat_divides_mul_r {n m} l : (n | m) -> (n | l * m) + := fun H => nat_divides_trans _ _. + +(** Divisibility of the sum is implied by divisibility of the summands. *) +Global Instance nat_divides_add n m l : (n | m) -> (n | l) -> (n | m + l). +Proof. + intros [x p] [y q]. + exists (x + y). + destruct p, q. + nrapply nat_dist_r. +Defined. + +(** If [n] divides a sum and the left summand, then [n] divides the right summand. *) +Definition nat_divides_add_r n m l : (n | m) -> (n | m + l) -> (n | l). +Proof. + intros [x p] [y q]. + exists (y - x). + lhs nrapply nat_dist_sub_r. + apply nat_moveR_nV. + lhs nrapply q. + lhs nrapply nat_add_comm. + exact (ap _ p^). +Defined. + +(** If [n] divides a sum and the right summand, then [n] divides the left summand. *) +Definition nat_divides_add_l n m l : (n | l) -> (n | m + l) -> (n | m). +Proof. + rewrite nat_add_comm; apply nat_divides_add_r. +Defined. + +(** Divisibility of the difference is implied by divisibility of the minuend and subtrahend. *) +Global Instance nat_divides_sub n m l : (n | m) -> (n | l) -> (n | m - l). +Proof. + intros [x p] [y q]. + exists (x - y). + destruct p, q. + nrapply nat_dist_sub_r. +Defined. + +(** Divisibility implies that the divisor is less than or equal to the dividend. *) +Definition leq_divides n m : 0 < m -> (n | m) -> n <= m. +Proof. + intros H1 [x p]. + destruct p, x. + 1: contradiction (not_lt_zero_r _ H1). + rapply (leq_mul_l _ _ 0). +Defined. + +(** Divisibility is antisymmetric *) +Definition nat_divides_antisym n m : (n | m) -> (m | n) -> n = m. +Proof. + intros H1 H2. + destruct m; only 1: exact (H2.2^ @ nat_mul_zero_r _). + destruct n; only 1: exact ((nat_mul_zero_r _)^ @ H1.2). + snrapply leq_antisym; nrapply leq_divides; exact _. +Defined. + +Global Instance antisymmetric_divides : AntiSymmetric NatDivides + := nat_divides_antisym. + +(** ** Properties of division *) + +Local Definition nat_div_mod_unique_helper b q1 q2 r1 r2 + : r1 < b -> r2 < b -> q1 < q2 -> b * q1 + r1 <> b * q2 + r2. +Proof. + intros H1 H2 H3 p. + rewrite 2 (nat_add_comm (b * _)) in p. + apply nat_moveL_nV in p. + rewrite nat_sub_l_add_r in p; only 2: rapply nat_mul_l_monotone. + rewrite <- nat_dist_sub_l in p. + rewrite nat_add_comm in p. + apply nat_moveR_nV in p. + nrapply (snd (@leq_iff_not_gt b (r1 - r2))). + 2: exact (lt_leq_lt_trans _ H1). + rewrite p. + snrapply (leq_mul_r _ _ 0). + by apply equiv_lt_lt_sub. +Defined. + +(** Quotients and remainders are uniquely determined. *) +Definition nat_div_mod_unique d q1 q2 r1 r2 + : r1 < d -> r2 < d -> d * q1 + r1 = d * q2 + r2 + -> q1 = q2 /\ r1 = r2. +Proof. + intros H1 H2 p. + destruct (nat_trichotomy q1 q2) as [[q | q] | q]. + - contradiction (nat_div_mod_unique_helper d q1 q2 r1 r2). + - split; trivial. + destruct q. + by apply isinj_nat_add_l in p. + - by contradiction (nat_div_mod_unique_helper d q2 q1 r2 r1). +Defined. + +(** Divisibility by a positive natural number is a hprop. *) +Global Instance ishprop_nat_divides n m : 0 < n -> IsHProp (n | m). +Proof. + intros H. + apply hprop_allpath. + intros [x p] [y q]. + rapply path_sigma_hprop. + destruct H as [|n]; simpl. + 1: exact ((nat_mul_one_r _)^ @ p @ q^ @ nat_mul_one_r _). + refine (fst (nat_div_mod_unique n.+1 x y 0 0 _ _ _)). + lhs nrapply nat_add_zero_r. + rhs nrapply nat_add_zero_r. + rewrite 2 (nat_mul_comm n.+1). + exact (p @ q^). +Defined. + +(** This specifies the behaviour of [nat_div_mod_helper] when [u <= y]. *) +Definition nat_div_mod_helper_spec x y q u (H : u <= y) + : let (q', u') := nat_div_mod x y q u in + x + y.+1 * q + (y - u) = y.+1 * q' + (y - u') /\ u' <= y. +Proof. + intros d r. + induction x as [|x IHx] in y, q, u, H, d, r |- *; only 1: by split. + destruct u as [|u]. + - destruct (IHx y q.+1 y _) as [p H']. + split; trivial. + rewrite <- p, nat_sub_zero_r, nat_sub_cancel, nat_add_zero_r. + simpl. + by rewrite nat_add_succ_r, <- 2 nat_add_assoc, nat_mul_succ_r. + - destruct (IHx y q u _) as [p H']. + split; trivial. + rewrite <- p, 2 nat_add_succ_l, <- nat_add_succ_r. + snrapply ap. + rewrite nat_sub_succ_r. + apply nat_succ_pred. + rapply lt_moveL_nV. +Defined. + +(** Division and modulo can be put in quotient-remainder form. *) +Definition nat_div_mod_spec x y : x = y * (x / y) + x mod y. +Proof. + destruct y as [|y]; only 1: reflexivity. + pose proof (p := fst (nat_div_mod_helper_spec x y 0 y _)). + by rewrite nat_mul_zero_r, nat_sub_cancel, 2 nat_add_zero_r in p. +Defined. + +Definition nat_div_mod_spec' x y : x - y * (x / y) = x mod y. +Proof. + apply nat_moveR_nV. + rhs nrapply nat_add_comm. + apply nat_div_mod_spec. +Defined. + +Definition nat_mod_lt_r' n m r : r < m -> n mod m < m. +Proof. + intros H; destruct H; only 1: exact _. + rapply (lt_leq_lt_trans (m:=m)). +Defined. +Hint Immediate nat_mod_lt_r' : typeclass_instances. + +(** [n] modulo [m] is less than [m]. *) +Global Instance nat_mod_lt_r n m : 0 < m -> n mod m < m + := nat_mod_lt_r' n m 0. + +(** [n] modulo [m] is less than or equal to [m]. *) +Global Instance nat_mod_leq_l n m : n mod m <= n. +Proof. + rewrite <- nat_div_mod_spec'. + rapply leq_moveR_nV. +Defined. + +(** Division is unique. *) +Definition nat_div_unique x y q r (H : r < y) (p : y * q + r = x) : q = x / y + := fst (nat_div_mod_unique y q (x / y) r (x mod y) _ _ (p @ nat_div_mod_spec x y)). + +(** Modulo is unique. *) +Definition nat_mod_unique x y q r (H : r < y) (p : y * q + r = x) : r = x mod y + := snd (nat_div_mod_unique y q (x / y) r (x mod y) _ _ (p @ nat_div_mod_spec x y)). + +(** [0] divided by any number is [0]. *) +Definition nat_div_zero_l n : 0 / n = 0. +Proof. + by induction n. +Defined. + +(** [n] divided by [0] is [0] by convention. *) +Definition nat_div_zero_r n : n / 0 = 0 := idpath. + +(** [n] divided by [1] is [n]. *) +Definition nat_div_one_r n : n / 1 = n. +Proof. + lhs_V nrapply nat_mul_one_l. + lhs_V nrapply nat_add_zero_r. + symmetry; apply nat_div_mod_spec. +Defined. + +(** [n] divided by [n] is [1]. *) +Definition nat_div_cancel n : 0 < n -> n / n = 1. +Proof. + intros [|m _]; trivial; symmetry. + nrapply (nat_div_unique _ _ _ 0); only 1: exact _. + lhs nrapply nat_add_zero_r. + nrapply nat_mul_one_r. +Defined. + +(** [n * m] divided by [n] is [m]. *) +Definition nat_div_mul_cancel_l n m : 0 < n -> (n * m) / n = m. +Proof. + intros H. + symmetry. + nrapply (nat_div_unique _ _ _ _ H). + apply nat_add_zero_r. +Defined. + +(** [n * m] divided by [n] is [m]. *) +Definition nat_div_mul_cancel_r n m : 0 < m -> (n * m) / m = n. +Proof. + rewrite nat_mul_comm. + apply nat_div_mul_cancel_l. +Defined. + +(** When [d] divides [n] and [m], division distributes over addition. *) +Definition nat_div_dist n m d + : 0 < d -> (d | n) -> (d | m) -> (n + m) / d = n / d + m / d. +Proof. + intros H1 [x p] [y q]. + destruct p, q. + rewrite <- nat_dist_r. + by rewrite 3 nat_div_mul_cancel_r. +Defined. + +(** [0] modulo [n] is [0]. *) +Definition nat_mod_zero_l n : 0 mod n = 0. +Proof. + induction n; trivial. + apply nat_sub_cancel. +Defined. + +(** [n] modulo [0] is [n]. *) +Definition nat_mod_zero_r n : n mod 0 = n := idpath. + +(** TODO: generalise for all small n *) +Definition nat_mod_one_l n : 1 < n -> 1 mod n = 1. +Proof. + intros H. + destruct H; trivial. + destruct m. + 1: contradiction (not_lt_zero_r _ H). + cbn; clear H. + by induction m. +Defined. + +(** [n] modulo [1] is [0]. *) +Definition nat_mod_one_r n : n mod 1 = 0. +Proof. + by induction n. +Defined. + +(** If [m] divides [n], then [n mod m = 0]. *) +Definition nat_mod_divides n m : 0 < m -> (m | n) -> n mod m = 0. +Proof. + intros H [x p]. + destruct p. + lhs_V nrapply nat_div_mod_spec'. + pose (nat_div_cancel m H). + rewrite nat_div_mul_cancel_r; only 2: exact _. + apply nat_moveR_nV, nat_mul_comm. +Defined. + +(** [n] modulo [n] is [0]. *) +Definition nat_mod_cancel n : n mod n = 0. +Proof. + destruct n; trivial. + symmetry. + snrapply (nat_mod_unique _ _ 1); only 1: exact _. + lhs nrapply nat_add_zero_r. + nrapply nat_mul_one_r. +Defined. + +(** ** Greatest Common Divisor *) + +(** The greatest common divisor of [0] and a number is the number itself. *) +Definition nat_gcd_zero_l n : nat_gcd 0 n = n := idpath. + +(** The greatest common divisor of a number and [0] is the number itself. *) +Definition nat_gcd_zero_r n : nat_gcd n 0 = n. +Proof. + induction n; simpl; only 1: reflexivity. + by rewrite nat_sub_cancel. +Defined. + +(** The greatest common divisor of [1] and any number is [1]. *) +Definition nat_gcd_one_l n : nat_gcd 1 n = 1 := idpath. + +(** The greatest common divisor of any number and [1] is [1]. *) +Definition nat_gcd_one_r n : nat_gcd n 1 = 1. +Proof. + destruct n; trivial. + simpl. + destruct n; trivial. + rewrite nat_sub_succ_l; only 2: exact _. + by rewrite nat_sub_cancel. +Defined. + +(** Idempotency. *) +Definition nat_gcd_idem n : nat_gcd n n = n. +Proof. + induction n. + 1: reflexivity. + unfold nat_gcd; fold nat_gcd. + by rewrite nat_mod_cancel. +Defined. + +(** We can prove that the greatest common divisor of [n] and [m] divides both [n] and [m]. This proof requires strong induction. *) +Definition nat_divides_l_gcd n m : (nat_gcd n m | n) /\ (nat_gcd n m | m). +Proof. + revert n m; snrapply nat_ind_strong; intros n IHn m. + destruct n. + 1: split; exact _. + destruct (IHn (m mod n.+1) _ n.+1) as [H1 H2]. + unfold nat_gcd; fold nat_gcd. + set (n' := n.+1) in *. + split; only 1: exact H2. + set (r := m mod n'); rewrite (nat_div_mod_spec m n'); unfold r; clear r. + exact _. +Defined. + +(** The greatest common divisor of [n] and [m] divides [n]. *) +Global Instance nat_divides_l_gcd_l n m : (nat_gcd n m | n) + := fst (nat_divides_l_gcd n m). + +(** The greatest common divisor of [n] and [m] divides [m]. *) +Global Instance divides_l_nat_gcd_r n m : (nat_gcd n m | m) + := snd (nat_divides_l_gcd n m). + +(** We can prove that any common divisor of [n] and [m] divides the greatest common divisor of [n] and [m]. It is in that sense the greatest. *) +Global Instance nat_divides_r_gcd n m p : (p | n) -> (p | m) -> (p | nat_gcd n m). +Proof. + revert n m p; snrapply nat_ind_strong; intros n IHn m p H1 H2. + destruct n; only 1: exact _. + unfold nat_gcd; fold nat_gcd. + apply IHn; only 1,3: exact _. + rewrite (nat_div_mod_spec m n.+1) in H2. + apply nat_divides_add_r in H2; exact _. +Defined. + +Definition nat_divides_r_iff_divides_r_gcd n m p + : (p | n) * (p | m) <-> (p | nat_gcd n m). +Proof. + split; [intros [H1 H2] | intros H; split]; exact _. +Defined. + +(** If [p] is divisible by all common divisors of [n] and [m], and [p] is also a common divisor, then it must necesserily be equal to the greatest common divisor. *) +Definition nat_gcd_unique n m p + (H : forall q, (q | n) -> (q | m) -> (q | p)) + : (p | n) -> (p | m) -> nat_gcd n m = p. +Proof. + intros H1 H2. + rapply nat_divides_antisym. +Defined. + +(** As a corollary of uniquness, we get that the greatest common divisor operation is commutative. *) +Definition nat_gcd_comm n m : nat_gcd n m = nat_gcd m n. +Proof. + rapply nat_gcd_unique. +Defined. + +(** [nat_gcd] is associative. *) +Definition nat_gcd_assoc n m k : nat_gcd n (nat_gcd m k) = nat_gcd (nat_gcd n m) k. +Proof. + nrapply nat_gcd_unique. + - intros q H1 H2. + rapply nat_divides_r_gcd. + - rapply (nat_divides_trans (nat_divides_l_gcd_l _ _)). + - apply nat_divides_r_gcd; rapply nat_divides_trans. +Defined. + +(** If [nat_gcd n m] is [0], then [n] must also be [0]. *) +Definition nat_gcd_is_zero_l n m : nat_gcd n m = 0 -> n = 0. +Proof. + intros H. + generalize (nat_divides_l_gcd_l n m). + rewrite H. + intros [x p]. + exact (p^ @ nat_mul_zero_r _). +Defined. + +(** If [nat_gcd n m] is [0], then [m] must also be [0]. *) +Definition nat_gcd_is_zero_r n m : nat_gcd n m = 0 -> m = 0. +Proof. + rewrite nat_gcd_comm. + apply nat_gcd_is_zero_l. +Defined. + +(** [nat_gcd n m] is [0] if and only if both [n] and [m] are [0]. *) +Definition nat_gcd_zero_iff_zero n m : nat_gcd n m = 0 <-> n = 0 /\ m = 0. +Proof. + split. + - split. + + by apply (nat_gcd_is_zero_l _ m). + + by apply (nat_gcd_is_zero_r n). + - intros [-> ->]. + reflexivity. +Defined. + +(** [nat_gcd] is positive for positive inputs. *) +Global Instance nat_gcd_pos n m : 0 < n -> 0 < m -> 0 < nat_gcd n m. +Proof. + intros H1 H2. + apply lt_iff_not_geq. + intros H3; hnf in H3. + apply path_zero_leq_zero_r in H3. + apply nat_gcd_zero_iff_zero in H3. + destruct H3 as [->]. + contradiction (not_lt_zero_r _ H1). +Defined. From 418a2fa470d013dacff9d79ca65af73ef7feb2fa Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 8 Aug 2024 17:36:07 +0100 Subject: [PATCH 275/282] Nat: add Bezout's identity Signed-off-by: Ali Caglayan --- theories/Spaces/Nat/Core.v | 15 +++- theories/Spaces/Nat/Division.v | 150 ++++++++++++++++++++++++++++++++- 2 files changed, 160 insertions(+), 5 deletions(-) diff --git a/theories/Spaces/Nat/Core.v b/theories/Spaces/Nat/Core.v index d8b308abc9d..e2cb9dffb45 100644 --- a/theories/Spaces/Nat/Core.v +++ b/theories/Spaces/Nat/Core.v @@ -1,6 +1,6 @@ Require Import Basics.Overture Basics.Tactics Basics.PathGroupoids - Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat Basics.Classes - Types.Sum Types.Sigma. + Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat + Basics.Classes Types.Prod Types.Sum Types.Sigma. Export Basics.Nat. Local Set Universe Minimization ToSet. @@ -329,6 +329,17 @@ Proof. nrapply nat_add_comm. Defined. +(** A sum being zero is equivalent to both summands being zero. *) +Definition equiv_nat_add_zero n m : n = 0 /\ m = 0 <~> n + m = 0. +Proof. + srapply equiv_iff_hprop. + - intros [-> ->]; reflexivity. + - destruct n. + + by split. + + intros H; symmetry in H. + by apply neq_nat_zero_succ in H. +Defined. + (** ** Properties of multiplication *) (** Multiplication by [0] on the left is [0]. *) diff --git a/theories/Spaces/Nat/Division.v b/theories/Spaces/Nat/Division.v index 22a8f6f24f2..4331e58fa47 100644 --- a/theories/Spaces/Nat/Division.v +++ b/theories/Spaces/Nat/Division.v @@ -1,9 +1,10 @@ -Require Import Basics.Overture Basics.Tactics Basics.PathGroupoids - Basics.Decidable Basics.Trunc Basics.Equivalences Basics.Nat Basics.Classes - Types.Sum Types.Sigma Spaces.Nat.Core. +Require Import Basics.Overture Basics.Tactics Basics.Trunc Basics.Classes + Basics.PathGroupoids Types.Sigma Spaces.Nat.Core. Local Open Scope nat_scope. +(** * Division of natural numbers *) + (** ** Divisibility *) (** We define divisibility as a relation between natural numbers. *) @@ -460,3 +461,146 @@ Proof. destruct H3 as [->]. contradiction (not_lt_zero_r _ H1). Defined. + +Definition nat_gcd_l_add_r_mul n m k : nat_gcd (n + k * m) m = nat_gcd n m. +Proof. + rapply nat_gcd_unique. + intros q H1 H2. + rapply nat_divides_r_gcd. + rapply (nat_divides_add_l _ _ (k * m)). +Defined. + +Definition nat_gcd_r_add_r_mul n m k : nat_gcd n (m + k * n) = nat_gcd n m. +Proof. + lhs nrapply nat_gcd_comm. + rhs nrapply nat_gcd_comm. + nrapply nat_gcd_l_add_r_mul. +Defined. + +Definition nat_gcd_l_add_r n m : nat_gcd (n + m) m = nat_gcd n m. +Proof. + rhs_V nrapply (nat_gcd_l_add_r_mul n m 1). + by rewrite nat_mul_one_l. +Defined. + +Definition nat_gcd_r_add_r n m : nat_gcd n (m + n) = nat_gcd n m. +Proof. + lhs nrapply nat_gcd_comm. + rhs nrapply nat_gcd_comm. + nrapply nat_gcd_l_add_r. +Defined. + +Definition nat_gcd_l_sub n m : m <= n -> nat_gcd (n - m) m = nat_gcd n m. +Proof. + intros H. + lhs_V nrapply nat_gcd_l_add_r. + by rewrite (nat_add_sub_l_cancel H). +Defined. + +Definition nat_gcd_r_sub n m : n <= m -> nat_gcd n (m - n) = nat_gcd n m. +Proof. + intros H. + lhs nrapply nat_gcd_comm. + rhs nrapply nat_gcd_comm. + rapply nat_gcd_l_sub. +Defined. + +(** ** Bezout's Identity *) + +(** Bezout's identity states that for any two numbers [n] and [m], their greatest common divisor can be written as a linear combination of [n] and [m]. This is easy to state for the integers, however since we are working with the natural numbers, we need to be more careful. This is why we write the linear combination as [a * n = d + b * m] rather than the usual [a * n + b * m = d]. *) + +(** We define a predicate for triples of integers satisfying Bezout's identity. *) +Definition NatBezout n m d : Type0 + := exists a b, a * n = d + b * m. +Existing Class NatBezout. + +Global Instance nat_bezout_refl_l n k : NatBezout n k n. +Proof. + by exists 1, 0. +Defined. + +(** If [a * n = 1 + b * m], then the gcd of [n] and [m] is [1]. *) +Definition nat_bezout_coprime n m : NatBezout n m 1 -> nat_gcd n m = 1. +Proof. + intros [a [b p]]. + rapply nat_gcd_unique. + intros q H1 H2. + rapply (nat_divides_add_l _ _ (b * m)). + destruct p; exact _. +Defined. + +Definition nat_bezout_comm n m d + : 0 < m -> NatBezout n m d -> NatBezout m n d. +Proof. + intros H [a [b p]]. + destruct (@equiv_leq_lt_or_eq 0 a _) as [|q]. + - exists (n * a.+1 * b.+1 - b), (m * a.+1 * b.+1 - a). + rewrite 2 nat_dist_sub_r. + apply nat_moveR_nV. + rewrite <- nat_add_comm, nat_add_assoc, <- (nat_add_comm d). + rewrite <- nat_sub_l_add_r. + 2: { apply nat_mul_r_monotone. + rewrite 2 nat_mul_succ_r. + nrapply (leq_trans _ (leq_add_l _ _)). + rapply (leq_trans _ (leq_add_r _ _)). } + apply nat_moveL_nV. + rewrite nat_add_comm. + snrapply (ap011 nat_add p). + lhs nrapply nat_mul_comm. + rhs_V nrapply nat_mul_assoc. + rhs_V nrapply nat_mul_assoc. + snrapply ap. + lhs_V nrapply nat_mul_assoc. + rhs nrapply nat_mul_assoc. + apply nat_mul_comm. + - destruct q. + exists 0, 0. + rewrite 2 nat_mul_zero_l, nat_add_zero_r in *. + symmetry in p; symmetry. + apply equiv_nat_add_zero in p. + by destruct p. +Defined. +Hint Immediate nat_bezout_comm : typeclass_instances. + +Global Instance nat_bezout_pos_l n m : 0 < n -> NatBezout n m (nat_gcd n m). +Proof. + pose (k := n + m); assert (p : n + m = k) by reflexivity; clearbody k. + revert k n m p; snrapply nat_ind_strong; hnf; intros k IHk n m q H. + (** Given a sum [n + m], we can always find another pair [n' + m'] equal to that sum such that [n' < m']. This extra hypothesis lets us prove the statement more directly. *) + assert (H' : forall n' m', n + m = n' + m' -> 0 < n' -> n' < m' + -> NatBezout n' m' (nat_gcd n' m')). + { intros n' m' p H1 H2; destruct q. + assert (m' < n + m) by (rewrite p; change (0 + m' < n' + m'); exact _). + destruct (IHk m' _ n' (m' - n') (nat_add_sub_r_cancel _) _) as [a [b r]]. + exists (a + b), b. + rewrite nat_dist_r, r, nat_dist_sub_l, <- nat_add_assoc. + rewrite nat_add_sub_l_cancel; only 2: rapply nat_mul_l_monotone. + snrapply (ap (fun x => x + _)). + rapply nat_gcd_r_sub. } + destruct (nat_trichotomy n m) as [[l | p] | r]. + - by apply H'. + - destruct p. + rewrite nat_gcd_idem; exact _. + - destruct (@equiv_leq_lt_or_eq 0 m _). + + rewrite nat_gcd_comm. + rapply nat_bezout_comm. + rapply H'. + apply nat_add_comm. + + destruct p. + rewrite nat_gcd_zero_r; exact _. +Defined. + +(** For strictly positive numbers, we have Bezout's identity in both directions. *) +Definition nat_bezout_pos n m + : 0 < n -> 0 < m + -> NatBezout n m (nat_gcd n m) /\ NatBezout m n (nat_gcd n m). +Proof. + intros H1 H2; split; [| apply nat_bezout_comm]; exact _. +Defined. + +(** For arbitrary natural numbers, we have Bezout's identity in at least one direction. *) +Definition nat_bezout n m + : NatBezout n m (nat_gcd n m) + NatBezout m n (nat_gcd n m). +Proof. + destruct n; [ right | left ]; exact _. +Defined. From 17430ac2ffa1b6d60f698a7c77aedd6a52def9e5 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 8 Aug 2024 18:44:30 +0100 Subject: [PATCH 276/282] move free group construction of integers to FreeInt.v Signed-off-by: Ali Caglayan --- theories/Algebra/AbGroups/Cyclic.v | 75 +----------------------------- theories/Algebra/AbSES/SixTerm.v | 2 +- theories/Spaces/FreeInt.v | 71 ++++++++++++++++++++++++++++ 3 files changed, 73 insertions(+), 75 deletions(-) create mode 100644 theories/Spaces/FreeInt.v diff --git a/theories/Algebra/AbGroups/Cyclic.v b/theories/Algebra/AbGroups/Cyclic.v index 517129dfa6a..01774a32439 100644 --- a/theories/Algebra/AbGroups/Cyclic.v +++ b/theories/Algebra/AbGroups/Cyclic.v @@ -1,80 +1,7 @@ -Require Import Basics Types WildCat.Core Truncations.Core - AbelianGroup AbHom Centralizer AbProjective - Groups.FreeGroup AbGroups.Z Spaces.Int. +Require Import Basics.Overture AbelianGroup AbGroups.Z Spaces.Int. (** * Cyclic groups *) -(** ** The free group on one generator *) - -(** We can define the integers as the free group on one generator, which we denote [Z1] below. Results from Centralizer.v and Groups.FreeGroup let us show that [Z1] is abelian. *) - -(** We define [Z] as the free group with a single generator. *) -Definition Z1 := FreeGroup Unit. -Definition Z1_gen : Z1 := freegroup_in tt. (* The generator *) - -(** The recursion principle of [Z1] and its computation rule. *) -Definition Z1_rec {G : Group} (g : G) : Z1 $-> G - := FreeGroup_rec Unit G (unit_name g). - -Definition Z1_rec_beta {G : Group} (g : G) : Z1_rec g Z1_gen = g - := FreeGroup_rec_beta _ _ _. - -(** The free group [Z] on one generator is isomorphic to the subgroup of [Z] generated by the generator. And such cyclic subgroups are known to be commutative, by [commutative_cyclic_subgroup]. *) -Global Instance Z1_commutative `{Funext} : Commutative (@group_sgop Z1) - := commutative_iso_commutative iso_subgroup_incl_freegroupon. -(* TODO: [Funext] is used in [isfreegroupon_freegroup], but there is a comment there saying that it can be removed. If that is done, can remove it from many results in this file. A different proof of this result, directly using the construction of the free group, could probably also avoid [Funext]. *) - -Definition ab_Z1 `{Funext} : AbGroup - := Build_AbGroup Z1 _. - -(** The universal property of [ab_Z1]. *) -Lemma equiv_Z1_hom@{u v | u < v} `{Funext} (A : AbGroup@{u}) - : GroupIsomorphism (ab_hom@{u v} ab_Z1@{u v} A) A. -Proof. - snrapply Build_GroupIsomorphism'. - - refine (_ oE (equiv_freegroup_rec@{u u u v} A Unit)^-1). - symmetry. refine (Build_Equiv _ _ (fun a => unit_name a) _). - - intros f g. cbn. reflexivity. -Defined. - -Definition nat_to_Z1 : nat -> Z1 - := fun n => grp_pow Z1_gen n. - -Definition Z1_mul_nat `{Funext} (n : nat) : ab_Z1 $-> ab_Z1 - := Z1_rec (nat_to_Z1 n). - -Lemma Z1_mul_nat_beta {A : AbGroup} (a : A) (n : nat) - : Z1_rec a (nat_to_Z1 n) = ab_mul n a. -Proof. - induction n as [|n H]. - 1: easy. - exact (grp_pow_natural _ _ _). -Defined. - -(** [ab_Z1] is projective. *) -Global Instance ab_Z1_projective `{Funext} - : IsAbProjective ab_Z1. -Proof. - intros A B p f H1. - pose proof (a := @center _ (H1 (f Z1_gen))). - strip_truncations. - snrefine (tr (Z1_rec a.1; _)). - cbn beta. apply ap10. - apply ap. (* of the coercion [grp_homo_map] *) - apply path_homomorphism_from_free_group. - simpl. - intros []. - exact a.2. -Defined. - -(** The map sending the generator to [1 : Int]. *) -Definition Z1_to_Z `{Funext} : ab_Z1 $-> abgroup_Z - := Z1_rec (G:=abgroup_Z) 1%int. - -(** TODO: Prove that [Z1_to_Z] is a group isomorphism. *) - -(** * Finite cyclic groups *) - (** The [n]-th cyclic group is the cokernel of [ab_mul n]. *) Definition cyclic (n : nat) : AbGroup := ab_cokernel (ab_mul (A:=abgroup_Z) n). diff --git a/theories/Algebra/AbSES/SixTerm.v b/theories/Algebra/AbSES/SixTerm.v index 560effb9eae..46b8e1d96fb 100644 --- a/theories/Algebra/AbSES/SixTerm.v +++ b/theories/Algebra/AbSES/SixTerm.v @@ -1,6 +1,6 @@ Require Import Basics Types WildCat HSet Pointed.Core Pointed.pTrunc Pointed.pEquiv Modalities.ReflectiveSubuniverse Truncations.Core Truncations.SeparatedTrunc - AbGroups Homotopy.ExactSequence Spaces.Int + AbGroups Homotopy.ExactSequence Spaces.Int Spaces.FreeInt AbSES.Core AbSES.Pullback AbSES.Pushout BaerSum Ext PullbackFiberSequence. (** * The contravariant six-term sequence of Ext *) diff --git a/theories/Spaces/FreeInt.v b/theories/Spaces/FreeInt.v new file mode 100644 index 00000000000..a835fbfebf4 --- /dev/null +++ b/theories/Spaces/FreeInt.v @@ -0,0 +1,71 @@ +Require Import Basics Types WildCat.Core Truncations.Core Spaces.Int + AbelianGroup AbHom Centralizer AbProjective Groups.FreeGroup AbGroups.Z. + +(** * The free group on one generator *) + +(** We can define the integers as the free group on one generator, which we denote [Z1] below. Results from Centralizer.v and Groups.FreeGroup let us show that [Z1] is abelian. *) + +(** We define [Z] as the free group with a single generator. *) +Definition Z1 := FreeGroup Unit. +Definition Z1_gen : Z1 := freegroup_in tt. (* The generator *) + +(** The recursion principle of [Z1] and its computation rule. *) +Definition Z1_rec {G : Group} (g : G) : Z1 $-> G + := FreeGroup_rec Unit G (unit_name g). + +Definition Z1_rec_beta {G : Group} (g : G) : Z1_rec g Z1_gen = g + := FreeGroup_rec_beta _ _ _. + +(** The free group [Z] on one generator is isomorphic to the subgroup of [Z] generated by the generator. And such cyclic subgroups are known to be commutative, by [commutative_cyclic_subgroup]. *) +Global Instance Z1_commutative `{Funext} : Commutative (@group_sgop Z1) + := commutative_iso_commutative iso_subgroup_incl_freegroupon. +(* TODO: [Funext] is used in [isfreegroupon_freegroup], but there is a comment there saying that it can be removed. If that is done, can remove it from many results in this file. A different proof of this result, directly using the construction of the free group, could probably also avoid [Funext]. *) + +Definition ab_Z1 `{Funext} : AbGroup + := Build_AbGroup Z1 _. + +(** The universal property of [ab_Z1]. *) +Lemma equiv_Z1_hom@{u v | u < v} `{Funext} (A : AbGroup@{u}) + : GroupIsomorphism (ab_hom@{u v} ab_Z1@{u v} A) A. +Proof. + snrapply Build_GroupIsomorphism'. + - refine (_ oE (equiv_freegroup_rec@{u u u v} A Unit)^-1). + symmetry. refine (Build_Equiv _ _ (fun a => unit_name a) _). + - intros f g. cbn. reflexivity. +Defined. + +Definition nat_to_Z1 : nat -> Z1 + := fun n => grp_pow Z1_gen n. + +Definition Z1_mul_nat `{Funext} (n : nat) : ab_Z1 $-> ab_Z1 + := Z1_rec (nat_to_Z1 n). + +Lemma Z1_mul_nat_beta {A : AbGroup} (a : A) (n : nat) + : Z1_rec a (nat_to_Z1 n) = ab_mul n a. +Proof. + induction n as [|n H]. + 1: easy. + exact (grp_pow_natural _ _ _). +Defined. + +(** [ab_Z1] is projective. *) +Global Instance ab_Z1_projective `{Funext} + : IsAbProjective ab_Z1. +Proof. + intros A B p f H1. + pose proof (a := @center _ (H1 (f Z1_gen))). + strip_truncations. + snrefine (tr (Z1_rec a.1; _)). + cbn beta. apply ap10. + apply ap. (* of the coercion [grp_homo_map] *) + apply path_homomorphism_from_free_group. + simpl. + intros []. + exact a.2. +Defined. + +(** The map sending the generator to [1 : Int]. *) +Definition Z1_to_Z `{Funext} : ab_Z1 $-> abgroup_Z + := Z1_rec (G:=abgroup_Z) 1%int. + +(** TODO: Prove that [Z1_to_Z] is a group isomorphism. *) From 8dd6d51018c694bea77d793c12c03afd44ec826e Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 8 Aug 2024 18:55:54 +0100 Subject: [PATCH 277/282] Int: conversion lemmas with nat Signed-off-by: Ali Caglayan --- theories/Spaces/Int.v | 43 +++++++++++++++++++++++++++++++++++++++++++ 1 file changed, 43 insertions(+) diff --git a/theories/Spaces/Int.v b/theories/Spaces/Int.v index a352b33d5a8..468acbb1d08 100644 --- a/theories/Spaces/Int.v +++ b/theories/Spaces/Int.v @@ -732,3 +732,46 @@ Proof. refine (ap (fun q => equiv_path A A (loopexp q z) a) _). apply eissect. Defined. + +(** ** Converting between integers and naturals *) + +Definition int_nat_succ (n : nat) + : (n.+1)%int = (n.+1)%nat :> Int. +Proof. + by induction n. +Defined. + +Definition int_nat_add (n m : nat) + : (n + m)%int = (n + m)%nat :> Int. +Proof. + induction n as [|n IHn]. + - reflexivity. + - rewrite <- 2 int_nat_succ. + rewrite int_add_succ_l. + exact (ap _ IHn). +Defined. + +Definition int_nat_sub (n m : nat) + : (m <= n)%nat -> (n - m)%int = (n - m)%nat :> Int. +Proof. + intros H. + induction H as [|n H IHn]. + - lhs nrapply int_add_neg_r. + by rewrite nat_sub_cancel. + - rewrite nat_sub_succ_l; only 2: exact _. + rewrite <- 2 int_nat_succ. + rewrite int_add_succ_l. + exact (ap _ IHn). +Defined. + +Definition int_nat_mul (n m : nat) + : (n * m)%int = (n * m)%nat :> Int. +Proof. + induction n as [|n IHn]. + - reflexivity. + - rewrite <- int_nat_succ. + rewrite int_mul_succ_l. + rewrite nat_mul_succ_l. + rhs_V nrapply int_nat_add. + exact (ap _ IHn). +Defined. From e7ba8cd4274c708f95017819c78f780295baa0f9 Mon Sep 17 00:00:00 2001 From: D Date: Thu, 8 Aug 2024 20:18:13 +0200 Subject: [PATCH 278/282] Added proof for ab_mul_cyclic_in --- theories/Algebra/AbGroups/Cyclic.v | 12 +++++++++++- 1 file changed, 11 insertions(+), 1 deletion(-) diff --git a/theories/Algebra/AbGroups/Cyclic.v b/theories/Algebra/AbGroups/Cyclic.v index 72130ccbc4a..f86233b368d 100644 --- a/theories/Algebra/AbGroups/Cyclic.v +++ b/theories/Algebra/AbGroups/Cyclic.v @@ -1,6 +1,6 @@ Require Import Basics Types WildCat.Core Truncations.Core AbelianGroup AbHom Centralizer AbProjective - Groups.FreeGroup AbGroups.Z Spaces.Int. + Groups.FreeGroup AbGroups.Z Spaces.Int Groups.QuotientGroup. (** * Cyclic groups *) @@ -84,3 +84,13 @@ Definition cyclic `{Funext} (n : nat) : AbGroup (** The [n]-th cyclic group is the cokernel of [ab_mul n]. *) Definition cyclic' (n : nat) : AbGroup := ab_cokernel (ab_mul (A:=abgroup_Z) n). + +Definition cyclic_in {n : nat} : abgroup_Z $-> cyclic' n + := grp_quotient_map. + +Definition ab_mul_cyclic_in (n : nat) (x y : abgroup_Z) + : ab_mul y (@cyclic_in n x) = cyclic_in (y * x)%int. +Proof. + lhs_V nrapply ab_mul_natural. + apply ap, abgroup_Z_ab_mul. +Defined. From 659d6fe2ae628d4aab3d68dda16efe140830cfec Mon Sep 17 00:00:00 2001 From: D Date: Thu, 8 Aug 2024 21:02:04 +0200 Subject: [PATCH 279/282] Making n in cyclic_in explicit --- theories/Algebra/AbGroups/Cyclic.v | 4 ++-- 1 file changed, 2 insertions(+), 2 deletions(-) diff --git a/theories/Algebra/AbGroups/Cyclic.v b/theories/Algebra/AbGroups/Cyclic.v index f86233b368d..8b9aab8e3ed 100644 --- a/theories/Algebra/AbGroups/Cyclic.v +++ b/theories/Algebra/AbGroups/Cyclic.v @@ -85,11 +85,11 @@ Definition cyclic `{Funext} (n : nat) : AbGroup Definition cyclic' (n : nat) : AbGroup := ab_cokernel (ab_mul (A:=abgroup_Z) n). -Definition cyclic_in {n : nat} : abgroup_Z $-> cyclic' n +Definition cyclic_in (n : nat) : abgroup_Z $-> cyclic' n := grp_quotient_map. Definition ab_mul_cyclic_in (n : nat) (x y : abgroup_Z) - : ab_mul y (@cyclic_in n x) = cyclic_in (y * x)%int. + : ab_mul y (cyclic_in n x) = cyclic_in n (y * x)%int. Proof. lhs_V nrapply ab_mul_natural. apply ap, abgroup_Z_ab_mul. From 7257e79077b2881b0e0c00ad340c88c71d2fd76d Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 8 Aug 2024 20:57:09 +0100 Subject: [PATCH 280/282] review comments Signed-off-by: Ali Caglayan --- theories/Algebra/AbGroups/AbelianGroup.v | 2 +- theories/Spaces/Int.v | 14 +++++--------- 2 files changed, 6 insertions(+), 10 deletions(-) diff --git a/theories/Algebra/AbGroups/AbelianGroup.v b/theories/Algebra/AbGroups/AbelianGroup.v index e6f2d4fb734..17e2bbe5ee7 100644 --- a/theories/Algebra/AbGroups/AbelianGroup.v +++ b/theories/Algebra/AbGroups/AbelianGroup.v @@ -312,7 +312,7 @@ Definition ab_sum_const {A : AbGroup} (n : nat) (a : A) Proof. induction n as [|n IHn] in f, p |- *. - reflexivity. - - rhs nrapply (ap@{Set _} _ (int_of_nat_succ_commute n)). + - rhs_V nrapply (ap@{Set _} _ (int_nat_succ n)). rhs nrapply grp_pow_succ. simpl. f_ap. apply IHn. diff --git a/theories/Spaces/Int.v b/theories/Spaces/Int.v index 468acbb1d08..ae7f0130a86 100644 --- a/theories/Spaces/Int.v +++ b/theories/Spaces/Int.v @@ -19,7 +19,7 @@ Inductive Int : Type0 := | zero : Int | posS : nat -> Int. -(** We can convert a [nat] to an [Int] by mapping [0] to [zero] and [S n] to [posS n]. *) +(** We can convert a [nat] to an [Int] by mapping [0] to [zero] and [S n] to [posS n]. Various operations on [nat] are preserved by this function. See the section on conversion functions starting with [int_nat_succ]. *) Definition int_of_nat (n : nat) := match n with | O => zero @@ -78,14 +78,6 @@ Definition int_pred (n : Int) : Int := Notation "n .-1" := (int_pred n) : int_scope. -(** [int_of_nat] is commutes with taking succesors *) - -Definition int_of_nat_succ_commute (n : nat) - : int_of_nat (S n) = (int_succ o int_of_nat) n :> Int. -Proof. - destruct n; reflexivity. -Defined. - (** *** Negation *) Definition int_neg@{} (x : Int) : Int := @@ -735,12 +727,14 @@ Defined. (** ** Converting between integers and naturals *) +(** [int_of_nat] preserves successors. *) Definition int_nat_succ (n : nat) : (n.+1)%int = (n.+1)%nat :> Int. Proof. by induction n. Defined. +(** [int_of_nat] preserves addition. Hence is a monoid homomorphism. *) Definition int_nat_add (n m : nat) : (n + m)%int = (n + m)%nat :> Int. Proof. @@ -751,6 +745,7 @@ Proof. exact (ap _ IHn). Defined. +(** [int_of_nat] preserves subtraction when not truncated. *) Definition int_nat_sub (n m : nat) : (m <= n)%nat -> (n - m)%int = (n - m)%nat :> Int. Proof. @@ -764,6 +759,7 @@ Proof. exact (ap _ IHn). Defined. +(** [int_of_nat] preserves multiplication. This makes [int_of_nat] a semiring homomorphism. *) Definition int_nat_mul (n m : nat) : (n * m)%int = (n * m)%nat :> Int. Proof. From a890cadd6491749f22d4d9f33962070a487b61fb Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 8 Aug 2024 21:04:24 +0100 Subject: [PATCH 281/282] Update theories/Spaces/FreeInt.v Co-authored-by: Dan Christensen --- theories/Spaces/FreeInt.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Spaces/FreeInt.v b/theories/Spaces/FreeInt.v index a835fbfebf4..7be13ea1124 100644 --- a/theories/Spaces/FreeInt.v +++ b/theories/Spaces/FreeInt.v @@ -16,7 +16,7 @@ Definition Z1_rec {G : Group} (g : G) : Z1 $-> G Definition Z1_rec_beta {G : Group} (g : G) : Z1_rec g Z1_gen = g := FreeGroup_rec_beta _ _ _. -(** The free group [Z] on one generator is isomorphic to the subgroup of [Z] generated by the generator. And such cyclic subgroups are known to be commutative, by [commutative_cyclic_subgroup]. *) +(** The free group [Z1] on one generator is isomorphic to the subgroup of [Z1] generated by the generator. And such cyclic subgroups are known to be commutative, by [commutative_cyclic_subgroup]. *) Global Instance Z1_commutative `{Funext} : Commutative (@group_sgop Z1) := commutative_iso_commutative iso_subgroup_incl_freegroupon. (* TODO: [Funext] is used in [isfreegroupon_freegroup], but there is a comment there saying that it can be removed. If that is done, can remove it from many results in this file. A different proof of this result, directly using the construction of the free group, could probably also avoid [Funext]. *) From aa6fc5362de35a8b18a900f51bab6c9b0025da38 Mon Sep 17 00:00:00 2001 From: Ali Caglayan Date: Thu, 8 Aug 2024 21:04:36 +0100 Subject: [PATCH 282/282] Update theories/Spaces/FreeInt.v Co-authored-by: Dan Christensen --- theories/Spaces/FreeInt.v | 2 +- 1 file changed, 1 insertion(+), 1 deletion(-) diff --git a/theories/Spaces/FreeInt.v b/theories/Spaces/FreeInt.v index 7be13ea1124..0e6e45233bd 100644 --- a/theories/Spaces/FreeInt.v +++ b/theories/Spaces/FreeInt.v @@ -5,7 +5,7 @@ Require Import Basics Types WildCat.Core Truncations.Core Spaces.Int (** We can define the integers as the free group on one generator, which we denote [Z1] below. Results from Centralizer.v and Groups.FreeGroup let us show that [Z1] is abelian. *) -(** We define [Z] as the free group with a single generator. *) +(** We define [Z1] as the free group with a single generator. *) Definition Z1 := FreeGroup Unit. Definition Z1_gen : Z1 := freegroup_in tt. (* The generator *)