1 (**************************************************************************)
4 (* ||A|| A project by Andrea Asperti *)
6 (* ||I|| Developers: *)
7 (* ||T|| The HELM team. *)
8 (* ||A|| http://helm.cs.unibo.it *)
10 (* \ / This file is distributed under the terms of the *)
11 (* v GNU General Public License Version 2 *)
13 (**************************************************************************)
15 include "algebra/groups.ma".
16 include "nat/relevant_equations.ma".
18 record finite_enumerable (T:Type) : Type≝
22 index_of_sur: ∀x.index_of x ≤ order;
23 index_of_repr: ∀n. n≤order → index_of (repr n) = n;
24 repr_index_of: ∀x. repr (index_of x) = x
27 notation "hvbox(C \sub i)" with precedence 89
30 (* CSC: multiple interpretations in the same file are not considered in the
32 interpretation "Finite_enumerable representation" 'repr C i = (repr C ? i).*)
34 interpretation "Finite_enumerable order" 'card C = (order C ?).
36 record finite_enumerable_SemiGroup : Type≝
37 { semigroup:> SemiGroup;
38 is_finite_enumerable:> finite_enumerable semigroup
41 interpretation "Finite_enumerable representation" 'repr S i =
42 (repr S (is_finite_enumerable S) i).
44 notation "hvbox(\iota e)" with precedence 60
45 for @{ 'index_of_finite_enumerable_semigroup $e }.
47 interpretation "Index_of_finite_enumerable representation"
48 'index_of_finite_enumerable_semigroup e
50 (index_of ? (is_finite_enumerable ?) e).
53 (* several definitions/theorems to be moved somewhere else *)
57 (∀x,y.x≤n → y≤n → f x = f y → x=y) →
58 (∀m. m ≤ n → f m ≤ n) →
59 ∀x. x≤n → ∃y.f y = x ∧ y ≤ n.
62 [ apply (ex_intro ? ? O);
64 [ rewrite < (le_n_O_to_eq ? H2);
65 rewrite < (le_n_O_to_eq ? (H1 O ?));
76 match ltb fSn1 fx with
80 cut (∀x,y. x ≤ n1 → y ≤ n1 → f' x = f' y → x=y);
81 [ cut (∀x. x ≤ n1 → f' x ≤ n1);
82 [ apply (nat_compare_elim (f (S n1)) x);
84 elim (H f' ? ? (pred x));
92 apply (ex_intro ? ? a);
94 [ generalize in match (eq_f ? ? S ? ? H6);
97 rewrite < S_pred in H5;
98 [ generalize in match H4;
102 apply (ltb_elim (f (S n1)) (f a));
107 | apply (ltn_to_ltO ? ? H4)
111 generalize in match (not_lt_to_le ? ? H4);
114 generalize in match (le_n_m_to_lt_m_Sn_to_eq_n_m ? ? H6 H5);
116 generalize in match (H1 ? ? ? ? H4);
118 generalize in match (le_n_m_to_lt_m_Sn_to_eq_n_m ? ? H6 H5);
120 generalize in match (H1 ? ? ? ? H9);
123 elim (not_le_Sn_n ? H7)
133 | apply (ltn_to_ltO ? ? H4)
140 | apply le_S_S_to_le;
143 | apply (ltn_to_ltO ? ? H4)
147 apply (ex_intro ? ? (S n1));
161 apply (ex_intro ? ? a);
163 [ generalize in match H4;
167 apply (ltb_elim (f (S n1)) (f a));
170 generalize in match (lt_to_lt_S_S ? ? H5);
172 rewrite < S_pred in H6;
173 [ elim (lt_n_m_to_not_lt_m_Sn ? ? H4 H6)
174 | apply (ltn_to_ltO ? ? H4)
185 | rewrite > (pred_Sn n1);
187 generalize in match (H2 (S n1));
189 generalize in match (lt_to_le_to_lt ? ? ? H4 (H5 (le_n ?)));
199 apply (ltb_elim (f (S n1)) (f x1));
202 [ generalize in match (H2 x1);
204 change in match n1 with (pred (S n1));
209 | generalize in match (H2 (S n1) (le_n ?));
211 generalize in match (not_lt_to_le ? ? H4);
213 generalize in match (transitive_le ? ? ? H7 H6);
215 cut (f x1 ≠ f (S n1));
216 [ generalize in match (not_eq_to_le_to_lt ? ? Hcut1 H7);
219 generalize in match (transitive_le ? ? ? H9 H6);
225 generalize in match (H1 ? ? ? ? H9);
228 apply (not_le_Sn_n ? H5)
239 apply (ltb_elim (f (S n1)) (f x1));
241 apply (ltb_elim (f (S n1)) (f y));
245 [ apply (H1 ? ? ? ? Hcut);
248 | apply eq_pred_to_eq;
249 [ apply (ltn_to_ltO ? ? H7)
250 | apply (ltn_to_ltO ? ? H6)
254 | (* pred (f x1) = f y absurd since y ≠ S n1 and thus f y ≠ f (S n1)
255 so that f y < f (S n1) < f x1; hence pred (f x1) = f y is absurd *)
257 [ generalize in match (lt_to_not_eq ? ? Hcut);
259 cut (f y ≠ f (S n1));
260 [ cut (f y < f (S n1));
261 [ rewrite < H8 in Hcut2;
264 generalize in match (le_S_S ? ? Hcut2);
266 generalize in match (transitive_le ? ? ? H10 H7);
268 rewrite < (S_pred (f x1)) in H11;
269 [ elim (not_le_Sn_n ? H11)
270 | fold simplify ((f (S n1)) < (f x1)) in H7;
271 apply (ltn_to_ltO ? ? H7)
273 | apply not_eq_to_le_to_lt;
275 | apply not_lt_to_le;
282 apply (H1 ? ? ? ? H10);
292 | (* f x1 = pred (f y) absurd since it implies S (f x1) = f y and
293 f x1 ≤ f (S n1) < f y = S (f x1) so that f x1 = f (S n1); by
294 injectivity x1 = S n1 that is absurd since x1 ≤ n1 *)
295 generalize in match (eq_f ? ? S ? ? H8);
297 rewrite < S_pred in H9;
298 [ rewrite < H9 in H6;
299 generalize in match (not_lt_to_le ? ? H7);
302 generalize in match (le_S_S ? ? H10);
304 generalize in match (antisym_le ? ? H11 H6);
306 generalize in match (inj_S ? ? H12);
308 generalize in match (H1 ? ? ? ? H13);
311 elim (not_le_Sn_n ? H4)
316 | apply (ltn_to_ltO ? ? H6)
318 | apply (H1 ? ? ? ? H8);
326 theorem finite_enumerable_SemiGroup_to_left_cancellable_to_right_cancellable_to_isMonoid:
327 ∀G:finite_enumerable_SemiGroup.
328 left_cancellable ? (op G) →
329 right_cancellable ? (op G) →
330 ∃e:G. isMonoid (mk_PreMonoid G e).
332 letin f ≝(λn.ι(G \sub O · G \sub n));
333 cut (∀n.n ≤ order ? (is_finite_enumerable G) → ∃m.f m = n);
334 [ letin EX ≝(Hcut O ?);
341 letin HH ≝(eq_f ? ? (repr ? (is_finite_enumerable G)) ? ? H2);
343 rewrite > (repr_index_of ? (is_finite_enumerable G)) in HH;
344 apply (ex_intro ? ? (G \sub a));
345 letin GOGO ≝(refl_eq ? (repr ? (is_finite_enumerable G) O));
347 rewrite < HH in GOGO;
348 rewrite < HH in GOGO:(? ? % ?);
349 rewrite > (op_associative ? G) in GOGO;
350 letin GaGa ≝(H ? ? ? GOGO);
355 apply (semigroup_properties G)
356 | unfold is_left_unit; intro;
357 letin GaxGax ≝(refl_eq ? (G \sub a ·x));
358 clearbody GaxGax; (* demo *)
359 rewrite < GaGa in GaxGax:(? ? % ?);
360 rewrite > (op_associative ? (semigroup_properties G)) in GaxGax;
361 apply (H ? ? ? GaxGax)
362 | unfold is_right_unit; intro;
363 letin GaxGax ≝(refl_eq ? (x·G \sub a));
365 rewrite < GaGa in GaxGax:(? ? % ?);
366 rewrite < (op_associative ? (semigroup_properties G)) in GaxGax;
367 apply (H1 ? ? ? GaxGax)
371 elim (pigeonhole (order ? G) f ? ? ? H2);
372 [ apply (ex_intro ? ? a);
377 cut (G \sub (ι(G \sub O · G \sub x)) = G \sub (ι(G \sub O · G \sub y)));
378 [ rewrite > (repr_index_of ? ? (G \sub O · G \sub x)) in Hcut;
379 rewrite > (repr_index_of ? ? (G \sub O · G \sub y)) in Hcut;
380 generalize in match (H ? ? ? Hcut);
382 generalize in match (eq_f ? ? (index_of ? G) ? ? H6);
384 rewrite > index_of_repr in H7;
385 rewrite > index_of_repr in H7;