]> matita.cs.unibo.it Git - helm.git/blobdiff - helm/matita/library/algebra/groups.ma
added links to svn tarballs
[helm.git] / helm / matita / library / algebra / groups.ma
index ae6b525e3906d9b6f5b153b0c767ed6c47faed08..5cefa0db156a209cc9af9f26f05cf77857b87041 100644 (file)
@@ -16,6 +16,8 @@ set "baseuri" "cic:/matita/algebra/groups/".
 
 include "algebra/monoids.ma".
 include "nat/le_arith.ma".
+include "datatypes/bool.ma".
+include "nat/compare.ma".
 
 record PreGroup : Type ≝
  { premonoid:> PreMonoid;
@@ -115,7 +117,7 @@ for @{ 'repr $C $i }.
 interpretation "Finite_enumerable representation" 'repr C i =
  (cic:/matita/algebra/groups/repr.con C _ i).*)
  
-notation "hvbox(|C|)" with precedence 89
+notation "hvbox(|C|)" with precedence 89
 for @{ 'card $C }.
 
 interpretation "Finite_enumerable order" 'card C =
@@ -162,7 +164,323 @@ interpretation "Index_of_finite_enumerable representation"
 =
  (cic:/matita/algebra/groups/index_of.con _
   (cic:/matita/algebra/groups/is_finite_enumerable.con _) e).
+
+
+(* 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<m.
+intros;
+elim (le_to_or_lt_eq ? ? H1);
+[ assumption
+| elim (H H2)
+].
+qed.
+
+theorem ltb_to_Prop :
+ ∀n,m.
+  match ltb n m with
+  [ true ⇒ n < m
+  | false ⇒ n ≮ m
+  ].
+intros;
+unfold ltb;
+apply leb_elim;
+apply eqb_elim;
+intros;
+simplify;
+[ rewrite < H;
+  apply le_to_not_lt;
+  constructor 1
+| apply (not_eq_to_le_to_lt ? ? H H1)
+| rewrite < H;
+  apply le_to_not_lt;
+  constructor 1
+| apply le_to_not_lt;
+  generalize in match (not_le_to_lt ? ? H1);
+  clear H1;
+  intro;
+  apply lt_to_le;
+  assumption
+].
+qed.
+
+theorem ltb_elim: \forall n,m:nat. \forall P:bool \to Prop.
+(n < m \to (P true)) \to (n ≮ m \to (P false)) \to
+P (ltb n m).
+intros.
+cut
+(match (ltb n m) with
+[ true  \Rightarrow n < m
+| false \Rightarrow n ≮ m] \to (P (ltb n m))).
+apply Hcut.apply ltb_to_Prop.
+elim (ltb n m).
+apply ((H H2)).
+apply ((H1 H2)).
+qed.
+
+theorem Not_lt_n_n: ∀n. n ≮ n.
+intro;
+unfold Not;
+intro;
+unfold lt in H;
+apply (not_le_Sn_n ? H).
+qed.
+
+theorem eq_pred_to_eq:
+ ∀n,m. O < n → O < m → pred n = pred m → n = m.
+intros 2;
+elim n;
+[ elim (Not_lt_n_n ? H)
+| generalize in match H3;
+  clear H3;
+  generalize in match H2;
+  clear H2;
+  elim m;
+  [ elim (Not_lt_n_n ? H2)
+  | simplify in H4;
+    apply (eq_f ? ? S);
+    assumption
+  ]
+].
+qed.
+
+theorem le_pred_to_le:
+ ∀n,m. O < m → pred n ≤ pred m \to n ≤ m.
+intros 2;
+elim n;
+[ apply le_O_n
+| simplify in H2;
+  rewrite > (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 pigeonhole:
+ ∀n:nat.∀f:nat→nat.
+  (∀x,y.x≤n → y≤n → f x = f y → x=y) →
+  (∀m. m ≤ n → f m ≤ n) →
+   ∀x. x≤n \to ∃y.f y = x ∧ y ≤ n.
+intro;
+elim n;
+[ apply (ex_intro ? ? O);
+  split;
+  [ rewrite < (le_n_O_to_eq ? H2);
+    rewrite < (le_n_O_to_eq ? (H1 O ?));
+    [ reflexivity
+    | apply le_n
+    ]
+  | apply le_n
+  ]
+| clear n;
+  apply (nat_compare_elim (f (S n1)) x);
+  [ (* TODO: caso complicato, ma simile al terzo *) 
+  | intros;
+    apply (ex_intro ? ? (S n1));
+    split;
+    [ assumption
+    | constructor 1
+    ] 
+  | intro;
+    letin f' ≝
+     (λx.
+       let fSn1 ≝ f (S n1) in
+       let fx ≝ f x in
+        match ltb fSn1 fx with
+        [ true ⇒ pred fx
+        | false ⇒ fx
+        ]);
+    cut (∀x,y. x ≤ n1 → y ≤ n1 → f' x = f' y → x=y);
+    [ cut (∀x. x ≤ n1 → f' x ≤ n1);
+      [ elim (H f' ? ? x);
+        [ simplify in H5;
+          clear Hcut;
+          clear Hcut1;
+          clear f';
+          elim H5;
+          clear H5;
+          apply (ex_intro ? ? a);
+          split;
+          [ generalize in match H4;
+            clear H4;
+            rewrite < H6;
+            clear H6;
+            apply (ltb_elim (f (S n1)) (f a));
+            [ (* TODO: caso impossibile (uso l'iniettivita') *)
+              simplify;
+            | simplify;
+              intros;
+              reflexivity
+            ]        
+          | apply le_S;
+            assumption
+          ]
+        | apply Hcut    
+        | apply Hcut1
+        | rewrite > (pred_Sn n1);
+          simplify;
+          generalize in match (H2 (S n1));
+          intro;
+          generalize in match (lt_to_le_to_lt ? ? ? H4 (H5 (le_n ?)));
+          intro;
+          unfold lt in H6;
+          apply le_S_S_to_le;
+          assumption
+        ]
+      | unfold f';
+        simplify;
+        intro;
+        apply (ltb_elim (f (S n1)) (f x1));
+        simplify;
+        intros;
+        [ generalize in match (H2 x1);
+          intro;
+          change in match n1 with (pred (S n1));
+          apply le_to_le_pred;
+          apply H7;
+          apply le_S;
+          assumption
+        | generalize in match (H2 (S n1) (le_n ?));
+          intro;
+          generalize in match (not_lt_to_le ? ? H5);
+          intro;
+          generalize in match (transitive_le ? ? ? H8 H7);
+          intro;
+          cut (f x1 ≠ f (S n1));
+          [ generalize in match (not_eq_to_le_to_lt ? ? Hcut1 H8);
+            intro;
+            unfold lt in H10;
+            generalize in match (transitive_le ? ? ? H10 H7);
+            intro;
+            apply le_S_S_to_le;
+            assumption
+          | unfold Not;
+            intro;
+            generalize in match (H1 ? ? ? ? H10);
+            [ intro;
+              rewrite > H11 in H6;
+              apply (not_le_Sn_n ? H6)
+            | apply le_S;
+              assumption
+            | apply le_n
+            ]
+          ] 
+        ]
+      ]
+    | intros 4;
+      unfold f';
+      simplify;
+      apply (ltb_elim (f (S n1)) (f x1));
+      simplify;
+      apply (ltb_elim (f (S n1)) (f y));
+      simplify;
+      intros;
+      [ cut (f x1 = f y);
+        [ apply (H1 ? ? ? ? Hcut);
+          apply le_S;
+          assumption
+        | apply eq_pred_to_eq;
+          [ apply (ltn_to_ltO ? ? H8)
+          | apply (ltn_to_ltO ? ? H7)
+          | assumption
+          ]
+        ]         
+      | (* pred (f x1) = f y absurd since y ≠ S n1 and thus f y ≠ f (S n1)
+           so that f y < f (S n1) < f x1; hence pred (f x1) = f y is absurd *)
+         cut (y < S n1);
+         [ generalize in match (lt_to_not_eq ? ? Hcut);
+           intro;
+           cut (f y ≠ f (S n1));
+           [ cut (f y < f (S n1));
+             [ rewrite < H9 in Hcut2;
+               unfold lt in Hcut2;
+               unfold lt in H8;
+               generalize in match (le_S_S ? ? Hcut2);
+               intro;
+               generalize in match (transitive_le ? ? ? H11 H8);
+               intros;
+               rewrite < (S_pred (f x1)) in H12;
+                [ elim (not_le_Sn_n ? H12)
+                | fold simplify ((f (S n1)) < (f x1)) in H8;
+                 apply (ltn_to_ltO ? ? H8)
+                ]
+             | apply not_eq_to_le_to_lt;
+               [ assumption
+               | apply not_lt_to_le;
+                 assumption
+               ]
+             ]
+           | unfold Not;
+             intro;
+             apply H10;
+             apply (H1 ? ? ? ? H11);
+             [ apply lt_to_le;
+               assumption
+             | constructor 1
+             ]
+           ]
+         | unfold lt;
+           apply le_S_S;
+           assumption
+         ]
+      | (* f x1 = pred (f y) absurd since it implies S (f x1) = f y and
+           f x1 ≤ f (S n1) < f y = S (f x1) so that f x1 = f (S n1); by
+           injectivity x1 = S n1 that is absurd since x1 ≤ n1 *)
+         generalize in match (eq_f ? ? S ? ? H9);
+         intro;
+         rewrite < S_pred in H10;
+         [ rewrite < H10 in H7;
+           generalize in match (not_lt_to_le ? ? H8);
+           intro;
+           unfold lt in H7;
+           generalize in match (le_S_S ? ? H11);
+           intro;
+           generalize in match (antisym_le ? ? H12 H7);
+           intro;
+           generalize in match (inj_S ? ? H13);
+           intro;
+           generalize in match (H1 ? ? ? ? H14);
+           [ intro;
+             rewrite > H15 in H5;
+             elim (not_le_Sn_n ? H5)
+           | apply le_S;
+             assumption
+           | apply le_n
+           ]
+         | apply (ltn_to_ltO ?? H7) 
+         ]
+      | apply (H1 ? ? ? ? H9);
+        apply le_S;
+        assumption
+      ]
+    ]
+  ]
+].
+qed.
+
 theorem foo:
  ∀G:finite_enumerable_SemiGroup.
   left_cancellable ? (op G) →
@@ -207,5 +525,5 @@ cut (∀n.n ≤ order ? (is_finite_enumerable G) → ∃m.f m = n);
       apply (H1 ? ? ? GaxGax)
     ]
   ]
-|
+| apply pigeonhole
 ].