src/HOL/Library/Permutations.thy
author nipkow
Sun, 18 Oct 2009 12:07:56 +0200
changeset 32989 c28279b29ff1
parent 32962 69916a850301
parent 32988 d1d4d7a08a66
child 33057 764547b68538
permissions -rw-r--r--
merged
chaieb@29777
     1
(* Title:      Library/Permutations
chaieb@29777
     2
   Author:     Amine Chaieb, University of Cambridge
chaieb@29777
     3
*)
chaieb@29777
     4
chaieb@29777
     5
header {* Permutations, both general and specifically on finite sets.*}
chaieb@29777
     6
chaieb@29777
     7
theory Permutations
nipkow@32988
     8
imports Finite_Cartesian_Product Parity Fact
chaieb@29777
     9
begin
chaieb@29777
    10
chaieb@29777
    11
definition permutes (infixr "permutes" 41) where
chaieb@29777
    12
  "(p permutes S) \<longleftrightarrow> (\<forall>x. x \<notin> S \<longrightarrow> p x = x) \<and> (\<forall>y. \<exists>!x. p x = y)"
chaieb@29777
    13
chaieb@29777
    14
(* ------------------------------------------------------------------------- *)
chaieb@29777
    15
(* Transpositions.                                                           *)
chaieb@29777
    16
(* ------------------------------------------------------------------------- *)
chaieb@29777
    17
chaieb@29777
    18
declare swap_self[simp]
huffman@30474
    19
lemma swapid_sym: "Fun.swap a b id = Fun.swap b a id"
chaieb@29777
    20
  by (auto simp add: expand_fun_eq swap_def fun_upd_def)
chaieb@29777
    21
lemma swap_id_refl: "Fun.swap a a id = id" by simp
chaieb@29777
    22
lemma swap_id_sym: "Fun.swap a b id = Fun.swap b a id"
chaieb@29777
    23
  by (rule ext, simp add: swap_def)
chaieb@29777
    24
lemma swap_id_idempotent[simp]: "Fun.swap a b id o Fun.swap a b id = id"
chaieb@29777
    25
  by (rule ext, auto simp add: swap_def)
chaieb@29777
    26
chaieb@29777
    27
lemma inv_unique_comp: assumes fg: "f o g = id" and gf: "g o f = id"
chaieb@29777
    28
  shows "inv f = g"
chaieb@29777
    29
  using fg gf inv_equality[of g f] by (auto simp add: expand_fun_eq)
chaieb@29777
    30
chaieb@29777
    31
lemma inverse_swap_id: "inv (Fun.swap a b id) = Fun.swap a b id"
chaieb@29777
    32
  by (rule inv_unique_comp, simp_all)
chaieb@29777
    33
chaieb@29777
    34
lemma swap_id_eq: "Fun.swap a b id x = (if x = a then b else if x = b then a else x)"
chaieb@29777
    35
  by (simp add: swap_def)
chaieb@29777
    36
chaieb@29777
    37
(* ------------------------------------------------------------------------- *)
chaieb@29777
    38
(* Basic consequences of the definition.                                     *)
chaieb@29777
    39
(* ------------------------------------------------------------------------- *)
chaieb@29777
    40
chaieb@29777
    41
lemma permutes_in_image: "p permutes S \<Longrightarrow> p x \<in> S \<longleftrightarrow> x \<in> S"
chaieb@29777
    42
  unfolding permutes_def by metis
chaieb@29777
    43
chaieb@29777
    44
lemma permutes_image: assumes pS: "p permutes S" shows "p ` S = S"
chaieb@29777
    45
  using pS
huffman@30474
    46
  unfolding permutes_def
huffman@30474
    47
  apply -
huffman@30474
    48
  apply (rule set_ext)
chaieb@29777
    49
  apply (simp add: image_iff)
chaieb@29777
    50
  apply metis
chaieb@29777
    51
  done
chaieb@29777
    52
huffman@30474
    53
lemma permutes_inj: "p permutes S ==> inj p "
huffman@30474
    54
  unfolding permutes_def inj_on_def by blast
chaieb@29777
    55
huffman@30474
    56
lemma permutes_surj: "p permutes s ==> surj p"
huffman@30474
    57
  unfolding permutes_def surj_def by metis
chaieb@29777
    58
chaieb@29777
    59
lemma permutes_inv_o: assumes pS: "p permutes S"
chaieb@29777
    60
  shows " p o inv p = id"
chaieb@29777
    61
  and "inv p o p = id"
chaieb@29777
    62
  using permutes_inj[OF pS] permutes_surj[OF pS]
chaieb@29777
    63
  unfolding inj_iff[symmetric] surj_iff[symmetric] by blast+
chaieb@29777
    64
chaieb@29777
    65
huffman@30474
    66
lemma permutes_inverses:
chaieb@29777
    67
  fixes p :: "'a \<Rightarrow> 'a"
chaieb@29777
    68
  assumes pS: "p permutes S"
chaieb@29777
    69
  shows "p (inv p x) = x"
chaieb@29777
    70
  and "inv p (p x) = x"
chaieb@29777
    71
  using permutes_inv_o[OF pS, unfolded expand_fun_eq o_def] by auto
chaieb@29777
    72
chaieb@29777
    73
lemma permutes_subset: "p permutes S \<Longrightarrow> S \<subseteq> T ==> p permutes T"
chaieb@29777
    74
  unfolding permutes_def by blast
chaieb@29777
    75
chaieb@29777
    76
lemma permutes_empty[simp]: "p permutes {} \<longleftrightarrow> p = id"
huffman@30474
    77
  unfolding expand_fun_eq permutes_def apply simp by metis
chaieb@29777
    78
chaieb@29777
    79
lemma permutes_sing[simp]: "p permutes {a} \<longleftrightarrow> p = id"
chaieb@29777
    80
  unfolding expand_fun_eq permutes_def apply simp by metis
huffman@30474
    81
chaieb@29777
    82
lemma permutes_univ: "p permutes UNIV \<longleftrightarrow> (\<forall>y. \<exists>!x. p x = y)"
chaieb@29777
    83
  unfolding permutes_def by simp
chaieb@29777
    84
chaieb@29777
    85
lemma permutes_inv_eq: "p permutes S ==> inv p y = x \<longleftrightarrow> p x = y"
nipkow@32988
    86
  unfolding permutes_def inv_onto_def apply auto
chaieb@29777
    87
  apply (erule allE[where x=y])
chaieb@29777
    88
  apply (erule allE[where x=y])
chaieb@29777
    89
  apply (rule someI_ex) apply blast
chaieb@29777
    90
  apply (rule some1_equality)
chaieb@29777
    91
  apply blast
chaieb@29777
    92
  apply blast
chaieb@29777
    93
  done
chaieb@29777
    94
chaieb@29777
    95
lemma permutes_swap_id: "a \<in> S \<Longrightarrow> b \<in> S ==> Fun.swap a b id permutes S"
nipkow@32988
    96
  unfolding permutes_def swap_def fun_upd_def  by auto metis
chaieb@29777
    97
nipkow@32988
    98
lemma permutes_superset:
nipkow@32988
    99
  "p permutes S \<Longrightarrow> (\<forall>x \<in> S - T. p x = x) \<Longrightarrow> p permutes T"
nipkow@32988
   100
by (simp add: Ball_def permutes_def Diff_iff) metis
chaieb@29777
   101
chaieb@29777
   102
(* ------------------------------------------------------------------------- *)
chaieb@29777
   103
(* Group properties.                                                         *)
chaieb@29777
   104
(* ------------------------------------------------------------------------- *)
chaieb@29777
   105
huffman@30474
   106
lemma permutes_id: "id permutes S" unfolding permutes_def by simp
chaieb@29777
   107
chaieb@29777
   108
lemma permutes_compose: "p permutes S \<Longrightarrow> q permutes S ==> q o p permutes S"
chaieb@29777
   109
  unfolding permutes_def o_def by metis
chaieb@29777
   110
chaieb@29777
   111
lemma permutes_inv: assumes pS: "p permutes S" shows "inv p permutes S"
huffman@30474
   112
  using pS unfolding permutes_def permutes_inv_eq[OF pS] by metis
chaieb@29777
   113
chaieb@29777
   114
lemma permutes_inv_inv: assumes pS: "p permutes S" shows "inv (inv p) = p"
chaieb@29777
   115
  unfolding expand_fun_eq permutes_inv_eq[OF pS] permutes_inv_eq[OF permutes_inv[OF pS]]
