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