Theory KAD.Range_Semiring
section ‹Range and Antirange Semirings›
theory Range_Semiring
imports Antidomain_Semiring
begin
subsection ‹Range Semirings›
text ‹We set up the duality between domain and antidomain semirings on the one hand and range and antirange semirings on the other hand.›
class range_op =
fixes range_op :: "'a ⇒ 'a" ("r")
class range_semiring = semiring_one_zero + plus_ord + range_op +
assumes rsr1 [simp]: "x + (x ⋅ r x) = x ⋅ r x"
and rsr2 [simp]: "r (r x ⋅ y) = r(x ⋅ y)"
and rsr3 [simp]: "r x + 1 = 1"
and rsr4 [simp]: "r 0 = 0"
and rsr5 [simp]: "r (x + y) = r x + r y"
begin
definition bd :: "'a ⇒ 'a ⇒ 'a" ("⟨_| _" [61,81] 82) where
"⟨x| y = r (y ⋅ x)"
no_notation range_op ("r")
end