chaieb@29777
   116
  by blast
chaieb@29777
   117
chaieb@29777
   118
(* ------------------------------------------------------------------------- *)
chaieb@29777
   119
(* The number of permutations on a finite set.                               *)
chaieb@29777
   120
(* ------------------------------------------------------------------------- *)
chaieb@29777
   121
huffman@30474
   122
lemma permutes_insert_lemma:
chaieb@29777
   123
  assumes pS: "p permutes (insert a S)"
chaieb@29777
   124
  shows "Fun.swap a (p a) id o p permutes S"
chaieb@29777
   125
  apply (rule permutes_superset[where S = "insert a S"])
chaieb@29777
   126
  apply (rule permutes_compose[OF pS])
chaieb@29777
   127
  apply (rule permutes_swap_id, simp)
chaieb@29777
   128
  using permutes_in_image[OF pS, of a] apply simp
chaieb@29777
   129
  apply (auto simp add: Ball_def Diff_iff swap_def)
chaieb@29777
   130
  done
chaieb@29777
   131
chaieb@29777
   132
lemma permutes_insert: "{p. p permutes (insert a S)} =
chaieb@29777
   133
        (\<lambda>(b,p). Fun.swap a b id o p) ` {(b,p). b \<in> insert a S \<and> p \<in> {p. p permutes S}}"
chaieb@29777
   134
proof-
chaieb@29777
   135
huffman@30474
   136
  {fix p
chaieb@29777
   137
    {assume pS: "p permutes insert a S"
chaieb@29777
   138
      let ?b = "p a"
chaieb@29777
   139
      let ?q = "Fun.swap a (p a) id o p"
huffman@30474
   140
      have th0: "p = Fun.swap a ?b id o ?q" unfolding expand_fun_eq o_assoc by simp
huffman@30474
   141
      have th1: "?b \<in> insert a S " unfolding permutes_in_image[OF pS] by simp
chaieb@29777
   142
      from permutes_insert_lemma[OF pS] th0 th1
chaieb@29777
   143
      have "\<exists> b q. p = Fun.swap a b id o q \<and> b \<in> insert a S \<and> q permutes S" by blast}
chaieb@29777
   144
    moreover
chaieb@29777
   145
    {fix b q assume bq: "p = Fun.swap a b id o q" "b \<in> insert a S" "q permutes S"
huffman@30474
   146
      from permutes_subset[OF bq(3), of "insert a S"]
chaieb@29777
   147
      have qS: "q permutes insert a S" by auto
chaieb@29777
   148
      have aS: "a \<in> insert a S" by simp
chaieb@29777
   149
      from bq(1) permutes_compose[OF qS permutes_swap_id[OF aS bq(2)]]
chaieb@29777
   150
      have "p permutes insert a S"  by simp }
chaieb@29777
   151
    ultimately have "p permutes insert a S \<longleftrightarrow> (\<exists> b q. p = Fun.swap a b id o q \<and> b \<in> insert a S \<and> q permutes S)" by blast}
chaieb@29777
   152
  thus ?thesis by auto
chaieb@29777
   153
qed
chaieb@29777
   154
chaieb@29777
   155
lemma hassize_insert: "a \<notin> F \<Longrightarrow> insert a F hassize n \<Longrightarrow> F hassize (n - 1)"
chaieb@29777
   156
  by (auto simp add: hassize_def)
chaieb@29777
   157
chaieb@29777
   158
lemma hassize_permutations: assumes Sn: "S hassize n"
chaieb@29777
   159
  shows "{p. p permutes S} hassize (fact n)"
chaieb@29777
   160
proof-
chaieb@29777
   161
  from Sn have fS:"finite S" by (simp add: hassize_def)
chaieb@29777
   162
chaieb@29777
   163
  have "\<forall>n. (S hassize n) \<longrightarrow> ({p. p permutes S} hassize (fact n))"
chaieb@29777
   164
  proof(rule finite_induct[where F = S])
chaieb@29777
   165
    from fS show "finite S" .
chaieb@29777
   166
  next
chaieb@29777
   167
    show "\<forall>n. ({} hassize n) \<longrightarrow> ({p. p permutes {}} hassize fact n)"
chaieb@29777
   168
      by (simp add: hassize_def permutes_empty)
chaieb@29777
   169
  next
huffman@30474
   170
    fix x F
huffman@30474
   171
    assume fF: "finite F" and xF: "x \<notin> F"
chaieb@29777
   172
      and H: "\<forall>n. (F hassize n) \<longrightarrow> ({p. p permutes F} hassize fact n)"
chaieb@29777
   173
    {fix n assume H0: "insert x F hassize n"
chaieb@29777
   174
      let ?xF = "{p. p permutes insert x F}"
chaieb@29777
   175
      let ?pF = "{p. p permutes F}"
chaieb@29777
   176
      let ?pF' = "{(b, p). b \<in> insert x F \<and> p \<in> ?pF}"
chaieb@29777
   177
      let ?g = "(\<lambda>(b, p). Fun.swap x b id \<circ> p)"
chaieb@29777
   178
      from permutes_insert[of x F]
chaieb@29777
   179
      have xfgpF': "?xF = ?g ` ?pF'" .
chaieb@29777
   180
      from hassize_insert[OF xF H0] have Fs: "F hassize (n - 1)" .
chaieb@29777
   181
      from H Fs have pFs: "?pF hassize fact (n - 1)" by blast
huffman@30474
   182
      hence pF'f: "finite ?pF'" using H0 unfolding hassize_def
wenzelm@32962
   183
        apply (simp only: Collect_split Collect_mem_eq)
wenzelm@32962
   184
        apply (rule finite_cartesian_product)
wenzelm@32962
   185
        apply simp_all
wenzelm@32962
   186
        done
chaieb@29777
   187
chaieb@29777
   188
      have ginj: "inj_on ?g ?pF'"
chaieb@29777
   189
      proof-
wenzelm@32962
   190
        {
wenzelm@32962
   191
          fix b p c q assume bp: "(b,p) \<in> ?pF'" and cq: "(c,q) \<in> ?pF'"
wenzelm@32962
   192
            and eq: "?g (b,p) = ?g (c,q)"
wenzelm@32962
   193
          from bp cq have ths: "b \<in> insert x F" "c \<in> insert x F" "x \<in> insert x F" "p permutes F" "q permutes F" by auto
wenzelm@32962
   194
          from ths(4) xF eq have "b = ?g (b,p) x" unfolding permutes_def
wenzelm@32962
   195
            by (auto simp add: swap_def fun_upd_def expand_fun_eq)
wenzelm@32962
   196
          also have "\<dots> = ?g (c,q) x" using ths(5) xF eq
wenzelm@32962
   197
            by (auto simp add: swap_def fun_upd_def expand_fun_eq)
wenzelm@32962
   198
          also have "\<dots> = c"using ths(5) xF unfolding permutes_def
wenzelm@32962
   199
            by (auto simp add: swap_def fun_upd_def expand_fun_eq)
wenzelm@32962
   200
          finally have bc: "b = c" .
wenzelm@32962
   201
          hence "Fun.swap x b id = Fun.swap x c id" by simp
wenzelm@32962
   202
          with eq have "Fun.swap x b id o p = Fun.swap x b id o q" by simp
wenzelm@32962
   203
          hence "Fun.swap x b id o (Fun.swap x b id o p) = Fun.swap x b id o (Fun.swap x b id o q)" by simp
wenzelm@32962
   204
          hence "p = q" by (simp add: o_assoc)
wenzelm@32962
   205
          with bc have "(b,p) = (c,q)" by simp }
wenzelm@32962
   206
        thus ?thesis  unfolding inj_on_def by blast
chaieb@29777
   207
      qed
chaieb@29777
   208
      from xF H0 have n0: "n \<noteq> 0 " by (auto simp add: hassize_def)
chaieb@29777
   209
      hence "\<exists>m. n = Suc m" by arith
huffman@30474
   210
      then obtain m where n[simp]: "n = Suc m" by blast
huffman@30474
   211
      from pFs H0 have xFc: "card ?xF = fact n"
wenzelm@32962
   212
        unfolding xfgpF' card_image[OF ginj] hassize_def
wenzelm@32962
   213
        apply (simp only: Collect_split Collect_mem_eq card_cartesian_product)
wenzelm@32962
   214
        by simp
