author | nipkow |
Tue, 17 Jun 2025 14:11:40 +0200 | |
changeset 82733 | 8b537e1af2ec |
parent 80914 | d97fdabd9e2b |
permissions | -rw-r--r-- |
47295 | 1 |
(* Title: HOL/Isar_Examples/Group_Notepad.thy |
2 |
Author: Makarius |
|
3 |
*) |
|
4 |
||
58882 | 5 |
section \<open>Some algebraic identities derived from group axioms -- proof notepad version\<close> |
47295 | 6 |
|
7 |
theory Group_Notepad |
|
63585 | 8 |
imports Main |
47295 | 9 |
begin |
10 |
||
11 |
notepad |
|
12 |
begin |
|
58614 | 13 |
txt \<open>hypothetical group axiomatization\<close> |
47295 | 14 |
|
80914
d97fdabd9e2b
standardize mixfix annotations via "isabelle update -a -u mixfix_cartouches" --- to simplify systematic editing;
wenzelm
parents:
63585
diff
changeset
|
15 |
fix prod :: "'a \<Rightarrow> 'a \<Rightarrow> 'a" (infixl \<open>\<odot>\<close> 70) |
47295 | 16 |
and one :: "'a" |
55656 | 17 |
and inverse :: "'a \<Rightarrow> 'a" |
61797 | 18 |
assume assoc: "(x \<odot> y) \<odot> z = x \<odot> (y \<odot> z)" |
19 |
and left_one: "one \<odot> x = x" |
|
20 |
and left_inverse: "inverse x \<odot> x = one" |
|
21 |
for x y z |
|
47295 | 22 |
|
58614 | 23 |
txt \<open>some consequences\<close> |
47295 | 24 |
|
61797 | 25 |
have right_inverse: "x \<odot> inverse x = one" for x |
47295 | 26 |
proof - |
61797 | 27 |
have "x \<odot> inverse x = one \<odot> (x \<odot> inverse x)" |
47295 | 28 |
by (simp only: left_one) |
61797 | 29 |
also have "\<dots> = one \<odot> x \<odot> inverse x" |
47295 | 30 |
by (simp only: assoc) |
61797 | 31 |
also have "\<dots> = inverse (inverse x) \<odot> inverse x \<odot> x \<odot> inverse x" |
47295 | 32 |
by (simp only: left_inverse) |
61797 | 33 |
also have "\<dots> = inverse (inverse x) \<odot> (inverse x \<odot> x) \<odot> inverse x" |
47295 | 34 |
by (simp only: assoc) |
61797 | 35 |
also have "\<dots> = inverse (inverse x) \<odot> one \<odot> inverse x" |
47295 | 36 |
by (simp only: left_inverse) |
61797 | 37 |
also have "\<dots> = inverse (inverse x) \<odot> (one \<odot> inverse x)" |
47295 | 38 |
by (simp only: assoc) |
61797 | 39 |
also have "\<dots> = inverse (inverse x) \<odot> inverse x" |
47295 | 40 |
by (simp only: left_one) |
41 |
also have "\<dots> = one" |
|
42 |
by (simp only: left_inverse) |
|
61797 | 43 |
finally show ?thesis . |
47295 | 44 |
qed |
45 |
||
61797 | 46 |
have right_one: "x \<odot> one = x" for x |
47295 | 47 |
proof - |
61797 | 48 |
have "x \<odot> one = x \<odot> (inverse x \<odot> x)" |
47295 | 49 |
by (simp only: left_inverse) |
61797 | 50 |
also have "\<dots> = x \<odot> inverse x \<odot> x" |
47295 | 51 |
by (simp only: assoc) |
61797 | 52 |
also have "\<dots> = one \<odot> x" |
47295 | 53 |
by (simp only: right_inverse) |
54 |
also have "\<dots> = x" |
|
55 |
by (simp only: left_one) |
|
61797 | 56 |
finally show ?thesis . |
47295 | 57 |
qed |
58 |
||
61797 | 59 |
have one_equality: "one = e" if eq: "e \<odot> x = x" for e x |
47295 | 60 |
proof - |
61797 | 61 |
have "one = x \<odot> inverse x" |
47295 | 62 |
by (simp only: right_inverse) |
61797 | 63 |
also have "\<dots> = (e \<odot> x) \<odot> inverse x" |
47295 | 64 |
by (simp only: eq) |
61797 | 65 |
also have "\<dots> = e \<odot> (x \<odot> inverse x)" |
47295 | 66 |
by (simp only: assoc) |
61797 | 67 |
also have "\<dots> = e \<odot> one" |
47295 | 68 |
by (simp only: right_inverse) |
69 |
also have "\<dots> = e" |
|
70 |
by (simp only: right_one) |
|
61797 | 71 |
finally show ?thesis . |
47295 | 72 |
qed |
73 |
||
61797 | 74 |
have inverse_equality: "inverse x = x'" if eq: "x' \<odot> x = one" for x x' |
47295 | 75 |
proof - |
61797 | 76 |
have "inverse x = one \<odot> inverse x" |
47295 | 77 |
by (simp only: left_one) |
61797 | 78 |
also have "\<dots> = (x' \<odot> x) \<odot> inverse x" |
47295 | 79 |
by (simp only: eq) |
61797 | 80 |
also have "\<dots> = x' \<odot> (x \<odot> inverse x)" |
47295 | 81 |
by (simp only: assoc) |
61797 | 82 |
also have "\<dots> = x' \<odot> one" |
47295 | 83 |
by (simp only: right_inverse) |
84 |
also have "\<dots> = x'" |
|
85 |
by (simp only: right_one) |
|
61797 | 86 |
finally show ?thesis . |
47295 | 87 |
qed |
88 |
||
89 |
end |
|
90 |
||
91 |
end |