![]() |
Metamath
Proof Explorer Theorem List (p. 181 of 429) | < Previous Next > |
Bad symbols? Try the
GIF version. |
||
Mirrors > Metamath Home Page > MPE Home Page > Theorem List Contents > Recent Proofs This page: Page List |
Color key: | ![]() (1-27903) |
![]() (27904-29428) |
![]() (29429-42879) |
Type | Label | Description |
---|---|---|
Statement | ||
Theorem | odcl 18001 | The order of a group element is always a nonnegative integer. (Contributed by Mario Carneiro, 14-Jan-2015.) (Revised by Stefan O'Rear, 5-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) ⇒ ⊢ (𝐴 ∈ 𝑋 → (𝑂‘𝐴) ∈ ℕ0) | ||
Theorem | odf 18002 | Functionality of the group element order. (Contributed by Stefan O'Rear, 5-Sep-2015.) (Proof shortened by AV, 5-Oct-2020.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) ⇒ ⊢ 𝑂:𝑋⟶ℕ0 | ||
Theorem | odid 18003 | Any element to the power of its order is the identity. (Contributed by Mario Carneiro, 14-Jan-2015.) (Revised by Stefan O'Rear, 5-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ (𝐴 ∈ 𝑋 → ((𝑂‘𝐴) · 𝐴) = 0 ) | ||
Theorem | odlem2 18004 | Any positive annihilator of a group element is an upper bound on the (positive) order of the element. (Contributed by Mario Carneiro, 14-Jan-2015.) (Revised by Stefan O'Rear, 5-Sep-2015.) (Proof shortened by AV, 5-Oct-2020.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ ((𝐴 ∈ 𝑋 ∧ 𝑁 ∈ ℕ ∧ (𝑁 · 𝐴) = 0 ) → (𝑂‘𝐴) ∈ (1...𝑁)) | ||
Theorem | odmodnn0 18005 | Reduce the argument of a group multiple by modding out the order of the element. (Contributed by Mario Carneiro, 23-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ (((𝐺 ∈ Mnd ∧ 𝐴 ∈ 𝑋 ∧ 𝑁 ∈ ℕ0) ∧ (𝑂‘𝐴) ∈ ℕ) → ((𝑁 mod (𝑂‘𝐴)) · 𝐴) = (𝑁 · 𝐴)) | ||
Theorem | mndodconglem 18006 | Lemma for mndodcong 18007. (Contributed by Mario Carneiro, 23-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Mnd) & ⊢ (𝜑 → 𝐴 ∈ 𝑋) & ⊢ (𝜑 → (𝑂‘𝐴) ∈ ℕ) & ⊢ (𝜑 → 𝑀 ∈ ℕ0) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → 𝑀 < (𝑂‘𝐴)) & ⊢ (𝜑 → 𝑁 < (𝑂‘𝐴)) & ⊢ (𝜑 → (𝑀 · 𝐴) = (𝑁 · 𝐴)) ⇒ ⊢ ((𝜑 ∧ 𝑀 ≤ 𝑁) → 𝑀 = 𝑁) | ||
Theorem | mndodcong 18007 | If two multipliers are congruent relative to the base point's order, the corresponding multiples are the same. (Contributed by Mario Carneiro, 23-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ (((𝐺 ∈ Mnd ∧ 𝐴 ∈ 𝑋) ∧ (𝑀 ∈ ℕ0 ∧ 𝑁 ∈ ℕ0) ∧ (𝑂‘𝐴) ∈ ℕ) → ((𝑂‘𝐴) ∥ (𝑀 − 𝑁) ↔ (𝑀 · 𝐴) = (𝑁 · 𝐴))) | ||
Theorem | mndodcongi 18008 | If two multipliers are congruent relative to the base point's order, the corresponding multiples are the same. For monoids, the reverse implication is false for elements with infinite order. For example, the powers of 2 mod 10 are 1,2,4,8,6,2,4,8,6,... so that the identity 1 never repeats, which is infinite order by our definition, yet other numbers like 6 appear many times in the sequence. (Contributed by Mario Carneiro, 23-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ ((𝐺 ∈ Mnd ∧ 𝐴 ∈ 𝑋 ∧ (𝑀 ∈ ℕ0 ∧ 𝑁 ∈ ℕ0)) → ((𝑂‘𝐴) ∥ (𝑀 − 𝑁) → (𝑀 · 𝐴) = (𝑁 · 𝐴))) | ||
Theorem | oddvdsnn0 18009 | The only multiples of 𝐴 that are equal to the identity are the multiples of the order of 𝐴. (Contributed by Mario Carneiro, 23-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ ((𝐺 ∈ Mnd ∧ 𝐴 ∈ 𝑋 ∧ 𝑁 ∈ ℕ0) → ((𝑂‘𝐴) ∥ 𝑁 ↔ (𝑁 · 𝐴) = 0 )) | ||
Theorem | odnncl 18010 | If a nonzero multiple of an element is zero, the element has positive order. (Contributed by Stefan O'Rear, 5-Sep-2015.) (Revised by Mario Carneiro, 22-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ (((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ 𝑁 ∈ ℤ) ∧ (𝑁 ≠ 0 ∧ (𝑁 · 𝐴) = 0 )) → (𝑂‘𝐴) ∈ ℕ) | ||
Theorem | odmod 18011 | Reduce the argument of a group multiple by modding out the order of the element. (Contributed by Mario Carneiro, 14-Jan-2015.) (Revised by Mario Carneiro, 6-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ (((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ 𝑁 ∈ ℤ) ∧ (𝑂‘𝐴) ∈ ℕ) → ((𝑁 mod (𝑂‘𝐴)) · 𝐴) = (𝑁 · 𝐴)) | ||
Theorem | oddvds 18012 | The only multiples of 𝐴 that are equal to the identity are the multiples of the order of 𝐴. (Contributed by Mario Carneiro, 14-Jan-2015.) (Revised by Mario Carneiro, 23-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ 𝑁 ∈ ℤ) → ((𝑂‘𝐴) ∥ 𝑁 ↔ (𝑁 · 𝐴) = 0 )) | ||
Theorem | oddvdsi 18013 | Any group element is annihilated by any multiple of its order. (Contributed by Stefan O'Rear, 5-Sep-2015.) (Revised by Mario Carneiro, 23-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ (𝑂‘𝐴) ∥ 𝑁) → (𝑁 · 𝐴) = 0 ) | ||
Theorem | odcong 18014 | If two multipliers are congruent relative to the base point's order, the corresponding multiples are the same. (Contributed by Stefan O'Rear, 5-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ (𝑀 ∈ ℤ ∧ 𝑁 ∈ ℤ)) → ((𝑂‘𝐴) ∥ (𝑀 − 𝑁) ↔ (𝑀 · 𝐴) = (𝑁 · 𝐴))) | ||
Theorem | odeq 18015* | The oddvds 18012 property uniquely defines the group order. (Contributed by Stefan O'Rear, 6-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ 𝑁 ∈ ℕ0) → (𝑁 = (𝑂‘𝐴) ↔ ∀𝑦 ∈ ℕ0 (𝑁 ∥ 𝑦 ↔ (𝑦 · 𝐴) = 0 ))) | ||
Theorem | odval2 18016* | A non-conditional definition of the group order. (Contributed by Stefan O'Rear, 6-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋) → (𝑂‘𝐴) = (℩𝑥 ∈ ℕ0 ∀𝑦 ∈ ℕ0 (𝑥 ∥ 𝑦 ↔ (𝑦 · 𝐴) = 0 ))) | ||
Theorem | odmulgid 18017 | A relationship between the order of a multiple and the order of the basepoint. (Contributed by Stefan O'Rear, 6-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) ⇒ ⊢ (((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ 𝑁 ∈ ℤ) ∧ 𝐾 ∈ ℤ) → ((𝑂‘(𝑁 · 𝐴)) ∥ 𝐾 ↔ (𝑂‘𝐴) ∥ (𝐾 · 𝑁))) | ||
Theorem | odmulg2 18018 | The order of a multiple divides the order of the base point. (Contributed by Stefan O'Rear, 6-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ 𝑁 ∈ ℤ) → (𝑂‘(𝑁 · 𝐴)) ∥ (𝑂‘𝐴)) | ||
Theorem | odmulg 18019 | Relationship between the order of an element and that of a multiple. (Contributed by Stefan O'Rear, 6-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ 𝑁 ∈ ℤ) → (𝑂‘𝐴) = ((𝑁 gcd (𝑂‘𝐴)) · (𝑂‘(𝑁 · 𝐴)))) | ||
Theorem | odmulgeq 18020 | A multiple of a point of finite order only has the same order if the multiplier is relatively prime. (Contributed by Stefan O'Rear, 12-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) ⇒ ⊢ (((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ 𝑁 ∈ ℤ) ∧ (𝑂‘𝐴) ∈ ℕ) → ((𝑂‘(𝑁 · 𝐴)) = (𝑂‘𝐴) ↔ (𝑁 gcd (𝑂‘𝐴)) = 1)) | ||
Theorem | odbezout 18021* | If 𝑁 is coprime to the order of 𝐴, there is a modular inverse 𝑥 to cancel multiplication by 𝑁. (Contributed by Mario Carneiro, 27-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) ⇒ ⊢ (((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ 𝑁 ∈ ℤ) ∧ (𝑁 gcd (𝑂‘𝐴)) = 1) → ∃𝑥 ∈ ℤ (𝑥 · (𝑁 · 𝐴)) = 𝐴) | ||
Theorem | od1 18022 | The order of the group identity is one. (Contributed by Mario Carneiro, 14-Jan-2015.) (Revised by Mario Carneiro, 23-Sep-2015.) |
⊢ 𝑂 = (od‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ (𝐺 ∈ Grp → (𝑂‘ 0 ) = 1) | ||
Theorem | odeq1 18023 | The group identity is the unique element of a group with order one. (Contributed by Mario Carneiro, 14-Jan-2015.) (Revised by Mario Carneiro, 23-Sep-2015.) |
⊢ 𝑂 = (od‘𝐺) & ⊢ 0 = (0g‘𝐺) & ⊢ 𝑋 = (Base‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋) → ((𝑂‘𝐴) = 1 ↔ 𝐴 = 0 )) | ||
Theorem | odinv 18024 | The order of the inverse of a group element. (Contributed by Mario Carneiro, 20-Oct-2015.) |
⊢ 𝑂 = (od‘𝐺) & ⊢ 𝐼 = (invg‘𝐺) & ⊢ 𝑋 = (Base‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋) → (𝑂‘(𝐼‘𝐴)) = (𝑂‘𝐴)) | ||
Theorem | odf1 18025* | The multiples of an element with infinite order form an infinite cyclic subgroup of 𝐺. (Contributed by Mario Carneiro, 14-Jan-2015.) (Revised by Mario Carneiro, 23-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 𝐹 = (𝑥 ∈ ℤ ↦ (𝑥 · 𝐴)) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋) → ((𝑂‘𝐴) = 0 ↔ 𝐹:ℤ–1-1→𝑋)) | ||
Theorem | odinf 18026* | The multiples of an element with infinite order form an infinite cyclic subgroup of 𝐺. (Contributed by Mario Carneiro, 14-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 𝐹 = (𝑥 ∈ ℤ ↦ (𝑥 · 𝐴)) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ (𝑂‘𝐴) = 0) → ¬ ran 𝐹 ∈ Fin) | ||
Theorem | dfod2 18027* | An alternative definition of the order of a group element is as the cardinality of the cyclic subgroup generated by the element. (Contributed by Mario Carneiro, 14-Jan-2015.) (Revised by Mario Carneiro, 2-Oct-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 𝐹 = (𝑥 ∈ ℤ ↦ (𝑥 · 𝐴)) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋) → (𝑂‘𝐴) = if(ran 𝐹 ∈ Fin, (#‘ran 𝐹), 0)) | ||
Theorem | odcl2 18028 | The order of an element of a finite group is finite. (Contributed by Mario Carneiro, 14-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝑋 ∈ Fin ∧ 𝐴 ∈ 𝑋) → (𝑂‘𝐴) ∈ ℕ) | ||
Theorem | oddvds2 18029 | The order of an element of a finite group divides the order (cardinality) of the group. Corollary of Lagrange's theorem for the order of a subgroup. (Contributed by Mario Carneiro, 14-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝑋 ∈ Fin ∧ 𝐴 ∈ 𝑋) → (𝑂‘𝐴) ∥ (#‘𝑋)) | ||
Theorem | submod 18030 | The order of an element is the same in a subgroup. (Contributed by Stefan O'Rear, 12-Sep-2015.) (Proof shortened by AV, 5-Oct-2020.) |
⊢ 𝐻 = (𝐺 ↾s 𝑌) & ⊢ 𝑂 = (od‘𝐺) & ⊢ 𝑃 = (od‘𝐻) ⇒ ⊢ ((𝑌 ∈ (SubMnd‘𝐺) ∧ 𝐴 ∈ 𝑌) → (𝑂‘𝐴) = (𝑃‘𝐴)) | ||
Theorem | subgod 18031 | The order of an element is the same in a subgroup. (Contributed by Mario Carneiro, 14-Jan-2015.) (Proof shortened by Stefan O'Rear, 12-Sep-2015.) |
⊢ 𝐻 = (𝐺 ↾s 𝑌) & ⊢ 𝑂 = (od‘𝐺) & ⊢ 𝑃 = (od‘𝐻) ⇒ ⊢ ((𝑌 ∈ (SubGrp‘𝐺) ∧ 𝐴 ∈ 𝑌) → (𝑂‘𝐴) = (𝑃‘𝐴)) | ||
Theorem | odsubdvds 18032 | The order of an element of a subgroup divides the order of the subgroup. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ 𝑂 = (od‘𝐺) ⇒ ⊢ ((𝑆 ∈ (SubGrp‘𝐺) ∧ 𝑆 ∈ Fin ∧ 𝐴 ∈ 𝑆) → (𝑂‘𝐴) ∥ (#‘𝑆)) | ||
Theorem | odf1o1 18033* | An element with zero order has infinitely many multiples. (Contributed by Stefan O'Rear, 6-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ 𝐾 = (mrCls‘(SubGrp‘𝐺)) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ (𝑂‘𝐴) = 0) → (𝑥 ∈ ℤ ↦ (𝑥 · 𝐴)):ℤ–1-1-onto→(𝐾‘{𝐴})) | ||
Theorem | odf1o2 18034* | An element with nonzero order has as many multiples as its order. (Contributed by Stefan O'Rear, 6-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ 𝐾 = (mrCls‘(SubGrp‘𝐺)) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ (𝑂‘𝐴) ∈ ℕ) → (𝑥 ∈ (0..^(𝑂‘𝐴)) ↦ (𝑥 · 𝐴)):(0..^(𝑂‘𝐴))–1-1-onto→(𝐾‘{𝐴})) | ||
Theorem | odhash 18035 | An element of zero order generates an infinite subgroup. (Contributed by Stefan O'Rear, 12-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ 𝐾 = (mrCls‘(SubGrp‘𝐺)) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ (𝑂‘𝐴) = 0) → (#‘(𝐾‘{𝐴})) = +∞) | ||
Theorem | odhash2 18036 | If an element has nonzero order, it generates a subgroup with size equal to the order. (Contributed by Stefan O'Rear, 12-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ 𝐾 = (mrCls‘(SubGrp‘𝐺)) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ (𝑂‘𝐴) ∈ ℕ) → (#‘(𝐾‘{𝐴})) = (𝑂‘𝐴)) | ||
Theorem | odhash3 18037 | An element which generates a finite subgroup has order the size of that subgroup. (Contributed by Stefan O'Rear, 12-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ 𝐾 = (mrCls‘(SubGrp‘𝐺)) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ (𝐾‘{𝐴}) ∈ Fin) → (𝑂‘𝐴) = (#‘(𝐾‘{𝐴}))) | ||
Theorem | odngen 18038* | A cyclic subgroup of size (𝑂‘𝐴) has (ϕ‘(𝑂‘𝐴)) generators. (Contributed by Stefan O'Rear, 12-Sep-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) & ⊢ 𝐾 = (mrCls‘(SubGrp‘𝐺)) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ (𝑂‘𝐴) ∈ ℕ) → (#‘{𝑥 ∈ (𝐾‘{𝐴}) ∣ (𝑂‘𝑥) = (𝑂‘𝐴)}) = (ϕ‘(𝑂‘𝐴))) | ||
Theorem | gexval 18039* | Value of the exponent of a group. (Contributed by Mario Carneiro, 23-Apr-2016.) (Revised by AV, 26-Sep-2020.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) & ⊢ 𝐼 = {𝑦 ∈ ℕ ∣ ∀𝑥 ∈ 𝑋 (𝑦 · 𝑥) = 0 } ⇒ ⊢ (𝐺 ∈ 𝑉 → 𝐸 = if(𝐼 = ∅, 0, inf(𝐼, ℝ, < ))) | ||
Theorem | gexlem1 18040* | The group element order is either zero or a nonzero multiplier that annihilates the element. (Contributed by Mario Carneiro, 23-Apr-2016.) (Proof shortened by AV, 26-Sep-2020.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) & ⊢ 𝐼 = {𝑦 ∈ ℕ ∣ ∀𝑥 ∈ 𝑋 (𝑦 · 𝑥) = 0 } ⇒ ⊢ (𝐺 ∈ 𝑉 → ((𝐸 = 0 ∧ 𝐼 = ∅) ∨ 𝐸 ∈ 𝐼)) | ||
Theorem | gexcl 18041 | The exponent of a group is a nonnegative integer. (Contributed by Mario Carneiro, 23-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) ⇒ ⊢ (𝐺 ∈ 𝑉 → 𝐸 ∈ ℕ0) | ||
Theorem | gexid 18042 | Any element to the power of the group exponent. (Contributed by Mario Carneiro, 24-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ (𝐴 ∈ 𝑋 → (𝐸 · 𝐴) = 0 ) | ||
Theorem | gexlem2 18043* | Any positive annihilator of all the group elements is an upper bound on the group exponent. (Contributed by Mario Carneiro, 24-Apr-2016.) (Proof shortened by AV, 26-Sep-2020.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ ((𝐺 ∈ 𝑉 ∧ 𝑁 ∈ ℕ ∧ ∀𝑥 ∈ 𝑋 (𝑁 · 𝑥) = 0 ) → 𝐸 ∈ (1...𝑁)) | ||
Theorem | gexdvdsi 18044 | Any group element is annihilated by any multiple of the group exponent. (Contributed by Mario Carneiro, 24-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋 ∧ 𝐸 ∥ 𝑁) → (𝑁 · 𝐴) = 0 ) | ||
Theorem | gexdvds 18045* | The only 𝑁 that annihilate all the elements of the group are the multiples of the group exponent. (Contributed by Mario Carneiro, 24-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) & ⊢ · = (.g‘𝐺) & ⊢ 0 = (0g‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝑁 ∈ ℤ) → (𝐸 ∥ 𝑁 ↔ ∀𝑥 ∈ 𝑋 (𝑁 · 𝑥) = 0 )) | ||
Theorem | gexdvds2 18046* | An integer divides the group exponent iff it divides all the group orders. In other words, the group exponent is the LCM of the orders of all the elements. (Contributed by Mario Carneiro, 24-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) & ⊢ 𝑂 = (od‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝑁 ∈ ℤ) → (𝐸 ∥ 𝑁 ↔ ∀𝑥 ∈ 𝑋 (𝑂‘𝑥) ∥ 𝑁)) | ||
Theorem | gexod 18047 | Any group element is annihilated by any multiple of the group exponent. (Contributed by Mario Carneiro, 24-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) & ⊢ 𝑂 = (od‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐴 ∈ 𝑋) → (𝑂‘𝐴) ∥ 𝐸) | ||
Theorem | gexcl3 18048* | If the order of every group element is bounded by 𝑁, the group has finite exponent. (Contributed by Mario Carneiro, 24-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) & ⊢ 𝑂 = (od‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ ∀𝑥 ∈ 𝑋 (𝑂‘𝑥) ∈ (1...𝑁)) → 𝐸 ∈ ℕ) | ||
Theorem | gexnnod 18049 | Every group element has finite order if the exponent is finite. (Contributed by Mario Carneiro, 24-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) & ⊢ 𝑂 = (od‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝐸 ∈ ℕ ∧ 𝐴 ∈ 𝑋) → (𝑂‘𝐴) ∈ ℕ) | ||
Theorem | gexcl2 18050 | The exponent of a finite group is finite. (Contributed by Mario Carneiro, 24-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝑋 ∈ Fin) → 𝐸 ∈ ℕ) | ||
Theorem | gexdvds3 18051 | The exponent of a finite group divides the order (cardinality) of the group. Corollary of Lagrange's theorem for the order of a subgroup. (Contributed by Mario Carneiro, 24-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝑋 ∈ Fin) → 𝐸 ∥ (#‘𝑋)) | ||
Theorem | gex1 18052 | A group or monoid has exponent 1 iff it is trivial. (Contributed by Mario Carneiro, 24-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝐸 = (gEx‘𝐺) ⇒ ⊢ (𝐺 ∈ Mnd → (𝐸 = 1 ↔ 𝑋 ≈ 1𝑜)) | ||
Theorem | ispgp 18053* | A group is a 𝑃-group if every element has some power of 𝑃 as its order. (Contributed by Mario Carneiro, 15-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) ⇒ ⊢ (𝑃 pGrp 𝐺 ↔ (𝑃 ∈ ℙ ∧ 𝐺 ∈ Grp ∧ ∀𝑥 ∈ 𝑋 ∃𝑛 ∈ ℕ0 (𝑂‘𝑥) = (𝑃↑𝑛))) | ||
Theorem | pgpprm 18054 | Reverse closure for the first argument of pGrp. (Contributed by Mario Carneiro, 15-Jan-2015.) |
⊢ (𝑃 pGrp 𝐺 → 𝑃 ∈ ℙ) | ||
Theorem | pgpgrp 18055 | Reverse closure for the second argument of pGrp. (Contributed by Mario Carneiro, 15-Jan-2015.) |
⊢ (𝑃 pGrp 𝐺 → 𝐺 ∈ Grp) | ||
Theorem | pgpfi1 18056 | A finite group with order a power of a prime 𝑃 is a 𝑃-group. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝑃 ∈ ℙ ∧ 𝑁 ∈ ℕ0) → ((#‘𝑋) = (𝑃↑𝑁) → 𝑃 pGrp 𝐺)) | ||
Theorem | pgp0 18057 | The identity subgroup is a 𝑃-group for every prime 𝑃. (Contributed by Mario Carneiro, 19-Jan-2015.) |
⊢ 0 = (0g‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝑃 ∈ ℙ) → 𝑃 pGrp (𝐺 ↾s { 0 })) | ||
Theorem | subgpgp 18058 | A subgroup of a p-group is a p-group. (Contributed by Mario Carneiro, 27-Apr-2016.) |
⊢ ((𝑃 pGrp 𝐺 ∧ 𝑆 ∈ (SubGrp‘𝐺)) → 𝑃 pGrp (𝐺 ↾s 𝑆)) | ||
Theorem | sylow1lem1 18059* | Lemma for sylow1 18064. The p-adic valuation of the size of 𝑆 is equal to the number of excess powers of 𝑃 in (#‘𝑋) / (𝑃↑𝑁). (Contributed by Mario Carneiro, 15-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → (𝑃↑𝑁) ∥ (#‘𝑋)) & ⊢ + = (+g‘𝐺) & ⊢ 𝑆 = {𝑠 ∈ 𝒫 𝑋 ∣ (#‘𝑠) = (𝑃↑𝑁)} ⇒ ⊢ (𝜑 → ((#‘𝑆) ∈ ℕ ∧ (𝑃 pCnt (#‘𝑆)) = ((𝑃 pCnt (#‘𝑋)) − 𝑁))) | ||
Theorem | sylow1lem2 18060* | Lemma for sylow1 18064. The function ⊕ is a group action on 𝑆. (Contributed by Mario Carneiro, 15-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → (𝑃↑𝑁) ∥ (#‘𝑋)) & ⊢ + = (+g‘𝐺) & ⊢ 𝑆 = {𝑠 ∈ 𝒫 𝑋 ∣ (#‘𝑠) = (𝑃↑𝑁)} & ⊢ ⊕ = (𝑥 ∈ 𝑋, 𝑦 ∈ 𝑆 ↦ ran (𝑧 ∈ 𝑦 ↦ (𝑥 + 𝑧))) ⇒ ⊢ (𝜑 → ⊕ ∈ (𝐺 GrpAct 𝑆)) | ||
Theorem | sylow1lem3 18061* | Lemma for sylow1 18064. One of the orbits of the group action has p-adic valuation less than the prime count of the set 𝑆. (Contributed by Mario Carneiro, 15-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → (𝑃↑𝑁) ∥ (#‘𝑋)) & ⊢ + = (+g‘𝐺) & ⊢ 𝑆 = {𝑠 ∈ 𝒫 𝑋 ∣ (#‘𝑠) = (𝑃↑𝑁)} & ⊢ ⊕ = (𝑥 ∈ 𝑋, 𝑦 ∈ 𝑆 ↦ ran (𝑧 ∈ 𝑦 ↦ (𝑥 + 𝑧))) & ⊢ ∼ = {〈𝑥, 𝑦〉 ∣ ({𝑥, 𝑦} ⊆ 𝑆 ∧ ∃𝑔 ∈ 𝑋 (𝑔 ⊕ 𝑥) = 𝑦)} ⇒ ⊢ (𝜑 → ∃𝑤 ∈ 𝑆 (𝑃 pCnt (#‘[𝑤] ∼ )) ≤ ((𝑃 pCnt (#‘𝑋)) − 𝑁)) | ||
Theorem | sylow1lem4 18062* | Lemma for sylow1 18064. The stabilizer subgroup of any element of 𝑆 is at most 𝑃↑𝑁 in size. (Contributed by Mario Carneiro, 15-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → (𝑃↑𝑁) ∥ (#‘𝑋)) & ⊢ + = (+g‘𝐺) & ⊢ 𝑆 = {𝑠 ∈ 𝒫 𝑋 ∣ (#‘𝑠) = (𝑃↑𝑁)} & ⊢ ⊕ = (𝑥 ∈ 𝑋, 𝑦 ∈ 𝑆 ↦ ran (𝑧 ∈ 𝑦 ↦ (𝑥 + 𝑧))) & ⊢ ∼ = {〈𝑥, 𝑦〉 ∣ ({𝑥, 𝑦} ⊆ 𝑆 ∧ ∃𝑔 ∈ 𝑋 (𝑔 ⊕ 𝑥) = 𝑦)} & ⊢ (𝜑 → 𝐵 ∈ 𝑆) & ⊢ 𝐻 = {𝑢 ∈ 𝑋 ∣ (𝑢 ⊕ 𝐵) = 𝐵} ⇒ ⊢ (𝜑 → (#‘𝐻) ≤ (𝑃↑𝑁)) | ||
Theorem | sylow1lem5 18063* | Lemma for sylow1 18064. Using Lagrange's theorem and the orbit-stabilizer theorem, show that there is a subgroup with size exactly 𝑃↑𝑁. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → (𝑃↑𝑁) ∥ (#‘𝑋)) & ⊢ + = (+g‘𝐺) & ⊢ 𝑆 = {𝑠 ∈ 𝒫 𝑋 ∣ (#‘𝑠) = (𝑃↑𝑁)} & ⊢ ⊕ = (𝑥 ∈ 𝑋, 𝑦 ∈ 𝑆 ↦ ran (𝑧 ∈ 𝑦 ↦ (𝑥 + 𝑧))) & ⊢ ∼ = {〈𝑥, 𝑦〉 ∣ ({𝑥, 𝑦} ⊆ 𝑆 ∧ ∃𝑔 ∈ 𝑋 (𝑔 ⊕ 𝑥) = 𝑦)} & ⊢ (𝜑 → 𝐵 ∈ 𝑆) & ⊢ 𝐻 = {𝑢 ∈ 𝑋 ∣ (𝑢 ⊕ 𝐵) = 𝐵} & ⊢ (𝜑 → (𝑃 pCnt (#‘[𝐵] ∼ )) ≤ ((𝑃 pCnt (#‘𝑋)) − 𝑁)) ⇒ ⊢ (𝜑 → ∃ℎ ∈ (SubGrp‘𝐺)(#‘ℎ) = (𝑃↑𝑁)) | ||
Theorem | sylow1 18064* | Sylow's first theorem. If 𝑃↑𝑁 is a prime power that divides the cardinality of 𝐺, then 𝐺 has a supgroup with size 𝑃↑𝑁. This is part of Metamath 100 proof #72. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ (𝜑 → 𝑁 ∈ ℕ0) & ⊢ (𝜑 → (𝑃↑𝑁) ∥ (#‘𝑋)) ⇒ ⊢ (𝜑 → ∃𝑔 ∈ (SubGrp‘𝐺)(#‘𝑔) = (𝑃↑𝑁)) | ||
Theorem | odcau 18065* | Cauchy's theorem for the order of an element in a group. A finite group whose order divides a prime 𝑃 contains an element of order 𝑃. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑂 = (od‘𝐺) ⇒ ⊢ (((𝐺 ∈ Grp ∧ 𝑋 ∈ Fin ∧ 𝑃 ∈ ℙ) ∧ 𝑃 ∥ (#‘𝑋)) → ∃𝑔 ∈ 𝑋 (𝑂‘𝑔) = 𝑃) | ||
Theorem | pgpfi 18066* | The converse to pgpfi1 18056. A finite group is a 𝑃-group iff it has size some power of 𝑃. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝑋 ∈ Fin) → (𝑃 pGrp 𝐺 ↔ (𝑃 ∈ ℙ ∧ ∃𝑛 ∈ ℕ0 (#‘𝑋) = (𝑃↑𝑛)))) | ||
Theorem | pgpfi2 18067 | Alternate version of pgpfi 18066. (Contributed by Mario Carneiro, 27-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝑋 ∈ Fin) → (𝑃 pGrp 𝐺 ↔ (𝑃 ∈ ℙ ∧ (#‘𝑋) = (𝑃↑(𝑃 pCnt (#‘𝑋)))))) | ||
Theorem | pgphash 18068 | The order of a p-group. (Contributed by Mario Carneiro, 27-Apr-2016.) |
⊢ 𝑋 = (Base‘𝐺) ⇒ ⊢ ((𝑃 pGrp 𝐺 ∧ 𝑋 ∈ Fin) → (#‘𝑋) = (𝑃↑(𝑃 pCnt (#‘𝑋)))) | ||
Theorem | isslw 18069* | The property of being a Sylow subgroup. A Sylow 𝑃-subgroup is a 𝑃-group which has no proper supersets that are also 𝑃-groups. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ (𝐻 ∈ (𝑃 pSyl 𝐺) ↔ (𝑃 ∈ ℙ ∧ 𝐻 ∈ (SubGrp‘𝐺) ∧ ∀𝑘 ∈ (SubGrp‘𝐺)((𝐻 ⊆ 𝑘 ∧ 𝑃 pGrp (𝐺 ↾s 𝑘)) ↔ 𝐻 = 𝑘))) | ||
Theorem | slwprm 18070 | Reverse closure for the first argument of a Sylow 𝑃-subgroup. (Contributed by Mario Carneiro, 16-Jan-2015.) (Revised by Mario Carneiro, 2-May-2015.) |
⊢ (𝐻 ∈ (𝑃 pSyl 𝐺) → 𝑃 ∈ ℙ) | ||
Theorem | slwsubg 18071 | A Sylow 𝑃-subgroup is a subgroup. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ (𝐻 ∈ (𝑃 pSyl 𝐺) → 𝐻 ∈ (SubGrp‘𝐺)) | ||
Theorem | slwispgp 18072 | Defining property of a Sylow 𝑃-subgroup. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ 𝑆 = (𝐺 ↾s 𝐾) ⇒ ⊢ ((𝐻 ∈ (𝑃 pSyl 𝐺) ∧ 𝐾 ∈ (SubGrp‘𝐺)) → ((𝐻 ⊆ 𝐾 ∧ 𝑃 pGrp 𝑆) ↔ 𝐻 = 𝐾)) | ||
Theorem | slwpss 18073 | A proper superset of a Sylow subgroup is not a 𝑃-group. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ 𝑆 = (𝐺 ↾s 𝐾) ⇒ ⊢ ((𝐻 ∈ (𝑃 pSyl 𝐺) ∧ 𝐾 ∈ (SubGrp‘𝐺) ∧ 𝐻 ⊊ 𝐾) → ¬ 𝑃 pGrp 𝑆) | ||
Theorem | slwpgp 18074 | A Sylow 𝑃-subgroup is a 𝑃-group. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ 𝑆 = (𝐺 ↾s 𝐻) ⇒ ⊢ (𝐻 ∈ (𝑃 pSyl 𝐺) → 𝑃 pGrp 𝑆) | ||
Theorem | pgpssslw 18075* | Every 𝑃-subgroup is contained in a Sylow 𝑃-subgroup. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ 𝑆 = (𝐺 ↾s 𝐻) & ⊢ 𝐹 = (𝑥 ∈ {𝑦 ∈ (SubGrp‘𝐺) ∣ (𝑃 pGrp (𝐺 ↾s 𝑦) ∧ 𝐻 ⊆ 𝑦)} ↦ (#‘𝑥)) ⇒ ⊢ ((𝐻 ∈ (SubGrp‘𝐺) ∧ 𝑋 ∈ Fin ∧ 𝑃 pGrp 𝑆) → ∃𝑘 ∈ (𝑃 pSyl 𝐺)𝐻 ⊆ 𝑘) | ||
Theorem | slwn0 18076 | Every finite group contains a Sylow 𝑃-subgroup. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝑋 ∈ Fin ∧ 𝑃 ∈ ℙ) → (𝑃 pSyl 𝐺) ≠ ∅) | ||
Theorem | subgslw 18077 | A Sylow subgroup that is contained in a larger subgroup is also Sylow with respect to the subgroup. (The converse need not be true.) (Contributed by Mario Carneiro, 19-Jan-2015.) |
⊢ 𝐻 = (𝐺 ↾s 𝑆) ⇒ ⊢ ((𝑆 ∈ (SubGrp‘𝐺) ∧ 𝐾 ∈ (𝑃 pSyl 𝐺) ∧ 𝐾 ⊆ 𝑆) → 𝐾 ∈ (𝑃 pSyl 𝐻)) | ||
Theorem | sylow2alem1 18078* | Lemma for sylow2a 18080. An equivalence class of fixed points is a singleton. (Contributed by Mario Carneiro, 17-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → ⊕ ∈ (𝐺 GrpAct 𝑌)) & ⊢ (𝜑 → 𝑃 pGrp 𝐺) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑌 ∈ Fin) & ⊢ 𝑍 = {𝑢 ∈ 𝑌 ∣ ∀ℎ ∈ 𝑋 (ℎ ⊕ 𝑢) = 𝑢} & ⊢ ∼ = {〈𝑥, 𝑦〉 ∣ ({𝑥, 𝑦} ⊆ 𝑌 ∧ ∃𝑔 ∈ 𝑋 (𝑔 ⊕ 𝑥) = 𝑦)} ⇒ ⊢ ((𝜑 ∧ 𝐴 ∈ 𝑍) → [𝐴] ∼ = {𝐴}) | ||
Theorem | sylow2alem2 18079* | Lemma for sylow2a 18080. All the orbits which are not for fixed points have size ∣ 𝐺 ∣ / ∣ 𝐺𝑥 ∣ (where 𝐺𝑥 is the stabilizer subgroup) and thus are powers of 𝑃. And since they are all nontrivial (because any orbit which is a singleton is a fixed point), they all divide 𝑃, and so does the sum of all of them. (Contributed by Mario Carneiro, 17-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → ⊕ ∈ (𝐺 GrpAct 𝑌)) & ⊢ (𝜑 → 𝑃 pGrp 𝐺) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑌 ∈ Fin) & ⊢ 𝑍 = {𝑢 ∈ 𝑌 ∣ ∀ℎ ∈ 𝑋 (ℎ ⊕ 𝑢) = 𝑢} & ⊢ ∼ = {〈𝑥, 𝑦〉 ∣ ({𝑥, 𝑦} ⊆ 𝑌 ∧ ∃𝑔 ∈ 𝑋 (𝑔 ⊕ 𝑥) = 𝑦)} ⇒ ⊢ (𝜑 → 𝑃 ∥ Σ𝑧 ∈ ((𝑌 / ∼ ) ∖ 𝒫 𝑍)(#‘𝑧)) | ||
Theorem | sylow2a 18080* | A named lemma of Sylow's second and third theorems. If 𝐺 is a finite 𝑃-group that acts on the finite set 𝑌, then the set 𝑍 of all points of 𝑌 fixed by every element of 𝐺 has cardinality equivalent to the cardinality of 𝑌, mod 𝑃. (Contributed by Mario Carneiro, 17-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → ⊕ ∈ (𝐺 GrpAct 𝑌)) & ⊢ (𝜑 → 𝑃 pGrp 𝐺) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑌 ∈ Fin) & ⊢ 𝑍 = {𝑢 ∈ 𝑌 ∣ ∀ℎ ∈ 𝑋 (ℎ ⊕ 𝑢) = 𝑢} & ⊢ ∼ = {〈𝑥, 𝑦〉 ∣ ({𝑥, 𝑦} ⊆ 𝑌 ∧ ∃𝑔 ∈ 𝑋 (𝑔 ⊕ 𝑥) = 𝑦)} ⇒ ⊢ (𝜑 → 𝑃 ∥ ((#‘𝑌) − (#‘𝑍))) | ||
Theorem | sylow2blem1 18081* | Lemma for sylow2b 18084. Evaluate the group action on a left coset. (Contributed by Mario Carneiro, 17-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝐻 ∈ (SubGrp‘𝐺)) & ⊢ (𝜑 → 𝐾 ∈ (SubGrp‘𝐺)) & ⊢ + = (+g‘𝐺) & ⊢ ∼ = (𝐺 ~QG 𝐾) & ⊢ · = (𝑥 ∈ 𝐻, 𝑦 ∈ (𝑋 / ∼ ) ↦ ran (𝑧 ∈ 𝑦 ↦ (𝑥 + 𝑧))) ⇒ ⊢ ((𝜑 ∧ 𝐵 ∈ 𝐻 ∧ 𝐶 ∈ 𝑋) → (𝐵 · [𝐶] ∼ ) = [(𝐵 + 𝐶)] ∼ ) | ||
Theorem | sylow2blem2 18082* | Lemma for sylow2b 18084. Left multiplication in a subgroup 𝐻 is a group action on the set of all left cosets of 𝐾. (Contributed by Mario Carneiro, 17-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝐻 ∈ (SubGrp‘𝐺)) & ⊢ (𝜑 → 𝐾 ∈ (SubGrp‘𝐺)) & ⊢ + = (+g‘𝐺) & ⊢ ∼ = (𝐺 ~QG 𝐾) & ⊢ · = (𝑥 ∈ 𝐻, 𝑦 ∈ (𝑋 / ∼ ) ↦ ran (𝑧 ∈ 𝑦 ↦ (𝑥 + 𝑧))) ⇒ ⊢ (𝜑 → · ∈ ((𝐺 ↾s 𝐻) GrpAct (𝑋 / ∼ ))) | ||
Theorem | sylow2blem3 18083* | Sylow's second theorem. Putting together the results of sylow2a 18080 and the orbit-stabilizer theorem to show that 𝑃 does not divide the set of all fixed points under the group action, we get that there is a fixed point of the group action, so that there is some 𝑔 ∈ 𝑋 with ℎ𝑔𝐾 = 𝑔𝐾 for all ℎ ∈ 𝐻. This implies that invg(𝑔)ℎ𝑔 ∈ 𝐾, so ℎ is in the conjugated subgroup 𝑔𝐾invg(𝑔). (Contributed by Mario Carneiro, 18-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝐻 ∈ (SubGrp‘𝐺)) & ⊢ (𝜑 → 𝐾 ∈ (SubGrp‘𝐺)) & ⊢ + = (+g‘𝐺) & ⊢ ∼ = (𝐺 ~QG 𝐾) & ⊢ · = (𝑥 ∈ 𝐻, 𝑦 ∈ (𝑋 / ∼ ) ↦ ran (𝑧 ∈ 𝑦 ↦ (𝑥 + 𝑧))) & ⊢ (𝜑 → 𝑃 pGrp (𝐺 ↾s 𝐻)) & ⊢ (𝜑 → (#‘𝐾) = (𝑃↑(𝑃 pCnt (#‘𝑋)))) & ⊢ − = (-g‘𝐺) ⇒ ⊢ (𝜑 → ∃𝑔 ∈ 𝑋 𝐻 ⊆ ran (𝑥 ∈ 𝐾 ↦ ((𝑔 + 𝑥) − 𝑔))) | ||
Theorem | sylow2b 18084* | Sylow's second theorem. Any 𝑃-group 𝐻 is a subgroup of a conjugated 𝑃-group 𝐾 of order 𝑃↑𝑛 ∥ (#‘𝑋) with 𝑛 maximal. This is usually stated under the assumption that 𝐾 is a Sylow subgroup, but we use a slightly different definition, whose equivalence to this one requires this theorem. This is part of Metamath 100 proof #72. (Contributed by Mario Carneiro, 18-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝐻 ∈ (SubGrp‘𝐺)) & ⊢ (𝜑 → 𝐾 ∈ (SubGrp‘𝐺)) & ⊢ + = (+g‘𝐺) & ⊢ (𝜑 → 𝑃 pGrp (𝐺 ↾s 𝐻)) & ⊢ (𝜑 → (#‘𝐾) = (𝑃↑(𝑃 pCnt (#‘𝑋)))) & ⊢ − = (-g‘𝐺) ⇒ ⊢ (𝜑 → ∃𝑔 ∈ 𝑋 𝐻 ⊆ ran (𝑥 ∈ 𝐾 ↦ ((𝑔 + 𝑥) − 𝑔))) | ||
Theorem | slwhash 18085 | A sylow subgroup has cardinality equal to the maximum power of 𝑃 dividing the group. (Contributed by Mario Carneiro, 18-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝐻 ∈ (𝑃 pSyl 𝐺)) ⇒ ⊢ (𝜑 → (#‘𝐻) = (𝑃↑(𝑃 pCnt (#‘𝑋)))) | ||
Theorem | fislw 18086 | The sylow subgroups of a finite group are exactly the groups which have cardinality equal to the maximum power of 𝑃 dividing the group. (Contributed by Mario Carneiro, 16-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) ⇒ ⊢ ((𝐺 ∈ Grp ∧ 𝑋 ∈ Fin ∧ 𝑃 ∈ ℙ) → (𝐻 ∈ (𝑃 pSyl 𝐺) ↔ (𝐻 ∈ (SubGrp‘𝐺) ∧ (#‘𝐻) = (𝑃↑(𝑃 pCnt (#‘𝑋)))))) | ||
Theorem | sylow2 18087* | Sylow's second theorem. See also sylow2b 18084 for the "hard" part of the proof. Any two Sylow 𝑃-subgroups are conjugate to one another, and hence the same size, namely 𝑃↑(𝑃 pCnt ∣ 𝑋 ∣ ) (see fislw 18086). This is part of Metamath 100 proof #72. (Contributed by Mario Carneiro, 18-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝐻 ∈ (𝑃 pSyl 𝐺)) & ⊢ (𝜑 → 𝐾 ∈ (𝑃 pSyl 𝐺)) & ⊢ + = (+g‘𝐺) & ⊢ − = (-g‘𝐺) ⇒ ⊢ (𝜑 → ∃𝑔 ∈ 𝑋 𝐻 = ran (𝑥 ∈ 𝐾 ↦ ((𝑔 + 𝑥) − 𝑔))) | ||
Theorem | sylow3lem1 18088* | Lemma for sylow3 18094, first part. (Contributed by Mario Carneiro, 19-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ + = (+g‘𝐺) & ⊢ − = (-g‘𝐺) & ⊢ ⊕ = (𝑥 ∈ 𝑋, 𝑦 ∈ (𝑃 pSyl 𝐺) ↦ ran (𝑧 ∈ 𝑦 ↦ ((𝑥 + 𝑧) − 𝑥))) ⇒ ⊢ (𝜑 → ⊕ ∈ (𝐺 GrpAct (𝑃 pSyl 𝐺))) | ||
Theorem | sylow3lem2 18089* | Lemma for sylow3 18094, first part. The stabilizer of a given Sylow subgroup 𝐾 in the group action ⊕ acting on all of 𝐺 is the normalizer NG(K). (Contributed by Mario Carneiro, 19-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ + = (+g‘𝐺) & ⊢ − = (-g‘𝐺) & ⊢ ⊕ = (𝑥 ∈ 𝑋, 𝑦 ∈ (𝑃 pSyl 𝐺) ↦ ran (𝑧 ∈ 𝑦 ↦ ((𝑥 + 𝑧) − 𝑥))) & ⊢ (𝜑 → 𝐾 ∈ (𝑃 pSyl 𝐺)) & ⊢ 𝐻 = {𝑢 ∈ 𝑋 ∣ (𝑢 ⊕ 𝐾) = 𝐾} & ⊢ 𝑁 = {𝑥 ∈ 𝑋 ∣ ∀𝑦 ∈ 𝑋 ((𝑥 + 𝑦) ∈ 𝐾 ↔ (𝑦 + 𝑥) ∈ 𝐾)} ⇒ ⊢ (𝜑 → 𝐻 = 𝑁) | ||
Theorem | sylow3lem3 18090* | Lemma for sylow3 18094, first part. The number of Sylow subgroups is the same as the index (number of cosets) of the normalizer of the Sylow subgroup 𝐾. (Contributed by Mario Carneiro, 19-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ + = (+g‘𝐺) & ⊢ − = (-g‘𝐺) & ⊢ ⊕ = (𝑥 ∈ 𝑋, 𝑦 ∈ (𝑃 pSyl 𝐺) ↦ ran (𝑧 ∈ 𝑦 ↦ ((𝑥 + 𝑧) − 𝑥))) & ⊢ (𝜑 → 𝐾 ∈ (𝑃 pSyl 𝐺)) & ⊢ 𝐻 = {𝑢 ∈ 𝑋 ∣ (𝑢 ⊕ 𝐾) = 𝐾} & ⊢ 𝑁 = {𝑥 ∈ 𝑋 ∣ ∀𝑦 ∈ 𝑋 ((𝑥 + 𝑦) ∈ 𝐾 ↔ (𝑦 + 𝑥) ∈ 𝐾)} ⇒ ⊢ (𝜑 → (#‘(𝑃 pSyl 𝐺)) = (#‘(𝑋 / (𝐺 ~QG 𝑁)))) | ||
Theorem | sylow3lem4 18091* | Lemma for sylow3 18094, first part. The number of Sylow subgroups is a divisor of the size of 𝐺 reduced by the size of a Sylow subgroup of 𝐺. (Contributed by Mario Carneiro, 19-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ + = (+g‘𝐺) & ⊢ − = (-g‘𝐺) & ⊢ ⊕ = (𝑥 ∈ 𝑋, 𝑦 ∈ (𝑃 pSyl 𝐺) ↦ ran (𝑧 ∈ 𝑦 ↦ ((𝑥 + 𝑧) − 𝑥))) & ⊢ (𝜑 → 𝐾 ∈ (𝑃 pSyl 𝐺)) & ⊢ 𝐻 = {𝑢 ∈ 𝑋 ∣ (𝑢 ⊕ 𝐾) = 𝐾} & ⊢ 𝑁 = {𝑥 ∈ 𝑋 ∣ ∀𝑦 ∈ 𝑋 ((𝑥 + 𝑦) ∈ 𝐾 ↔ (𝑦 + 𝑥) ∈ 𝐾)} ⇒ ⊢ (𝜑 → (#‘(𝑃 pSyl 𝐺)) ∥ ((#‘𝑋) / (𝑃↑(𝑃 pCnt (#‘𝑋))))) | ||
Theorem | sylow3lem5 18092* | Lemma for sylow3 18094, second part. Reduce the group action of sylow3lem1 18088 to a given Sylow subgroup. (Contributed by Mario Carneiro, 19-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ + = (+g‘𝐺) & ⊢ − = (-g‘𝐺) & ⊢ (𝜑 → 𝐾 ∈ (𝑃 pSyl 𝐺)) & ⊢ ⊕ = (𝑥 ∈ 𝐾, 𝑦 ∈ (𝑃 pSyl 𝐺) ↦ ran (𝑧 ∈ 𝑦 ↦ ((𝑥 + 𝑧) − 𝑥))) ⇒ ⊢ (𝜑 → ⊕ ∈ ((𝐺 ↾s 𝐾) GrpAct (𝑃 pSyl 𝐺))) | ||
Theorem | sylow3lem6 18093* | Lemma for sylow3 18094, second part. Using the lemma sylow2a 18080, show that the number of sylow subgroups is equivalent mod 𝑃 to the number of fixed points under the group action. But 𝐾 is the unique element of the set of Sylow subgroups that is fixed under the group action, so there is exactly one fixed point and so ((#‘(𝑃 pSyl 𝐺)) mod 𝑃) = 1. (Contributed by Mario Carneiro, 19-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ + = (+g‘𝐺) & ⊢ − = (-g‘𝐺) & ⊢ (𝜑 → 𝐾 ∈ (𝑃 pSyl 𝐺)) & ⊢ ⊕ = (𝑥 ∈ 𝐾, 𝑦 ∈ (𝑃 pSyl 𝐺) ↦ ran (𝑧 ∈ 𝑦 ↦ ((𝑥 + 𝑧) − 𝑥))) & ⊢ 𝑁 = {𝑥 ∈ 𝑋 ∣ ∀𝑦 ∈ 𝑋 ((𝑥 + 𝑦) ∈ 𝑠 ↔ (𝑦 + 𝑥) ∈ 𝑠)} ⇒ ⊢ (𝜑 → ((#‘(𝑃 pSyl 𝐺)) mod 𝑃) = 1) | ||
Theorem | sylow3 18094 | Sylow's third theorem. The number of Sylow subgroups is a divisor of ∣ 𝐺 ∣ / 𝑑, where 𝑑 is the common order of a Sylow subgroup, and is equivalent to 1 mod 𝑃. This is part of Metamath 100 proof #72. (Contributed by Mario Carneiro, 19-Jan-2015.) |
⊢ 𝑋 = (Base‘𝐺) & ⊢ (𝜑 → 𝐺 ∈ Grp) & ⊢ (𝜑 → 𝑋 ∈ Fin) & ⊢ (𝜑 → 𝑃 ∈ ℙ) & ⊢ 𝑁 = (#‘(𝑃 pSyl 𝐺)) ⇒ ⊢ (𝜑 → (𝑁 ∥ ((#‘𝑋) / (𝑃↑(𝑃 pCnt (#‘𝑋)))) ∧ (𝑁 mod 𝑃) = 1)) | ||
Syntax | clsm 18095 | Extend class notation with subgroup sum. |
class LSSum | ||
Syntax | cpj1 18096 | Extend class notation with left projection. |
class proj1 | ||
Definition | df-lsm 18097* | Define subgroup sum (inner direct product of subgroups). (Contributed by NM, 28-Jan-2014.) |
⊢ LSSum = (𝑤 ∈ V ↦ (𝑡 ∈ 𝒫 (Base‘𝑤), 𝑢 ∈ 𝒫 (Base‘𝑤) ↦ ran (𝑥 ∈ 𝑡, 𝑦 ∈ 𝑢 ↦ (𝑥(+g‘𝑤)𝑦)))) | ||
Definition | df-pj1 18098* | Define the left projection function, which takes two subgroups 𝑡, 𝑢 with trivial intersection and returns a function mapping the elements of the subgroup sum 𝑡 + 𝑢 to their projections onto 𝑡. (The other projection function can be obtained by swapping the roles of 𝑡 and 𝑢.) (Contributed by Mario Carneiro, 15-Oct-2015.) |
⊢ proj1 = (𝑤 ∈ V ↦ (𝑡 ∈ 𝒫 (Base‘𝑤), 𝑢 ∈ 𝒫 (Base‘𝑤) ↦ (𝑧 ∈ (𝑡(LSSum‘𝑤)𝑢) ↦ (℩𝑥 ∈ 𝑡 ∃𝑦 ∈ 𝑢 𝑧 = (𝑥(+g‘𝑤)𝑦))))) | ||
Theorem | lsmfval 18099* | The subgroup sum function (for a group or vector space). (Contributed by NM, 28-Jan-2014.) (Revised by Mario Carneiro, 19-Apr-2016.) |
⊢ 𝐵 = (Base‘𝐺) & ⊢ + = (+g‘𝐺) & ⊢ ⊕ = (LSSum‘𝐺) ⇒ ⊢ (𝐺 ∈ 𝑉 → ⊕ = (𝑡 ∈ 𝒫 𝐵, 𝑢 ∈ 𝒫 𝐵 ↦ ran (𝑥 ∈ 𝑡, 𝑦 ∈ 𝑢 ↦ (𝑥 + 𝑦)))) | ||
Theorem | lsmvalx 18100* | Subspace sum value (for a group or vector space). Extended domain version of lsmval 18109. (Contributed by NM, 28-Jan-2014.) (Revised by Mario Carneiro, 19-Apr-2016.) |
⊢ 𝐵 = (Base‘𝐺) & ⊢ + = (+g‘𝐺) & ⊢ ⊕ = (LSSum‘𝐺) ⇒ ⊢ ((𝐺 ∈ 𝑉 ∧ 𝑇 ⊆ 𝐵 ∧ 𝑈 ⊆ 𝐵) → (𝑇 ⊕ 𝑈) = ran (𝑥 ∈ 𝑇, 𝑦 ∈ 𝑈 ↦ (𝑥 + 𝑦))) |
< Previous Next > |
Copyright terms: Public domain | < Previous Next > |