huffman@30474
   215
      from finite_imageI[OF pF'f, of ?g] have xFf: "finite ?xF" unfolding xfgpF' by simp
chaieb@29777
   216
      have "?xF hassize fact n"
wenzelm@32962
   217
        using xFf xFc
wenzelm@32962
   218
        unfolding hassize_def  xFf by blast }
huffman@30474
   219
    thus "\<forall>n. (insert x F hassize n) \<longrightarrow> ({p. p permutes insert x F} hassize fact n)"
chaieb@29777
   220
      by blast
chaieb@29777
   221
  qed
chaieb@29777
   222
  with Sn show ?thesis by blast
chaieb@29777
   223
qed
chaieb@29777
   224
chaieb@29777
   225
lemma finite_permutations: "finite S ==> finite {p. p permutes S}"
chaieb@29777
   226
  using hassize_permutations[of S] unfolding hassize_def by blast
chaieb@29777
   227
chaieb@29777
   228
(* ------------------------------------------------------------------------- *)
chaieb@29777
   229
(* Permutations of index set for iterated operations.                        *)
chaieb@29777
   230
(* ------------------------------------------------------------------------- *)
chaieb@29777
   231
huffman@30474
   232
lemma (in ab_semigroup_mult) fold_image_permute: assumes fS: "finite S" and pS: "p permutes S"
chaieb@29777
   233
  shows "fold_image times f z S = fold_image times (f o p) z S"
chaieb@29777
   234
  using fold_image_reindex[OF fS subset_inj_on[OF permutes_inj[OF pS], of S, simplified], of f z]
chaieb@29777
   235
  unfolding permutes_image[OF pS] .
huffman@30474
   236
lemma (in ab_semigroup_add) fold_image_permute: assumes fS: "finite S" and pS: "p permutes S"
chaieb@29777
   237
  shows "fold_image plus f z S = fold_image plus (f o p) z S"
chaieb@29777
   238
proof-
chaieb@29777
   239
  interpret ab_semigroup_mult plus apply unfold_locales apply (simp add: add_assoc)
chaieb@29777
   240
    apply (simp add: add_commute) done
chaieb@29777
   241
  from fold_image_reindex[OF fS subset_inj_on[OF permutes_inj[OF pS], of S, simplified], of f z]
chaieb@29777
   242
  show ?thesis
chaieb@29777
   243
  unfolding permutes_image[OF pS] .
chaieb@29777
   244
qed
chaieb@29777
   245
huffman@30474
   246
lemma setsum_permute: assumes pS: "p permutes S"
chaieb@29777
   247
  shows "setsum f S = setsum (f o p) S"
chaieb@29777
   248
  unfolding setsum_def using fold_image_permute[of S p f 0] pS by clarsimp
chaieb@29777
   249
huffman@30474
   250
lemma setsum_permute_natseg:assumes pS: "p permutes {m .. n}"
chaieb@29777
   251
  shows "setsum f {m .. n} = setsum (f o p) {m .. n}"
huffman@30474
   252
  using setsum_permute[OF pS, of f ] pS by blast
chaieb@29777
   253
huffman@30474
   254
lemma setprod_permute: assumes pS: "p permutes S"
chaieb@29777
   255
  shows "setprod f S = setprod (f o p) S"
huffman@30474
   256
  unfolding setprod_def
chaieb@29777
   257
  using ab_semigroup_mult_class.fold_image_permute[of S p f 1] pS by clarsimp
chaieb@29777
   258
huffman@30474
   259
lemma setprod_permute_natseg:assumes pS: "p permutes {m .. n}"
chaieb@29777
   260
  shows "setprod f {m .. n} = setprod (f o p) {m .. n}"
huffman@30474
   261
  using setprod_permute[OF pS, of f ] pS by blast
chaieb@29777
   262
chaieb@29777
   263
(* ------------------------------------------------------------------------- *)
chaieb@29777
   264
(* Various combinations of transpositions with 2, 1 and 0 common elements.   *)
chaieb@29777
   265
(* ------------------------------------------------------------------------- *)
chaieb@29777
   266
chaieb@29777
   267
lemma swap_id_common:" a \<noteq> c \<Longrightarrow> b \<noteq> c \<Longrightarrow>  Fun.swap a b id o Fun.swap a c id = Fun.swap b c id o Fun.swap a b id" by (simp add: expand_fun_eq swap_def)
chaieb@29777
   268
chaieb@29777
   269
lemma swap_id_common': "~(a = b) \<Longrightarrow> ~(a = c) \<Longrightarrow> Fun.swap a c id o Fun.swap b c id = Fun.swap b c id o Fun.swap a b id" by (simp add: expand_fun_eq swap_def)
chaieb@29777
   270
chaieb@29777
   271
lemma swap_id_independent: "~(a = c) \<Longrightarrow> ~(a = d) \<Longrightarrow> ~(b = c) \<Longrightarrow> ~(b = d) ==> Fun.swap a b id o Fun.swap c d id = Fun.swap c d id o Fun.swap a b id"
chaieb@29777
   272
  by (simp add: swap_def expand_fun_eq)
chaieb@29777
   273
chaieb@29777
   274
(* ------------------------------------------------------------------------- *)
chaieb@29777
   275
(* Permutations as transposition sequences.                                  *)
chaieb@29777
   276
(* ------------------------------------------------------------------------- *)
chaieb@29777
   277
chaieb@29777
   278
chaieb@29777
   279
inductive swapidseq :: "nat \<Rightarrow> ('a \<Rightarrow> 'a) \<Rightarrow> bool" where
chaieb@29777
   280
  id[simp]: "swapidseq 0 id"
chaieb@29777
   281
| comp_Suc: "swapidseq n p \<Longrightarrow> a \<noteq> b \<Longrightarrow> swapidseq (Suc n) (Fun.swap a b id o p)"
chaieb@29777
   282
chaieb@29777
   283
declare id[unfolded id_def, simp]
chaieb@29777
   284
definition "permutation p \<longleftrightarrow> (\<exists>n. swapidseq n p)"
chaieb@29777
   285
chaieb@29777
   286
(* ------------------------------------------------------------------------- *)
chaieb@29777
   287
(* Some closure properties of the set of permutations, with lengths.         *)
chaieb@29777
   288
(* ------------------------------------------------------------------------- *)
chaieb@29777
   289
chaieb@29777
   290
lemma permutation_id[simp]: "permutation id"unfolding permutation_def
chaieb@29777
   291
  by (rule exI[where x=0], simp)
chaieb@29777
   292
declare permutation_id[unfolded id_def, simp]
chaieb@29777
   293
chaieb@29777
   294
lemma swapidseq_swap: "swapidseq (if a = b then 0 else 1) (Fun.swap a b id)"
chaieb@29777
   295
  apply clarsimp
chaieb@29777
   296
  using comp_Suc[of 0 id a b] by simp
chaieb@29777
   297
chaieb@29777
   298
lemma permutation_swap_id: "permutation (Fun.swap a b id)"
chaieb@29777
   299
  apply (cases "a=b", simp_all)
huffman@30474
   300
  unfolding permutation_def using swapidseq_swap[of a b] by blast
chaieb@29777
   301
chaieb@29777
   302
lemma swapidseq_comp_add: "swapidseq n p \<Longrightarrow> swapidseq m q ==> swapidseq (n + m) (p o q)"
chaieb@29777
   303
  proof (induct n p arbitrary: m q rule: swapidseq.induct)
chaieb@29777
   304
    case (id m q) thus ?case by simp
chaieb@29777
   305
  next
huffman@30474
   306
    case (comp_Suc n p a b m q)
chaieb@29777
   307
    have th: "Suc n + m = Suc (n + m)" by arith
huffman@30474
   308
    show ?case unfolding th o_assoc[symmetric]
huffman@30474
   309
      apply (rule swapidseq.comp_Suc) using comp_Suc.hyps(2)[OF comp_Suc.prems]  comp_Suc.hyps(3) by blast+
chaieb@29777
   310
qed
chaieb@29777
   311
chaieb@29777
   312
lemma permutation_compose: "permutation p \<Longrightarrow> permutation q ==> permutation(p o q)"
chaieb@29777
   313
  unfolding permutation_def using swapidseq_comp_add[of _ p _ q] by metis
chaieb@29777
   314
chaieb@29777
   315
lemma swapidseq_endswap: "swapidseq n p \<Longrightarrow> a \<noteq> b ==> swapidseq (Suc n) (p o Fun.swap a b id)"
chaieb@29777
   316
  apply (induct n p rule: swapidseq.induct)
