src/HOL/Library/Perm.thy
author wenzelm
Sat, 27 May 2017 13:01:25 +0200
changeset 65946 5dd3974cf0bc
parent 64911 f0e07600de47
child 67399 eab6ce8368fa
permissions -rw-r--r--
tuned signature;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen *)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
     2
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
     3
section \<open>Permutations as abstract type\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
     4
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
     5
theory Perm
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
     6
imports Main
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
     7
begin
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
     8
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
     9
text \<open>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    10
  This theory introduces basics about permutations, i.e. almost
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    11
  everywhere fix bijections.  But it is by no means complete.
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    12
  Grieviously missing are cycles since these would require more
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    13
  elaboration, e.g. the concept of distinct lists equivalent
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    14
  under rotation, which maybe would also deserve its own theory.
64911
f0e07600de47 isabelle update_cartouches -c -t;
wenzelm
parents: 63993
diff changeset
    15
  But see theory \<open>src/HOL/ex/Perm_Fragments.thy\<close> for
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    16
  fragments on that.
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    17
\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    18
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    19
subsection \<open>Abstract type of permutations\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    20
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    21
typedef 'a perm = "{f :: 'a \<Rightarrow> 'a. bij f \<and> finite {a. f a \<noteq> a}}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    22
  morphisms "apply" Perm
63433
wenzelm
parents: 63375
diff changeset
    23
proof
wenzelm
parents: 63375
diff changeset
    24
  show "id \<in> ?perm" by simp
wenzelm
parents: 63375
diff changeset
    25
