X-Git-Url: http://matita.cs.unibo.it/gitweb/?a=blobdiff_plain;f=helm%2Fsoftware%2Fmatita%2Flibrary%2Falgebra%2Ffinite_groups.ma;h=8dfa31342e468a2a45676217e7cbf93e3b3101c1;hb=6d297b12c480352eb2f156ab4515f73921ea2e81;hp=7f76ae0c852f62c5b88d50818c6821894bb84508;hpb=683978a2627cf1ce15673360f26806593d22f7b5;p=helm.git diff --git a/helm/software/matita/library/algebra/finite_groups.ma b/helm/software/matita/library/algebra/finite_groups.ma index 7f76ae0c8..8dfa31342 100644 --- a/helm/software/matita/library/algebra/finite_groups.ma +++ b/helm/software/matita/library/algebra/finite_groups.ma @@ -12,8 +12,6 @@ (* *) (**************************************************************************) -set "baseuri" "cic:/matita/algebra/finite_groups/". - include "algebra/groups.ma". record finite_enumerable (T:Type) : Type≝ @@ -33,9 +31,6 @@ for @{ 'repr $C $i }. interpretation "Finite_enumerable representation" 'repr C i = (cic:/matita/algebra/finite_groups/repr.con C _ i).*) -notation < "hvbox(|C|)" with precedence 89 -for @{ 'card $C }. - interpretation "Finite_enumerable order" 'card C = (cic:/matita/algebra/finite_groups/order.con C _). @@ -48,7 +43,7 @@ interpretation "Finite_enumerable representation" 'repr S i = (cic:/matita/algebra/finite_groups/repr.con S (cic:/matita/algebra/finite_groups/is_finite_enumerable.con S) i). -notation "hvbox(ι e)" with precedence 60 +notation "hvbox(\iota e)" with precedence 60 for @{ 'index_of_finite_enumerable_semigroup $e }. interpretation "Index_of_finite_enumerable representation" @@ -60,140 +55,6 @@ 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). - -theorem not_eq_to_le_to_lt: ∀n,m. n≠m → n≤m → n (S_pred m); - [ apply le_S_S; - assumption - | assumption - ] -]. -qed. - -theorem le_to_le_pred: - ∀n,m. n ≤ m → pred n ≤ pred m. -intros 2; -elim n; -[ simplify; - apply le_O_n -| simplify; - generalize in match H1; - clear H1; - elim m; - [ elim (not_le_Sn_O ? H1) - | simplify; - apply le_S_S_to_le; - assumption - ] -]. -qed. - -theorem lt_n_m_to_not_lt_m_Sn: ∀n,m. n < m → m ≮ S n. -intros; -unfold Not; -intro; -unfold lt in H; -unfold lt in H1; -generalize in match (le_S_S ? ? H); -intro; -generalize in match (transitive_le ? ? ? H2 H1); -intro; -apply (not_le_Sn_n ? H3). -qed. - - - -(* moved to nat/order.ma -theorem lt_O_S: ∀n. O < S n. -intro; -unfold lt; -apply le_S_S; -apply le_O_n. -qed. *) - -theorem le_n_m_to_lt_m_Sn_to_eq_n_m: ∀n,m. n ≤ m → m < S n → n=m. -intros; -unfold lt in H1; -generalize in match (le_S_S_to_le ? ? H1); -intro; -apply cic:/matita/nat/orders/antisym_le.con; -assumption. -qed. - theorem pigeonhole: ∀n:nat.∀f:nat→nat. (∀x,y.x≤n → y≤n → f x = f y → x=y) → @@ -227,6 +88,7 @@ elim n; [ simplify in H5; clear Hcut; clear Hcut1; + unfold f' in H5; clear f'; elim H5; clear H5; @@ -295,6 +157,7 @@ elim n; [ simplify in H5; clear Hcut; clear Hcut1; + unfold f' in H5; clear f'; elim H5; clear H5; @@ -307,7 +170,7 @@ elim n; apply (ltb_elim (f (S n1)) (f a)); [ simplify; intros; - generalize in match (lt_S_S ? ? H5); + generalize in match (lt_to_lt_S_S ? ? H5); intro; rewrite < S_pred in H6; [ elim (lt_n_m_to_not_lt_m_Sn ? ? H4 H6) @@ -385,7 +248,8 @@ elim n; [ apply (H1 ? ? ? ? Hcut); apply le_S; assumption - | apply eq_pred_to_eq; + | alias id "eq_pred_to_eq" = "cic:/matita/nat/relevant_equations/eq_pred_to_eq.con". +apply eq_pred_to_eq; [ apply (ltn_to_ltO ? ? H7) | apply (ltn_to_ltO ? ? H6) | assumption @@ -528,6 +392,7 @@ cut (∀n.n ≤ order ? (is_finite_enumerable G) → ∃m.f m = n); assumption ] | intros; + unfold f; apply index_of_sur ] ].