chaieb@29777
   317
  using swapidseq_swap[of a b]
chaieb@29777
   318
  by (auto simp add: o_assoc[symmetric] intro: swapidseq.comp_Suc)
chaieb@29777
   319
chaieb@29777
   320
lemma swapidseq_inverse_exists: "swapidseq n p ==> \<exists>q. swapidseq n q \<and> p o q = id \<and> q o p = id"
chaieb@29777
   321
proof(induct n p rule: swapidseq.induct)
chaieb@29777
   322
  case id  thus ?case by (rule exI[where x=id], simp)
huffman@30474
   323
next
chaieb@29777
   324
  case (comp_Suc n p a b)
chaieb@29777
   325
  from comp_Suc.hyps obtain q where q: "swapidseq n q" "p \<circ> q = id" "q \<circ> p = id" by blast
chaieb@29777
   326
  let ?q = "q o Fun.swap a b id"
chaieb@29777
   327
  note H = comp_Suc.hyps
chaieb@29777
   328
  from swapidseq_swap[of a b] H(3)  have th0: "swapidseq 1 (Fun.swap a b id)" by simp
huffman@30474
   329
  from swapidseq_comp_add[OF q(1) th0] have th1:"swapidseq (Suc n) ?q" by simp
chaieb@29777
   330
  have "Fun.swap a b id o p o ?q = Fun.swap a b id o (p o q) o Fun.swap a b id" by (simp add: o_assoc)
chaieb@29777
   331
  also have "\<dots> = id" by (simp add: q(2))
chaieb@29777
   332
  finally have th2: "Fun.swap a b id o p o ?q = id" .
huffman@30474
   333
  have "?q \<circ> (Fun.swap a b id \<circ> p) = q \<circ> (Fun.swap a b id o Fun.swap a b id) \<circ> p" by (simp only: o_assoc)
chaieb@29777
   334
  hence "?q \<circ> (Fun.swap a b id \<circ> p) = id" by (simp add: q(3))
chaieb@29777
   335
  with th1 th2 show ?case by blast
chaieb@29777
   336
qed
chaieb@29777
   337
chaieb@29777
   338
chaieb@29777
   339
lemma swapidseq_inverse: assumes H: "swapidseq n p" shows "swapidseq n (inv p)"
chaieb@29777
   340
  using swapidseq_inverse_exists[OF H] inv_unique_comp[of p] by auto
chaieb@29777
   341
chaieb@29777
   342
lemma permutation_inverse: "permutation p ==> permutation (inv p)"
chaieb@29777
   343
  using permutation_def swapidseq_inverse by blast
chaieb@29777
   344
chaieb@29777
   345
(* ------------------------------------------------------------------------- *)
chaieb@29777
   346
(* The identity map only has even transposition sequences.                   *)
chaieb@29777
   347
(* ------------------------------------------------------------------------- *)
chaieb@29777
   348
chaieb@29777
   349
lemma symmetry_lemma:"(\<And>a b c d. P a b c d ==> P a b d c) \<Longrightarrow>
chaieb@29777
   350
   (\<And>a b c d. a \<noteq> b \<Longrightarrow> c \<noteq> d \<Longrightarrow> (a = c \<and> b = d \<or>  a = c \<and> b \<noteq> d \<or> a \<noteq> c \<and> b = d \<or> a \<noteq> c \<and> a \<noteq> d \<and> b \<noteq> c \<and> b \<noteq> d) ==> P a b c d)
chaieb@29777
   351
   ==> (\<And>a b c d. a \<noteq> b --> c \<noteq> d \<longrightarrow>  P a b c d)" by metis
chaieb@29777
   352
huffman@30474
   353
lemma swap_general: "a \<noteq> b \<Longrightarrow> c \<noteq> d \<Longrightarrow> Fun.swap a b id o Fun.swap c d id = id \<or>
huffman@30474
   354
  (\<exists>x y z. x \<noteq> a \<and> y \<noteq> a \<and> z \<noteq> a \<and> x \<noteq> y \<and> Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id)"
chaieb@29777
   355
proof-
chaieb@29777
   356
  assume H: "a\<noteq>b" "c\<noteq>d"
huffman@30474
   357
have "a \<noteq> b \<longrightarrow> c \<noteq> d \<longrightarrow>
huffman@30474
   358
(  Fun.swap a b id o Fun.swap c d id = id \<or>
huffman@30474
   359
  (\<exists>x y z. x \<noteq> a \<and> y \<noteq> a \<and> z \<noteq> a \<and> x \<noteq> y \<and> Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id))"
chaieb@29777
   360
  apply (rule symmetry_lemma[where a=a and b=b and c=c and d=d])
huffman@30474
   361
  apply (simp_all only: swapid_sym)
chaieb@29777
   362
  apply (case_tac "a = c \<and> b = d", clarsimp simp only: swapid_sym swap_id_idempotent)
chaieb@29777
   363
  apply (case_tac "a = c \<and> b \<noteq> d")
chaieb@29777
   364
  apply (rule disjI2)
chaieb@29777
   365
  apply (rule_tac x="b" in exI)
chaieb@29777
   366
  apply (rule_tac x="d" in exI)
chaieb@29777
   367
  apply (rule_tac x="b" in exI)
chaieb@29777
   368
  apply (clarsimp simp add: expand_fun_eq swap_def)
chaieb@29777
   369
  apply (case_tac "a \<noteq> c \<and> b = d")
chaieb@29777
   370
  apply (rule disjI2)
chaieb@29777
   371
  apply (rule_tac x="c" in exI)
chaieb@29777
   372
  apply (rule_tac x="d" in exI)
chaieb@29777
   373
  apply (rule_tac x="c" in exI)
chaieb@29777
   374
  apply (clarsimp simp add: expand_fun_eq swap_def)
chaieb@29777
   375
  apply (rule disjI2)
chaieb@29777
   376
  apply (rule_tac x="c" in exI)
chaieb@29777
   377
  apply (rule_tac x="d" in exI)
chaieb@29777
   378
  apply (rule_tac x="b" in exI)
chaieb@29777
   379
  apply (clarsimp simp add: expand_fun_eq swap_def)
chaieb@29777
   380
  done
huffman@30474
   381
with H show ?thesis by metis
chaieb@29777
   382
qed
chaieb@29777
   383
chaieb@29777
   384
lemma swapidseq_id_iff[simp]: "swapidseq 0 p \<longleftrightarrow> p = id"
chaieb@29777
   385
  using swapidseq.cases[of 0 p "p = id"]
chaieb@29777
   386
  by auto
chaieb@29777
   387
chaieb@29777
   388
lemma swapidseq_cases: "swapidseq n p \<longleftrightarrow> (n=0 \<and> p = id \<or> (\<exists>a b q m. n = Suc m \<and> p = Fun.swap a b id o q \<and> swapidseq m q \<and> a\<noteq> b))"
chaieb@29777
   389
  apply (rule iffI)
chaieb@29777
   390
  apply (erule swapidseq.cases[of n p])
chaieb@29777
   391
  apply simp
chaieb@29777
   392
  apply (rule disjI2)
chaieb@29777
   393
  apply (rule_tac x= "a" in exI)
chaieb@29777
   394
  apply (rule_tac x= "b" in exI)
chaieb@29777
   395
  apply (rule_tac x= "pa" in exI)
chaieb@29777
   396
  apply (rule_tac x= "na" in exI)
chaieb@29777
   397
  apply simp
chaieb@29777
   398
  apply auto
chaieb@29777
   399
  apply (rule comp_Suc, simp_all)
chaieb@29777
   400
  done
chaieb@29777
   401
lemma fixing_swapidseq_decrease:
chaieb@29777
   402
  assumes spn: "swapidseq n p" and ab: "a\<noteq>b" and pa: "(Fun.swap a b id o p) a = a"
chaieb@29777
   403
  shows "n \<noteq> 0 \<and> swapidseq (n - 1) (Fun.swap a b id o p)"
chaieb@29777
   404
  using spn ab pa
chaieb@29777
   405
proof(induct n arbitrary: p a b)
chaieb@29777
   406
  case 0 thus ?case by (auto simp add: swap_def fun_upd_def)
chaieb@29777
   407
next
chaieb@29777
   408
  case (Suc n p a b)
chaieb@29777
   409
  from Suc.prems(1) swapidseq_cases[of "Suc n" p] obtain
chaieb@29777
   410
    c d q m where cdqm: "Suc n = Suc m" "p = Fun.swap c d id o q" "swapidseq m q" "c \<noteq> d" "n = m"
