Theory Partial_Semigroups
section ‹Partial Semigroups›
theory Partial_Semigroups
imports Main
begin
notation times (infixl ‹⋅› 70)
and times (infixl ‹⊕› 70)
subsection ‹Partial Semigroups›
text ‹In this context, partiality is modelled by a definedness constraint $D$ instead of a bottom element,
which would make the algebra total. This is common practice in mathematics.›
class partial_times = times +
fixes D :: "'a ⇒ 'a ⇒ bool"
text ‹The definedness constraints for associativity state that the right-hand side of the associativity
law is defined if and only if the left-hand side is and that, in this case, both sides are equal. This and
slightly different constraints can be found in the literature.›
class partial_semigroup = partial_times +
assumes add_assocD: "D y z ∧ D x (y ⋅ z) ⟷ D x y ∧ D (x ⋅ y) z"
and add_assoc: "D x y ∧ D (x ⋅ y) z ⟹ (x ⋅ y) ⋅ z = x ⋅ (y ⋅ z)"
text ‹Every semigroup is a partial semigroup.›
sublocale semigroup_mult ⊆ sg: partial_semigroup _ "λx y. True"
by standard (simp_all add: mult_assoc)
context partial_semigroup
begin
text ‹The following abbreviation is useful for sublocale statements.›
abbreviation (input) "R x y z ≡ D y z ∧ x = y ⋅ z"
lemma add_assocD_var1: "D y z ∧ D x (y ⋅ z) ⟹ D x y ∧ D (x ⋅ y) z"
by (simp add: add_assocD)
lemma add_assocD_var2: " D x y ∧ D (x ⋅ y) z ⟹ D y z ∧ D x (y ⋅ z)"
by (simp add: add_assocD)
lemma add_assoc_var: " D y z ∧ D x (y ⋅ z) ⟹ (x ⋅ y) ⋅ z = x ⋅ (y ⋅ z)"
by (simp add: add_assoc add_assocD)
subsection ‹Green's Preorders and Green's Relations›
text ‹We define the standard Green's preorders and Green's relations. They are usually defined on monoids.
On (partial) semigroups, we only obtain transitive relations.›
definition gR_rel :: "'a ⇒ 'a ⇒ bool" (infix ‹≼⇩R› 50) where
"x ≼⇩R y = (∃z. D x z ∧ x ⋅ z = y)"
definition strict_gR_rel :: "'a ⇒ 'a ⇒ bool" (infix ‹≺⇩R› 50) where
"x ≺⇩R y = (x ≼⇩R y ∧ ¬ y ≼⇩R x)"
definition gL_rel :: "'a ⇒ 'a ⇒ bool" (infix ‹≼⇩L› 50) where
"x ≼⇩L y = (∃z. D z x ∧ z ⋅ x = y)"
definition strict_gL_rel :: "'a ⇒ 'a ⇒ bool" (infix ‹≺⇩L› 50) where
"x ≺⇩L y = (x ≼⇩L y ∧ ¬ y ≼⇩L x)"
definition gH_rel :: "'a ⇒ 'a ⇒ bool" (infix ‹≼⇩H› 50) where
"x ≼⇩H y = (x ≼⇩L y ∧ x ≼⇩R y)"
definition gJ_rel :: "'a ⇒ 'a ⇒ bool" (infix ‹≼⇩J› 50) where
"x ≼⇩J y = (∃v w. D v x ∧ D (v ⋅ x) w ∧ (v ⋅ x) ⋅ w = y)"
definition "gR x y = (x ≼⇩R y ∧ y ≼⇩R x)"
definition "gL x y = (x ≼⇩L y ∧ y ≼⇩L x)"
definition "gH x y = (x ≼⇩H y ∧ y ≼⇩H x)"
definition "gJ x y = (x ≼⇩J y ∧ y ≼⇩J x)"
definition gR_downset :: "'a ⇒ 'a set" (‹_↓› [100]100) where
"x↓ ≡ {y. y ≼⇩R x}"
text ‹The following counterexample rules out reflexivity.›
lemma "x ≼⇩R x"
oops
lemma gR_rel_trans: "x ≼⇩R y ⟹ y ≼⇩R z ⟹ x ≼⇩R z"
by (metis gR_rel_def add_assoc add_assocD_var2)
lemma gL_rel_trans: "x ≼⇩L y ⟹ y ≼⇩L z ⟹ x ≼⇩L z"
by (metis gL_rel_def add_assocD_var1 add_assoc_var)
lemma gR_add_isol: "D z y ⟹ x ≼⇩R y ⟹ z ⋅ x ≼⇩R z ⋅ y"
apply (simp add: gR_rel_def)
using add_assocD_var1 add_assoc_var by blast
lemma gL_add_isor: "D y z ⟹ x ≼⇩L y ⟹ x ⋅ z ≼⇩L y ⋅ z"
apply (simp add: gL_rel_def)
by (metis add_assoc add_assocD_var2)
definition annil :: "'a ⇒ bool" where
"annil x = (∀y. D x y ∧ x ⋅ y = x)"
definition annir :: "'a ⇒ bool" where
"annir x = (∀y. D y x ∧ y ⋅ x = x)"
end
subsection ‹Morphisms›
definition ps_morphism :: "('a::partial_semigroup ⇒ 'b::partial_semigroup) ⇒ bool" where
"ps_morphism f = (∀x y. D x y ⟶ D (f x) (f y) ∧ f (x ⋅ y) = (f x) ⋅ (f y))"
definition strong_ps_morphism :: "('a::partial_semigroup ⇒ 'b::partial_semigroup) ⇒ bool" where
"strong_ps_morphism f = (ps_morphism f ∧ (∀x y. D (f x) (f y) ⟶ D x y))"
subsection ‹ Locally Finite Partial Semigroups›
text ‹In locally finite partial semigroups, elements can only be split in finitely many ways.›
class locally_finite_partial_semigroup = partial_semigroup +
assumes loc_fin: "finite (x↓)"
subsection ‹Cancellative Partial Semigroups›
class cancellative_partial_semigroup = partial_semigroup +
assumes add_cancl: "D z x ⟹ D z y ⟹ z ⋅ x = z ⋅ y ⟹ x = y"
and add_cancr: "D x z ⟹ D y z ⟹ x ⋅ z = y ⋅ z ⟹ x = y"
begin
lemma unique_resl: "D x z ⟹ D x z' ⟹ x ⋅ z = y ⟹ x ⋅ z' = y ⟹ z = z'"
by (simp add: add_cancl)
lemma unique_resr: "D z x ⟹ D z' x ⟹ z ⋅ x = y ⟹ z' ⋅ x = y ⟹ z = z'"
by (simp add: add_cancr)
lemma gR_rel_mult: "D x y ⟹ x ≼⇩R x ⋅ y"
using gR_rel_def by force
lemma gL_rel_mult: "D x y ⟹ y ≼⇩L x ⋅ y"
using gL_rel_def by force
text ‹By cancellation, the element z is uniquely defined for each pair x y, provided it exists.
In both cases, z is therefore a function of x and y; it is a quotient or residual of x y.›
lemma quotr_unique: "x ≼⇩R y ⟹ (∃!z. D x z ∧ y = x ⋅ z)"
using gR_rel_def add_cancl by force
lemma quotl_unique: "x ≼⇩L y ⟹ (∃!z. D z x ∧ y = z ⋅ x)"
using gL_rel_def unique_resr by force
definition "rquot y x = (THE z. D x z ∧ x ⋅ z = y)"
definition "lquot y x = (THE z. D z x ∧ z ⋅ x = y)"
lemma rquot_prop: "D x z ∧ y = x ⋅ z ⟹ z = rquot y x"
by (metis (mono_tags, lifting) rquot_def the_equality unique_resl)
lemma rquot_mult: "x ≼⇩R y ⟹ z = rquot y x ⟹ x ⋅ z = y"
using gR_rel_def rquot_prop by force
lemma rquot_D: "x ≼⇩R y ⟹ z = rquot y x ⟹ D x z"
using gR_rel_def rquot_prop by force
lemma add_rquot: "x ≼⇩R y ⟹ (D x z ∧ x ⊕ z = y ⟷ z = rquot y x)"
using gR_rel_def rquot_prop by fastforce
lemma add_canc1: "D x y ⟹ rquot (x ⋅ y) x = y"
using rquot_prop by simp
lemma add_canc2: "x ≼⇩R y ⟹ x ⋅ (rquot y x) = y"
using gR_rel_def add_canc1 by force
lemma add_canc2_prop: "x ≼⇩R y ⟹ rquot y x ≼⇩L y"
using gL_rel_mult rquot_D rquot_mult by fastforce
text ‹The next set of lemmas establishes standard Galois connections for cancellative partial semigroups.›
lemma gR_galois_imp1: "D x z ⟹ x ⋅ z ≼⇩R y ⟹ z ≼⇩R rquot y x"
by (metis gR_rel_def add_assoc add_assocD_var2 rquot_prop)
lemma gR_galois_imp21: "x ≼⇩R y ⟹ z ≼⇩R rquot y x ⟹ x ⋅ z ≼⇩R y"
using gR_add_isol rquot_D rquot_mult by fastforce
lemma gR_galois_imp22: "x ≼⇩R y ⟹ z ≼⇩R rquot y x ⟹ D x z"
using gR_rel_def add_assocD add_canc1 by fastforce
lemma gR_galois: "x ≼⇩R y ⟹ (D x z ∧ x ⋅ z ≼⇩R y ⟷ z ≼⇩R rquot y x)"
using gR_galois_imp1 gR_galois_imp21 gR_galois_imp22 by blast
lemma gR_rel_defined: "x ≼⇩R y ⟹ D x (rquot y x)"
by (simp add: rquot_D)
lemma ex_add_galois: "D x z ⟹ (∃y. x ⋅ z = y ⟷ rquot y x = z)"
using add_canc1 by force
end
subsection ‹Partial Monoids›
text ‹We allow partial monoids with multiple units. This is similar to and inspired by small categories.›
class partial_monoid = partial_semigroup +
fixes E :: "'a set"
assumes unitl_ex: "∃e ∈ E. D e x ∧ e ⋅ x = x"
and unitr_ex: "∃e ∈ E. D x e ∧ x ⋅ e = x"
and units_eq: "e1 ∈ E ⟹ e2 ∈ E ⟹ D e1 e2 ⟹ e1 = e2"
text ‹Every monoid is a partial monoid.›
sublocale monoid_mult ⊆ mon: partial_monoid _ "λx y. True" "{1}"
by (standard; simp_all)
context partial_monoid
begin
lemma units_eq_var: "e1 ∈ E ⟹ e2 ∈ E ⟹ e1 ≠ e2 ⟹ ¬ D e1 e2"
using units_eq by force
text ‹In partial monoids, Green's relations become preorders, but need not be partial orders.›