Theory LocaleTest2
section ‹Test of Locale Interpretation›
theory LocaleTest2
imports Main
begin
section ‹Interpretation of Defined Concepts›
text ‹Naming convention for global objects: prefixes D and d›
subsection ‹Lattices›
text ‹Much of the lattice proofs are from HOL/Lattice.›
subsubsection ‹Definitions›
locale dpo =
fixes le :: "['a, 'a] => bool" (infixl "⊑" 50)
assumes refl [intro, simp]: "x ⊑ x"
and antisym [intro]: "[| x ⊑ y; y ⊑ x |] ==> x = y"
and trans [trans]: "[| x ⊑ y; y ⊑ z |] ==> x ⊑ z"
begin
theorem circular:
"[| x ⊑ y; y ⊑ z; z ⊑ x |] ==> x = y & y = z"
by (blast intro: trans)
definition
less :: "['a, 'a] => bool" (infixl "⊏" 50)
where "(x ⊏ y) = (x ⊑ y & x ~= y)"
theorem abs_test:
"(⊏) = (λx y. x ⊏ y)"
by simp
definition
is_inf :: "['a, 'a, 'a] => bool"
where "is_inf x y i = (i ⊑ x ∧ i ⊑ y ∧ (∀z. z ⊑ x ∧ z ⊑ y ⟶ z ⊑ i))"
definition
is_sup :: "['a, 'a, 'a] => bool"
where "is_sup x y s = (x ⊑ s ∧ y ⊑ s ∧ (∀z. x ⊑ z ∧ y ⊑ z ⟶ s ⊑ z))"
end
locale dlat = dpo +
assumes ex_inf: "∃inf. dpo.is_inf le x y inf"
and ex_sup: "∃sup. dpo.is_sup le x y sup"
begin
definition
meet :: "['a, 'a] => 'a" (infixl "⊓" 70)
where "x ⊓ y = (THE inf. is_inf x y inf)"
definition
join :: "['a, 'a] => 'a" (infixl "⊔" 65)
where "x ⊔ y = (THE sup. is_sup x y sup)"
lemma is_infI [intro?]: "i ⊑ x ⟹ i ⊑ y ⟹
(⋀z. z ⊑ x ⟹ z ⊑ y ⟹ z ⊑ i) ⟹ is_inf x y i"
by (unfold is_inf_def) blast
lemma is_inf_lower [elim?]:
"is_inf x y i ⟹ (i ⊑ x ⟹ i ⊑ y ⟹ C) ⟹ C"
by (unfold is_inf_def) blast
lemma is_inf_greatest [elim?]:
"is_inf x y i ⟹ z ⊑ x ⟹ z ⊑ y ⟹ z ⊑ i"
by (unfold is_inf_def) blast
theorem is_inf_uniq: "is_inf x y i ⟹ is_inf x y i' ⟹ i = i'"
proof -
assume inf: "is_inf x y i"
assume inf': "is_inf x y i'"
show ?thesis
proof (rule antisym)
from inf' show "i ⊑ i'"
proof (rule is_inf_greatest)
from inf show "i ⊑ x" ..
from inf show "i ⊑ y" ..
qed
from inf show "i' ⊑ i"
proof (rule is_inf_greatest)
from inf' show "i' ⊑ x" ..
from inf' show "i' ⊑ y" ..
qed
qed
qed
theorem is_inf_related [elim?]: "x ⊑ y ⟹ is_inf x y x"
proof -
assume "x ⊑ y"
show ?thesis
proof
show "x ⊑ x" ..
show "x ⊑ y" by fact
fix z assume "z ⊑ x" and "z ⊑ y" show "z ⊑ x" by fact
qed
qed
lemma meet_equality [elim?]: "is_inf x y i ⟹ x ⊓ y = i"
proof (unfold meet_def)
assume "is_inf x y i"
then show "(THE i. is_inf x y i) = i"
by (rule the_equality) (rule is_inf_uniq [OF _ ‹is_inf x y i›])
qed
lemma meetI [intro?]:
"i ⊑ x ⟹ i ⊑ y ⟹ (⋀z. z ⊑ x ⟹ z ⊑ y ⟹ z ⊑ i) ⟹ x ⊓ y = i"
by (rule meet_equality, rule is_infI) blast+
lemma is_inf_meet [intro?]: "is_inf x y (x ⊓ y)"
proof (unfold meet_def)
from ex_inf obtain i where "is_inf x y i" ..
then show "is_inf x y (THE i. is_inf x y i)"
by (rule theI) (rule is_inf_uniq [OF _ ‹is_inf x y i›])
qed
lemma meet_left [intro?]:
"x ⊓ y ⊑ x"
by (rule is_inf_lower) (rule is_inf_meet)
lemma meet_right [intro?]:
"x ⊓ y ⊑ y"
by (rule is_inf_lower) (rule is_inf_meet)
lemma meet_le [intro?]:
"[| z ⊑ x; z ⊑ y |] ==> z ⊑ x ⊓ y"
by (rule is_inf_greatest) (rule is_inf_meet)
lemma is_supI [intro?]: "x ⊑ s ⟹ y ⊑ s ⟹
(⋀z. x ⊑ z ⟹ y ⊑ z ⟹ s ⊑ z) ⟹ is_sup x y s"
by (unfold is_sup_def) blast
lemma is_sup_least [elim?]:
"is_sup x y s ⟹ x ⊑ z ⟹ y ⊑ z ⟹ s ⊑ z"
by (unfold is_sup_def) blast
lemma is_sup_upper [elim?]:
"is_sup x y s ⟹ (x ⊑ s ⟹ y ⊑ s ⟹ C) ⟹ C"
by (unfold is_sup_def) blast
theorem is_sup_uniq: "is_sup x y s ⟹ is_sup x y s' ⟹ s = s'"
proof -
assume sup: "is_sup x y s"
assume sup': "is_sup x y s'"
show ?thesis
proof (rule antisym)
from sup show "s ⊑ s'"
proof (rule is_sup_least)
from sup' show "x ⊑ s'" ..
from sup' show "y ⊑ s'" ..
qed
from sup' show "s' ⊑ s"
proof (rule is_sup_least)
from sup show "x ⊑ s" ..
from sup show "y ⊑ s" ..
qed
qed
qed
theorem is_sup_related [elim?]: "x ⊑ y ⟹ is_sup x y y"
proof -
assume "x ⊑ y"
show ?thesis
proof
show "x ⊑ y" by fact
show "y ⊑ y" ..
fix z assume "x ⊑ z" and "y ⊑ z"
show "y ⊑ z" by fact
qed
qed
lemma join_equality [elim?]: "is_sup x y s ⟹ x ⊔ y = s"
proof (unfold join_def)
assume "is_sup x y s"
then show "(THE s. is_sup x y s) = s"
by (rule the_equality) (rule is_sup_uniq [OF _ ‹is_sup x y s›])
qed
lemma joinI [intro?]: "x ⊑ s ⟹ y ⊑ s ⟹
(⋀z. x ⊑ z ⟹ y ⊑ z ⟹ s ⊑ z) ⟹ x ⊔ y = s"
by (rule join_equality, rule is_supI) blast+
lemma is_sup_join [intro?]: "is_sup x y (x ⊔ y)"
proof (unfold join_def)
from ex_sup obtain s where "is_sup x y s" ..
then show "is_sup x y (THE s. is_sup x y s)"
by (rule theI) (rule is_sup_uniq [OF _ ‹is_sup x y s›])
qed
lemma join_left [intro?]:
"x ⊑ x ⊔ y"
by (rule is_sup_upper) (rule is_sup_join)
lemma join_right [intro?]:
"y ⊑ x ⊔ y"
by (rule is_sup_upper) (rule is_sup_join)
lemma join_le [intro?]:
"[| x ⊑ z; y ⊑ z |] ==> x ⊔ y ⊑ z"
by (rule is_sup_least) (rule is_sup_join)
theorem meet_assoc: "(x ⊓ y) ⊓ z = x ⊓ (y ⊓ z)"
proof (rule meetI)
show "x ⊓ (y ⊓ z) ⊑ x ⊓ y"
proof
show "x ⊓ (y ⊓ z) ⊑ x" ..
show "x ⊓ (y ⊓ z) ⊑ y"
proof -
have "x ⊓ (y ⊓ z) ⊑ y ⊓ z" ..
also have "… ⊑ y" ..
finally show ?thesis .
qed
qed
show "x ⊓ (y ⊓ z) ⊑ z"
proof -
have "x ⊓ (y ⊓ z) ⊑ y ⊓ z" ..
also have "… ⊑ z" ..
finally show ?thesis .
qed
fix w assume "w ⊑ x ⊓ y" and "w ⊑ z"
show "w ⊑ x ⊓ (y ⊓ z)"
proof
show "w ⊑ x"
proof -
have "w ⊑ x ⊓ y" by fact
also have "… ⊑ x" ..
finally show ?thesis .
qed
show "w ⊑ y ⊓ z"
proof
show "w ⊑ y"
proof -
have "w ⊑ x ⊓ y" by fact
also have "… ⊑ y" ..
finally show ?thesis .
qed
show "w ⊑ z" by fact
qed
qed
qed
theorem meet_commute: "x ⊓ y = y ⊓ x"
proof (rule meetI)
show "y ⊓ x ⊑ x" ..
show "y ⊓ x ⊑ y" ..
fix z assume "z ⊑ y" and "z ⊑ x"
then show "z ⊑ y ⊓ x" ..
qed
theorem meet_join_absorb: "x ⊓ (x ⊔ y) = x"
proof (rule meetI)
show "x ⊑ x" ..
show "x ⊑ x ⊔ y" ..
fix z assume "z ⊑ x" and "z ⊑ x ⊔ y"
show "z ⊑ x" by fact
qed
theorem join_assoc: "(x ⊔ y) ⊔ z = x ⊔ (y ⊔ z)"
proof (rule joinI)
show "x ⊔ y ⊑ x ⊔ (y ⊔ z)"
proof
show "x ⊑ x ⊔ (y ⊔ z)" ..
show "y ⊑ x ⊔ (y ⊔ z)"
proof -
have "y ⊑ y ⊔ z" ..
also have "... ⊑ x ⊔ (y ⊔ z)" ..
finally show ?thesis .
qed
qed
show "z ⊑ x ⊔ (y ⊔ z)"
proof -
have "z ⊑ y ⊔ z" ..
also have "... ⊑ x ⊔ (y ⊔ z)" ..
finally show ?thesis .
qed
fix w assume "x ⊔ y ⊑ w" and "z ⊑ w"
show "x ⊔ (y ⊔ z) ⊑ w"
proof
show "x ⊑ w"
proof -
have "x ⊑ x ⊔ y" ..
also have "… ⊑ w" by fact
finally show ?thesis .
qed
show "y ⊔ z ⊑ w"
proof
show "y ⊑ w"
proof -
have "y ⊑ x ⊔ y" ..
also have "... ⊑ w" by fact
finally show ?thesis .
qed
show "z ⊑ w" by fact
qed
qed
qed
theorem join_commute: "x ⊔ y = y ⊔ x"
proof (rule joinI)
show "x ⊑ y ⊔ x" ..
show "y ⊑ y ⊔ x" ..
fix z assume "y ⊑ z" and "x ⊑ z"
then show "y ⊔ x ⊑ z" ..
qed
theorem join_meet_absorb: "x ⊔ (x ⊓ y) = x"
proof (rule joinI)
show "x ⊑ x" ..
show "x ⊓ y ⊑ x" ..
fix z assume "x ⊑ z" and "x ⊓ y ⊑ z"
show "x ⊑ z" by fact
qed
theorem meet_idem: "x ⊓ x = x"
proof -
have "x ⊓ (x ⊔ (x ⊓ x)) = x" by (rule meet_join_absorb)
also have "x ⊔ (x ⊓ x) = x" by (rule join_meet_absorb)
finally show ?thesis .
qed
theorem meet_related [elim?]: "x ⊑ y ⟹ x ⊓ y = x"
proof (rule meetI)
assume "x ⊑ y"
show "x ⊑ x" ..
show "x ⊑ y" by fact
fix z assume "z ⊑ x" and "z ⊑ y"
show "z ⊑ x" by fact
qed
theorem meet_related2 [elim?]: "y ⊑ x ⟹ x ⊓ y = y"
by (drule meet_related) (simp add: meet_commute)
theorem join_related [elim?]: "x ⊑ y ⟹ x ⊔ y = y"
proof (rule joinI)
assume "x ⊑ y"
show "y ⊑ y" ..
show "x ⊑ y" by fact
fix z assume "x ⊑ z" and "y ⊑ z"
show "y ⊑ z" by fact
qed
theorem join_related2 [elim?]: "y ⊑ x ⟹ x ⊔ y = x"
by (drule join_related) (simp add: join_commute)
text ‹Additional theorems›
theorem meet_connection: "(x ⊑ y) = (x ⊓ y = x)"
proof
assume "x ⊑ y"
then have "is_inf x y x" ..
then show "x ⊓ y = x" ..
next
have "x ⊓ y ⊑ y" ..
also assume "x ⊓ y = x"
finally show "x ⊑ y" .
qed
theorem meet_connection2: "(x ⊑ y) = (y ⊓ x = x)"
using meet_commute meet_connection by simp
theorem join_connection: "(x ⊑ y) = (x ⊔ y = y)"
proof
assume "x ⊑ y"
then have "is_sup x y y" ..
then show "x ⊔ y = y" ..
next
have "x ⊑ x ⊔ y" ..
also assume "x ⊔ y = y"
finally show "x ⊑ y" .
qed
theorem join_connection2: "(x ⊑ y) = (x ⊔ y = y)"
using join_commute join_connection by simp
text ‹Naming according to Jacobson I, p.\ 459.›
lemmas L1 = join_commute meet_commute
lemmas L2 = join_assoc meet_assoc
lemmas L4 = join_meet_absorb meet_join_absorb
end
locale ddlat = dlat +
assumes meet_distr:
"dlat.meet le x (dlat.join le y z) =
dlat.join le (dlat.meet le x y) (dlat.meet le x z)"
begin
lemma join_distr:
"x ⊔ (y ⊓ z) = (x ⊔ y) ⊓ (x ⊔ z)"
txt ‹Jacobson I, p.\ 462›
proof -
have "x ⊔ (y ⊓ z) = (x ⊔ (x ⊓ z)) ⊔ (y ⊓ z)" by (simp add: L4)
also have "... = x ⊔ ((x ⊓ z) ⊔ (y ⊓ z))" by (simp add: L2)
also have "... = x ⊔ ((x ⊔ y) ⊓ z)" by (simp add: L1 meet_distr)
also have "... = ((x ⊔ y) ⊓ x) ⊔ ((x ⊔ y) ⊓ z)" by (simp add: L1 L4)
also have "... = (x ⊔ y) ⊓ (x ⊔ z)" by (simp add: meet_distr)
finally show ?thesis .
qed
end
locale dlo = dpo +
assumes total: "x ⊑ y | y ⊑ x"
begin
lemma less_total: "x ⊏ y | x = y | y ⊏ x"
using total
by (unfold less_def) blast
end
sublocale dlo < dlat
proof
fix x y
from total have "is_inf x y (if x ⊑ y then x else y)" by (auto simp: is_inf_def)
then show "∃inf. is_inf x y inf" by blast
next
fix x y
from total have "is_sup x y (if x ⊑ y then y else x)" by (auto simp: is_sup_def)
then show "∃sup. is_sup x y sup" by blast
qed
sublocale dlo < ddlat
proof
fix x y z
show "x ⊓ (y ⊔ z) = x ⊓ y ⊔ x ⊓ z" (is "?l = ?r")
txt ‹Jacobson I, p.\ 462›
proof -
{ assume c: "y ⊑ x" "z ⊑ x"
from c have "?l = y ⊔ z"
by (metis c join_connection2 join_related2 meet_connection meet_related2 total)
also from c have "... = ?r" by (metis meet_related2)
finally have "?l = ?r" . }
moreover
{ assume c: "x ⊑ y | x ⊑ z"
from c have "?l = x"
by (metis join_connection2 join_related2 meet_connection total trans)
also from c have "... = ?r"
by (metis join_commute join_related2 meet_connection meet_related2 total)
finally have "?l = ?r" . }
moreover note total
ultimately show ?thesis by blast
qed
qed
subsubsection ‹Total order ‹<=› on \<^typ>‹int››
interpretation int: dpo "(<=) :: [int, int] => bool"
rewrites "(dpo.less (<=) (x::int) y) = (x < y)"
txt ‹We give interpretation for less, but not ‹is_inf› and ‹is_sub›.›
proof -
show "dpo ((<=) :: [int, int] => bool)"
proof qed auto
then interpret int: dpo "(<=) :: [int, int] => bool" .
txt ‹Gives interpreted version of ‹less_def› (without condition).›
show "(dpo.less (<=) (x::int) y) = (x < y)"
by (unfold int.less_def) auto
qed
thm int.circular
lemma "⟦ (x::int) ≤ y; y ≤ z; z ≤ x⟧ ⟹ x = y ∧ y = z"
apply (rule int.circular) apply assumption apply assumption apply assumption done
thm int.abs_test
lemma "((<) :: [int, int] => bool) = (<)"
apply (rule int.abs_test) done
interpretation int: dlat "(<=) :: [int, int] => bool"
rewrites meet_eq: "dlat.meet (<=) (x::int) y = min x y"
and join_eq: "dlat.join (<=) (x::int) y = max x y"
proof -
show "dlat ((<=) :: [int, int] => bool)"
apply unfold_locales
apply (unfold int.is_inf_def int.is_sup_def)
apply arith+
done
then interpret int: dlat "(<=) :: [int, int] => bool" .
txt ‹Interpretation to ease use of definitions, which are
conditional in general but unconditional after interpretation.›
show "dlat.meet (<=) (x::int) y = min x y"
apply (unfold int.meet_def)
apply (rule the_equality)
apply (unfold int.is_inf_def)
by auto
show "dlat.join (<=) (x::int) y = max x y"
apply (unfold int.join_def)
apply (rule the_equality)
apply (unfold int.is_sup_def)
by auto
qed
interpretation int: dlo "(<=) :: [int, int] => bool"
proof qed arith
text ‹Interpreted theorems from the locales, involving defined terms.›
thm int.less_def text ‹from dpo›
thm int.meet_left text ‹from dlat›
thm int.meet_distr text ‹from ddlat›
thm int.less_total text ‹from dlo›
subsubsection ‹Total order ‹<=› on \<^typ>‹nat››
interpretation nat: dpo "(<=) :: [nat, nat] => bool"
rewrites "dpo.less (<=) (x::nat) y = (x < y)"
txt ‹We give interpretation for less, but not ‹is_inf› and ‹is_sub›.›
proof -
show "dpo ((<=) :: [nat, nat] => bool)"
proof qed auto
then interpret nat: dpo "(<=) :: [nat, nat] => bool" .
txt ‹Gives interpreted version of ‹less_def› (without condition).›
show "dpo.less (<=) (x::nat) y = (x < y)"
apply (unfold nat.less_def)
apply auto
done
qed
interpretation nat: dlat "(<=) :: [nat, nat] => bool"
rewrites "dlat.meet (<=) (x::nat) y = min x y"
and "dlat.join (<=) (x::nat) y = max x y"
proof -
show "dlat ((<=) :: [nat, nat] => bool)"
apply unfold_locales
apply (unfold nat.is_inf_def nat.is_sup_def)
apply arith+
done
then interpret nat: dlat "(<=) :: [nat, nat] => bool" .
txt ‹Interpretation to ease use of definitions, which are
conditional in general but unconditional after interpretation.›
show "dlat.meet (<=) (x::nat) y = min x y"
apply (unfold nat.meet_def)
apply (rule the_equality)
apply (unfold nat.is_inf_def)
by auto
show "dlat.join (<=) (x::nat) y = max x y"
apply (unfold nat.join_def)
apply (rule the_equality)
apply (unfold nat.is_sup_def)
by auto
qed
interpretation nat: dlo "(<=) :: [nat, nat] => bool"
proof qed arith
text ‹Interpreted theorems from the locales, involving defined terms.›
thm nat.less_def text ‹from dpo›
thm nat.meet_left text ‹from dlat›
thm nat.meet_distr text ‹from ddlat›
thm nat.less_total text ‹from ldo›
subsubsection ‹Lattice ‹dvd› on \<^typ>‹nat››
interpretation nat_dvd: dpo "(dvd) :: [nat, nat] => bool"
rewrites "dpo.less (dvd) (x::nat) y = (x dvd y & x ~= y)"
txt ‹We give interpretation for less, but not ‹is_inf› and ‹is_sub›.›
proof -
show "dpo ((dvd) :: [nat, nat] => bool)"
proof qed (auto simp: dvd_def)
then interpret nat_dvd: dpo "(dvd) :: [nat, nat] => bool" .
txt ‹Gives interpreted version of ‹less_def› (without condition).›
show "dpo.less (dvd) (x::nat) y = (x dvd y & x ~= y)"
apply (unfold nat_dvd.less_def)
apply auto
done
qed
interpretation nat_dvd: dlat "(dvd) :: [nat, nat] => bool"
rewrites "dlat.meet (dvd) (x::nat) y = gcd x y"
and "dlat.join (dvd) (x::nat) y = lcm x y"
proof -
show "dlat ((dvd) :: [nat, nat] => bool)"
apply unfold_locales
apply (unfold nat_dvd.is_inf_def nat_dvd.is_sup_def)
apply (rule_tac x = "gcd x y" in exI)
apply auto [1]
apply (rule_tac x = "lcm x y" in exI)
apply (auto intro: dvd_lcm1 dvd_lcm2 lcm_least)
done
then interpret nat_dvd: dlat "(dvd) :: [nat, nat] => bool" .
txt ‹Interpretation to ease use of definitions, which are
conditional in general but unconditional after interpretation.›
show "dlat.meet (dvd) (x::nat) y = gcd x y"
apply (unfold nat_dvd.meet_def)
apply (rule the_equality)
apply (unfold nat_dvd.is_inf_def)
by auto
show "dlat.join (dvd) (x::nat) y = lcm x y"
apply (unfold nat_dvd.join_def)
apply (rule the_equality)
apply (unfold nat_dvd.is_sup_def)
by (auto intro: dvd_lcm1 dvd_lcm2 lcm_least)
qed
text ‹Interpreted theorems from the locales, involving defined terms.›
thm nat_dvd.less_def text ‹from dpo›
lemma "((x::nat) dvd y & x ~= y) = (x dvd y & x ~= y)"
apply (rule nat_dvd.less_def) done
thm nat_dvd.meet_left text ‹from dlat›
lemma "gcd x y dvd (x::nat)"
apply (rule nat_dvd.meet_left) done
subsection ‹Group example with defined operations ‹inv› and ‹unit››
subsubsection ‹Locale declarations and lemmas›
locale Dsemi =
fixes prod (infixl "**" 65)
assumes assoc: "(x ** y) ** z = x ** (y ** z)"
locale Dmonoid = Dsemi +
fixes one
assumes l_one [simp]: "one ** x = x"
and r_one [simp]: "x ** one = x"
begin
definition
inv where "inv x = (THE y. x ** y = one & y ** x = one)"
definition
unit where "unit x = (∃y. x ** y = one & y ** x = one)"
lemma inv_unique:
assumes eq: "y ** x = one" "x ** y' = one"
shows "y = y'"
proof -
from eq have "y = y ** (x ** y')" by (simp add: r_one)
also have "... = (y ** x) ** y'" by (simp add: assoc)
also from eq have "... = y'" by (simp add: l_one)
finally show ?thesis .
qed
lemma unit_one [intro, simp]:
"unit one"
by (unfold unit_def) auto
lemma unit_l_inv_ex:
"unit x ==> ∃y. y ** x = one"
by (unfold unit_def) auto
lemma unit_r_inv_ex:
"unit x ==> ∃y. x ** y = one"
by (unfold unit_def) auto
lemma unit_l_inv:
"unit x ==> inv x ** x = one"
apply (simp add: unit_def inv_def) apply (erule exE)
apply (rule theI2, fast)
apply (rule inv_unique)
apply fast+
done
lemma unit_r_inv:
"unit x ==> x ** inv x = one"
apply (simp add: unit_def inv_def) apply (erule exE)
apply (rule theI2, fast)
apply (rule inv_unique)
apply fast+
done
lemma unit_inv_unit [intro, simp]:
"unit x ==> unit (inv x)"
proof -
assume x: "unit x"
show "unit (inv x)"
by (auto simp add: unit_def
intro: unit_l_inv unit_r_inv x)
qed
lemma unit_l_cancel [simp]:
"unit x ==> (x ** y = x ** z) = (y = z)"
proof
assume eq: "x ** y = x ** z"
and G: "unit x"
then have "(inv x ** x) ** y = (inv x ** x) ** z"
by (simp add: assoc)
with G show "y = z" by (simp add: unit_l_inv)
next
assume eq: "y = z"
and G: "unit x"
then show "x ** y = x ** z" by simp
qed
lemma unit_inv_inv [simp]:
"unit x ==> inv (inv x) = x"
proof -
assume x: "unit x"
then have "inv x ** inv (inv x) = inv x ** x"
by (simp add: unit_l_inv unit_r_inv)
with x show ?thesis by simp
qed
lemma inv_inj_on_unit:
"inj_on inv {x. unit x}"
proof (rule inj_onI, simp)
fix x y
assume G: "unit x" "unit y" and eq: "inv x = inv y"
then have "inv (inv x) = inv (inv y)" by simp
with G show "x = y" by simp
qed
lemma unit_inv_comm:
assumes inv: "x ** y = one"
and G: "unit x" "unit y"
shows "y ** x = one"
proof -
from G have "x ** y ** x = x ** one" by (auto simp add: inv)
with G show ?thesis by (simp del: r_one add: assoc)
qed
end
locale Dgrp = Dmonoid +
assumes unit [intro, simp]: "Dmonoid.unit ((**)) one x"
begin
lemma l_inv_ex [simp]:
"∃y. y ** x = one"
using unit_l_inv_ex by simp
lemma r_inv_ex [simp]:
"∃y. x ** y = one"
using unit_r_inv_ex by simp
lemma l_inv [simp]:
"inv x ** x = one"
using unit_l_inv by simp
lemma l_cancel [simp]:
"(x ** y = x ** z) = (y = z)"
using unit_l_inv by simp
lemma r_inv [simp]:
"x ** inv x = one"
proof -
have "inv x ** (x ** inv x) = inv x ** one"
by (simp add: assoc [symmetric] l_inv)
then show ?thesis by (simp del: r_one)
qed
lemma r_cancel [simp]:
"(y ** x = z ** x) = (y = z)"
proof
assume eq: "y ** x = z ** x"
then have "y ** (x ** inv x) = z ** (x ** inv x)"
by (simp add: assoc [symmetric] del: r_inv)
then show "y = z" by simp
qed simp
lemma inv_one [simp]:
"inv one = one"
proof -
have "inv one = one ** (inv one)" by (simp del: r_inv)
moreover have "... = one" by simp
finally show ?thesis .
qed
lemma inv_inv [simp]:
"inv (inv x) = x"
using unit_inv_inv by simp
lemma inv_inj:
"inj_on inv UNIV"
using inv_inj_on_unit by simp
lemma inv_mult_group:
"inv (x ** y) = inv y ** inv x"
proof -
have "inv (x ** y) ** (x ** y) = (inv y ** inv x) ** (x ** y)"
by (simp add: assoc l_inv) (simp add: assoc [symmetric])
then show ?thesis by (simp del: l_inv)
qed
lemma inv_comm:
"x ** y = one ==> y ** x = one"
by (rule unit_inv_comm) auto
lemma inv_equality:
"y ** x = one ==> inv x = y"
apply (simp add: inv_def)
apply (rule the_equality)
apply (simp add: inv_comm [of y x])
apply (rule r_cancel [THEN iffD1], auto)
done
end
locale Dhom = prod: Dgrp prod one + sum: Dgrp sum zero
for prod (infixl "**" 65) and one and sum (infixl "+++" 60) and zero +
fixes hom
assumes hom_mult [simp]: "hom (x ** y) = hom x +++ hom y"
begin
lemma hom_one [simp]:
"hom one = zero"
proof -
have "hom one +++ zero = hom one +++ hom one"
by (simp add: hom_mult [symmetric] del: hom_mult)
then show ?thesis by (simp del: sum.r_one)
qed
end
subsubsection ‹Interpretation of Functions›
interpretation Dfun: Dmonoid "(o)" "id :: 'a => 'a"
rewrites "Dmonoid.unit (o) id f = bij (f::'a => 'a)"
proof -
show "Dmonoid (o) (id :: 'a => 'a)" proof qed (simp_all add: o_assoc)
note Dmonoid = this
show "Dmonoid.unit (o) (id :: 'a => 'a) f = bij f"
apply (unfold Dmonoid.unit_def [OF Dmonoid])
apply rule apply clarify
proof -
fix f g
assume id1: "f o g = id" and id2: "g o f = id"
show "bij f"
proof (rule bijI)
show "inj f"
proof (rule inj_onI)
fix x y
assume "f x = f y"
then have "(g o f) x = (g o f) y" by simp
with id2 show "x = y" by simp
qed
next
show "surj f"
proof (rule surjI)
fix x
from id1 have "(f o g) x = x" by simp
then show "f (g x) = x" by simp
qed
qed
next
fix f
assume bij: "bij f"
then
have inv: "f o Hilbert_Choice.inv f = id & Hilbert_Choice.inv f o f = id"
by (simp add: bij_def surj_iff inj_iff)
show "∃g. f ∘ g = id ∧ g ∘ f = id" by rule (rule inv)
qed
qed
thm Dmonoid.unit_def Dfun.unit_def
thm Dmonoid.inv_inj_on_unit Dfun.inv_inj_on_unit
lemma unit_id:
"(f :: unit => unit) = id"
by rule simp
interpretation Dfun: Dgrp "(o)" "id :: unit => unit"
rewrites "Dmonoid.inv (o) id f = inv (f :: unit => unit)"
proof -
have "Dmonoid (o) (id :: 'a => 'a)" ..
note Dmonoid = this
show "Dgrp (o) (id :: unit => unit)"
apply unfold_locales
apply (unfold Dmonoid.unit_def [OF Dmonoid])
apply (insert unit_id)
apply simp
done
show "Dmonoid.inv (o) id f = inv (f :: unit ⇒ unit)"
apply (unfold Dmonoid.inv_def [OF Dmonoid])
apply (insert unit_id)
apply simp
apply (rule the_equality)
apply rule
apply rule
apply simp
done
qed
thm Dfun.unit_l_inv Dfun.l_inv
thm Dfun.inv_equality
thm Dfun.inv_equality
end