chaieb@29777
   411
    by auto
chaieb@29777
   412
  {assume H: "Fun.swap a b id o Fun.swap c d id = id"
huffman@30474
   413
huffman@30474
   414
    have ?case apply (simp only: cdqm o_assoc H)
chaieb@29777
   415
      by (simp add: cdqm)}
chaieb@29777
   416
  moreover
chaieb@29777
   417
  { fix x y z
huffman@30474
   418
    assume H: "x\<noteq>a" "y\<noteq>a" "z \<noteq>a" "x \<noteq>y"
chaieb@29777
   419
      "Fun.swap a b id o Fun.swap c d id = Fun.swap x y id o Fun.swap a z id"
chaieb@29777
   420
    from H have az: "a \<noteq> z" by simp
chaieb@29777
   421
chaieb@29777
   422
    {fix h have "(Fun.swap x y id o h) a = a \<longleftrightarrow> h a = a"
chaieb@29777
   423
      using H by (simp add: swap_def)}
chaieb@29777
   424
    note th3 = this
chaieb@29777
   425
    from cdqm(2) have "Fun.swap a b id o p = Fun.swap a b id o (Fun.swap c d id o q)" by simp
chaieb@29777
   426
    hence "Fun.swap a b id o p = Fun.swap x y id o (Fun.swap a z id o q)" by (simp add: o_assoc H)
chaieb@29777
   427
    hence "(Fun.swap a b id o p) a = (Fun.swap x y id o (Fun.swap a z id o q)) a" by simp
chaieb@29777
   428
    hence "(Fun.swap x y id o (Fun.swap a z id o q)) a  = a" unfolding Suc by metis
chaieb@29777
   429
    hence th1: "(Fun.swap a z id o q) a = a" unfolding th3 .
chaieb@29777
   430
    from Suc.hyps[OF cdqm(3)[ unfolded cdqm(5)[symmetric]] az th1]
chaieb@29777
   431
    have th2: "swapidseq (n - 1) (Fun.swap a z id o q)" "n \<noteq> 0" by blast+
huffman@30474
   432
    have th: "Suc n - 1 = Suc (n - 1)" using th2(2) by auto
chaieb@29777
   433
    have ?case unfolding cdqm(2) H o_assoc th
chaieb@29777
   434
      apply (simp only: Suc_not_Zero simp_thms o_assoc[symmetric])
chaieb@29777
   435
      apply (rule comp_Suc)
chaieb@29777
   436
      using th2 H apply blast+
chaieb@29777
   437
      done}
huffman@30474
   438
  ultimately show ?case using swap_general[OF Suc.prems(2) cdqm(4)] by metis
chaieb@29777
   439
qed
chaieb@29777
   440
huffman@30474
   441
lemma swapidseq_identity_even:
chaieb@29777
   442
  assumes "swapidseq n (id :: 'a \<Rightarrow> 'a)" shows "even n"
chaieb@29777
   443
  using `swapidseq n id`
chaieb@29777
   444
proof(induct n rule: nat_less_induct)
chaieb@29777
   445
  fix n
chaieb@29777
   446
  assume H: "\<forall>m<n. swapidseq m (id::'a \<Rightarrow> 'a) \<longrightarrow> even m" "swapidseq n (id :: 'a \<Rightarrow> 'a)"
huffman@30474
   447
  {assume "n = 0" hence "even n" by arith}
huffman@30474
   448
  moreover
chaieb@29777
   449
  {fix a b :: 'a and q m
chaieb@29777
   450
    assume h: "n = Suc m" "(id :: 'a \<Rightarrow> 'a) = Fun.swap a b id \<circ> q" "swapidseq m q" "a \<noteq> b"
chaieb@29777
   451
    from fixing_swapidseq_decrease[OF h(3,4), unfolded h(2)[symmetric]]
chaieb@29777
   452
    have m: "m \<noteq> 0" "swapidseq (m - 1) (id :: 'a \<Rightarrow> 'a)" by auto
chaieb@29777
   453
    from h m have mn: "m - 1 < n" by arith
chaieb@29777
   454
    from H(1)[rule_format, OF mn m(2)] h(1) m(1) have "even n" apply arith done}
chaieb@29777
   455
  ultimately show "even n" using H(2)[unfolded swapidseq_cases[of n id]] by auto
chaieb@29777
   456
qed
chaieb@29777
   457
chaieb@29777
   458
(* ------------------------------------------------------------------------- *)
chaieb@29777
   459
(* Therefore we have a welldefined notion of parity.                         *)
chaieb@29777
   460
(* ------------------------------------------------------------------------- *)
chaieb@29777
   461
chaieb@29777
   462
definition "evenperm p = even (SOME n. swapidseq n p)"
chaieb@29777
   463
huffman@30474
   464
lemma swapidseq_even_even: assumes
chaieb@29777
   465
  m: "swapidseq m p" and n: "swapidseq n p"
chaieb@29777
   466
  shows "even m \<longleftrightarrow> even n"
chaieb@29777
   467
proof-
chaieb@29777
   468
  from swapidseq_inverse_exists[OF n]
chaieb@29777
   469
  obtain q where q: "swapidseq n q" "p \<circ> q = id" "q \<circ> p = id" by blast
huffman@30474
   470
chaieb@29777
   471
  from swapidseq_identity_even[OF swapidseq_comp_add[OF m q(1), unfolded q]]
chaieb@29777
   472
  show ?thesis by arith
chaieb@29777
   473
qed
chaieb@29777
   474
chaieb@29777
   475
lemma evenperm_unique: assumes p: "swapidseq n p" and n:"even n = b"
chaieb@29777
   476
  shows "evenperm p = b"
chaieb@29777
   477
  unfolding n[symmetric] evenperm_def
chaieb@29777
   478
  apply (rule swapidseq_even_even[where p = p])
chaieb@29777
   479
  apply (rule someI[where x = n])
chaieb@29777
   480
  using p by blast+
chaieb@29777
   481
chaieb@29777
   482
(* ------------------------------------------------------------------------- *)
chaieb@29777
   483
(* And it has the expected composition properties.                           *)
chaieb@29777
   484
(* ------------------------------------------------------------------------- *)
chaieb@29777
   485
chaieb@29777
   486
lemma evenperm_id[simp]: "evenperm id = True"
chaieb@29777
   487
  apply (rule evenperm_unique[where n = 0]) by simp_all
chaieb@29777
   488
chaieb@29777
   489
lemma evenperm_swap: "evenperm (Fun.swap a b id) = (a = b)"
chaieb@29777
   490
apply (rule evenperm_unique[where n="if a = b then 0 else 1"])
chaieb@29777
   491
by (simp_all add: swapidseq_swap)
chaieb@29777
   492
huffman@30474
   493
lemma evenperm_comp:
chaieb@29777
   494
  assumes p: "permutation p" and q:"permutation q"
chaieb@29777
   495
  shows "evenperm (p o q) = (evenperm p = evenperm q)"
chaieb@29777
   496
proof-
huffman@30474
   497
  from p q obtain
huffman@30474
   498
    n m where n: "swapidseq n p" and m: "swapidseq m q"
chaieb@29777
   499
    unfolding permutation_def by blast
chaieb@29777
   500
  note nm =  swapidseq_comp_add[OF n m]
chaieb@29777
   501
  have th: "even (n + m) = (even n \<longleftrightarrow> even m)" by arith
chaieb@29777
   502
  from evenperm_unique[OF n refl] evenperm_unique[OF m refl]
chaieb@29777
   503
    evenperm_unique[OF nm th]
chaieb@29777
   504
  show ?thesis by blast
chaieb@29777
   505
qed
chaieb@29777
   506
chaieb@29777
   507
lemma evenperm_inv: assumes p: "permutation p"
chaieb@29777
   508
  shows "evenperm (inv p) = evenperm p"
chaieb@29777
   509
proof-
chaieb@29777
   510
  from p obtain n where n: "swapidseq n p" unfolding permutation_def by blast
chaieb@29777
   511
  from evenperm_unique[OF swapidseq_inverse[OF n] evenperm_unique[OF n refl, symmetric]]
chaieb@29777
   512
  show ?thesis .
chaieb@29777
   513
qed
chaieb@29777
   514
chaieb@29777
   515
(* ------------------------------------------------------------------------- *)
chaieb@29777
   516
(* A more abstract characterization of permutations.                         *)
chaieb@29777
   517
