Theory Classpackage

Up to index of Isabelle/HOL/ex

theory Classpackage
imports List
begin

(*  ID:         $Id: Classpackage.thy,v 1.41 2007/10/16 21:12:54 haftmann Exp $
    Author:     Florian Haftmann, TU Muenchen
*)

header {* Test and examples for Isar class package *}

theory Classpackage
imports List
begin

class semigroup = type +
  fixes mult :: "'a => 'a => 'a" (infixl "⊗" 70)
  assumes assoc: "x ⊗ y ⊗ z = x ⊗ (y ⊗ z)"

instance nat :: semigroup
  "m ⊗ n ≡ m + n"
proof
  fix m n q :: nat 
  from mult_nat_def show "m ⊗ n ⊗ q = m ⊗ (n ⊗ q)" by simp
qed

instance int :: semigroup
  "k ⊗ l ≡ k + l"
proof
  fix k l j :: int
  from mult_int_def show "k ⊗ l ⊗ j = k ⊗ (l ⊗ j)" by simp
qed

instance * :: (semigroup, semigroup) semigroup
  mult_prod_def: "x ⊗ y ≡ let (x1, x2) = x; (y1, y2) = y in
              (x1 ⊗ y1, x2 ⊗ y2)"
by default (simp_all add: split_paired_all mult_prod_def assoc)

instance list :: (type) semigroup
  "xs ⊗ ys ≡ xs @ ys"
proof
  fix xs ys zs :: "'a list"
  show "xs ⊗ ys ⊗ zs = xs ⊗ (ys ⊗ zs)"
  proof -
    from mult_list_def have "!!xs ys::'a list. xs ⊗ ys ≡ xs @ ys" .
    thus ?thesis by simp
  qed
qed

class monoidl = semigroup +
  fixes one :: 'a ("\<one>")
  assumes neutl: "\<one> ⊗ x = x"

instance nat :: monoidl and int :: monoidl
  "\<one> ≡ 0"
  "\<one> ≡ 0"
proof
  fix n :: nat
  from mult_nat_def one_nat_def show "\<one> ⊗ n = n" by simp
next
  fix k :: int
  from mult_int_def one_int_def show "\<one> ⊗ k = k" by simp
qed

instance * :: (monoidl, monoidl) monoidl
  one_prod_def: "\<one> ≡ (\<one>, \<one>)"
by default (simp_all add: split_paired_all mult_prod_def one_prod_def neutl)

instance list :: (type) monoidl
  "\<one> ≡ []"
proof
  fix xs :: "'a list"
  show "\<one> ⊗ xs = xs"
  proof -
    from mult_list_def have "!!xs ys::'a list. xs ⊗ ys ≡ xs @ ys" .
    moreover from one_list_def have "\<one> ≡ []::'a list" by simp
    ultimately show ?thesis by simp
  qed
qed  

class monoid = monoidl +
  assumes neutr: "x ⊗ \<one> = x"
begin

definition
  units :: "'a set" where
  "units = {y. ∃x. x ⊗ y = \<one> ∧ y ⊗ x = \<one>}"

lemma inv_obtain:
  assumes "x ∈ units"
  obtains y where "y ⊗ x = \<one>" and "x ⊗ y = \<one>"
proof -
  from assms units_def obtain y
    where "y ⊗ x = \<one>" and "x ⊗ y = \<one>" by auto
  thus ?thesis ..
qed

lemma inv_unique:
  assumes "y ⊗ x = \<one>" "x ⊗ y' = \<one>"
  shows "y = y'"
proof -
  from assms neutr have "y = y ⊗ (x ⊗ y')" by simp
  also with assoc have "... = (y ⊗ x) ⊗ y'" by simp
  also with assms neutl have "... = y'" by simp
  finally show ?thesis .
qed

lemma units_inv_comm:
  assumes inv: "x ⊗ y = \<one>"
    and G: "x ∈ units"
  shows "y ⊗ x = \<one>"
proof -
  from G inv_obtain obtain z
    where z_choice: "z ⊗ x = \<one>" by blast
  from inv neutl neutr have "x ⊗ y ⊗ x = x ⊗ \<one>" by simp
  with assoc have "z ⊗ x ⊗ y ⊗ x = z ⊗ x ⊗ \<one>" by simp
  with neutl z_choice show ?thesis by simp
qed

fun
  npow :: "nat => 'a => 'a"
where
  "npow 0 x = \<one>"
  | "npow (Suc n) x = x ⊗ npow n x"

abbreviation
  npow_syn :: "'a => nat => 'a" (infix "\<up>" 75) where
  "x \<up> n ≡ npow n x"

