src/HOL/Combinatorics/Orbits.thy
author wenzelm
Tue, 02 Aug 2022 12:57:04 +0200
changeset 75734 7671f9fc66d7
parent 73555 92783562ab78
permissions -rw-r--r--
removed somewhat pointless operations (see a6c69599ab99);
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
73555
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
     1
(*  Author:     Lars Noschinski
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
     2
*)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
     3
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
     4
section \<open>Permutation orbits\<close>
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
     5
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
     6
theory Orbits
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
     7
imports
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
     8
  "HOL-Library.FuncSet"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
     9
  "HOL-Combinatorics.Permutations"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    10
begin
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    11
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    12
subsection \<open>Orbits and cyclic permutations\<close>
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    13
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    14
inductive_set orbit :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a set" for f x where
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    15
  base: "f x \<in> orbit f x" |
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    16
  step: "y \<in> orbit f x \<Longrightarrow> f y \<in> orbit f x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    17
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    18
definition cyclic_on :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a set \<Rightarrow> bool" where
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    19
  "cyclic_on f S \<longleftrightarrow> (\<exists>s\<in>S. S = orbit f s)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    20
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    21
lemma orbit_altdef: "orbit f x = {(f ^^ n) x | n. 0 < n}" (is "?L = ?R")
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    22
proof (intro set_eqI iffI)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    23
  fix y assume "y \<in> ?L" then show "y \<in> ?R"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    24
    by (induct rule: orbit.induct) (auto simp: exI[where x=1] exI[where x="Suc n" for n])
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    25
next
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    26
  fix y assume "y \<in> ?R"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    27
  then obtain n where "y = (f ^^ n) x" "0 < n" by blast
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    28
  then show "y \<in> ?L"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    29
  proof (induction n arbitrary: y)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    30
    case (Suc n) then show ?case by (cases "n = 0") (auto intro: orbit.intros)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    31
  qed simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    32
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    33
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    34
lemma orbit_trans:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    35
  assumes "s \<in> orbit f t" "t \<in> orbit f u" shows "s \<in> orbit f u"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    36
  using assms by induct (auto intro: orbit.intros)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    37
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    38
lemma orbit_subset:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    39
  assumes "s \<in> orbit f (f t)" shows "s \<in> orbit f t"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    40
  using assms by (induct) (auto intro: orbit.intros)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    41
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    42
lemma orbit_sim_step:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    43
  assumes "s \<in> orbit f t" shows "f s \<in> orbit f (f t)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    44
  using assms by induct (auto intro: orbit.intros)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    45
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    46
lemma orbit_step:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    47
  assumes "y \<in> orbit f x" "f x \<noteq> y" shows "y \<in> orbit f (f x)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    48
  using assms
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    49
proof induction
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    50
  case (step y) then show ?case by (cases "x = y") (auto intro: orbit.intros)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    51
qed simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    52
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    53
lemma self_in_orbit_trans:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    54
  assumes "s \<in> orbit f s" "t \<in> orbit f s" shows "t \<in> orbit f t"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    55
  using assms(2,1) by induct (auto intro: orbit_sim_step)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    56
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    57
lemma orbit_swap:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    58
  assumes "s \<in> orbit f s" "t \<in> orbit f s" shows "s \<in> orbit f t"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    59
  using assms(2,1)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    60
proof induction
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    61
  case base then show ?case by (cases "f s = s") (auto intro: orbit_step)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    62
next
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    63
  case (step x) then show ?case by (cases "f x = s") (auto intro: orbit_step)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    64
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    65
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    66
lemma permutation_self_in_orbit:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    67
  assumes "permutation f" shows "s \<in> orbit f s"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    68
  unfolding orbit_altdef using permutation_self[OF assms, of s] by simp metis
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    69
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    70
lemma orbit_altdef_self_in:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    71
  assumes "s \<in> orbit f s" shows "orbit f s = {(f ^^ n) s | n. True}"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    72
proof (intro set_eqI iffI)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    73
  fix x assume "x \<in> {(f ^^ n) s | n. True}"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    74
  then obtain n where "x = (f ^^ n) s" by auto
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    75
  then show "x \<in> orbit f s" using assms by (cases "n = 0") (auto simp: orbit_altdef)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    76