(* ------------------------------------------------------------------------- *)
chaieb@29777
   518
chaieb@29777
   519
chaieb@29777
   520
lemma bij_iff: "bij f \<longleftrightarrow> (\<forall>x. \<exists>!y. f y = x)"
chaieb@29777
   521
  unfolding bij_def inj_on_def surj_def
chaieb@29777
   522
  apply auto
chaieb@29777
   523
  apply metis
chaieb@29777
   524
  apply metis
chaieb@29777
   525
  done
chaieb@29777
   526
huffman@30474
   527
lemma permutation_bijective:
huffman@30474
   528
  assumes p: "permutation p"
chaieb@29777
   529
  shows "bij p"
chaieb@29777
   530
proof-
chaieb@29777
   531
  from p obtain n where n: "swapidseq n p" unfolding permutation_def by blast
huffman@30474
   532
  from swapidseq_inverse_exists[OF n] obtain q where
chaieb@29777
   533
    q: "swapidseq n q" "p \<circ> q = id" "q \<circ> p = id" by blast
chaieb@29777
   534
  thus ?thesis unfolding bij_iff  apply (auto simp add: expand_fun_eq) apply metis done
huffman@30474
   535
qed
chaieb@29777
   536
chaieb@29777
   537
lemma permutation_finite_support: assumes p: "permutation p"
chaieb@29777
   538
  shows "finite {x. p x \<noteq> x}"
chaieb@29777
   539
proof-
chaieb@29777
   540
  from p obtain n where n: "swapidseq n p" unfolding permutation_def by blast
chaieb@29777
   541
  from n show ?thesis
chaieb@29777
   542
  proof(induct n p rule: swapidseq.induct)
chaieb@29777
   543
    case id thus ?case by simp
chaieb@29777
   544
  next
chaieb@29777
   545
    case (comp_Suc n p a b)
chaieb@29777
   546
    let ?S = "insert a (insert b {x. p x \<noteq> x})"
chaieb@29777
   547
    from comp_Suc.hyps(2) have fS: "finite ?S" by simp
chaieb@29777
   548
    from `a \<noteq> b` have th: "{x. (Fun.swap a b id o p) x \<noteq> x} \<subseteq> ?S"
chaieb@29777
   549
      by (auto simp add: swap_def)
chaieb@29777
   550
    from finite_subset[OF th fS] show ?case  .
chaieb@29777
   551
qed
chaieb@29777
   552
qed
chaieb@29777
   553
chaieb@29777
   554
lemma bij_inv_eq_iff: "bij p ==> x = inv p y \<longleftrightarrow> p x = y"
chaieb@29777
   555
  using surj_f_inv_f[of p] inv_f_f[of f] by (auto simp add: bij_def)
chaieb@29777
   556
huffman@30474
   557
lemma bij_swap_comp:
chaieb@29777
   558
  assumes bp: "bij p" shows "Fun.swap a b id o p = Fun.swap (inv p a) (inv p b) p"
chaieb@29777
   559
  using surj_f_inv_f[OF bij_is_surj[OF bp]]
chaieb@29777
   560
  by (simp add: expand_fun_eq swap_def bij_inv_eq_iff[OF bp])
chaieb@29777
   561
chaieb@29777
   562
lemma bij_swap_ompose_bij: "bij p \<Longrightarrow> bij (Fun.swap a b id o p)"
chaieb@29777
   563
proof-
chaieb@29777
   564
  assume H: "bij p"
huffman@30474
   565
  show ?thesis
chaieb@29777
   566
    unfolding bij_swap_comp[OF H] bij_swap_iff
chaieb@29777
   567
    using H .
chaieb@29777
   568
qed
chaieb@29777
   569
huffman@30474
   570
lemma permutation_lemma:
chaieb@29777
   571
  assumes fS: "finite S" and p: "bij p" and pS: "\<forall>x. x\<notin> S \<longrightarrow> p x = x"
chaieb@29777
   572
  shows "permutation p"
chaieb@29777
   573
using fS p pS
chaieb@29777
   574
proof(induct S arbitrary: p rule: finite_induct)
chaieb@29777
   575
  case (empty p) thus ?case by simp
chaieb@29777
   576
next
chaieb@29777
   577
  case (insert a F p)
chaieb@29777
   578
  let ?r = "Fun.swap a (p a) id o p"
chaieb@29777
   579
  let ?q = "Fun.swap a (p a) id o ?r "
chaieb@29777
   580
  have raa: "?r a = a" by (simp add: swap_def)
chaieb@29777
   581
  from bij_swap_ompose_bij[OF insert(4)]
huffman@30474
   582
  have br: "bij ?r"  .
huffman@30474
   583
huffman@30474
   584
  from insert raa have th: "\<forall>x. x \<notin> F \<longrightarrow> ?r x = x"
chaieb@29777
   585
    apply (clarsimp simp add: swap_def)
chaieb@29777
   586
    apply (erule_tac x="x" in allE)
chaieb@29777
   587
    apply auto
chaieb@29777
   588
    unfolding bij_iff apply metis
chaieb@29777
   589
    done
chaieb@29777
   590
  from insert(3)[OF br th]
chaieb@29777
   591
  have rp: "permutation ?r" .
chaieb@29777
   592
  have "permutation ?q" by (simp add: permutation_compose permutation_swap_id rp)
chaieb@29777
   593
  thus ?case by (simp add: o_assoc)
chaieb@29777
   594
qed
chaieb@29777
   595
huffman@30474
   596
lemma permutation: "permutation p \<longleftrightarrow> bij p \<and> finite {x. p x \<noteq> x}"
chaieb@29777
   597
  (is "?lhs \<longleftrightarrow> ?b \<and> ?f")
chaieb@29777
   598
proof
chaieb@29777
   599
  assume p: ?lhs
chaieb@29777
   600
  from p permutation_bijective permutation_finite_support show "?b \<and> ?f" by auto
chaieb@29777
   601
next
chaieb@29777
   602
  assume bf: "?b \<and> ?f"
chaieb@29777
   603
  hence bf: "?f" "?b" by blast+
chaieb@29777
   604
  from permutation_lemma[OF bf] show ?lhs by blast
chaieb@29777
   605
qed
chaieb@29777
   606
chaieb@29777
   607
lemma permutation_inverse_works: assumes p: "permutation p"
chaieb@29777
   608
  shows "inv p o p = id" "p o inv p = id"
chaieb@29777
   609
using permutation_bijective[OF p] surj_iff bij_def inj_iff by auto
chaieb@29777
   610
chaieb@29777
   611
lemma permutation_inverse_compose:
chaieb@29777
   612
  assumes p: "permutation p" and q: "permutation q"
chaieb@29777
   613
  shows "inv (p o q) = inv q o inv p"
chaieb@29777
   614
proof-
chaieb@29777
   615
  note ps = permutation_inverse_works[OF p]
chaieb@29777
   616
  note qs = permutation_inverse_works[OF q]
chaieb@29777
   617
  have "p o q o (inv q o inv p) = p o (q o inv q) o inv p" by (simp add: o_assoc)
chaieb@29777
   618
  also have "\<dots> = id" by (simp add: ps qs)
chaieb@29777
   619
  finally have th0: "p o q o (inv q o inv p) = id" .
chaieb@29777
   620
  have "inv q o inv p o (p o q) = inv q o (inv p o p) o q" by (simp add: o_assoc)
chaieb@29777
   621
  also have "\<dots> = id" by (simp add: ps qs)
huffman@30474
   622
  finally have th1: "inv q o inv p o (p o q) = id" .
chaieb@29777
   623
  from inv_unique_comp[OF th0 th1] show ?thesis .
chaieb@29777
   624
qed
chaieb@29777
   625
chaieb@29777
   626
(* ------------------------------------------------------------------------- *)
chaieb@29777
   627
(* Relation to "permutes".                                                   *)
chaieb@29777
   628
(* ------------------------------------------------------------------------- *)
chaieb@29777
   629
chaieb@29777
   630
lemma permutation_permutes: "permutation p \<longleftrightarrow> (\<exists>S. finite S \<and> p permutes S)"
chaieb@29777
   631
unfolding permutation permutes_def bij_iff[symmetric]
chaieb@29777
   632
apply (rule iffI, clarify)
chaieb@29777
   633
apply (rule exI[where x="{x. p x \<noteq> x}"])
chaieb@29777
   634
apply simp
chaieb@29777
   635
apply clarsimp
chaieb@29777
   636
