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