qed (auto simp: orbit_altdef)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    77
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    78
lemma orbit_altdef_permutation:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    79
  assumes "permutation f" shows "orbit f s = {(f ^^ n) s | n. True}"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    80
  using assms by (intro orbit_altdef_self_in permutation_self_in_orbit)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    81
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    82
lemma orbit_altdef_bounded:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    83
  assumes "(f ^^ n) s = s" "0 < n" shows "orbit f s = {(f ^^ m) s| m. m < n}"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    84
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    85
  from assms have "s \<in> orbit f s"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    86
    by (auto simp add: orbit_altdef) metis 
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    87
  then have "orbit f s = {(f ^^ m) s|m. True}" by (rule orbit_altdef_self_in)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    88
  also have "\<dots> = {(f ^^ m) s| m. m < n}"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    89
    using assms
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    90
    by (auto simp: funpow_mod_eq intro: exI[where x="m mod n" for m])
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    91
  finally show ?thesis .
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    92
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    93
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    94
lemma funpow_in_orbit:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    95
  assumes "s \<in> orbit f t" shows "(f ^^ n) s \<in> orbit f t"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    96
  using assms by (induct n) (auto intro: orbit.intros)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    97
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    98
lemma finite_orbit:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
    99
  assumes "s \<in> orbit f s" shows "finite (orbit f s)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   100
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   101
  from assms obtain n where n: "0 < n" "(f ^^ n) s = s"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   102
    by (auto simp: orbit_altdef)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   103
  then show ?thesis by (auto simp: orbit_altdef_bounded)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   104
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   105
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   106
lemma self_in_orbit_step:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   107
  assumes "s \<in> orbit f s" shows "orbit f (f s) = orbit f s"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   108
proof (intro set_eqI iffI)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   109
  fix t assume "t \<in> orbit f s" then show "t \<in> orbit f (f s)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   110
    using assms by (auto intro: orbit_step orbit_sim_step)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   111
qed (auto intro: orbit_subset)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   112
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   113
lemma permutation_orbit_step:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   114
  assumes "permutation f" shows "orbit f (f s) = orbit f s"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   115
  using assms by (intro self_in_orbit_step permutation_self_in_orbit)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   116
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   117
lemma orbit_nonempty:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   118
  "orbit f s \<noteq> {}"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   119
  using orbit.base by fastforce
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   120
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   121
lemma orbit_inv_eq:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   122
  assumes "permutation f"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   123
  shows "orbit (inv f) x = orbit f x" (is "?L = ?R")
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   124
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   125
  { fix g y assume A: "permutation g" "y \<in> orbit (inv g) x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   126
    have "y \<in> orbit g x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   127
    proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   128
      have inv_g: "\<And>y. x = g y \<Longrightarrow> inv g x = y" "\<And>y. inv g (g y) = y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   129
        by (metis A(1) bij_inv_eq_iff permutation_bijective)+
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   130
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   131
      { fix y assume "y \<in> orbit g x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   132
        then have "inv g y \<in> orbit g x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   133
          by (cases) (simp_all add: inv_g A(1) permutation_self_in_orbit)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   134
      } note inv_g_in_orb = this
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   135
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   136
      from A(2) show ?thesis
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   137
        by induct (simp_all add: inv_g_in_orb A permutation_self_in_orbit)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   138
    qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   139
  } note orb_inv_ss = this
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   140
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   141
  have "inv (inv f) = f"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   142
    by (simp add: assms inv_inv_eq permutation_bijective)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   143
  then show ?thesis
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   144
    using orb_inv_ss[OF assms] orb_inv_ss[OF permutation_inverse[OF assms]] by auto
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   145
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   146
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   147
lemma cyclic_on_alldef:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   148
  "cyclic_on f S \<longleftrightarrow> S \<noteq> {} \<and> (\<forall>s\<in>S. S = orbit f s)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   149
  unfolding cyclic_on_def by (auto intro: orbit.step orbit_swap orbit_trans)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   150
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   151
lemma cyclic_on_funpow_in:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   152
  assumes "cyclic_on f S" "s \<in> S" shows "(f^^n) s \<in> S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   153
  using assms unfolding cyclic_on_def by (auto intro: funpow_in_orbit)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   154
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   155
lemma finite_cyclic_on:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   156
  assumes "cyclic_on f S" shows "finite S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   157
  using assms by (auto simp: cyclic_on_def finite_orbit)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   158
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   159
lemma cyclic_on_singleI:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   160
  assumes "s \<in> S" "S = orbit f s" shows "cyclic_on f S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   161
  using assms unfolding cyclic_on_def by blast
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   162
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   163
lemma cyclic_on_inI:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   164
  assumes "cyclic_on f S" "s \<in> S" shows "f s \<in> S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   165
  using assms by (auto simp: cyclic_on_def intro: orbit.intros)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   166
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   167
lemma orbit_inverse:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   168
  assumes self: "a \<in> orbit g a"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   169
    and eq: "\<And>x. x \<in> orbit g a \<Longrightarrow> g' (f x) = f (g x)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   170
  shows "f ` orbit g a = orbit g' (f a)" (is "?L = ?R")
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   171
proof (intro set_eqI iffI)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   172
  fix x assume "x \<in> ?L"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   173
  then obtain x0 where "x0 \<in> orbit g a" "x = f x0" by auto
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   174
  then show "x \<in> ?R"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   175
  proof (induct arbitrary: x)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   176
    case base then show ?case by (auto simp: self orbit.base eq[symmetric])
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   177
  next
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   178
    case step then show ?case by cases (auto simp: eq[symmetric] orbit.intros)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   179
  qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   180