apply (rule_tac B="S" in finite_subset)
chaieb@29777
   637
apply auto
chaieb@29777
   638
done
chaieb@29777
   639
chaieb@29777
   640
(* ------------------------------------------------------------------------- *)
chaieb@29777
   641
(* Hence a sort of induction principle composing by swaps.                   *)
chaieb@29777
   642
(* ------------------------------------------------------------------------- *)
chaieb@29777
   643
chaieb@29777
   644
lemma permutes_induct: "finite S \<Longrightarrow>  P id  \<Longrightarrow> (\<And> a b p. a \<in> S \<Longrightarrow> b \<in> S \<Longrightarrow> P p \<Longrightarrow> P p \<Longrightarrow> permutation p ==> P (Fun.swap a b id o p))
chaieb@29777
   645
         ==> (\<And>p. p permutes S ==> P p)"
chaieb@29777
   646
proof(induct S rule: finite_induct)
chaieb@29777
   647
  case empty thus ?case by auto
huffman@30474
   648
next
chaieb@29777
   649
  case (insert x F p)
chaieb@29777
   650
  let ?r = "Fun.swap x (p x) id o p"
chaieb@29777
   651
  let ?q = "Fun.swap x (p x) id o ?r"
chaieb@29777
   652
  have qp: "?q = p" by (simp add: o_assoc)
chaieb@29777
   653
  from permutes_insert_lemma[OF insert.prems(3)] insert have Pr: "P ?r" by blast
huffman@30474
   654
  from permutes_in_image[OF insert.prems(3), of x]
chaieb@29777
   655
  have pxF: "p x \<in> insert x F" by simp
chaieb@29777
   656
  have xF: "x \<in> insert x F" by simp
chaieb@29777
   657
  have rp: "permutation ?r"
huffman@30474
   658
    unfolding permutation_permutes using insert.hyps(1)
chaieb@29777
   659
      permutes_insert_lemma[OF insert.prems(3)] by blast
huffman@30474
   660
  from insert.prems(2)[OF xF pxF Pr Pr rp]
huffman@30474
   661
  show ?case  unfolding qp .
chaieb@29777
   662
qed
chaieb@29777
   663
chaieb@29777
   664
(* ------------------------------------------------------------------------- *)
chaieb@29777
   665
(* Sign of a permutation as a real number.                                   *)
chaieb@29777
   666
(* ------------------------------------------------------------------------- *)
chaieb@29777
   667
chaieb@29777
   668
definition "sign p = (if evenperm p then (1::int) else -1)"
chaieb@29777
   669
huffman@30474
   670
lemma sign_nz: "sign p \<noteq> 0" by (simp add: sign_def)
chaieb@29777
   671
lemma sign_id: "sign id = 1" by (simp add: sign_def)
chaieb@29777
   672
lemma sign_inverse: "permutation p ==> sign (inv p) = sign p"
chaieb@29777
   673
  by (simp add: sign_def evenperm_inv)
chaieb@29777
   674
lemma sign_compose: "permutation p \<Longrightarrow> permutation q ==> sign (p o q) = sign(p) * sign(q)" by (simp add: sign_def evenperm_comp)
chaieb@29777
   675
lemma sign_swap_id: "sign (Fun.swap a b id) = (if a = b then 1 else -1)"
chaieb@29777
   676
  by (simp add: sign_def evenperm_swap)
chaieb@29777
   677
lemma sign_idempotent: "sign p * sign p = 1" by (simp add: sign_def)
chaieb@29777
   678
chaieb@29777
   679
(* ------------------------------------------------------------------------- *)
chaieb@29777
   680
(* More lemmas about permutations.                                           *)
chaieb@29777
   681
(* ------------------------------------------------------------------------- *)
chaieb@29777
   682
chaieb@29777
   683
lemma permutes_natset_le:
huffman@29974
   684
  assumes p: "p permutes (S::'a::wellorder set)" and le: "\<forall>i \<in> S.  p i <= i" shows "p = id"
chaieb@29777
   685
proof-
chaieb@29777
   686
  {fix n
huffman@30474
   687
    have "p n = n"
chaieb@29777
   688
      using p le
huffman@29974
   689
    proof(induct n arbitrary: S rule: less_induct)
huffman@30474
   690
      fix n S assume H: "\<And>m S. \<lbrakk>m < n; p permutes S; \<forall>i\<in>S. p i \<le> i\<rbrakk> \<Longrightarrow> p m = m"
wenzelm@32962
   691
        "p permutes S" "\<forall>i \<in>S. p i \<le> i"
chaieb@29777
   692
      {assume "n \<notin> S"
wenzelm@32962
   693
        with H(2) have "p n = n" unfolding permutes_def by metis}
chaieb@29777
   694
      moreover
chaieb@29777
   695
      {assume ns: "n \<in> S"
wenzelm@32962
   696
        from H(3)  ns have "p n < n \<or> p n = n" by auto
wenzelm@32962
   697
        moreover{assume h: "p n < n"
wenzelm@32962
   698
          from H h have "p (p n) = p n" by metis
wenzelm@32962
   699
          with permutes_inj[OF H(2)] have "p n = n" unfolding inj_on_def by blast
wenzelm@32962
   700
          with h have False by simp}
wenzelm@32962
   701
        ultimately have "p n = n" by blast }
chaieb@29777
   702
      ultimately show "p n = n"  by blast
chaieb@29777
   703
    qed}
chaieb@29777
   704
  thus ?thesis by (auto simp add: expand_fun_eq)
chaieb@29777
   705
qed
chaieb@29777
   706
chaieb@29777
   707
lemma permutes_natset_ge:
huffman@29974
   708
  assumes p: "p permutes (S::'a::wellorder set)" and le: "\<forall>i \<in> S.  p i \<ge> i" shows "p = id"
chaieb@29777
   709
proof-
chaieb@29777
   710
  {fix i assume i: "i \<in> S"
chaieb@29777
   711
    from i permutes_in_image[OF permutes_inv[OF p]] have "inv p i \<in> S" by simp
chaieb@29777
   712
    with le have "p (inv p i) \<ge> inv p i" by blast
chaieb@29777
   713
    with permutes_inverses[OF p] have "i \<ge> inv p i" by simp}
chaieb@29777
   714
  then have th: "\<forall>i\<in>S. inv p i \<le> i"  by blast
huffman@30474
   715
  from permutes_natset_le[OF permutes_inv[OF p] th]
chaieb@29777
   716
  have "inv p = inv id" by simp
huffman@30474
   717
  then show ?thesis
chaieb@29777
   718
    apply (subst permutes_inv_inv[OF p, symmetric])
chaieb@29777
   719
    apply (rule inv_unique_comp)
chaieb@29777
   720
    apply simp_all
chaieb@29777
   721
    done
chaieb@29777
   722
qed
chaieb@29777
   723
chaieb@29777
   724
lemma image_inverse_permutations: "{inv p |p. p permutes S} = {p. p permutes S}"
chaieb@29777
   725
apply (rule set_ext)
chaieb@29777
   726
apply auto
chaieb@29777
   727
  using permutes_inv_inv permutes_inv apply auto
chaieb@29777
   728
  apply (rule_tac x="inv x" in exI)
chaieb@29777
   729
  apply auto
chaieb@29777
   730
  done
chaieb@29777
   731
huffman@30474
   732
lemma image_compose_permutations_left:
chaieb@29777
   733
  assumes q: "q permutes S" shows "{q o p | p. p permutes S} = {p . p permutes S}"
chaieb@29777
   734
apply (rule set_ext)
chaieb@29777
   735
apply auto
chaieb@29777
   736
apply (rule permutes_compose)
chaieb@29777
   737
using q apply auto
chaieb@29777
   738
apply (rule_tac x = "inv q o x" in exI)
chaieb@29777
   739
by (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o)
chaieb@29777
   740
chaieb@29777
   741
lemma image_compose_permutations_right:
chaieb@29777
   742
  assumes q: "q permutes S"
chaieb@29777
   743
  shows "{p o q | p. p permutes S} = {p . p permutes S}"
chaieb@29777
   744
apply (rule set_ext)
chaieb@29777
   745
apply auto
chaieb@29777
   746
apply (rule permutes_compose)
chaieb@29777
   747
using q apply auto
chaieb@29777
   748
apply (rule_tac x = "x o inv q" in exI)
chaieb@29777
   749
by (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o o_assoc[symmetric])
chaieb@29777
   750
