diff -r 008b7858f3c0 -r b77980afc975 src/HOL/Isar_Examples/Group_Context.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Isar_Examples/Group_Context.thy Tue Apr 03 16:53:32 2012 +0200 @@ -0,0 +1,94 @@ +(* Title: HOL/Isar_Examples/Group_Context.thy + Author: Makarius +*) + +header {* Some algebraic identities derived from group axioms -- theory context version *} + +theory Group_Context +imports Main +begin + +text {* hypothetical group axiomatization *} + +context + fixes prod :: "'a \ 'a \ 'a" (infixl "**" 70) + and one :: "'a" + and inverse :: "'a => 'a" + assumes assoc: "\x y z. (x ** y) ** z = x ** (y ** z)" + and left_one: "\x. one ** x = x" + and left_inverse: "\x. inverse x ** x = one" +begin + +text {* some consequences *} + +lemma right_inverse: "x ** inverse x = one" +proof - + have "x ** inverse x = one ** (x ** inverse x)" + by (simp only: left_one) + also have "\ = one ** x ** inverse x" + by (simp only: assoc) + also have "\ = inverse (inverse x) ** inverse x ** x ** inverse x" + by (simp only: left_inverse) + also have "\ = inverse (inverse x) ** (inverse x ** x) ** inverse x" + by (simp only: assoc) + also have "\ = inverse (inverse x) ** one ** inverse x" + by (simp only: left_inverse) + also have "\ = inverse (inverse x) ** (one ** inverse x)" + by (simp only: assoc) + also have "\ = inverse (inverse x) ** inverse x" + by (simp only: left_one) + also have "\ = one" + by (simp only: left_inverse) + finally show "x ** inverse x = one" . +qed + +lemma right_one: "x ** one = x" +proof - + have "x ** one = x ** (inverse x ** x)" + by (simp only: left_inverse) + also have "\ = x ** inverse x ** x" + by (simp only: assoc) + also have "\ = one ** x" + by (simp only: right_inverse) + also have "\ = x" + by (simp only: left_one) + finally show "x ** one = x" . +qed + +lemma one_equality: "e ** x = x \ one = e" +proof - + fix e x + assume eq: "e ** x = x" + have "one = x ** inverse x" + by (simp only: right_inverse) + also have "\ = (e ** x) ** inverse x" + by (simp only: eq) + also have "\ = e ** (x ** inverse x)" + by (simp only: assoc) + also have "\ = e ** one" + by (simp only: right_inverse) + also have "\ = e" + by (simp only: right_one) + finally show "one = e" . +qed + +lemma inverse_equality: "x' ** x = one \ inverse x = x'" +proof - + fix x x' + assume eq: "x' ** x = one" + have "inverse x = one ** inverse x" + by (simp only: left_one) + also have "\ = (x' ** x) ** inverse x" + by (simp only: eq) + also have "\ = x' ** (x ** inverse x)" + by (simp only: assoc) + also have "\ = x' ** one" + by (simp only: right_inverse) + also have "\ = x'" + by (simp only: right_one) + finally show "inverse x = x'" . +qed + +end + +end