next
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   181
  fix x assume "x \<in> ?R"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   182
  then show "x \<in> ?L"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   183
  proof (induct arbitrary: )
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   184
    case base then show ?case by (auto simp: self orbit.base eq)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   185
  next
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   186
    case step then show ?case by cases (auto simp: eq orbit.intros)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   187
  qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   188
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   189
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   190
lemma cyclic_on_image:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   191
  assumes "cyclic_on f S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   192
  assumes "\<And>x. x \<in> S \<Longrightarrow> g (h x) = h (f x)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   193
  shows "cyclic_on g (h ` S)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   194
  using assms by (auto simp: cyclic_on_def) (meson orbit_inverse)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   195
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   196
lemma cyclic_on_f_in:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   197
  assumes "f permutes S" "cyclic_on f A" "f x \<in> A"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   198
  shows "x \<in> A"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   199
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   200
  from assms have fx_in_orb: "f x \<in> orbit f (f x)" by (auto simp: cyclic_on_alldef)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   201
  from assms have "A = orbit f (f x)" by (auto simp: cyclic_on_alldef)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   202
  moreover
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   203
  then have "\<dots> = orbit f x" using \<open>f x \<in> A\<close> by (auto intro: orbit_step orbit_subset)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   204
  ultimately
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   205
    show ?thesis by (metis (no_types) orbit.simps permutes_inverses(2)[OF assms(1)])
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   206
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   207
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   208
lemma orbit_cong0:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   209
  assumes "x \<in> A" "f \<in> A \<rightarrow> A" "\<And>y. y \<in> A \<Longrightarrow> f y = g y" shows "orbit f x = orbit g x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   210
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   211
  { fix n have "(f ^^ n) x = (g ^^ n) x \<and> (f ^^ n) x \<in> A"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   212
      by (induct n rule: nat.induct) (insert assms, auto)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   213
  } then show ?thesis by (auto simp: orbit_altdef)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   214
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   215
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   216
lemma orbit_cong:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   217
  assumes self_in: "t \<in> orbit f t" and eq: "\<And>s. s \<in> orbit f t \<Longrightarrow> g s = f s"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   218
  shows "orbit g t = orbit f t"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   219
  using assms(1) _ assms(2) by (rule orbit_cong0) (auto simp: orbit.step eq)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   220
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   221
lemma cyclic_cong:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   222
  assumes "\<And>s. s \<in> S \<Longrightarrow> f s = g s" shows "cyclic_on f S = cyclic_on g S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   223
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   224
  have "(\<exists>s\<in>S. orbit f s = orbit g s) \<Longrightarrow> cyclic_on f S = cyclic_on g S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   225
    by (metis cyclic_on_alldef cyclic_on_def)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   226
  then show ?thesis by (metis assms orbit_cong cyclic_on_def)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   227
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   228
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   229
lemma permutes_comp_preserves_cyclic1:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   230
  assumes "g permutes B" "cyclic_on f C"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   231
  assumes "A \<inter> B = {}" "C \<subseteq> A"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   232
  shows "cyclic_on (f o g) C"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   233
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   234
  have *: "\<And>c. c \<in> C \<Longrightarrow> f (g c) = f c"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   235
    using assms by (subst permutes_not_in [of g]) auto
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   236
  with assms(2) show ?thesis by (simp cong: cyclic_cong)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   237
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   238
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   239
lemma permutes_comp_preserves_cyclic2:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   240
  assumes "f permutes A" "cyclic_on g C"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   241
  assumes "A \<inter> B = {}" "C \<subseteq> B"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   242
  shows "cyclic_on (f o g) C"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   243
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   244
  obtain c where c: "c \<in> C" "C = orbit g c" "c \<in> orbit g c"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   245
    using \<open>cyclic_on g C\<close> by (auto simp: cyclic_on_def)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   246
  then have "\<And>c. c \<in> C \<Longrightarrow> f (g c) = g c"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   247
    using assms c by (subst permutes_not_in [of f]) (auto intro: orbit.intros)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   248
  with assms(2) show ?thesis by (simp cong: cyclic_cong)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   249
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   250
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   251
lemma permutes_orbit_subset:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   252
  assumes "f permutes S" "x \<in> S" shows "orbit f x \<subseteq> S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   253