lemma nat_pow_one:
  "\<one> \<up> n = \<one>"
using neutl by (induct n) simp_all

lemma nat_pow_mult: "x \<up> n ⊗ x \<up> m = x \<up> (n + m)"
proof (induct n)
  case 0 with neutl show ?case by simp
next
  case (Suc n) with Suc.hyps assoc show ?case by simp
qed

lemma nat_pow_pow: "(x \<up> m) \<up> n = x \<up> (n * m)"
using nat_pow_mult by (induct n) simp_all

end

instance * :: (monoid, monoid) monoid
by default (simp_all add: split_paired_all mult_prod_def one_prod_def neutr)

instance list :: (type) monoid
proof
  fix xs :: "'a list"
  show "xs ⊗ \<one> = xs"
  proof -
    from mult_list_def have "!!xs ys::'a list. xs ⊗ ys ≡ xs @ ys" .
    moreover from one_list_def have "\<one> ≡ []::'a list" by simp
    ultimately show ?thesis by simp
  qed
qed  

class monoid_comm = monoid +
  assumes comm: "x ⊗ y = y ⊗ x"

instance nat :: monoid_comm and int :: monoid_comm
proof
  fix n :: nat
  from mult_nat_def one_nat_def show "n ⊗ \<one> = n" by simp
next
  fix n m :: nat
  from mult_nat_def show "n ⊗ m = m ⊗ n" by simp
next
  fix k :: int
  from mult_int_def one_int_def show "k ⊗ \<one> = k" by simp
next
  fix k l :: int
  from mult_int_def show "k ⊗ l = l ⊗ k" by simp
qed

instance * :: (monoid_comm, monoid_comm) monoid_comm
by default (simp_all add: split_paired_all mult_prod_def comm)

class group = monoidl +
  fixes inv :: "'a => 'a" ("÷ _" [81] 80)
  assumes invl: "÷ x ⊗ x = \<one>"
begin

lemma cancel:
  "(x ⊗ y = x ⊗ z) = (y = z)"
proof
  fix x y z :: 'a
  assume eq: "x ⊗ y = x ⊗ z"
  hence "÷ x ⊗ (x ⊗ y) = ÷ x ⊗ (x ⊗ z)" by simp
  with assoc have "÷ x ⊗ x ⊗ y = ÷ x ⊗ x ⊗ z" by simp
  with neutl invl show "y = z" by simp
next
  fix x y z :: 'a
  assume eq: "y = z"
  thus "x ⊗ y = x ⊗ z" by simp
qed

subclass monoid
proof unfold_locales
  fix x
  from invl have "÷ x ⊗ x = \<one>" by simp
  with assoc [symmetric] neutl invl have "÷ x ⊗ (x ⊗ \<one>) = ÷ x ⊗ x" by simp
  with cancel show "x ⊗ \<one> = x" by simp
qed

end context group begin

find_theorems name: neut

lemma invr:
  "x ⊗ ÷ x = \<one>"
proof -
  from neutl invl have "÷ x ⊗ x ⊗ ÷ x = ÷ x" by simp
  with neutr have "÷ x ⊗ x ⊗ ÷ x = ÷ x ⊗ \<one> " by simp
  with assoc have "÷ x ⊗ (x ⊗ ÷ x) = ÷ x ⊗ \<one> " by simp
  with cancel show ?thesis ..
qed

lemma all_inv [intro]:
  "(x::'a) ∈ units"
  unfolding units_def
proof -
  fix x :: "'a"
  from invl invr have "÷ x ⊗ x = \<one>" and "x ⊗ ÷ x = \<one>" . 
  then obtain y where "y ⊗ x = \<one>" and "x ⊗ y = \<one>" ..
  hence "∃y::'a. y ⊗ x = \<one> ∧ x ⊗ y = \<one>" by blast
  thus "x ∈ {y::'a. ∃x::'a. x ⊗ y = \<one> ∧ y ⊗ x = \<one>}" by simp
qed

lemma cancer:
  "(y ⊗ x = z ⊗ x) = (y = z)"
proof
  assume eq: "y ⊗ x = z ⊗ x"
  with assoc [symmetric] have "y ⊗ (x ⊗ ÷ x) = z ⊗ (x ⊗ ÷ x)" by (simp del: invr)
  with invr neutr show "y = z" by simp
next
  assume eq: "y = z"
  thus "y ⊗ x = z ⊗ x" by simp
qed

lemma inv_one:
  "÷ \<one> = \<one>"