chaieb@29777
   751
lemma permutes_in_seg: "p permutes {1 ..n} \<Longrightarrow> i \<in> {1..n} ==> 1 <= p i \<and> p i <= n"
chaieb@29777
   752
chaieb@29777
   753
apply (simp add: permutes_def)
chaieb@29777
   754
apply metis
chaieb@29777
   755
done
chaieb@29777
   756
chaieb@29777
   757
term setsum
huffman@29973
   758
lemma setsum_permutations_inverse: "setsum f {p. p permutes S} = setsum (\<lambda>p. f(inv p)) {p. p permutes S}" (is "?lhs = ?rhs")
chaieb@29777
   759
proof-
huffman@29973
   760
  let ?S = "{p . p permutes S}"
huffman@30474
   761
have th0: "inj_on inv ?S"
chaieb@29777
   762
proof(auto simp add: inj_on_def)
chaieb@29777
   763
  fix q r
huffman@29973
   764
  assume q: "q permutes S" and r: "r permutes S" and qr: "inv q = inv r"
chaieb@29777
   765
  hence "inv (inv q) = inv (inv r)" by simp
chaieb@29777
   766
  with permutes_inv_inv[OF q] permutes_inv_inv[OF r]
chaieb@29777
   767
  show "q = r" by metis
chaieb@29777
   768
qed
chaieb@29777
   769
  have th1: "inv ` ?S = ?S" using image_inverse_permutations by blast
chaieb@29777
   770
  have th2: "?rhs = setsum (f o inv) ?S" by (simp add: o_def)
chaieb@29777
   771
  from setsum_reindex[OF th0, of f]  show ?thesis unfolding th1 th2 .
chaieb@29777
   772
qed
chaieb@29777
   773
chaieb@29777
   774
lemma setum_permutations_compose_left:
huffman@29973
   775
  assumes q: "q permutes S"
huffman@29973
   776
  shows "setsum f {p. p permutes S} =
huffman@29973
   777
            setsum (\<lambda>p. f(q o p)) {p. p permutes S}" (is "?lhs = ?rhs")
chaieb@29777
   778
proof-
huffman@29973
   779
  let ?S = "{p. p permutes S}"
chaieb@29777
   780
  have th0: "?rhs = setsum (f o (op o q)) ?S" by (simp add: o_def)
chaieb@29777
   781
  have th1: "inj_on (op o q) ?S"
chaieb@29777
   782
    apply (auto simp add: inj_on_def)
chaieb@29777
   783
  proof-
chaieb@29777
   784
    fix p r
huffman@29973
   785
    assume "p permutes S" and r:"r permutes S" and rp: "q \<circ> p = q \<circ> r"
chaieb@29777
   786
    hence "inv q o q o p = inv q o q o r" by (simp add: o_assoc[symmetric])
chaieb@29777
   787
    with permutes_inj[OF q, unfolded inj_iff]
chaieb@29777
   788
chaieb@29777
   789
    show "p = r" by simp
chaieb@29777
   790
  qed
chaieb@29777
   791
  have th3: "(op o q) ` ?S = ?S" using image_compose_permutations_left[OF q] by auto
chaieb@29777
   792
  from setsum_reindex[OF th1, of f]
chaieb@29777
   793
  show ?thesis unfolding th0 th1 th3 .
chaieb@29777
   794
qed
chaieb@29777
   795
chaieb@29777
   796
lemma sum_permutations_compose_right:
huffman@29973
   797
  assumes q: "q permutes S"
huffman@29973
   798
  shows "setsum f {p. p permutes S} =
huffman@29973
   799
            setsum (\<lambda>p. f(p o q)) {p. p permutes S}" (is "?lhs = ?rhs")
chaieb@29777
   800
proof-
huffman@29973
   801
  let ?S = "{p. p permutes S}"
chaieb@29777
   802
  have th0: "?rhs = setsum (f o (\<lambda>p. p o q)) ?S" by (simp add: o_def)
chaieb@29777
   803
  have th1: "inj_on (\<lambda>p. p o q) ?S"
chaieb@29777
   804
    apply (auto simp add: inj_on_def)
chaieb@29777
   805
  proof-
chaieb@29777
   806
    fix p r
huffman@29973
   807
    assume "p permutes S" and r:"r permutes S" and rp: "p o q = r o q"
chaieb@29777
   808
    hence "p o (q o inv q)  = r o (q o inv q)" by (simp add: o_assoc)
chaieb@29777
   809
    with permutes_surj[OF q, unfolded surj_iff]
chaieb@29777
   810
chaieb@29777
   811
    show "p = r" by simp
chaieb@29777
   812
  qed
chaieb@29777
   813
  have th3: "(\<lambda>p. p o q) ` ?S = ?S" using image_compose_permutations_right[OF q] by auto
chaieb@29777
   814
  from setsum_reindex[OF th1, of f]
chaieb@29777
   815
  show ?thesis unfolding th0 th1 th3 .
chaieb@29777
   816
qed
chaieb@29777
   817
chaieb@29777
   818
(* ------------------------------------------------------------------------- *)
chaieb@29777
   819
(* Sum over a set of permutations (could generalize to iteration).           *)
chaieb@29777
   820
(* ------------------------------------------------------------------------- *)
chaieb@29777
   821
chaieb@29777
   822
lemma setsum_over_permutations_insert:
chaieb@29777
   823
  assumes fS: "finite S" and aS: "a \<notin> S"
chaieb@29777
   824
  shows "setsum f {p. p permutes (insert a S)} = setsum (\<lambda>b. setsum (\<lambda>q. f (Fun.swap a b id o q)) {p. p permutes S}) (insert a S)"
chaieb@29777
   825
proof-
chaieb@29777
   826
  have th0: "\<And>f a b. (\<lambda>(b,p). f (Fun.swap a b id o p)) = f o (\<lambda>(b,p). Fun.swap a b id o p)"
chaieb@29777
   827
    by (simp add: expand_fun_eq)
chaieb@29777
   828
  have th1: "\<And>P Q. P \<times> Q = {(a,b). a \<in> P \<and> b \<in> Q}" by blast
chaieb@29777
   829
  have th2: "\<And>P Q. P \<Longrightarrow> (P \<Longrightarrow> Q) \<Longrightarrow> P \<and> Q" by blast
huffman@30474
   830
  show ?thesis
huffman@30474
   831
    unfolding permutes_insert
chaieb@29777
   832
    unfolding setsum_cartesian_product
chaieb@29777
   833
    unfolding  th1[symmetric]
chaieb@29777
   834
    unfolding th0
chaieb@29777
   835
  proof(rule setsum_reindex)
chaieb@29777
   836
    let ?f = "(\<lambda>(b, y). Fun.swap a b id \<circ> y)"
chaieb@29777
   837
    let ?P = "{p. p permutes S}"
huffman@30474
   838
    {fix b c p q assume b: "b \<in> insert a S" and c: "c \<in> insert a S"
huffman@30474
   839
      and p: "p permutes S" and q: "q permutes S"
chaieb@29777
   840
      and eq: "Fun.swap a b id o p = Fun.swap a c id o q"
chaieb@29777
   841
      from p q aS have pa: "p a = a" and qa: "q a = a"
wenzelm@32962
   842
        unfolding permutes_def by metis+
chaieb@29777
   843
      from eq have "(Fun.swap a b id o p) a  = (Fun.swap a c id o q) a" by simp
chaieb@29777
   844
      hence bc: "b = c"
wenzelm@32962
   845
        by (simp add: permutes_def pa qa o_def fun_upd_def swap_def id_def cong del: if_weak_cong split: split_if_asm)
chaieb@29777
   846
      from eq[unfolded bc] have "(\<lambda>p. Fun.swap a c id o p) (Fun.swap a c id o p) = (\<lambda>p. Fun.swap a c id o p) (Fun.swap a c id o q)" by simp
chaieb@29777
   847
      hence "p = q" unfolding o_assoc swap_id_idempotent
wenzelm@32962
   848
        by (simp add: o_def)
chaieb@29777
   849
      with bc have "b = c \<and> p = q" by blast
chaieb@29777
   850
    }
huffman@30474
   851
huffman@30474
   852
    then show "inj_on ?f (insert a S \<times> ?P)"
chaieb@29777
   853
      unfolding inj_on_def
chaieb@29777
   854
      apply clarify by metis
chaieb@29777
   855
  qed
chaieb@29777
   856
qed
chaieb@29777
   857
chaieb@29777
   858
end