proof
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   254
  fix y assume "y \<in> orbit f x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   255
  then show "y \<in> S" by induct (auto simp: permutes_in_image assms)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   256
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   257
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   258
lemma cyclic_on_orbit':
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   259
  assumes "permutation f" shows "cyclic_on f (orbit f x)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   260
  unfolding cyclic_on_alldef using orbit_nonempty[of f x]
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   261
  by (auto intro: assms orbit_swap orbit_trans permutation_self_in_orbit)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   262
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   263
lemma cyclic_on_orbit:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   264
  assumes "f permutes S" "finite S" shows "cyclic_on f (orbit f x)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   265
  using assms by (intro cyclic_on_orbit') (auto simp: permutation_permutes)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   266
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   267
lemma orbit_cyclic_eq3:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   268
  assumes "cyclic_on f S" "y \<in> S" shows "orbit f y = S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   269
  using assms unfolding cyclic_on_alldef by simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   270
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   271
lemma orbit_eq_singleton_iff: "orbit f x = {x} \<longleftrightarrow> f x = x" (is "?L \<longleftrightarrow> ?R")
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   272
proof
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   273
  assume A: ?R
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   274
  { fix y assume "y \<in> orbit f x" then have "y = x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   275
      by induct (auto simp: A)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   276
  } then show ?L by (metis orbit_nonempty singletonI subsetI subset_singletonD)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   277
next
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   278
  assume A: ?L
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   279
  then have "\<And>y. y \<in> orbit f x \<Longrightarrow> f x = y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   280
    by - (erule orbit.cases, simp_all)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   281
  then show ?R using A by blast
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   282
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   283
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   284
lemma eq_on_cyclic_on_iff1:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   285
  assumes "cyclic_on f S" "x \<in> S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   286
  obtains "f x \<in> S" "f x = x \<longleftrightarrow> card S = 1"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   287
proof
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   288
  from assms show "f x \<in> S" by (auto simp: cyclic_on_def intro: orbit.intros)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   289
  from assms have "S = orbit f x" by (auto simp: cyclic_on_alldef)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   290
  then have "f x = x \<longleftrightarrow> S = {x}" by (metis orbit_eq_singleton_iff)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   291
  then show "f x = x \<longleftrightarrow> card S = 1" using \<open>x \<in> S\<close> by (auto simp: card_Suc_eq)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   292
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   293
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   294
lemma orbit_eqI:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   295
  "y = f x \<Longrightarrow> y \<in> orbit f x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   296
  "z = f y \<Longrightarrow>y \<in> orbit f x \<Longrightarrow>z \<in> orbit f x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   297
  by (metis orbit.base) (metis orbit.step)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   298
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   299
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   300
subsection \<open>Decomposition of arbitrary permutations\<close>
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   301
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   302
definition perm_restrict :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a set \<Rightarrow> ('a \<Rightarrow> 'a)" where
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   303
  "perm_restrict f S x \<equiv> if x \<in> S then f x else x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   304
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   305
lemma perm_restrict_comp:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   306
  assumes "A \<inter> B = {}" "cyclic_on f B"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   307
  shows "perm_restrict f A o perm_restrict f B = perm_restrict f (A \<union> B)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   308
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   309
  have "\<And>x. x \<in> B \<Longrightarrow> f x \<in> B" using \<open>cyclic_on f B\<close> by (rule cyclic_on_inI)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   310
  with assms show ?thesis by (auto simp: perm_restrict_def fun_eq_iff)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   311
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   312
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   313
lemma perm_restrict_simps:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   314
  "x \<in> S \<Longrightarrow> perm_restrict f S x = f x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   315
  "x \<notin> S \<Longrightarrow> perm_restrict f S x = x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   316
  by (auto simp: perm_restrict_def)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   317
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   318
lemma perm_restrict_perm_restrict:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   319
  "perm_restrict (perm_restrict f A) B = perm_restrict f (A \<inter> B)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   320
  by (auto simp: perm_restrict_def)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   321
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   322
lemma perm_restrict_union:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   323
  assumes "perm_restrict f A permutes A" "perm_restrict f B permutes B" "A \<inter> B = {}"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   324
  shows "perm_restrict f A o perm_restrict f B = perm_restrict f (A \<union> B)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   325
  using assms by (auto simp: fun_eq_iff perm_restrict_def permutes_def) (metis Diff_iff Diff_triv)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   326
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   327
lemma perm_restrict_id[simp]:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   328
  assumes "f permutes S" shows "perm_restrict f S = f"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   329
  using assms by (auto simp: permutes_def perm_restrict_def)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   330
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   331
lemma cyclic_on_perm_restrict:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   332
  "cyclic_on (perm_restrict f S) S \<longleftrightarrow> cyclic_on f S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   333
  by (simp add: perm_restrict_def cong: cyclic_cong)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   334
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   335
lemma perm_restrict_diff_cyclic:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   336
  assumes "f permutes S" "cyclic_on f A"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   337
  shows "perm_restrict f (S - A) permutes (S - A)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   338
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   339
  { fix y
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   340
    have "\<exists>x. perm_restrict f (S - A) x = y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   341
    proof cases
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   342
      assume A: "y \<in> S - A"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   343
      with \<open>f permutes S\<close> obtain x where "f x = y" "x \<in> S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   344
        unfolding permutes_def by auto metis
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   345
      moreover
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   346
      with A have "x \<notin> A" by (metis Diff_iff assms(2) cyclic_on_inI)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   347
      ultimately
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   348
      have "perm_restrict f (S - A) x = y"  by (simp add: perm_restrict_simps)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   349
      then show ?thesis ..
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   350
    next
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   351
      assume "y \<notin> S - A"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   352
      then have "perm_restrict f (S - A) y = y" by (simp add: perm_restrict_simps)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   353
      then show ?thesis ..
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   354
    qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   355
  } note X = this
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   356
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   357
  { fix x y assume "perm_restrict f (S - A) x = perm_restrict f (S - A) y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   358
    with assms have "x = y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   359
      by (auto simp: perm_restrict_def permutes_def split: if_splits intro: cyclic_on_f_in)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   360
  } note Y = this
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   361
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   362
  show ?thesis by (auto simp: permutes_def perm_restrict_simps X intro: Y)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   363
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   364
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   365
lemma permutes_decompose:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   366
  assumes "f permutes S" "finite S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   367
  shows "\<exists>C. (\<forall>c \<in> C. cyclic_on f c) \<and> \<Union>C = S \<and> (\<forall>c1 \<in> C. \<forall>c2 \<in> C. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {})"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   368
  using assms(2,1)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   369
proof (induction arbitrary: f rule: finite_psubset_induct)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   370
  case (psubset S)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   371
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   372
  show ?case
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   373
  proof (cases "S = {}")
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   374
    case True then show ?thesis by (intro exI[where x="{}"]) auto
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   375
  next
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   376
    case False
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   377
    then obtain s where "s \<in> S" by auto
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   378
    with \<open>f permutes S\<close> have "orbit f s \<subseteq> S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   379
      by (rule permutes_orbit_subset)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   380
    have cyclic_orbit: "cyclic_on f (orbit f s)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   381
      using \<open>f permutes S\<close> \<open>finite S\<close> by (rule cyclic_on_orbit)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   382
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   383
    let ?f' = "perm_restrict f (S - orbit f s)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   384
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   385
    have "f s \<in> S" using \<open>f permutes S\<close> \<open>s \<in> S\<close> by (auto simp: permutes_in_image)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   386
    then have "S - orbit f s \<subset> S" using orbit.base[of f s] \<open>s \<in> S\<close> by blast
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   387
    moreover
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   388
    have "?f' permutes (S - orbit f s)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   389
      using \<open>f permutes S\<close> cyclic_orbit by (rule perm_restrict_diff_cyclic)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   390
    ultimately
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   391
    obtain C where C: "\<And>c. c \<in> C \<Longrightarrow> cyclic_on ?f' c" "\<Union>C = S - orbit f s"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   392
        "\<forall>c1 \<in> C. \<forall>c2 \<in> C. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {}"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   393
      using psubset.IH by metis
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   394
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   395
    { fix c assume "c \<in> C"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   396
      then have *: "\<And>x. x \<in> c \<Longrightarrow> perm_restrict f (S - orbit f s) x = f x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   397
        using C(2) \<open>f permutes S\<close> by (auto simp add: perm_restrict_def)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   398
      then have "cyclic_on f c" using C(1)[OF \<open>c \<in> C\<close>] by (simp cong: cyclic_cong add: *)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   399
    } note in_C_cyclic = this
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   400
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   401
    have Un_ins: "\<Union>(insert (orbit f s) C) = S"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   402
      using \<open>\<Union>C = _\<close>  \<open>orbit f s \<subseteq> S\<close> by blast
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   403
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   404
    have Disj_ins: "(\<forall>c1 \<in> insert (orbit f s) C. \<forall>c2 \<in> insert (orbit f s) C. c1 \<noteq> c2 \<longrightarrow> c1 \<inter> c2 = {})"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   405
      using C by auto
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   406
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   407
    show ?thesis
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   408
      by (intro conjI Un_ins Disj_ins exI[where x="insert (orbit f s) C"])
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   409
        (auto simp: cyclic_orbit in_C_cyclic)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   410
  qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   411
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   412
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   413
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   414
subsection \<open>Function-power distance between values\<close>
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   415
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   416
definition funpow_dist :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> nat" where
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   417
  "funpow_dist f x y \<equiv> LEAST n. (f ^^ n) x = y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   418
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   419
abbreviation funpow_dist1 :: "('a \<Rightarrow> 'a) \<Rightarrow> 'a \<Rightarrow> 'a \<Rightarrow> nat" where
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   420
  "funpow_dist1 f x y \<equiv> Suc (funpow_dist f (f x) y)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   421
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   422
lemma funpow_dist_0:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   423
  assumes "x = y" shows "funpow_dist f x y = 0"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   424
  using assms unfolding funpow_dist_def by (intro Least_eq_0) simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   425
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   426
lemma funpow_dist_least:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   427
  assumes "n < funpow_dist f x y" shows "(f ^^ n) x \<noteq> y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   428
proof (rule notI)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   429
  assume "(f ^^ n) x = y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   430
  then have "funpow_dist f x y \<le> n" unfolding funpow_dist_def by (rule Least_le)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   431
  with assms show False by linarith
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   432
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   433
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   434
lemma funpow_dist1_least:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   435
  assumes "0 < n" "n < funpow_dist1 f x y" shows "(f ^^ n) x \<noteq> y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   436
proof (rule notI)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   437
  assume "(f ^^ n) x = y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   438
  then have "(f ^^ (n - 1)) (f x) = y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   439
    using \<open>0 < n\<close> by (cases n) (simp_all add: funpow_swap1)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   440
  then have "funpow_dist f (f x) y \<le> n - 1" unfolding funpow_dist_def by (rule Least_le)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   441
  with assms show False by simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   442
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   443
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   444
lemma funpow_dist_prop:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   445
  "y \<in> orbit f x \<Longrightarrow> (f ^^ funpow_dist f x y) x = y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   446
  unfolding funpow_dist_def by (rule LeastI_ex) (auto simp: orbit_altdef)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   447
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   448
lemma funpow_dist_0_eq:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   449
  assumes "y \<in> orbit f x" shows "funpow_dist f x y = 0 \<longleftrightarrow> x = y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   450
  using assms by (auto simp: funpow_dist_0 dest: funpow_dist_prop)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   451
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   452
lemma funpow_dist_step:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   453
  assumes "x \<noteq> y" "y \<in> orbit f x" shows "funpow_dist f x y = Suc (funpow_dist f (f x) y)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   454
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   455
  from \<open>y \<in> _\<close> obtain n where "(f ^^ n) x = y" by (auto simp: orbit_altdef)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   456
  with \<open>x \<noteq> y\<close> obtain n' where [simp]: "n = Suc n'" by (cases n) auto
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   457
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   458
  show ?thesis
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   459
    unfolding funpow_dist_def
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   460
  proof (rule Least_Suc2)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   461
    show "(f ^^ n) x = y" by fact
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   462
    then show "(f ^^ n') (f x) = y" by (simp add: funpow_swap1)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   463
    show "(f ^^ 0) x \<noteq> y" using \<open>x \<noteq> y\<close> by simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   464
    show "\<forall>k. ((f ^^ Suc k) x = y) = ((f ^^ k) (f x) = y)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   465
      by (simp add: funpow_swap1)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   466
  qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   467
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   468
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   469
lemma funpow_dist1_prop:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   470
  assumes "y \<in> orbit f x" shows "(f ^^ funpow_dist1 f x y) x = y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   471
  by (metis assms funpow_dist_prop funpow_dist_step funpow_simps_right(2) o_apply self_in_orbit_step)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   472
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   473
(*XXX simplify? *)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   474
lemma funpow_neq_less_funpow_dist:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   475
  assumes "y \<in> orbit f x" "m \<le> funpow_dist f x y" "n \<le> funpow_dist f x y" "m \<noteq> n"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   476
  shows "(f ^^ m) x \<noteq> (f ^^ n) x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   477
proof (rule notI)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   478
  assume A: "(f ^^ m) x = (f ^^ n) x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   479
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   480
  define m' n' where "m' = min m n" and "n' = max m n"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   481
  with A assms have A': "m' < n'" "(f ^^ m') x = (f ^^ n') x" "n' \<le> funpow_dist f x y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   482
    by (auto simp: min_def max_def)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   483
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   484
  have "y = (f ^^ funpow_dist f x y) x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   485
    using \<open>y \<in> _\<close> by (simp only: funpow_dist_prop)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   486
  also have "\<dots> = (f ^^ ((funpow_dist f x y - n') + n')) x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   487
    using \<open>n' \<le> _\<close> by simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   488
  also have "\<dots> = (f ^^ ((funpow_dist f x y - n') + m')) x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   489
    by (simp add: funpow_add \<open>(f ^^ m') x = _\<close>)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   490
  also have "(f ^^ ((funpow_dist f x y - n') + m')) x \<noteq> y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   491
    using A' by (intro funpow_dist_least) linarith
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   492
  finally show "False" by simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   493
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   494
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   495
(* XXX reduce to funpow_neq_less_funpow_dist? *)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   496
lemma funpow_neq_less_funpow_dist1:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   497
  assumes "y \<in> orbit f x" "m < funpow_dist1 f x y" "n < funpow_dist1 f x y" "m \<noteq> n"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   498
  shows "(f ^^ m) x \<noteq> (f ^^ n) x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   499
proof (rule notI)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   500
  assume A: "(f ^^ m) x = (f ^^ n) x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   501
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   502
  define m' n' where "m' = min m n" and "n' = max m n"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   503
  with A assms have A': "m' < n'" "(f ^^ m') x = (f ^^ n') x" "n' < funpow_dist1 f x y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   504
    by (auto simp: min_def max_def)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   505
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   506
  have "y = (f ^^ funpow_dist1 f x y) x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   507
    using \<open>y \<in> _\<close> by (simp only: funpow_dist1_prop)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   508
  also have "\<dots> = (f ^^ ((funpow_dist1 f x y - n') + n')) x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   509
    using \<open>n' < _\<close> by simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   510
  also have "\<dots> = (f ^^ ((funpow_dist1 f x y - n') + m')) x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   511
    by (simp add: funpow_add \<open>(f ^^ m') x = _\<close>)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   512
  also have "(f ^^ ((funpow_dist1 f x y - n') + m')) x \<noteq> y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   513
    using A' by (intro funpow_dist1_least) linarith+
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   514
  finally show "False" by simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   515
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   516
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   517
lemma inj_on_funpow_dist:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   518
  assumes "y \<in> orbit f x" shows "inj_on (\<lambda>n. (f ^^ n) x) {0..funpow_dist f x y}"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   519
  using funpow_neq_less_funpow_dist[OF assms] by (intro inj_onI) auto
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   520
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   521
lemma inj_on_funpow_dist1:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   522
  assumes "y \<in> orbit f x" shows "inj_on (\<lambda>n. (f ^^ n) x) {0..<funpow_dist1 f x y}"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   523
  using funpow_neq_less_funpow_dist1[OF assms] by (intro inj_onI) auto
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   524
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   525
lemma orbit_conv_funpow_dist1:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   526
  assumes "x \<in> orbit f x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   527
  shows "orbit f x = (\<lambda>n. (f ^^ n) x) ` {0..<funpow_dist1 f x x}" (is "?L = ?R")
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   528
  using funpow_dist1_prop[OF assms]
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   529
  by (auto simp: orbit_altdef_bounded[where n="funpow_dist1 f x x"])
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   530
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   531
lemma funpow_dist1_prop1:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   532
  assumes "(f ^^ n) x = y" "0 < n" shows "(f ^^ funpow_dist1 f x y) x = y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   533
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   534
  from assms have "y \<in> orbit f x" by (auto simp: orbit_altdef)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   535
  then show ?thesis by (rule funpow_dist1_prop)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   536
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   537
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   538
lemma funpow_dist1_dist:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   539
  assumes "funpow_dist1 f x y < funpow_dist1 f x z"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   540
  assumes "{y,z} \<subseteq> orbit f x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   541
  shows "funpow_dist1 f x z = funpow_dist1 f x y + funpow_dist1 f y z" (is "?L = ?R")
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   542
proof -
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   543
  define n where \<open>n = funpow_dist1 f x z - funpow_dist1 f x y - 1\<close>
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   544
  with assms have *: \<open>funpow_dist1 f x z = Suc (funpow_dist1 f x y + n)\<close>
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   545
    by simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   546
  have x_z: "(f ^^ funpow_dist1 f x z) x = z" using assms by (blast intro: funpow_dist1_prop)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   547
  have x_y: "(f ^^ funpow_dist1 f x y) x = y" using assms by (blast intro: funpow_dist1_prop)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   548
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   549
  have "(f ^^ (funpow_dist1 f x z - funpow_dist1 f x y)) y
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   550
      = (f ^^ (funpow_dist1 f x z - funpow_dist1 f x y)) ((f ^^ funpow_dist1 f x y) x)"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   551
    using x_y by simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   552
  also have "\<dots> = z"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   553
    using assms x_z by (simp add: * funpow_add ac_simps funpow_swap1)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   554
  finally have y_z_diff: "(f ^^ (funpow_dist1 f x z - funpow_dist1 f x y)) y = z" .
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   555
  then have "(f ^^ funpow_dist1 f y z) y = z"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   556
    using assms by (intro funpow_dist1_prop1) auto
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   557
  then have "(f ^^ funpow_dist1 f y z) ((f ^^ funpow_dist1 f x y) x) = z"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   558
    using x_y by simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   559
  then have "(f ^^ (funpow_dist1 f y z + funpow_dist1 f x y)) x = z"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   560
    by (simp add: * funpow_add funpow_swap1)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   561
  show ?thesis
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   562
  proof (rule antisym)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   563
    from y_z_diff have "(f ^^ funpow_dist1 f y z) y = z"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   564
      using assms by (intro funpow_dist1_prop1) auto
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   565
    then have "(f ^^ funpow_dist1 f y z) ((f ^^ funpow_dist1 f x y) x) = z"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   566
      using x_y by simp
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   567
    then have "(f ^^ (funpow_dist1 f y z + funpow_dist1 f x y)) x = z"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   568
      by (simp add: * funpow_add funpow_swap1)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   569
    then have "funpow_dist1 f x z \<le> funpow_dist1 f y z + funpow_dist1 f x y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   570
      using funpow_dist1_least not_less by fastforce
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   571
    then show "?L \<le> ?R" by presburger
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   572
  next
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   573
    have "funpow_dist1 f y z \<le> funpow_dist1 f x z - funpow_dist1 f x y"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   574
      using y_z_diff assms(1) by (metis not_less zero_less_diff funpow_dist1_least)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   575
    then show "?R \<le> ?L" by linarith
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   576
  qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   577
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   578
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   579
lemma funpow_dist1_le_self:
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   580
  assumes "(f ^^ m) x = x" "0 < m" "y \<in> orbit f x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   581
  shows "funpow_dist1 f x y \<le> m"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   582
proof (cases "x = y")
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   583
  case True with assms show ?thesis by (auto dest!: funpow_dist1_least)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   584
next
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   585
  case False
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   586
  have "(f ^^ funpow_dist1 f x y) x = (f ^^ (funpow_dist1 f x y mod m)) x"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   587
    using assms by (simp add: funpow_mod_eq)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   588
  with False \<open>y \<in> orbit f x\<close> have "funpow_dist1 f x y \<le> funpow_dist1 f x y mod m"
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   589
    by auto (metis \<open>(f ^^ funpow_dist1 f x y) x = (f ^^ (funpow_dist1 f x y mod m)) x\<close> funpow_dist1_prop funpow_dist_least funpow_dist_step leI)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   590
  with \<open>m > 0\<close> show ?thesis
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   591
    by (auto intro: order_trans)
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   592
qed
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   593
92783562ab78 collected combinatorial material
haftmann
parents:
diff changeset
   594
end