proof -
  from neutl have "÷ \<one> = \<one> ⊗ (÷ \<one>)" ..
  moreover from invr have "... = \<one>" by simp
  finally show ?thesis .
qed

lemma inv_inv:
  "÷ (÷ x) = x"
proof -
  from invl invr neutr
    have "÷ (÷ x) ⊗ ÷ x ⊗ x = x ⊗ ÷ x ⊗ x" by simp
  with assoc [symmetric]
    have "÷ (÷ x) ⊗ (÷ x ⊗ x) = x ⊗ (÷ x ⊗ x)" by simp
  with invl neutr show ?thesis by simp
qed

lemma inv_mult_group:
  "÷ (x ⊗ y) = ÷ y ⊗ ÷ x"
proof -
  from invl have "÷ (x ⊗ y) ⊗ (x ⊗ y) = \<one>" by simp
  with assoc have "÷ (x ⊗ y) ⊗ x ⊗ y = \<one>" by simp
  with neutl have "÷ (x ⊗ y) ⊗ x ⊗ y ⊗ ÷ y ⊗ ÷ x = ÷ y ⊗ ÷ x" by simp
  with assoc have "÷ (x ⊗ y) ⊗ (x ⊗ (y ⊗ ÷ y) ⊗ ÷ x) = ÷ y ⊗ ÷ x" by simp
  with invr neutr show ?thesis by simp
qed

lemma inv_comm:
  "x ⊗ ÷ x = ÷ x ⊗ x"
using invr invl by simp

definition
  pow :: "int => 'a => 'a"
where
  "pow k x = (if k < 0 then ÷ (npow (nat (-k)) x)
    else (npow (nat k) x))"

abbreviation
  pow_syn :: "'a => int => 'a" (infix "\<up>\<up>" 75)
where
  "x \<up>\<up> k ≡ pow k x"

lemma int_pow_zero:
  "x \<up>\<up> (0::int) = \<one>"
using pow_def by simp

lemma int_pow_one:
  "\<one> \<up>\<up> (k::int) = \<one>"
using pow_def nat_pow_one inv_one by simp

end

instance * :: (group, group) group
  inv_prod_def: "÷ x ≡ let (x1, x2) = x in (÷ x1, ÷ x2)"
by default (simp_all add: split_paired_all mult_prod_def one_prod_def inv_prod_def invl)

class group_comm = group + monoid_comm

instance int :: group_comm
  "÷ k ≡ - (k::int)"
proof
  fix k :: int
  from mult_int_def one_int_def inv_int_def show "÷ k ⊗ k = \<one>" by simp
qed

instance * :: (group_comm, group_comm) group_comm
by default (simp_all add: split_paired_all mult_prod_def comm)

definition
  "X a b c = (a ⊗ \<one> ⊗ b, a ⊗ \<one> ⊗ b, npow 3 [a, b] ⊗ \<one> ⊗ [a, b, c])"

definition
  "Y a b c = (a, ÷ a) ⊗ \<one> ⊗ ÷ (b, ÷ pow (-3) c)"

definition "x1 = X (1::nat) 2 3"
definition "x2 = X (1::int) 2 3"
definition "y2 = Y (1::int) 2 3"

export_code x1 x2 y2 pow in SML module_name Classpackage
  in OCaml file -
  in Haskell file -

end

lemma inv_obtain:

  [| xunits; !!y. [| y  x = \<one>; x  y = \<one> |] ==> thesis |] ==> thesis

lemma inv_unique:

  [| y  x = \<one>; x  y' = \<one> |] ==> y = y'

lemma units_inv_comm:

  [| x  y = \<one>; xunits |] ==> y  x = \<one>

lemma nat_pow_one:

  \<one> \<up> n = \<one>

lemma nat_pow_mult:

  x \<up> n  x \<up> m = x \<up> (n + m)

lemma nat_pow_pow:

  (x \<up> m) \<up> n = x \<up> (n * m)

lemma cancel:

  (x  y = x  z) = (y = z)

lemma invr:

  x  ÷ x = \<one>

lemma all_inv:

  xunits

lemma cancer:

  (y  x = z  x) = (y = z)

lemma inv_one:

  ÷ \<one> = \<one>

lemma inv_inv:

  ÷ (÷ x) = x

lemma inv_mult_group:

  ÷ (x  y) = ÷ y  ÷ x

lemma inv_comm:

  x  ÷ x = ÷ x  x

lemma int_pow_zero:

  x \<up>\<up> 0 = \<one>

lemma int_pow_one:

  \<one> \<up>\<up> k = \<one>