src/HOL/ex/Perm_Fragments.thy
author wenzelm
Fri, 29 Nov 2024 17:40:15 +0100
changeset 81507 08574da77b4a
parent 73648 1bd3463e30b8
permissions -rw-r--r--
clarified signature: shorten common cases;
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>Fragments on permuations\<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_Fragments
73622
4dc3baf45d6a more appropriate location
haftmann
parents: 72536
diff changeset
     6
imports "HOL-Library.Dlist" "HOL-Combinatorics.Perm"
63375
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>On cycles\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    10
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    11
context
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    12
  includes permutation_syntax
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    13
begin
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    14
63882
018998c00003 renamed listsum -> sum_list, listprod ~> prod_list
nipkow
parents: 63375
diff changeset
    15
lemma cycle_prod_list:
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    16
  "\<langle>a # as\<rangle> = prod_list (map (\<lambda>b. \<langle>a \<leftrightarrow> b\<rangle>) (rev as))"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    17
  by (induct as) simp_all
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
lemma cycle_append [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    20
  "\<langle>a # as @ bs\<rangle> = \<langle>a # bs\<rangle> * \<langle>a # as\<rangle>"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    21
proof (induct as rule: cycle.induct)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    22
  case (3 b c as)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    23
  then have "\<langle>a # (b # as) @ bs\<rangle> = \<langle>a # bs\<rangle> * \<langle>a # b # as\<rangle>"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    24
    by simp
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    25
  then have "\<langle>a # as @ bs\<rangle> * \<langle>a \<leftrightarrow> b\<rangle> =
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    26
    \<langle>a # bs\<rangle> * \<langle>a # as\<rangle> * \<langle>a \<leftrightarrow> b\<rangle>"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    27
    by (simp add: ac_simps)
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    28
  then have "\<langle>a # as @ bs\<rangle> * \<langle>a \<leftrightarrow> b\<rangle> * \<langle>a \<leftrightarrow> b\<rangle> =
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    29
    \<langle>a # bs\<rangle> * \<langle>a # as\<rangle> * \<langle>a \<leftrightarrow> b\<rangle> * \<langle>a \<leftrightarrow> b\<rangle>"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    30
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    31
  then have "\<langle>a # as @ bs\<rangle> = \<langle>a # bs\<rangle> * \<langle>a # as\<rangle>"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    32
    by (simp add: ac_simps)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    33
  then show "\<langle>a # (b # c # as) @ bs\<rangle> =
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    34
    \<langle>a # bs\<rangle> * \<langle>a # b # c # as\<rangle>"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    35
    by (simp add: ac_simps)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    36
qed simp_all
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    37
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    38
lemma affected_cycle:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    39
  "affected \<langle>as\<rangle> \<subseteq> set as"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    40
proof (induct as rule: cycle.induct)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    41
  case (3 a b as)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    42
  from affected_times
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    43
  have "affected (\<langle>a # as\<rangle> * \<langle>a \<leftrightarrow> b\<rangle>)
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    44
    \<subseteq> affected \<langle>a # as\<rangle> \<union> affected \<langle>a \<leftrightarrow> b\<rangle>" .
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    45
  moreover from 3
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    46
  have "affected (\<langle>a # as\<rangle>) \<subseteq> insert a (set as)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    47
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    48
  moreover
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    49
  have "affected \<langle>a \<leftrightarrow> b\<rangle> \<subseteq> {a, b}"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    50
    by (cases "a = b") (simp_all add: affected_swap)
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    51
  ultimately have "affected (\<langle>a # as\<rangle> * \<langle>a \<leftrightarrow> b\<rangle>)
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    52
    \<subseteq> insert a (insert b (set as))"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    53
    by blast
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    54
  then show ?case by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    55
qed simp_all
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    56
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    57
lemma orbit_cycle_non_elem:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    58
  assumes "a \<notin> set as"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    59
  shows "orbit \<langle>as\<rangle> a = {a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    60
  unfolding not_in_affected_iff_orbit_eq_singleton [symmetric]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    61
  using assms affected_cycle [of as] by blast
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    62
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    63
lemma inverse_cycle:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    64
  assumes "distinct as"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    65
  shows "inverse \<langle>as\<rangle> = \<langle>rev as\<rangle>"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    66
using assms proof (induct as rule: cycle.induct)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    67
  case (3 a b as)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    68
  then have *: "inverse \<langle>a # as\<rangle> = \<langle>rev (a # as)\<rangle>"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    69
    and distinct: "distinct (a # b # as)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    70
    by simp_all
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    71
  show " inverse \<langle>a # b # as\<rangle> = \<langle>rev (a # b # as)\<rangle>"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    72
  proof (cases as rule: rev_cases)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    73
    case Nil with * show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    74
      by (simp add: swap_sym)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    75
  next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    76
    case (snoc cs c)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    77
    with distinct have "distinct (a # b # cs @ [c])"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    78
      by simp
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
    79
    then have **: "\<langle>a \<leftrightarrow> b\<rangle> * \<langle>c \<leftrightarrow> a\<rangle> = \<langle>c \<leftrightarrow> a\<rangle> * \<langle>c \<leftrightarrow> b\<rangle>"
73648
1bd3463e30b8 more elementary swap
haftmann
parents: 73622
diff changeset
    80
      by transfer (auto simp add: fun_eq_iff transpose_def)
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    81
    with snoc * show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    82
      by (simp add: mult.assoc [symmetric])
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    83
  qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    84
qed simp_all
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    85
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    86
lemma order_cycle_non_elem:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    87
  assumes "a \<notin> set as"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    88
  shows "order \<langle>as\<rangle> a = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    89
proof -
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    90
  from assms have "orbit \<langle>as\<rangle> a = {a}" 
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    91
    by (rule orbit_cycle_non_elem)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    92
  then have "card (orbit \<langle>as\<rangle> a) = card {a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    93
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    94
  then show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    95
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    96
qed
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
lemma orbit_cycle_elem:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
    99
  assumes "distinct as" and "a \<in> set as"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   100
  shows "orbit \<langle>as\<rangle> a = set as"
67226
ec32cdaab97b isabelle update_cartouches -c -t;
wenzelm
parents: 66453
diff changeset
   101
  oops \<comment> \<open>(we need rotation here\<close>
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   102
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   103
lemma order_cycle_elem:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   104
  assumes "distinct as" and "a \<in> set as"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   105
  shows "order \<langle>as\<rangle> a = length as"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   106
  oops
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   107
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   108
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   109
text \<open>Adding fixpoints\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   110
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   111
definition fixate :: "'a \<Rightarrow> 'a perm \<Rightarrow> 'a perm"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   112
where
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
   113
  "fixate a f = (if a \<in> affected f then f * \<langle>inverse f \<langle>$\<rangle> a \<leftrightarrow> a\<rangle> else f)"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   114
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   115
lemma affected_fixate_trivial:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   116
  assumes "a \<notin> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   117
  shows "affected (fixate a f) = affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   118
  using assms by (simp add: fixate_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   119
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   120
lemma affected_fixate_binary_circle:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   121
  assumes "order f a = 2"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   122
  shows "affected (fixate a f) = affected f - {a, apply f a}" (is "?A = ?B")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   123
proof (rule set_eqI)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   124
  interpret bijection "apply f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   125
    by standard simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   126
  fix b
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   127
  from assms order_greater_eq_two_iff [of f a] have "a \<in> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   128
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   129
  moreover have "apply (f ^ 2) a = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   130
    by (simp add: assms [symmetric])
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   131
  ultimately show "b \<in> ?A \<longleftrightarrow> b \<in> ?B"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   132
    by (cases "b \<in> {a, apply (inverse f) a}")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   133
      (auto simp add: in_affected power2_eq_square apply_inverse apply_times fixate_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   134
qed
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_fixate_no_binary_circle:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   137
  assumes "order f a > 2"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   138
  shows "affected (fixate a f) = affected f - {a}" (is "?A = ?B")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   139
proof (rule set_eqI)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   140
  interpret bijection "apply f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   141
    by standard simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   142
  fix b
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   143
  from assms order_greater_eq_two_iff [of f a]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   144
  have "a \<in> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   145
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   146
  moreover from assms
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   147
  have "apply f (apply f a) \<noteq> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   148
    using apply_power_eq_iff [of f 2 a 0]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   149
    by (simp add: power2_eq_square apply_times)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   150
  ultimately show "b \<in> ?A \<longleftrightarrow> b \<in> ?B"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   151
    by (cases "b \<in> {a, apply (inverse f) a}")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   152
      (auto simp add: in_affected apply_inverse apply_times fixate_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   153
qed
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
lemma affected_fixate:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   156
  "affected (fixate a f) \<subseteq> affected f - {a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   157
proof -
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   158
  have "a \<notin> affected f \<or> order f a = 2 \<or> order f a > 2"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   159
    by (auto simp add: not_less dest: affected_order_greater_eq_two)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   160
  then consider "a \<notin> affected f" | "order f a = 2" | "order f a > 2"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   161
    by blast
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   162
  then show ?thesis apply cases
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   163
  using affected_fixate_trivial [of a f]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   164
    affected_fixate_binary_circle [of f a]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   165
    affected_fixate_no_binary_circle [of f a]
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   166
    by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   167
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   168
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   169
lemma orbit_fixate_self [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   170
  "orbit (fixate a f) a = {a}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   171
proof -
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   172
  have "apply (f * inverse f) a = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   173
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   174
  then have "apply f (apply (inverse f) a) = a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   175
    by (simp only: apply_times comp_apply)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   176
  then show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   177
    by (simp add: fixate_def not_in_affected_iff_orbit_eq_singleton [symmetric] in_affected apply_times)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   178
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   179
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   180
lemma order_fixate_self [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   181
  "order (fixate a f) a = 1"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   182
proof -
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   183
  have "card (orbit (fixate a f) a) = card {a}"
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
  then show ?thesis
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   186
    by (simp only: card_orbit_eq) simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   187
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   188
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   189
lemma 
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   190
  assumes "b \<notin> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   191
  shows "orbit (fixate b f) a = orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   192
  oops
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   193
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   194
lemma
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   195
  assumes "b \<in> orbit f a" and "b \<noteq> a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   196
  shows "orbit (fixate b f) a = orbit f a - {b}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   197
  oops
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   198
    
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   199
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   200
text \<open>Distilling cycles from permutations\<close>
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   201
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   202
inductive_set orbits :: "'a perm \<Rightarrow> 'a set set" for f
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   203
where
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   204
  in_orbitsI: "a \<in> affected f \<Longrightarrow> orbit f a \<in> orbits f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   205
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   206
lemma orbits_unfold:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   207
  "orbits f = orbit f ` affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   208
  by (auto intro: in_orbitsI elim: orbits.cases)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   209
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   210
lemma in_orbit_affected:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   211
  assumes "b \<in> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   212
  assumes "a \<in> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   213
  shows "b \<in> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   214
proof (cases "a = b")
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   215
  case True with assms show ?thesis by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   216
next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   217
  case False with assms have "{a, b} \<subseteq> orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   218
    by auto
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   219
  also from assms have "orbit f a \<subseteq> affected f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   220
    by (blast intro!: orbit_subset_eq_affected)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   221
  finally show ?thesis by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   222
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   223
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   224
lemma Union_orbits [simp]:
69745
aec42cee2521 more canonical and less specialized syntax
nipkow
parents: 68084
diff changeset
   225
  "\<Union>(orbits f) = affected f"
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   226
  by (auto simp add: orbits.simps intro: in_orbitsI in_orbit_affected)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   227
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   228
lemma finite_orbits [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   229
  "finite (orbits f)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   230
  by (simp add: orbits_unfold)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   231
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   232
lemma card_in_orbits:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   233
  assumes "A \<in> orbits f"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   234
  shows "card A \<ge> 2"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   235
  using assms by cases
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   236
    (auto dest: affected_order_greater_eq_two)
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
lemma disjoint_orbits:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   239
  assumes "A \<in> orbits f" and "B \<in> orbits f" and "A \<noteq> B"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   240
  shows "A \<inter> B = {}"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   241
  using \<open>A \<in> orbits f\<close> apply cases
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   242
  using \<open>B \<in> orbits f\<close> apply cases
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   243
  using \<open>A \<noteq> B\<close> apply (simp_all add: orbit_disjoint)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   244
  done
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   245
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   246
definition trace :: "'a \<Rightarrow> 'a perm \<Rightarrow> 'a list"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   247
where
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   248
  "trace a f = map (\<lambda>n. apply (f ^ n) a) [0..<order f a]"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   249
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   250
lemma set_trace_eq [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   251
  "set (trace a f) = orbit f a"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   252
  by (auto simp add: trace_def orbit_unfold_image)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   253
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   254
definition seeds :: "'a perm \<Rightarrow> 'a::linorder list"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   255
where
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   256
  "seeds f = sorted_list_of_set (Min ` orbits f)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   257
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   258
definition cycles :: "'a perm \<Rightarrow> 'a::linorder list list"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   259
where
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   260
  "cycles f = map (\<lambda>a. trace a f) (seeds f)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   261
72536
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
   262
end
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
   263
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
   264
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
   265
text \<open>Misc\<close>
589645894305 bundle mixins for locale and class specifications
haftmann
parents: 69745
diff changeset
   266
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   267
lemma (in comm_monoid_list_set) sorted_list_of_set:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   268
  assumes "finite A"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   269
  shows "list.F (map h (sorted_list_of_set A)) = set.F h A"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   270
proof -
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   271
  from distinct_sorted_list_of_set
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   272
  have "set.F h (set (sorted_list_of_set A)) = list.F (map h (sorted_list_of_set A))"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   273
    by (rule distinct_set_conv_list)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   274
  with \<open>finite A\<close> show ?thesis
68084
152cc388cd1e updated to lemma name change
nipkow
parents: 67226
diff changeset
   275
    by (simp)
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   276
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   277
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   278
primrec subtract :: "'a list \<Rightarrow> 'a list \<Rightarrow> 'a list"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   279
where
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   280
  "subtract [] ys = ys"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   281
| "subtract (x # xs) ys = subtract xs (removeAll x ys)"
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 length_subtract_less_eq [simp]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   284
  "length (subtract xs ys) \<le> length ys"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   285
proof (induct xs arbitrary: ys)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   286
  case Nil then show ?case by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   287
next
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   288
  case (Cons x xs)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   289
  then have "length (subtract xs (removeAll x ys)) \<le> length (removeAll x ys)" .
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   290
  also have "length (removeAll x ys) \<le> length ys"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   291
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   292
  finally show ?case
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   293
    by simp
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   294
qed
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   295
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents:
diff changeset
   296
end