qed
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    26
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    27
setup_lifting type_definition_perm
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    28
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    29
notation "apply" (infixl "\<langle>$\<rangle>" 999)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    30
no_notation "apply" ("op \<langle>$\<rangle>")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    31
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    32
lemma bij_apply [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    33
  "bij (apply f)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    34
  using "apply" [of f] by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    35
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    36
lemma perm_eqI:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    37
  assumes "\<And>a. f \<langle>$\<rangle> a = g \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    38
  shows "f = g"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    39
  using assms by transfer (simp add: fun_eq_iff)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    40
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    41
lemma perm_eq_iff:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    42
  "f = g \<longleftrightarrow> (\<forall>a. f \<langle>$\<rangle> a = g \<langle>$\<rangle> a)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    43
  by (auto intro: perm_eqI)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    44
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    45
lemma apply_inj:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    46
  "f \<langle>$\<rangle> a = f \<langle>$\<rangle> b \<longleftrightarrow> a = b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    47
  by (rule inj_eq) (rule bij_is_inj, simp)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    48
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    49
lift_definition affected :: "'a perm \<Rightarrow> 'a set"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    50
  is "\<lambda>f. {a. f a \<noteq> a}" .
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    51
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    52
lemma in_affected:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    53
  "a \<in> affected f \<longleftrightarrow> f \<langle>$\<rangle> a \<noteq> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    54
  by transfer simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    55
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    56
lemma finite_affected [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    57
  "finite (affected f)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    58
  by transfer simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    59
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    60
lemma apply_affected [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    61
  "f \<langle>$\<rangle> a \<in> affected f \<longleftrightarrow> a \<in> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    62
proof transfer
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    63
  fix f :: "'a \<Rightarrow> 'a" and a :: 'a
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    64
  assume "bij f \<and> finite {b. f b \<noteq> b}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    65
  then have "bij f" by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    66
  interpret bijection f by standard (rule \<open>bij f\<close>)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    67
  have "f a \<in> {a. f a = a} \<longleftrightarrow> a \<in> {a. f a = a}" (is "?P \<longleftrightarrow> ?Q")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    68
    by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    69
  then show "f a \<in> {a. f a \<noteq> a} \<longleftrightarrow> a \<in> {a. f a \<noteq> a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    70
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    71
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    72
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    73
lemma card_affected_not_one:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    74
  "card (affected f) \<noteq> 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    75
proof
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    76
  interpret bijection "apply f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    77
    by standard (rule bij_apply)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    78
  assume "card (affected f) = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    79
  then obtain a where *: "affected f = {a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    80
    by (rule card_1_singletonE)
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63539
diff changeset
    81
  then have **: "f \<langle>$\<rangle> a \<noteq> a"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    82
    by (simp add: in_affected [symmetric])
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63539
diff changeset
    83
  with * have "f \<langle>$\<rangle> a \<notin> affected f"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    84
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    85
  then have "f \<langle>$\<rangle> (f \<langle>$\<rangle> a) = f \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    86
    by (simp add: in_affected)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    87
  then have "inv (apply f) (f \<langle>$\<rangle> (f \<langle>$\<rangle> a)) = inv (apply f) (f \<langle>$\<rangle> a)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    88
    by simp
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63539
diff changeset
    89
  with ** show False by simp
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    90
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    91
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    92
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    93
subsection \<open>Identity, composition and inversion\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    94
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    95
instantiation Perm.perm :: (type) "{monoid_mult, inverse}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    96
begin
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    97
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    98
lift_definition one_perm :: "'a perm"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    99
  is id
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   100
  by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   101
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   102
lemma apply_one [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   103
  "apply 1 = id"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   104
  by (fact one_perm.rep_eq)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   105
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   106
lemma affected_one [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   107
  "affected 1 = {}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   108
  by transfer simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   109
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   110
lemma affected_empty_iff [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   111
  "affected f = {} \<longleftrightarrow> f = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   112
  by transfer auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   113
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   114
lift_definition times_perm :: "'a perm \<Rightarrow> 'a perm \<Rightarrow> 'a perm"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   115
  is comp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   116
proof
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   117
  fix f g :: "'a \<Rightarrow> 'a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   118
  assume "bij f \<and> finite {a. f a \<noteq> a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   119
    "bij g \<and>finite {a. g a \<noteq> a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   120
  then have "finite ({a. f a \<noteq> a} \<union> {a. g a \<noteq> a})"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   121
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   122
  moreover have "{a. (f \<circ> g) a \<noteq> a} \<subseteq> {a. f a \<noteq> a} \<union> {a. g a \<noteq> a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   123
    by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   124
  ultimately show "finite {a. (f \<circ> g) a \<noteq> a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   125
    by (auto intro: finite_subset)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   126
qed (auto intro: bij_comp)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   127
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   128
lemma apply_times:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   129
  "apply (f * g) = apply f \<circ> apply g"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   130
  by (fact times_perm.rep_eq)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   131
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   132
lemma apply_sequence:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   133
  "f \<langle>$\<rangle> (g \<langle>$\<rangle> a) = apply (f * g) a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   134
  by (simp add: apply_times)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   135
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   136
lemma affected_times [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   137
  "affected (f * g) \<subseteq> affected f \<union> affected g"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   138
  by transfer auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   139
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   140
lift_definition inverse_perm :: "'a perm \<Rightarrow> 'a perm"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   141
  is inv
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   142
proof transfer
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   143
  fix f :: "'a \<Rightarrow> 'a" and a
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   144
  assume "bij f \<and> finite {b. f b \<noteq> b}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   145
  then have "bij f" and fin: "finite {b. f b \<noteq> b}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   146
    by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   147
  interpret bijection f by standard (rule \<open>bij f\<close>)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   148
  from fin show "bij (inv f) \<and> finite {a. inv f a \<noteq> a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   149
    by (simp add: bij_inv)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   150
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   151
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   152
instance
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   153
  by standard (transfer; simp add: comp_assoc)+
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   154
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   155
end
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   156
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   157
lemma apply_inverse:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   158
  "apply (inverse f) = inv (apply f)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   159
  by (fact inverse_perm.rep_eq)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   160
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   161
lemma affected_inverse [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   162
  "affected (inverse f) = affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   163
proof transfer
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   164
  fix f :: "'a \<Rightarrow> 'a" and a
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   165
  assume "bij f \<and> finite {b. f b \<noteq> b}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   166
  then have "bij f" by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   167
  interpret bijection f by standard (rule \<open>bij f\<close>)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   168
  show "{a. inv f a \<noteq> a} = {a. f a \<noteq> a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   169
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   170
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   171
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   172
global_interpretation perm: group times "1::'a perm" inverse
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   173
proof
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   174
  fix f :: "'a perm"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   175
  show "1 * f = f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   176
    by transfer simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   177
  show "inverse f * f = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   178
  proof transfer
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   179
    fix f :: "'a \<Rightarrow> 'a" and a
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   180
    assume "bij f \<and> finite {b. f b \<noteq> b}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   181
    then have "bij f" by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   182
    interpret bijection f by standard (rule \<open>bij f\<close>)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   183
    show "inv f \<circ> f = id"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   184
      by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   185
  qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   186
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   187
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   188
declare perm.inverse_distrib_swap [simp]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   189
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   190
lemma perm_mult_commute:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   191
  assumes "affected f \<inter> affected g = {}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   192
  shows "g * f = f * g"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   193
proof (rule perm_eqI)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   194
  fix a
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   195
  from assms have *: "a \<in> affected f \<Longrightarrow> a \<notin> affected g"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   196
    "a \<in> affected g \<Longrightarrow> a \<notin> affected f" for a
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   197
    by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   198
  consider "a \<in> affected f \<and> a \<notin> affected g
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   199
        \<and> f \<langle>$\<rangle> a \<in> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   200
    | "a \<notin> affected f \<and> a \<in> affected g
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   201
        \<and> f \<langle>$\<rangle> a \<notin> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   202
    | "a \<notin> affected f \<and> a \<notin> affected g"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   203
    using assms by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   204
  then show "(g * f) \<langle>$\<rangle> a = (f * g) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   205
  proof cases
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63539
diff changeset
   206
    case 1
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63539
diff changeset
   207
    with * have "f \<langle>$\<rangle> a \<notin> affected g"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   208
      by auto
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63539
diff changeset
   209
    with 1 show ?thesis by (simp add: in_affected apply_times)
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   210
  next
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63539
diff changeset
   211
    case 2
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63539
diff changeset
   212
    with * have "g \<langle>$\<rangle> a \<notin> affected f"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   213
      by auto
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63539
diff changeset
   214
    with 2 show ?thesis by (simp add: in_affected apply_times)
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   215
  next
63540
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63539
diff changeset
   216
    case 3
f8652d0534fa tuned proofs -- avoid unstructured calculation;
wenzelm
parents: 63539
diff changeset
   217
    then show ?thesis by (simp add: in_affected apply_times)
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   218
  qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   219
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   220
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   221
lemma apply_power:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   222
  "apply (f ^ n) = apply f ^^ n"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   223
  by (induct n) (simp_all add: apply_times)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   224
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   225
lemma perm_power_inverse:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   226
  "inverse f ^ n = inverse ((f :: 'a perm) ^ n)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   227
proof (induct n)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   228
  case 0 then show ?case by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   229
next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   230
  case (Suc n)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   231
  then show ?case
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   232
    unfolding power_Suc2 [of f] by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   233
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   234
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   235
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   236
subsection \<open>Orbit and order of elements\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   237
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   238
definition orbit :: "'a perm \<Rightarrow> 'a \<Rightarrow> 'a set"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   239
where
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   240
  "orbit f a = range (\<lambda>n. (f ^ n) \<langle>$\<rangle> a)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   241
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   242
lemma in_orbitI:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   243
  assumes "(f ^ n) \<langle>$\<rangle> a = b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   244
  shows "b \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   245
  using assms by (auto simp add: orbit_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   246
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   247
lemma apply_power_self_in_orbit [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   248
  "(f ^ n) \<langle>$\<rangle> a \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   249
  by (rule in_orbitI) rule
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   250
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   251
lemma in_orbit_self [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   252
  "a \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   253
  using apply_power_self_in_orbit [of _ 0] by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   254
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   255
lemma apply_self_in_orbit [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   256
  "f \<langle>$\<rangle> a \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   257
  using apply_power_self_in_orbit [of _ 1] by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   258
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   259
lemma orbit_not_empty [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   260
  "orbit f a \<noteq> {}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   261
  using in_orbit_self [of a f] by blast
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   262
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   263
lemma not_in_affected_iff_orbit_eq_singleton:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   264
  "a \<notin> affected f \<longleftrightarrow> orbit f a = {a}" (is "?P \<longleftrightarrow> ?Q")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   265
proof
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   266
  assume ?P
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   267
  then have "f \<langle>$\<rangle> a = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   268
    by (simp add: in_affected)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   269
  then have "(f ^ n) \<langle>$\<rangle> a = a" for n
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   270
    by (induct n) (simp_all add: apply_times)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   271
  then show ?Q
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   272
    by (auto simp add: orbit_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   273
next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   274
  assume ?Q
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   275
  then show ?P
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   276
    by (auto simp add: orbit_def in_affected dest: range_eq_singletonD [of _ _ 1])
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   277
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   278
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   279
definition order :: "'a perm \<Rightarrow> 'a \<Rightarrow> nat"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   280
where
63993
9c0ff0c12116 eliminated hard tabs;
wenzelm
parents: 63540
diff changeset
   281
  "order f = card \<circ> orbit f"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   282
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   283
lemma orbit_subset_eq_affected:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   284
  assumes "a \<in> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   285
  shows "orbit f a \<subseteq> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   286
proof (rule ccontr)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   287
  assume "\<not> orbit f a \<subseteq> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   288
  then obtain b where "b \<in> orbit f a" and "b \<notin> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   289
    by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   290
  then have "b \<in> range (\<lambda>n. (f ^ n) \<langle>$\<rangle> a)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   291
    by (simp add: orbit_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   292
  then obtain n where "b = (f ^ n) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   293
    by blast
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   294
  with \<open>b \<notin> affected f\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   295
  have "(f ^ n) \<langle>$\<rangle> a \<notin> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   296
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   297
  then have "f \<langle>$\<rangle> a \<notin> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   298
    by (induct n) (simp_all add: apply_times)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   299
  with assms show False
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   300
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   301
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   302
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   303
lemma finite_orbit [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   304
  "finite (orbit f a)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   305
proof (cases "a \<in> affected f")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   306
  case False then show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   307
    by (simp add: not_in_affected_iff_orbit_eq_singleton)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   308
next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   309
  case True then have "orbit f a \<subseteq> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   310
    by (rule orbit_subset_eq_affected)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   311
  then show ?thesis using finite_affected
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   312
    by (rule finite_subset)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   313
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   314
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   315
lemma orbit_1 [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   316
  "orbit 1 a = {a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   317
  by (auto simp add: orbit_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   318
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   319
lemma order_1 [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   320
  "order 1 a = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   321
  unfolding order_def by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   322
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   323
lemma card_orbit_eq [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   324
  "card (orbit f a) = order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   325
  by (simp add: order_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   326
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   327
lemma order_greater_zero [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   328
  "order f a > 0"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   329
  by (simp only: card_gt_0_iff order_def comp_def) simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   330
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   331
lemma order_eq_one_iff:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   332
  "order f a = Suc 0 \<longleftrightarrow> a \<notin> affected f" (is "?P \<longleftrightarrow> ?Q")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   333
proof
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   334
  assume ?P then have "card (orbit f a) = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   335
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   336
  then obtain b where "orbit f a = {b}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   337
    by (rule card_1_singletonE)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   338
  with in_orbit_self [of a f]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   339
    have "b = a" by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   340
  with \<open>orbit f a = {b}\<close> show ?Q
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   341
    by (simp add: not_in_affected_iff_orbit_eq_singleton)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   342
next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   343
  assume ?Q
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   344
  then have "orbit f a = {a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   345
    by (simp add: not_in_affected_iff_orbit_eq_singleton)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   346
  then have "card (orbit f a) = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   347
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   348
  then show ?P
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   349
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   350
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   351
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   352
lemma order_greater_eq_two_iff:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   353
  "order f a \<ge> 2 \<longleftrightarrow> a \<in> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   354
  using order_eq_one_iff [of f a]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   355
  apply (auto simp add: neq_iff)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   356
  using order_greater_zero [of f a]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   357
  apply simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   358
  done
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   359
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   360
lemma order_less_eq_affected:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   361
  assumes "f \<noteq> 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   362
  shows "order f a \<le> card (affected f)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   363
proof (cases "a \<in> affected f")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   364
  from assms have "affected f \<noteq> {}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   365
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   366
  then obtain B b where "affected f = insert b B"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   367
    by blast
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   368
  with finite_affected [of f] have "card (affected f) \<ge> 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   369
    by (simp add: card_insert)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   370
  case False then have "order f a = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   371
    by (simp add: order_eq_one_iff)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   372
  with \<open>card (affected f) \<ge> 1\<close> show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   373
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   374
next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   375
  case True
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   376
  have "card (orbit f a) \<le> card (affected f)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   377
    by (rule card_mono) (simp_all add: True orbit_subset_eq_affected card_mono)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   378
  then show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   379
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   380
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   381
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   382
lemma affected_order_greater_eq_two:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   383
  assumes "a \<in> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   384
  shows "order f a \<ge> 2"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   385
proof (rule ccontr)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   386
  assume "\<not> 2 \<le> order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   387
  then have "order f a < 2"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   388
    by (simp add: not_le)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   389
  with order_greater_zero [of f a] have "order f a = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   390
    by arith
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   391
  with assms show False
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   392
    by (simp add: order_eq_one_iff)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   393
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   394
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   395
lemma order_witness_unfold:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   396
  assumes "n > 0" and "(f ^ n) \<langle>$\<rangle> a = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   397
  shows "order f a = card ((\<lambda>m. (f ^ m) \<langle>$\<rangle> a) ` {0..<n})"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   398
proof  -
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   399
  have "orbit f a = (\<lambda>m. (f ^ m) \<langle>$\<rangle> a) ` {0..<n}" (is "_ = ?B")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   400
  proof (rule set_eqI, rule)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   401
    fix b
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   402
    assume "b \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   403
    then obtain m where "(f ^ m) \<langle>$\<rangle> a = b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   404
      by (auto simp add: orbit_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   405
    then have "b = (f ^ (m mod n + n * (m div n))) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   406
      by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   407
    also have "\<dots> = (f ^ (m mod n)) \<langle>$\<rangle> ((f ^ (n * (m div n))) \<langle>$\<rangle> a)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   408
      by (simp only: power_add apply_times) simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   409
    also have "(f ^ (n * q)) \<langle>$\<rangle> a = a" for q
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   410
      by (induct q)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   411
        (simp_all add: power_add apply_times assms)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   412
    finally have "b = (f ^ (m mod n)) \<langle>$\<rangle> a" .
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   413
    moreover from \<open>n > 0\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   414
    have "m mod n < n" 
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   415
      by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   416
    ultimately show "b \<in> ?B"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   417
      by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   418
  next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   419
    fix b
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   420
    assume "b \<in> ?B"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   421
    then obtain m where "(f ^ m) \<langle>$\<rangle> a = b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   422
      by blast
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   423
    then show "b \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   424
      by (rule in_orbitI)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   425
  qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   426
  then have "card (orbit f a) = card ?B"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   427
    by (simp only:)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   428
  then show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   429
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   430
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   431
    
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   432
lemma inj_on_apply_range:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   433
  "inj_on (\<lambda>m. (f ^ m) \<langle>$\<rangle> a) {..<order f a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   434
proof -
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   435
  have "inj_on (\<lambda>m. (f ^ m) \<langle>$\<rangle> a) {..<n}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   436
    if "n \<le> order f a" for n
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   437
  using that proof (induct n)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   438
    case 0 then show ?case by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   439
  next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   440
    case (Suc n)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   441
    then have prem: "n < order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   442
      by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   443
    with Suc.hyps have hyp: "inj_on (\<lambda>m. (f ^ m) \<langle>$\<rangle> a) {..<n}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   444
      by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   445
    have "(f ^ n) \<langle>$\<rangle> a \<notin> (\<lambda>m. (f ^ m) \<langle>$\<rangle> a) ` {..<n}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   446
    proof
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   447
      assume "(f ^ n) \<langle>$\<rangle> a \<in> (\<lambda>m. (f ^ m) \<langle>$\<rangle> a) ` {..<n}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   448
      then obtain m where *: "(f ^ m) \<langle>$\<rangle> a = (f ^ n) \<langle>$\<rangle> a" and "m < n"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   449
        by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   450
      interpret bijection "apply (f ^ m)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   451
        by standard simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   452
      from \<open>m < n\<close> have "n = m + (n - m)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   453
        and nm: "0 < n - m" "n - m \<le> n"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   454
        by arith+
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   455
      with * have "(f ^ m) \<langle>$\<rangle> a = (f ^ (m + (n - m))) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   456
        by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   457
      then have "(f ^ m) \<langle>$\<rangle> a = (f ^ m) \<langle>$\<rangle> ((f ^ (n - m)) \<langle>$\<rangle> a)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   458
        by (simp add: power_add apply_times)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   459
      then have "(f ^ (n - m)) \<langle>$\<rangle> a = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   460
        by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   461
      with \<open>n - m > 0\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   462
      have "order f a = card ((\<lambda>m. (f ^ m) \<langle>$\<rangle> a) ` {0..<n - m})"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   463
         by (rule order_witness_unfold)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   464
      also have "card ((\<lambda>m. (f ^ m) \<langle>$\<rangle> a) ` {0..<n - m}) \<le> card {0..<n - m}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   465
        by (rule card_image_le) simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   466
      finally have "order f a \<le> n - m"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   467
        by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   468
      with prem show False by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   469
    qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   470
    with hyp show ?case
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   471
      by (simp add: lessThan_Suc)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   472
  qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   473
  then show ?thesis by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   474
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   475
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   476
lemma orbit_unfold_image:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   477
  "orbit f a = (\<lambda>n. (f ^ n) \<langle>$\<rangle> a) ` {..<order f a}" (is "_ = ?A")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   478
proof (rule sym, rule card_subset_eq)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   479
  show "finite (orbit f a)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   480
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   481
  show "?A \<subseteq> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   482
    by (auto simp add: orbit_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   483
  from inj_on_apply_range [of f a]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   484
  have "card ?A = order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   485
    by (auto simp add: card_image)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   486
  then show "card ?A = card (orbit f a)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   487
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   488
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   489
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   490
lemma in_orbitE:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   491
  assumes "b \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   492
  obtains n where "b = (f ^ n) \<langle>$\<rangle> a" and "n < order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   493
  using assms unfolding orbit_unfold_image by blast
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   494
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   495
lemma apply_power_order [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   496
  "(f ^ order f a) \<langle>$\<rangle> a = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   497
proof -
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   498
  have "(f ^ order f a) \<langle>$\<rangle> a \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   499
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   500
  then obtain n where
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   501
    *: "(f ^ order f a) \<langle>$\<rangle> a = (f ^ n) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   502
    and "n < order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   503
    by (rule in_orbitE)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   504
  show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   505
  proof (cases n)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   506
    case 0 with * show ?thesis by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   507
  next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   508
    case (Suc m)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   509
    from order_greater_zero [of f a]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   510
      have "Suc (order f a - 1) = order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   511
      by arith
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   512
    from Suc \<open>n < order f a\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   513
      have "m < order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   514
      by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   515
    with Suc *
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   516
    have "(inverse f) \<langle>$\<rangle> ((f ^ Suc (order f a - 1)) \<langle>$\<rangle> a) =
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   517
      (inverse f) \<langle>$\<rangle> ((f ^ Suc m) \<langle>$\<rangle> a)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   518
      by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   519
    then have "(f ^ (order f a - 1)) \<langle>$\<rangle> a =
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   520
      (f ^ m) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   521
      by (simp only: power_Suc apply_times)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   522
        (simp add: apply_sequence mult.assoc [symmetric])
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   523
    with inj_on_apply_range
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   524
    have "order f a - 1 = m"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   525
      by (rule inj_onD)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   526
        (simp_all add: \<open>m < order f a\<close>)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   527
    with Suc have "n = order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   528
      by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   529
    with \<open>n < order f a\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   530
    show ?thesis by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   531
  qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   532
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   533
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   534
lemma apply_power_left_mult_order [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   535
  "(f ^ (n * order f a)) \<langle>$\<rangle> a = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   536
  by (induct n) (simp_all add: power_add apply_times)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   537
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   538
lemma apply_power_right_mult_order [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   539
  "(f ^ (order f a * n)) \<langle>$\<rangle> a = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   540
  by (simp add: ac_simps)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   541
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   542
lemma apply_power_mod_order_eq [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   543
  "(f ^ (n mod order f a)) \<langle>$\<rangle> a = (f ^ n) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   544
proof -
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   545
  have "(f ^ n) \<langle>$\<rangle> a = (f ^ (n mod order f a + order f a * (n div order f a))) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   546
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   547
  also have "\<dots> = (f ^ (n mod order f a) * f ^ (order f a * (n div order f a))) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   548
    by (simp add: power_add [symmetric])
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   549
  finally show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   550
    by (simp add: apply_times)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   551
qed  
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   552
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   553
lemma apply_power_eq_iff:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   554
  "(f ^ m) \<langle>$\<rangle> a = (f ^ n) \<langle>$\<rangle> a \<longleftrightarrow> m mod order f a = n mod order f a" (is "?P \<longleftrightarrow> ?Q")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   555
proof
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   556
  assume ?Q
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   557
  then have "(f ^ (m mod order f a)) \<langle>$\<rangle> a = (f ^ (n mod order f a)) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   558
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   559
  then show ?P
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   560
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   561
next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   562
  assume ?P
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   563
  then have "(f ^ (m mod order f a)) \<langle>$\<rangle> a = (f ^ (n mod order f a)) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   564
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   565
  with inj_on_apply_range
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   566
  show ?Q
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   567
    by (rule inj_onD) simp_all
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   568
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   569
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   570
lemma apply_inverse_eq_apply_power_order_minus_one:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   571
  "(inverse f) \<langle>$\<rangle> a = (f ^ (order f a - 1)) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   572
proof (cases "order f a")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   573
  case 0 with order_greater_zero [of f a] show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   574
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   575
next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   576
  case (Suc n)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   577
  moreover have "(f ^ order f a) \<langle>$\<rangle> a = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   578
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   579
  then have *: "(inverse f) \<langle>$\<rangle> ((f ^ order f a) \<langle>$\<rangle> a) = (inverse f) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   580
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   581
  ultimately show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   582
    by (simp add: apply_sequence mult.assoc [symmetric])
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   583
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   584
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   585
lemma apply_inverse_self_in_orbit [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   586
  "(inverse f) \<langle>$\<rangle> a \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   587
  using apply_inverse_eq_apply_power_order_minus_one [symmetric]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   588
  by (rule in_orbitI)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   589
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   590
lemma apply_inverse_power_eq:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   591
  "(inverse (f ^ n)) \<langle>$\<rangle> a = (f ^ (order f a - n mod order f a)) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   592
proof (induct n)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   593
  case 0 then show ?case by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   594
next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   595
  case (Suc n)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   596
  define m where "m = order f a - n mod order f a - 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   597
  moreover have "order f a - n mod order f a > 0"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   598
    by simp
63539
70d4d9e5707b tuned proofs -- avoid improper use of "this";
wenzelm
parents: 63433
diff changeset
   599
  ultimately have *: "order f a - n mod order f a = Suc m"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   600
    by arith
63539
70d4d9e5707b tuned proofs -- avoid improper use of "this";
wenzelm
parents: 63433
diff changeset
   601
  moreover from * have m2: "order f a - Suc n mod order f a = (if m = 0 then order f a else m)"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   602
    by (auto simp add: mod_Suc)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   603
  ultimately show ?case
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   604
    using Suc
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   605
      by (simp_all add: apply_times power_Suc2 [of _ n] power_Suc [of _ m] del: power_Suc)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   606
        (simp add: apply_sequence mult.assoc [symmetric])
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   607
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   608
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   609
lemma apply_power_eq_self_iff:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   610
  "(f ^ n) \<langle>$\<rangle> a = a \<longleftrightarrow> order f a dvd n"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   611
  using apply_power_eq_iff [of f n a 0]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   612
    by (simp add: mod_eq_0_iff_dvd)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   613
  
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   614
lemma orbit_equiv:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   615
  assumes "b \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   616
  shows "orbit f b = orbit f a" (is "?B = ?A")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   617
proof
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   618
  from assms obtain n where "n < order f a" and b: "b = (f ^ n) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   619
    by (rule in_orbitE)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   620
  then show "?B \<subseteq> ?A"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   621
    by (auto simp add: apply_sequence power_add [symmetric] intro: in_orbitI elim!: in_orbitE)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   622
  from b have "(inverse (f ^ n)) \<langle>$\<rangle> b = (inverse (f ^ n)) \<langle>$\<rangle> ((f ^ n) \<langle>$\<rangle> a)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   623
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   624
  then have a: "a = (inverse (f ^ n)) \<langle>$\<rangle> b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   625
    by (simp add: apply_sequence)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   626
  then show "?A \<subseteq> ?B"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   627
    apply (auto simp add: apply_sequence power_add [symmetric] intro: in_orbitI elim!: in_orbitE)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   628
    unfolding apply_times comp_def apply_inverse_power_eq
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   629
    unfolding apply_sequence power_add [symmetric]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   630
    apply (rule in_orbitI) apply rule
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   631
    done
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   632
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   633
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   634
lemma orbit_apply [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   635
  "orbit f (f \<langle>$\<rangle> a) = orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   636
  by (rule orbit_equiv) simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   637
  
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   638
lemma order_apply [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   639
  "order f (f \<langle>$\<rangle> a) = order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   640
  by (simp only: order_def comp_def orbit_apply)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   641
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   642
lemma orbit_apply_inverse [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   643
  "orbit f (inverse f \<langle>$\<rangle> a) = orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   644
  by (rule orbit_equiv) simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   645
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   646
lemma order_apply_inverse [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   647
  "order f (inverse f \<langle>$\<rangle> a) = order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   648
  by (simp only: order_def comp_def orbit_apply_inverse)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   649
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   650
lemma orbit_apply_power [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   651
  "orbit f ((f ^ n) \<langle>$\<rangle> a) = orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   652
  by (rule orbit_equiv) simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   653
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   654
lemma order_apply_power [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   655
  "order f ((f ^ n) \<langle>$\<rangle> a) = order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   656
  by (simp only: order_def comp_def orbit_apply_power)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   657
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   658
lemma orbit_inverse [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   659
  "orbit (inverse f) = orbit f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   660
proof (rule ext, rule set_eqI, rule)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   661
  fix b a
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   662
  assume "b \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   663
  then obtain n where b: "b = (f ^ n) \<langle>$\<rangle> a" "n < order f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   664
    by (rule in_orbitE)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   665
  then have "b = apply (inverse (inverse f) ^ n) a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   666
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   667
  then have "b = apply (inverse (inverse f ^ n)) a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   668
    by (simp add: perm_power_inverse)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   669
  then have "b = apply (inverse f ^ (n * (order (inverse f ^ n) a - 1))) a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   670
    by (simp add: apply_inverse_eq_apply_power_order_minus_one power_mult)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   671
  then show "b \<in> orbit (inverse f) a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   672
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   673
next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   674
  fix b a
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   675
  assume "b \<in> orbit (inverse f) a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   676
  then show "b \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   677
    by (rule in_orbitE)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   678
      (simp add: apply_inverse_eq_apply_power_order_minus_one
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   679
      perm_power_inverse power_mult [symmetric])
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   680
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   681
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   682
lemma order_inverse [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   683
  "order (inverse f) = order f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   684
  by (simp add: order_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   685
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   686
lemma orbit_disjoint:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   687
  assumes "orbit f a \<noteq> orbit f b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   688
  shows "orbit f a \<inter> orbit f b = {}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   689
proof (rule ccontr)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   690
  assume "orbit f a \<inter> orbit f b \<noteq> {}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   691
  then obtain c where "c \<in> orbit f a \<inter> orbit f b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   692
    by blast
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   693
  then have "c \<in> orbit f a" and "c \<in> orbit f b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   694
    by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   695
  then obtain m n where "c = (f ^ m) \<langle>$\<rangle> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   696
    and "c = apply (f ^ n) b" by (blast elim!: in_orbitE)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   697
  then have "(f ^ m) \<langle>$\<rangle> a = apply (f ^ n) b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   698
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   699
  then have "apply (inverse f ^ m) ((f ^ m) \<langle>$\<rangle> a) =
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   700
    apply (inverse f ^ m) (apply (f ^ n) b)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   701
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   702
  then have *: "apply (inverse f ^ m * f ^ n) b = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   703
    by (simp add: apply_sequence perm_power_inverse)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   704
  have "a \<in> orbit f b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   705
  proof (cases n m rule: linorder_cases)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   706
    case equal with * show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   707
      by (simp add: perm_power_inverse)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   708
  next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   709
    case less
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   710
    moreover define q where "q = m - n"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   711
    ultimately have "m = q + n" by arith
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   712
    with * have "apply (inverse f ^ q) b = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   713
      by (simp add: power_add mult.assoc perm_power_inverse)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   714
    then have "a \<in> orbit (inverse f) b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   715
      by (rule in_orbitI)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   716
    then show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   717
      by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   718
  next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   719
    case greater
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   720
    moreover define q where "q = n - m"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   721
    ultimately have "n = m + q" by arith
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   722
    with * have "apply (f ^ q) b = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   723
      by (simp add: power_add mult.assoc [symmetric] perm_power_inverse)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   724
    then show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   725
      by (rule in_orbitI)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   726
  qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   727
  with assms show False
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   728
    by (auto dest: orbit_equiv)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   729
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   730
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   731
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   732
subsection \<open>Swaps\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   733
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   734
lift_definition swap :: "'a \<Rightarrow> 'a \<Rightarrow> 'a perm"  ("\<langle>_\<leftrightarrow>_\<rangle>")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   735
  is "\<lambda>a b. Fun.swap a b id"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   736
proof
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   737
  fix a b :: 'a
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   738
  have "{c. Fun.swap a b id c \<noteq> c} \<subseteq> {a, b}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   739
    by (auto simp add: Fun.swap_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   740
  then show "finite {c. Fun.swap a b id c \<noteq> c}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   741
    by (rule finite_subset) simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   742
qed simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   743
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   744
lemma apply_swap_simp [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   745
  "\<langle>a\<leftrightarrow>b\<rangle> \<langle>$\<rangle> a = b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   746
  "\<langle>a\<leftrightarrow>b\<rangle> \<langle>$\<rangle> b = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   747
  by (transfer; simp)+
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   748
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   749
lemma apply_swap_same [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   750
  "c \<noteq> a \<Longrightarrow> c \<noteq> b \<Longrightarrow> \<langle>a\<leftrightarrow>b\<rangle> \<langle>$\<rangle> c = c"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   751
  by transfer simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   752
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   753
lemma apply_swap_eq_iff [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   754
  "\<langle>a\<leftrightarrow>b\<rangle> \<langle>$\<rangle> c = a \<longleftrightarrow> c = b"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   755
  "\<langle>a\<leftrightarrow>b\<rangle> \<langle>$\<rangle> c = b \<longleftrightarrow> c = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   756
  by (transfer; auto simp add: Fun.swap_def)+
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   757
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   758
lemma swap_1 [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   759
  "\<langle>a\<leftrightarrow>a\<rangle> = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   760
  by transfer simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   761
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   762
lemma swap_sym:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   763
  "\<langle>b\<leftrightarrow>a\<rangle> = \<langle>a\<leftrightarrow>b\<rangle>"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   764
  by (transfer; auto simp add: Fun.swap_def)+
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   765
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   766
lemma swap_self [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   767
  "\<langle>a\<leftrightarrow>b\<rangle> * \<langle>a\<leftrightarrow>b\<rangle> = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   768
  by transfer (simp add: Fun.swap_def fun_eq_iff)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   769
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   770
lemma affected_swap:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   771
  "a \<noteq> b \<Longrightarrow> affected \<langle>a\<leftrightarrow>b\<rangle> = {a, b}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   772
  by transfer (auto simp add: Fun.swap_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   773
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   774
lemma inverse_swap [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   775
  "inverse \<langle>a\<leftrightarrow>b\<rangle> = \<langle>a\<leftrightarrow>b\<rangle>"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   776
  by transfer (auto intro: inv_equality simp: Fun.swap_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   777
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   778
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   779
subsection \<open>Permutations specified by cycles\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   780
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   781
fun cycle :: "'a list \<Rightarrow> 'a perm"  ("\<langle>_\<rangle>")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   782
where
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   783
  "\<langle>[]\<rangle> = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   784
| "\<langle>[a]\<rangle> = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   785
| "\<langle>a # b # as\<rangle> = \<langle>a # as\<rangle> * \<langle>a\<leftrightarrow>b\<rangle>"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   786
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   787
text \<open>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   788
  We do not continue and restrict ourselves to syntax from here.
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   789
  See also introductory note.
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   790
\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   791
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   792
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   793
subsection \<open>Syntax\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   794
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   795
bundle no_permutation_syntax
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   796
begin
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   797
  no_notation swap    ("\<langle>_\<leftrightarrow>_\<rangle>")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   798
  no_notation cycle   ("\<langle>_\<rangle>")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   799
  no_notation "apply" (infixl "\<langle>$\<rangle>" 999)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   800
end
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   801
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   802
bundle permutation_syntax
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   803
begin
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   804
  notation swap       ("\<langle>_\<leftrightarrow>_\<rangle>")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   805
  notation cycle      ("\<langle>_\<rangle>")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   806
  notation "apply"    (infixl "\<langle>$\<rangle>" 999)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   807
  no_notation "apply" ("op \<langle>$\<rangle>")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   808
end
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   809
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   810
unbundle no_permutation_syntax
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   811
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   812
end