X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=matita%2Flibrary%2Falgebra%2Ffinite_groups.ma;h=64c65b266e50f47a06e879abeaedca66156a2f02;hb=2fa59f0450a2f1fe871a09fd9841ddc1bfd67080;hp=2078cf1d75cccc151280cc54877f25d1e41156e9;hpb=9e32b487a1343ead9069ea76e40515ceb19d26d6;p=helm.git diff --git a/matita/library/algebra/finite_groups.ma b/matita/library/algebra/finite_groups.ma index 2078cf1d7..64c65b266 100644 --- a/matita/library/algebra/finite_groups.ma +++ b/matita/library/algebra/finite_groups.ma @@ -16,7 +16,7 @@ set "baseuri" "cic:/matita/algebra/finite_groups/". include "algebra/groups.ma". -record finite_enumerable (T:Type) : Type ≝ +record finite_enumerable (T:Type) : Type≝ { order: nat; repr: nat → T; index_of: T → nat; @@ -39,7 +39,7 @@ for @{ 'card $C }. interpretation "Finite_enumerable order" 'card C = (cic:/matita/algebra/finite_groups/order.con C _). -record finite_enumerable_SemiGroup : Type ≝ +record finite_enumerable_SemiGroup : Type≝ { semigroup:> SemiGroup; is_finite_enumerable:> finite_enumerable semigroup }. @@ -60,7 +60,7 @@ interpretation "Index_of_finite_enumerable representation" (* several definitions/theorems to be moved somewhere else *) -definition ltb ≝ λn,m. leb n m ∧ notb (eqb n m). +definition ltb ≝λn,m. leb n m ∧ notb (eqb n m). theorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n (repr_index_of ? (is_finite_enumerable G)) in HH; apply (ex_intro ? ? (G \sub a)); - letin GOGO ≝ (refl_eq ? (repr ? (is_finite_enumerable G) O)); + letin GOGO ≝(refl_eq ? (repr ? (is_finite_enumerable G) O)); clearbody GOGO; rewrite < HH in GOGO; rewrite < HH in GOGO:(? ? % ?); rewrite > (op_associative ? G) in GOGO; - letin GaGa ≝ (H ? ? ? GOGO); + letin GaGa ≝(H ? ? ? GOGO); clearbody GaGa; clear GOGO; constructor 1; [ simplify; apply (semigroup_properties G) | unfold is_left_unit; intro; - letin GaxGax ≝ (refl_eq ? (G \sub a ·x)); - clearbody GaxGax; + letin GaxGax ≝(refl_eq ? (G \sub a ·x)); + clearbody GaxGax; (* demo *) rewrite < GaGa in GaxGax:(? ? % ?); rewrite > (op_associative ? (semigroup_properties G)) in GaxGax; apply (H ? ? ? GaxGax) | unfold is_right_unit; intro; - letin GaxGax ≝ (refl_eq ? (x·G \sub a)); + letin GaxGax ≝(refl_eq ? (x·G \sub a)); clearbody GaxGax; rewrite < GaGa in GaxGax:(? ? % ?); rewrite < (op_associative ? (semigroup_properties G)) in GaxGax; @@ -516,7 +516,7 @@ cut (∀n.n ≤ order ? (is_finite_enumerable G) → ∃m.f m = n); elim H3; assumption | intros; - change in H5 with (ι(G \sub O · G \sub x) = ι(G \sub O · G \sub y)); + simplify in H5; cut (G \sub (ι(G \sub O · G \sub x)) = G \sub (ι(G \sub O · G \sub y))); [ rewrite > (repr_index_of ? ? (G \sub O · G \sub x)) in Hcut; rewrite > (repr_index_of ? ? (G \sub O · G \sub y)) in Hcut;