| 47295 |      1 | (*  Title:      HOL/Isar_Examples/Group_Context.thy
 | 
|  |      2 |     Author:     Makarius
 | 
|  |      3 | *)
 | 
|  |      4 | 
 | 
| 58882 |      5 | section \<open>Some algebraic identities derived from group axioms -- theory context version\<close>
 | 
| 47295 |      6 | 
 | 
|  |      7 | theory Group_Context
 | 
| 63585 |      8 |   imports Main
 | 
| 47295 |      9 | begin
 | 
|  |     10 | 
 | 
| 58614 |     11 | text \<open>hypothetical group axiomatization\<close>
 | 
| 47295 |     12 | 
 | 
|  |     13 | context
 | 
| 61797 |     14 |   fixes prod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a"  (infixl "\<odot>" 70)
 | 
| 47295 |     15 |     and one :: "'a"
 | 
| 55656 |     16 |     and inverse :: "'a \<Rightarrow> 'a"
 | 
| 61797 |     17 |   assumes assoc: "(x \<odot> y) \<odot> z = x \<odot> (y \<odot> z)"
 | 
|  |     18 |     and left_one: "one \<odot> x = x"
 | 
|  |     19 |     and left_inverse: "inverse x \<odot> x = one"
 | 
| 47295 |     20 | begin
 | 
|  |     21 | 
 | 
| 58614 |     22 | text \<open>some consequences\<close>
 | 
| 47295 |     23 | 
 | 
| 61797 |     24 | lemma right_inverse: "x \<odot> inverse x = one"
 | 
| 47295 |     25 | proof -
 | 
| 61797 |     26 |   have "x \<odot> inverse x = one \<odot> (x \<odot> inverse x)"
 | 
| 47295 |     27 |     by (simp only: left_one)
 | 
| 61797 |     28 |   also have "\<dots> = one \<odot> x \<odot> inverse x"
 | 
| 47295 |     29 |     by (simp only: assoc)
 | 
| 61797 |     30 |   also have "\<dots> = inverse (inverse x) \<odot> inverse x \<odot> x \<odot> inverse x"
 | 
| 47295 |     31 |     by (simp only: left_inverse)
 | 
| 61797 |     32 |   also have "\<dots> = inverse (inverse x) \<odot> (inverse x \<odot> x) \<odot> inverse x"
 | 
| 47295 |     33 |     by (simp only: assoc)
 | 
| 61797 |     34 |   also have "\<dots> = inverse (inverse x) \<odot> one \<odot> inverse x"
 | 
| 47295 |     35 |     by (simp only: left_inverse)
 | 
| 61797 |     36 |   also have "\<dots> = inverse (inverse x) \<odot> (one \<odot> inverse x)"
 | 
| 47295 |     37 |     by (simp only: assoc)
 | 
| 61797 |     38 |   also have "\<dots> = inverse (inverse x) \<odot> inverse x"
 | 
| 47295 |     39 |     by (simp only: left_one)
 | 
|  |     40 |   also have "\<dots> = one"
 | 
|  |     41 |     by (simp only: left_inverse)
 | 
| 61797 |     42 |   finally show ?thesis .
 | 
| 47295 |     43 | qed
 | 
|  |     44 | 
 | 
| 61797 |     45 | lemma right_one: "x \<odot> one = x"
 | 
| 47295 |     46 | proof -
 | 
| 61797 |     47 |   have "x \<odot> one = x \<odot> (inverse x \<odot> x)"
 | 
| 47295 |     48 |     by (simp only: left_inverse)
 | 
| 61797 |     49 |   also have "\<dots> = x \<odot> inverse x \<odot> x"
 | 
| 47295 |     50 |     by (simp only: assoc)
 | 
| 61797 |     51 |   also have "\<dots> = one \<odot> x"
 | 
| 47295 |     52 |     by (simp only: right_inverse)
 | 
|  |     53 |   also have "\<dots> = x"
 | 
|  |     54 |     by (simp only: left_one)
 | 
| 61797 |     55 |   finally show ?thesis .
 | 
| 47295 |     56 | qed
 | 
|  |     57 | 
 | 
| 47872 |     58 | lemma one_equality:
 | 
| 61797 |     59 |   assumes eq: "e \<odot> x = x"
 | 
| 47872 |     60 |   shows "one = e"
 | 
| 47295 |     61 | proof -
 | 
| 61797 |     62 |   have "one = x \<odot> inverse x"
 | 
| 47295 |     63 |     by (simp only: right_inverse)
 | 
| 61797 |     64 |   also have "\<dots> = (e \<odot> x) \<odot> inverse x"
 | 
| 47295 |     65 |     by (simp only: eq)
 | 
| 61797 |     66 |   also have "\<dots> = e \<odot> (x \<odot> inverse x)"
 | 
| 47295 |     67 |     by (simp only: assoc)
 | 
| 61797 |     68 |   also have "\<dots> = e \<odot> one"
 | 
| 47295 |     69 |     by (simp only: right_inverse)
 | 
|  |     70 |   also have "\<dots> = e"
 | 
|  |     71 |     by (simp only: right_one)
 | 
| 61797 |     72 |   finally show ?thesis .
 | 
| 47295 |     73 | qed
 | 
|  |     74 | 
 | 
| 47872 |     75 | lemma inverse_equality:
 | 
| 61797 |     76 |   assumes eq: "x' \<odot> x = one"
 | 
| 47872 |     77 |   shows "inverse x = x'"
 | 
| 47295 |     78 | proof -
 | 
| 61797 |     79 |   have "inverse x = one \<odot> inverse x"
 | 
| 47295 |     80 |     by (simp only: left_one)
 | 
| 61797 |     81 |   also have "\<dots> = (x' \<odot> x) \<odot> inverse x"
 | 
| 47295 |     82 |     by (simp only: eq)
 | 
| 61797 |     83 |   also have "\<dots> = x' \<odot> (x \<odot> inverse x)"
 | 
| 47295 |     84 |     by (simp only: assoc)
 | 
| 61797 |     85 |   also have "\<dots> = x' \<odot> one"
 | 
| 47295 |     86 |     by (simp only: right_inverse)
 | 
|  |     87 |   also have "\<dots> = x'"
 | 
|  |     88 |     by (simp only: right_one)
 | 
| 61797 |     89 |   finally show ?thesis .
 | 
| 47295 |     90 | qed
 | 
|  |     91 | 
 | 
|  |     92 | end
 | 
|  |     93 | 
 | 
|  |     94 | end
 |