Theory Specifications_with_bundle_mixins
theory Specifications_with_bundle_mixins
imports "HOL-Combinatorics.Perm"
begin
locale involutory = opening permutation_syntax +
fixes f :: ‹'a perm›
assumes involutory: ‹⋀x. f ⟨$⟩ (f ⟨$⟩ x) = x›
begin
lemma
‹f * f = 1›
using involutory
by (simp add: perm_eq_iff apply_sequence)
end
context involutory
begin
thm involutory
end
class at_most_two_elems = opening permutation_syntax +
assumes swap_distinct: ‹a ≠ b ⟹ ⟨a ↔ b⟩ ⟨$⟩ c ≠ c›
begin
lemma
‹card (UNIV :: 'a set) ≤ 2›
proof (rule ccontr)
fix a :: 'a
assume ‹¬ card (UNIV :: 'a set) ≤ 2›
then have c0: ‹card (UNIV :: 'a set) ≥ 3›
by simp
then have [simp]: ‹finite (UNIV :: 'a set)›
using card.infinite by fastforce
from c0 card_Diff1_le [of UNIV a]
have ca: ‹card (UNIV - {a}) ≥ 2›
by simp
then obtain b where ‹b ∈ (UNIV - {a})›
by (metis all_not_in_conv card.empty card_2_iff' le_zero_eq)
with ca card_Diff1_le [of ‹UNIV - {a}› b]
have cb: ‹card (UNIV - {a, b}) ≥ 1› and ‹a ≠ b›
by simp_all
then obtain c where ‹c ∈ (UNIV - {a, b})›
by (metis One_nat_def all_not_in_conv card.empty le_zero_eq nat.simps(3))
then have ‹a ≠ c› ‹b ≠ c›
by auto
with swap_distinct [of a b c] show False
by (simp add: ‹a ≠ b›)
qed
end
thm swap_distinct
end