| author | wenzelm | 
| Fri, 26 Apr 2024 20:14:27 +0200 | |
| changeset 80157 | 6b9d5cae4579 | 
| parent 73622 | 4dc3baf45d6a | 
| child 80723 | ac6a69b0f634 | 
| permissions | -rw-r--r-- | 
| 72536 | 1  | 
theory Specifications_with_bundle_mixins  | 
| 73622 | 2  | 
imports "HOL-Combinatorics.Perm"  | 
| 72536 | 3  | 
begin  | 
4  | 
||
| 
72739
 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 
haftmann 
parents: 
72536 
diff
changeset
 | 
5  | 
locale involutory = opening permutation_syntax +  | 
| 72536 | 6  | 
fixes f :: \<open>'a perm\<close>  | 
7  | 
assumes involutory: \<open>\<And>x. f \<langle>$\<rangle> (f \<langle>$\<rangle> x) = x\<close>  | 
|
8  | 
begin  | 
|
9  | 
||
10  | 
lemma  | 
|
11  | 
\<open>f * f = 1\<close>  | 
|
12  | 
using involutory  | 
|
13  | 
by (simp add: perm_eq_iff apply_sequence)  | 
|
14  | 
||
15  | 
end  | 
|
16  | 
||
17  | 
context involutory  | 
|
18  | 
begin  | 
|
19  | 
||
20  | 
thm involutory (*syntax from permutation_syntax only present in locale specification and initial block*)  | 
|
21  | 
||
22  | 
end  | 
|
23  | 
||
24  | 
||
| 
72739
 
e7c2848b78e8
refined syntax for bundle mixins for locale and class specifications
 
haftmann 
parents: 
72536 
diff
changeset
 | 
25  | 
class at_most_two_elems = opening permutation_syntax +  | 
| 72536 | 26  | 
assumes swap_distinct: \<open>a \<noteq> b \<Longrightarrow> \<langle>a \<leftrightarrow> b\<rangle> \<langle>$\<rangle> c \<noteq> c\<close>  | 
27  | 
begin  | 
|
28  | 
||
29  | 
lemma  | 
|
30  | 
\<open>card (UNIV :: 'a set) \<le> 2\<close>  | 
|
31  | 
proof (rule ccontr)  | 
|
32  | 
fix a :: 'a  | 
|
33  | 
assume \<open>\<not> card (UNIV :: 'a set) \<le> 2\<close>  | 
|
34  | 
then have c0: \<open>card (UNIV :: 'a set) \<ge> 3\<close>  | 
|
35  | 
by simp  | 
|
36  | 
then have [simp]: \<open>finite (UNIV :: 'a set)\<close>  | 
|
37  | 
using card.infinite by fastforce  | 
|
38  | 
from c0 card_Diff1_le [of UNIV a]  | 
|
39  | 
  have ca: \<open>card (UNIV - {a}) \<ge> 2\<close>
 | 
|
40  | 
by simp  | 
|
41  | 
  then obtain b where \<open>b \<in> (UNIV - {a})\<close>
 | 
|
42  | 
by (metis all_not_in_conv card.empty card_2_iff' le_zero_eq)  | 
|
43  | 
  with ca card_Diff1_le [of \<open>UNIV - {a}\<close> b]
 | 
|
44  | 
  have cb: \<open>card (UNIV - {a, b}) \<ge> 1\<close> and \<open>a \<noteq> b\<close>
 | 
|
45  | 
by simp_all  | 
|
46  | 
  then obtain c where \<open>c \<in> (UNIV - {a, b})\<close>
 | 
|
47  | 
by (metis One_nat_def all_not_in_conv card.empty le_zero_eq nat.simps(3))  | 
|
48  | 
then have \<open>a \<noteq> c\<close> \<open>b \<noteq> c\<close>  | 
|
49  | 
by auto  | 
|
50  | 
with swap_distinct [of a b c] show False  | 
|
51  | 
by (simp add: \<open>a \<noteq> b\<close>)  | 
|
52  | 
qed  | 
|
53  | 
||
54  | 
end  | 
|
55  | 
||
56  | 
thm swap_distinct (*syntax from permutation_syntax only present in class specification and initial block*)  | 
|
57  | 
||
58  | 
end  |