products/sources/formale Sprachen/Isabelle/HOL/Library image not shown  

Quellcode-Bibliothek

© Kompilation durch diese Firma

[Weder Korrektheit noch Funktionsfähigkeit der Software werden zugesichert.]

Datei: Equipollence.thy   Sprache: Isabelle

Original von: Isabelle©

(*  Title:      HOL/Library/Permutations.thy
    Author:     Amine Chaieb, University of Cambridge
*)


section \<open>Permutations, both general and specifically on finite sets.\<close>

theory Permutations
  imports Multiset Disjoint_Sets
begin

subsection \<open>Transpositions\<close>

lemma swap_id_idempotent [simp]: "Fun.swap a b id \ Fun.swap a b id = id"
  by (rule ext) (auto simp add: Fun.swap_def)

lemma inv_swap_id: "inv (Fun.swap a b id) = Fun.swap a b id"
  by (rule inv_unique_comp) simp_all

lemma swap_id_eq: "Fun.swap a b id x = (if x = a then b else if x = b then a else x)"
  by (simp add: Fun.swap_def)

lemma bij_swap_comp:
  assumes "bij p"
  shows "Fun.swap a b id \ p = Fun.swap (inv p a) (inv p b) p"
  using surj_f_inv_f[OF bij_is_surj[OF \<open>bij p\<close>]]
  by (simp add: fun_eq_iff Fun.swap_def bij_inv_eq_iff[OF \<open>bij p\<close>])

lemma bij_swap_compose_bij:
  assumes "bij p"
  shows "bij (Fun.swap a b id \ p)"
  by (simp only: bij_swap_comp[OF \<open>bij p\<close>] bij_swap_iff \<open>bij p\<close>)


subsection \<open>Basic consequences of the definition\<close>

definition permutes  (infixr "permutes" 41)
  where "(p permutes S) \ (\x. x \ S \ p x = x) \ (\y. \!x. p x = y)"

lemma permutes_in_image: "p permutes S \ p x \ S \ x \ S"
  unfolding permutes_def by metis

lemma permutes_not_in: "f permutes S \ x \ S \ f x = x"
  by (auto simp: permutes_def)

lemma permutes_image: "p permutes S \ p ` S = S"
  unfolding permutes_def
  apply (rule set_eqI)
  apply (simp add: image_iff)
  apply metis
  done

lemma permutes_inj: "p permutes S \ inj p"
  unfolding permutes_def inj_def by blast

lemma permutes_inj_on: "f permutes S \ inj_on f A"
  by (auto simp: permutes_def inj_on_def)

lemma permutes_surj: "p permutes s \ surj p"
  unfolding permutes_def surj_def by metis

lemma permutes_bij: "p permutes s \ bij p"
  unfolding bij_def by (metis permutes_inj permutes_surj)

lemma permutes_imp_bij: "p permutes S \ bij_betw p S S"
  by (metis UNIV_I bij_betw_subset permutes_bij permutes_image subsetI)

lemma bij_imp_permutes: "bij_betw p S S \ (\x. x \ S \ p x = x) \ p permutes S"
  unfolding permutes_def bij_betw_def inj_on_def
  by auto (metis image_iff)+

lemma permutes_inv_o:
  assumes permutes: "p permutes S"
  shows "p \ inv p = id"
    and "inv p \ p = id"
  using permutes_inj[OF permutes] permutes_surj[OF permutes]
  unfolding inj_iff[symmetric] surj_iff[symmetric] by blast+

lemma permutes_inverses:
  fixes p :: "'a \ 'a"
  assumes permutes: "p permutes S"
  shows "p (inv p x) = x"
    and "inv p (p x) = x"
  using permutes_inv_o[OF permutes, unfolded fun_eq_iff o_def] by auto

lemma permutes_subset: "p permutes S \ S \ T \ p permutes T"
  unfolding permutes_def by blast

lemma permutes_empty[simp]: "p permutes {} \ p = id"
  by (auto simp add: fun_eq_iff permutes_def)

lemma permutes_sing[simp]: "p permutes {a} \ p = id"
  by (simp add: fun_eq_iff permutes_def) metis  (*somewhat slow*)

lemma permutes_univ: "p permutes UNIV \ (\y. \!x. p x = y)"
  by (simp add: permutes_def)

lemma permutes_inv_eq: "p permutes S \ inv p y = x \ p x = y"
  unfolding permutes_def inv_def
  apply auto
  apply (erule allE[where x=y])
  apply (erule allE[where x=y])
  apply (rule someI_ex)
  apply blast
  apply (rule some1_equality)
  apply blast
  apply blast
  done

lemma permutes_swap_id: "a \ S \ b \ S \ Fun.swap a b id permutes S"
  unfolding permutes_def Fun.swap_def fun_upd_def by auto metis

lemma permutes_superset: "p permutes S \ (\x \ S - T. p x = x) \ p permutes T"
  by (simp add: Ball_def permutes_def) metis

lemma permutes_bij_inv_into: \<^marker>\<open>contributor \<open>Lukas Bulwahn\<close>\<close>
  fixes A :: "'a set"
    and B :: "'b set"
  assumes "p permutes A"
    and "bij_betw f A B"
  shows "(\x. if x \ B then f (p (inv_into A f x)) else x) permutes B"
proof (rule bij_imp_permutes)
  from assms have "bij_betw p A A" "bij_betw f A B" "bij_betw (inv_into A f) B A"
    by (auto simp add: permutes_imp_bij bij_betw_inv_into)
  then have "bij_betw (f \ p \ inv_into A f) B B"
    by (simp add: bij_betw_trans)
  then show "bij_betw (\x. if x \ B then f (p (inv_into A f x)) else x) B B"
    by (subst bij_betw_cong[where g="f \ p \ inv_into A f"]) auto
next
  fix x
  assume "x \ B"
  then show "(if x \ B then f (p (inv_into A f x)) else x) = x" by auto
qed

lemma permutes_image_mset: \<^marker>\<open>contributor \<open>Lukas Bulwahn\<close>\<close>
  assumes "p permutes A"
  shows "image_mset p (mset_set A) = mset_set A"
  using assms by (metis image_mset_mset_set bij_betw_imp_inj_on permutes_imp_bij permutes_image)

lemma permutes_implies_image_mset_eq: \<^marker>\<open>contributor \<open>Lukas Bulwahn\<close>\<close>
  assumes "p permutes A" "\x. x \ A \ f x = f' (p x)"
  shows "image_mset f' (mset_set A) = image_mset f (mset_set A)"
proof -
  have "f x = f' (p x)" if "x \# mset_set A" for x
    using assms(2)[of x] that by (cases "finite A") auto
  with assms have "image_mset f (mset_set A) = image_mset (f' \ p) (mset_set A)"
    by (auto intro!: image_mset_cong)
  also have "\ = image_mset f' (image_mset p (mset_set A))"
    by (simp add: image_mset.compositionality)
  also have "\ = image_mset f' (mset_set A)"
  proof -
    from assms permutes_image_mset have "image_mset p (mset_set A) = mset_set A"
      by blast
    then show ?thesis by simp
  qed
  finally show ?thesis ..
qed


subsection \<open>Group properties\<close>

lemma permutes_id: "id permutes S"
  by (simp add: permutes_def)

lemma permutes_compose: "p permutes S \ q permutes S \ q \ p permutes S"
  unfolding permutes_def o_def by metis

lemma permutes_inv:
  assumes "p permutes S"
  shows "inv p permutes S"
  using assms unfolding permutes_def permutes_inv_eq[OF assms] by metis

lemma permutes_inv_inv:
  assumes "p permutes S"
  shows "inv (inv p) = p"
  unfolding fun_eq_iff permutes_inv_eq[OF assms] permutes_inv_eq[OF permutes_inv[OF assms]]
  by blast

lemma permutes_invI:
  assumes perm: "p permutes S"
    and inv: "\x. x \ S \ p' (p x) = x"
    and outside: "\x. x \ S \ p' x = x"
  shows "inv p = p'"
proof
  show "inv p x = p' x" for x
  proof (cases "x \ S")
    case True
    from assms have "p' x = p' (p (inv p x))"
      by (simp add: permutes_inverses)
    also from permutes_inv[OF perm] True have "\ = inv p x"
      by (subst inv) (simp_all add: permutes_in_image)
    finally show ?thesis ..
  next
    case False
    with permutes_inv[OF perm] show ?thesis
      by (simp_all add: outside permutes_not_in)
  qed
qed

lemma permutes_vimage: "f permutes A \ f -` A = A"
  by (simp add: bij_vimage_eq_inv_image permutes_bij permutes_image[OF permutes_inv])


subsection \<open>Mapping permutations with bijections\<close>

lemma bij_betw_permutations:
  assumes "bij_betw f A B"
  shows   "bij_betw (\\ x. if x \ B then f (\ (inv_into A f x)) else x)
             {\<pi>. \<pi> permutes A} {\<pi>. \<pi> permutes B}" (is "bij_betw ?f _ _")
proof -
  let ?g = "(\\ x. if x \ A then inv_into A f (\ (f x)) else x)"
  show ?thesis
  proof (rule bij_betw_byWitness [of _ ?g], goal_cases)
    case 3
    show ?case using permutes_bij_inv_into[OF _ assms] by auto
  next
    case 4
    have bij_inv: "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms)
    {
      fix \<pi> assume "\<pi> permutes B"
      from permutes_bij_inv_into[OF this bij_inv] and assms
        have "(\x. if x \ A then inv_into A f (\ (f x)) else x) permutes A"
        by (simp add: inv_into_inv_into_eq cong: if_cong)
    }
    from this show ?case by (auto simp: permutes_inv)
  next
    case 1
    thus ?case using assms
      by (auto simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_left
               dest: bij_betwE)
  next
    case 2
    moreover have "bij_betw (inv_into A f) B A"
      by (intro bij_betw_inv_into assms)
    ultimately show ?case using assms
      by (auto simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_right 
               dest: bij_betwE)
  qed
qed

lemma bij_betw_derangements:
  assumes "bij_betw f A B"
  shows   "bij_betw (\\ x. if x \ B then f (\ (inv_into A f x)) else x)
             {\<pi>. \<pi> permutes A \<and> (\<forall>x\<in>A. \<pi> x \<noteq> x)} {\<pi>. \<pi> permutes B \<and> (\<forall>x\<in>B. \<pi> x \<noteq> x)}" 
           (is "bij_betw ?f _ _")
proof -
  let ?g = "(\\ x. if x \ A then inv_into A f (\ (f x)) else x)"
  show ?thesis
  proof (rule bij_betw_byWitness [of _ ?g], goal_cases)
    case 3
    have "?f \ x \ x" if "\ permutes A" "\x. x \ A \ \ x \ x" "x \ B" for \ x
      using that and assms by (metis bij_betwE bij_betw_imp_inj_on bij_betw_imp_surj_on
                                     inv_into_f_f inv_into_into permutes_imp_bij)
    with permutes_bij_inv_into[OF _ assms] show ?case by auto
  next
    case 4
    have bij_inv: "bij_betw (inv_into A f) B A" by (intro bij_betw_inv_into assms)
    have "?g \ permutes A" if "\ permutes B" for \
      using permutes_bij_inv_into[OF that bij_inv] and assms
      by (simp add: inv_into_inv_into_eq cong: if_cong)
    moreover have "?g \ x \ x" if "\ permutes B" "\x. x \ B \ \ x \ x" "x \ A" for \ x
      using that and assms by (metis bij_betwE bij_betw_imp_surj_on f_inv_into_f permutes_imp_bij)
    ultimately show ?case by auto
  next
    case 1
    thus ?case using assms
      by (force simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_left
                dest: bij_betwE)
  next
    case 2
    moreover have "bij_betw (inv_into A f) B A"
      by (intro bij_betw_inv_into assms)
    ultimately show ?case using assms
      by (force simp: fun_eq_iff permutes_not_in permutes_in_image bij_betw_inv_into_right 
                dest: bij_betwE)
  qed
qed


subsection \<open>The number of permutations on a finite set\<close>

lemma permutes_insert_lemma:
  assumes "p permutes (insert a S)"
  shows "Fun.swap a (p a) id \ p permutes S"
  apply (rule permutes_superset[where S = "insert a S"])
  apply (rule permutes_compose[OF assms])
  apply (rule permutes_swap_id, simp)
  using permutes_in_image[OF assms, of a]
  apply simp
  apply (auto simp add: Ball_def Fun.swap_def)
  done

lemma permutes_insert: "{p. p permutes (insert a S)} =
  (\<lambda>(b, p). Fun.swap a b id \<circ> p) ` {(b, p). b \<in> insert a S \<and> p \<in> {p. p permutes S}}"
proof -
  have "p permutes insert a S \
    (\<exists>b q. p = Fun.swap a b id \<circ> q \<and> b \<in> insert a S \<and> q permutes S)" for p
  proof -
    have "\b q. p = Fun.swap a b id \ q \ b \ insert a S \ q permutes S"
      if p: "p permutes insert a S"
    proof -
      let ?b = "p a"
      let ?q = "Fun.swap a (p a) id \ p"
      have *: "p = Fun.swap a ?b id \ ?q"
        by (simp add: fun_eq_iff o_assoc)
      have **: "?b \ insert a S"
        unfolding permutes_in_image[OF p] by simp
      from permutes_insert_lemma[OF p] * ** show ?thesis
       by blast
    qed
    moreover have "p permutes insert a S"
      if bq: "p = Fun.swap a b id \ q" "b \ insert a S" "q permutes S" for b q
    proof -
      from permutes_subset[OF bq(3), of "insert a S"have q: "q permutes insert a S"
        by auto
      have a: "a \ insert a S"
        by simp
      from bq(1) permutes_compose[OF q permutes_swap_id[OF a bq(2)]] show ?thesis
        by simp
    qed
    ultimately show ?thesis by blast
  qed
  then show ?thesis by auto
qed

lemma card_permutations:
  assumes "card S = n"
    and "finite S"
  shows "card {p. p permutes S} = fact n"
  using assms(2,1)
proof (induct arbitrary: n)
  case empty
  then show ?case by simp
next
  case (insert x F)
  {
    fix n
    assume card_insert: "card (insert x F) = n"
    let ?xF = "{p. p permutes insert x F}"
    let ?pF = "{p. p permutes F}"
    let ?pF' = "{(b, p). b \ insert x F \ p \ ?pF}"
    let ?g = "(\(b, p). Fun.swap x b id \ p)"
    have xfgpF': "?xF = ?g ` ?pF'"
      by (rule permutes_insert[of x F])
    from \<open>x \<notin> F\<close> \<open>finite F\<close> card_insert have Fs: "card F = n - 1"
      by auto
    from \<open>finite F\<close> insert.hyps Fs have pFs: "card ?pF = fact (n - 1)"
      by auto
    then have "finite ?pF"
      by (auto intro: card_ge_0_finite)
    with \<open>finite F\<close> card.insert_remove have pF'f: "finite ?pF'"
      apply (simp only: Collect_case_prod Collect_mem_eq)
      apply (rule finite_cartesian_product)
      apply simp_all
      done

    have ginj: "inj_on ?g ?pF'"
    proof -
      {
        fix b p c q
        assume bp: "(b, p) \ ?pF'"
        assume cq: "(c, q) \ ?pF'"
        assume eq: "?g (b, p) = ?g (c, q)"
        from bp cq have pF: "p permutes F" and qF: "q permutes F"
          by auto
        from pF \<open>x \<notin> F\<close> eq have "b = ?g (b, p) x"
          by (auto simp: permutes_def Fun.swap_def fun_upd_def fun_eq_iff)
        also from qF \<open>x \<notin> F\<close> eq have "\<dots> = ?g (c, q) x"
          by (auto simp: swap_def fun_upd_def fun_eq_iff)
        also from qF \<open>x \<notin> F\<close> have "\<dots> = c"
          by (auto simp: permutes_def Fun.swap_def fun_upd_def fun_eq_iff)
        finally have "b = c" .
        then have "Fun.swap x b id = Fun.swap x c id"
          by simp
        with eq have "Fun.swap x b id \ p = Fun.swap x b id \ q"
          by simp
        then have "Fun.swap x b id \ (Fun.swap x b id \ p) = Fun.swap x b id \ (Fun.swap x b id \ q)"
          by simp
        then have "p = q"
          by (simp add: o_assoc)
        with \<open>b = c\<close> have "(b, p) = (c, q)"
          by simp
      }
      then show ?thesis
        unfolding inj_on_def by blast
    qed
    from \<open>x \<notin> F\<close> \<open>finite F\<close> card_insert have "n \<noteq> 0"
      by auto
    then have "\m. n = Suc m"
      by presburger
    then obtain m where n: "n = Suc m"
      by blast
    from pFs card_insert have *: "card ?xF = fact n"
      unfolding xfgpF' card_image[OF ginj]
      using \<open>finite F\<close> \<open>finite ?pF\<close>
      by (simp only: Collect_case_prod Collect_mem_eq card_cartesian_product) (simp add: n)
    from finite_imageI[OF pF'f, of ?g] have xFf: "finite ?xF"
      by (simp add: xfgpF' n)
    from * have "card ?xF = fact n"
      unfolding xFf by blast
  }
  with insert show ?case by simp
qed

lemma finite_permutations:
  assumes "finite S"
  shows "finite {p. p permutes S}"
  using card_permutations[OF refl assms] by (auto intro: card_ge_0_finite)


subsection \<open>Permutations of index set for iterated operations\<close>

lemma (in comm_monoid_set) permute:
  assumes "p permutes S"
  shows "F g S = F (g \ p) S"
proof -
  from \<open>p permutes S\<close> have "inj p"
    by (rule permutes_inj)
  then have "inj_on p S"
    by (auto intro: subset_inj_on)
  then have "F g (p ` S) = F (g \ p) S"
    by (rule reindex)
  moreover from \<open>p permutes S\<close> have "p ` S = S"
    by (rule permutes_image)
  ultimately show ?thesis
    by simp
qed


subsection \<open>Various combinations of transpositions with 2, 1 and 0 common elements\<close>

lemma swap_id_common:" a \ c \ b \ c \
  Fun.swap a b id \<circ> Fun.swap a c id = Fun.swap b c id \<circ> Fun.swap a b id"
  by (simp add: fun_eq_iff Fun.swap_def)

lemma swap_id_common': "a \ b \ a \ c \
  Fun.swap a c id \<circ> Fun.swap b c id = Fun.swap b c id \<circ> Fun.swap a b id"
  by (simp add: fun_eq_iff Fun.swap_def)

lemma swap_id_independent: "a \ c \ a \ d \ b \ c \ b \ d \
  Fun.swap a b id \<circ> Fun.swap c d id = Fun.swap c d id \<circ> Fun.swap a b id"
  by (simp add: fun_eq_iff Fun.swap_def)


subsection \<open>Permutations as transposition sequences\<close>

inductive swapidseq :: "nat \ ('a \ 'a) \ bool"
  where
    id[simp]: "swapidseq 0 id"
  | comp_Suc: "swapidseq n p \ a \ b \ swapidseq (Suc n) (Fun.swap a b id \ p)"

declare id[unfolded id_def, simp]

definition "permutation p \ (\n. swapidseq n p)"


subsection \<open>Some closure properties of the set of permutations, with lengths\<close>

lemma permutation_id[simp]: "permutation id"
  unfolding permutation_def by (rule exI[where x=0]) simp

declare permutation_id[unfolded id_def, simp]

lemma swapidseq_swap: "swapidseq (if a = b then 0 else 1) (Fun.swap a b id)"
  apply clarsimp
  using comp_Suc[of 0 id a b]
  apply simp
  done

lemma permutation_swap_id: "permutation (Fun.swap a b id)"
proof (cases "a = b")
  case True
  then show ?thesis by simp
next
  case False
  then show ?thesis
    unfolding permutation_def
    using swapidseq_swap[of a b] by blast
qed


lemma swapidseq_comp_add: "swapidseq n p \ swapidseq m q \ swapidseq (n + m) (p \ q)"
proof (induct n p arbitrary: m q rule: swapidseq.induct)
  case (id m q)
  then show ?case by simp
next
  case (comp_Suc n p a b m q)
  have eq: "Suc n + m = Suc (n + m)"
    by arith
  show ?case
    apply (simp only: eq comp_assoc)
    apply (rule swapidseq.comp_Suc)
    using comp_Suc.hyps(2)[OF comp_Suc.prems] comp_Suc.hyps(3)
     apply blast+
    done
qed

lemma permutation_compose: "permutation p \ permutation q \ permutation (p \ q)"
  unfolding permutation_def using swapidseq_comp_add[of _ p _ q] by metis

lemma swapidseq_endswap: "swapidseq n p \ a \ b \ swapidseq (Suc n) (p \ Fun.swap a b id)"
  by (induct n p rule: swapidseq.induct)
    (use swapidseq_swap[of a b] in \<open>auto simp add: comp_assoc intro: swapidseq.comp_Suc\<close>)

lemma swapidseq_inverse_exists: "swapidseq n p \ \q. swapidseq n q \ p \ q = id \ q \ p = id"
proof (induct n p rule: swapidseq.induct)
  case id
  then show ?case
    by (rule exI[where x=id]) simp
next
  case (comp_Suc n p a b)
  from comp_Suc.hyps obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id"
    by blast
  let ?q = "q \ Fun.swap a b id"
  note H = comp_Suc.hyps
  from swapidseq_swap[of a b] H(3) have *: "swapidseq 1 (Fun.swap a b id)"
    by simp
  from swapidseq_comp_add[OF q(1) *] have **: "swapidseq (Suc n) ?q"
    by simp
  have "Fun.swap a b id \ p \ ?q = Fun.swap a b id \ (p \ q) \ Fun.swap a b id"
    by (simp add: o_assoc)
  also have "\ = id"
    by (simp add: q(2))
  finally have ***: "Fun.swap a b id \ p \ ?q = id" .
  have "?q \ (Fun.swap a b id \ p) = q \ (Fun.swap a b id \ Fun.swap a b id) \ p"
    by (simp only: o_assoc)
  then have "?q \ (Fun.swap a b id \ p) = id"
    by (simp add: q(3))
  with ** *** show ?case
    by blast
qed

lemma swapidseq_inverse:
  assumes "swapidseq n p"
  shows "swapidseq n (inv p)"
  using swapidseq_inverse_exists[OF assms] inv_unique_comp[of p] by auto

lemma permutation_inverse: "permutation p \ permutation (inv p)"
  using permutation_def swapidseq_inverse by blast


subsection \<open>The identity map only has even transposition sequences\<close>

lemma symmetry_lemma:
  assumes "\a b c d. P a b c d \ P a b d c"
    and "\a b c d. a \ b \ c \ d \
      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 \<Longrightarrow>
      P a b c d"
  shows "\a b c d. a \ b \ c \ d \ P a b c d"
  using assms by metis

lemma swap_general: "a \ b \ c \ d \
  Fun.swap a b id \<circ> Fun.swap c d id = id \<or>
  (\<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 \<circ> Fun.swap c d id = Fun.swap x y id \<circ> Fun.swap a z id)"
proof -
  assume neq: "a \ b" "c \ d"
  have "a \ b \ c \ d \
    (Fun.swap a b id \<circ> Fun.swap c d id = id \<or>
      (\<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 \<circ> Fun.swap c d id = Fun.swap x y id \<circ> Fun.swap a z id))"
    apply (rule symmetry_lemma[where a=a and b=b and c=c and d=d])
     apply (simp_all only: swap_commute)
    apply (case_tac "a = c \ b = d")
     apply (clarsimp simp only: swap_commute swap_id_idempotent)
    apply (case_tac "a = c \ b \ d")
     apply (rule disjI2)
     apply (rule_tac x="b" in exI)
     apply (rule_tac x="d" in exI)
     apply (rule_tac x="b" in exI)
     apply (clarsimp simp add: fun_eq_iff Fun.swap_def)
    apply (case_tac "a \ c \ b = d")
     apply (rule disjI2)
     apply (rule_tac x="c" in exI)
     apply (rule_tac x="d" in exI)
     apply (rule_tac x="c" in exI)
     apply (clarsimp simp add: fun_eq_iff Fun.swap_def)
    apply (rule disjI2)
    apply (rule_tac x="c" in exI)
    apply (rule_tac x="d" in exI)
    apply (rule_tac x="b" in exI)
    apply (clarsimp simp add: fun_eq_iff Fun.swap_def)
    done
  with neq show ?thesis by metis
qed

lemma swapidseq_id_iff[simp]: "swapidseq 0 p \ p = id"
  using swapidseq.cases[of 0 p "p = id"by auto

lemma swapidseq_cases: "swapidseq n p \
    n = 0 \<and> p = id \<or> (\<exists>a b q m. n = Suc m \<and> p = Fun.swap a b id \<circ> q \<and> swapidseq m q \<and> a \<noteq> b)"
  apply (rule iffI)
   apply (erule swapidseq.cases[of n p])
    apply simp
   apply (rule disjI2)
   apply (rule_tac x= "a" in exI)
   apply (rule_tac x= "b" in exI)
   apply (rule_tac x= "pa" in exI)
   apply (rule_tac x= "na" in exI)
   apply simp
  apply auto
  apply (rule comp_Suc, simp_all)
  done

lemma fixing_swapidseq_decrease:
  assumes "swapidseq n p"
    and "a \ b"
    and "(Fun.swap a b id \ p) a = a"
  shows "n \ 0 \ swapidseq (n - 1) (Fun.swap a b id \ p)"
  using assms
proof (induct n arbitrary: p a b)
  case 0
  then show ?case
    by (auto simp add: Fun.swap_def fun_upd_def)
next
  case (Suc n p a b)
  from Suc.prems(1) swapidseq_cases[of "Suc n" p]
  obtain c d q m where
    cdqm: "Suc n = Suc m" "p = Fun.swap c d id \ q" "swapidseq m q" "c \ d" "n = m"
    by auto
  consider "Fun.swap a b id \ Fun.swap c d id = id"
    | x y z where "x \ a" "y \ a" "z \ a" "x \ y"
      "Fun.swap a b id \ Fun.swap c d id = Fun.swap x y id \ Fun.swap a z id"
    using swap_general[OF Suc.prems(2) cdqm(4)] by metis
  then show ?case
  proof cases
    case 1
    then show ?thesis
      by (simp only: cdqm o_assoc) (simp add: cdqm)
  next
    case prems: 2
    then have az: "a \ z"
      by simp
    from prems have *: "(Fun.swap x y id \ h) a = a \ h a = a" for h
      by (simp add: Fun.swap_def)
    from cdqm(2) have "Fun.swap a b id \ p = Fun.swap a b id \ (Fun.swap c d id \ q)"
      by simp
    then have "Fun.swap a b id \ p = Fun.swap x y id \ (Fun.swap a z id \ q)"
      by (simp add: o_assoc prems)
    then have "(Fun.swap a b id \ p) a = (Fun.swap x y id \ (Fun.swap a z id \ q)) a"
      by simp
    then have "(Fun.swap x y id \ (Fun.swap a z id \ q)) a = a"
      unfolding Suc by metis
    then have "(Fun.swap a z id \ q) a = a"
      by (simp only: *)
    from Suc.hyps[OF cdqm(3)[ unfolded cdqm(5)[symmetric]] az this]
    have **: "swapidseq (n - 1) (Fun.swap a z id \ q)" "n \ 0"
      by blast+
    from \<open>n \<noteq> 0\<close> have ***: "Suc n - 1 = Suc (n - 1)"
      by auto
    show ?thesis
      apply (simp only: cdqm(2) prems o_assoc ***)
      apply (simp only: Suc_not_Zero simp_thms comp_assoc)
      apply (rule comp_Suc)
      using ** prems
       apply blast+
      done
  qed
qed

lemma swapidseq_identity_even:
  assumes "swapidseq n (id :: 'a \ 'a)"
  shows "even n"
  using \<open>swapidseq n id\<close>
proof (induct n rule: nat_less_induct)
  case H: (1 n)
  consider "n = 0"
    | a b :: 'a and q m where "n = Suc m" "id = Fun.swap a b id \ q" "swapidseq m q" "a \ b"
    using H(2)[unfolded swapidseq_cases[of n id]] by auto
  then show ?case
  proof cases
    case 1
    then show ?thesis by presburger
  next
    case h: 2
    from fixing_swapidseq_decrease[OF h(3,4), unfolded h(2)[symmetric]]
    have m: "m \ 0" "swapidseq (m - 1) (id :: 'a \ 'a)"
      by auto
    from h m have mn: "m - 1 < n"
      by arith
    from H(1)[rule_format, OF mn m(2)] h(1) m(1) show ?thesis
      by presburger
  qed
qed


subsection \<open>Therefore we have a welldefined notion of parity\<close>

definition "evenperm p = even (SOME n. swapidseq n p)"

lemma swapidseq_even_even:
  assumes m: "swapidseq m p"
    and n: "swapidseq n p"
  shows "even m \ even n"
proof -
  from swapidseq_inverse_exists[OF n] obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id"
    by blast
  from swapidseq_identity_even[OF swapidseq_comp_add[OF m q(1), unfolded q]] show ?thesis
    by arith
qed

lemma evenperm_unique:
  assumes p: "swapidseq n p"
    and n:"even n = b"
  shows "evenperm p = b"
  unfolding n[symmetric] evenperm_def
  apply (rule swapidseq_even_even[where p = p])
   apply (rule someI[where x = n])
  using p
   apply blast+
  done


subsection \<open>And it has the expected composition properties\<close>

lemma evenperm_id[simp]: "evenperm id = True"
  by (rule evenperm_unique[where n = 0]) simp_all

lemma evenperm_swap: "evenperm (Fun.swap a b id) = (a = b)"
  by (rule evenperm_unique[where n="if a = b then 0 else 1"]) (simp_all add: swapidseq_swap)

lemma evenperm_comp:
  assumes "permutation p" "permutation q"
  shows "evenperm (p \ q) \ evenperm p = evenperm q"
proof -
  from assms obtain n m where n: "swapidseq n p" and m: "swapidseq m q"
    unfolding permutation_def by blast
  have "even (n + m) \ (even n \ even m)"
    by arith
  from evenperm_unique[OF n refl] evenperm_unique[OF m refl]
    and evenperm_unique[OF swapidseq_comp_add[OF n m] this] show ?thesis
    by blast
qed

lemma evenperm_inv:
  assumes "permutation p"
  shows "evenperm (inv p) = evenperm p"
proof -
  from assms obtain n where n: "swapidseq n p"
    unfolding permutation_def by blast
  show ?thesis
    by (rule evenperm_unique[OF swapidseq_inverse[OF n] evenperm_unique[OF n refl, symmetric]])
qed


subsection \<open>A more abstract characterization of permutations\<close>

lemma bij_iff: "bij f \ (\x. \!y. f y = x)"
  unfolding bij_def inj_def surj_def
  apply auto
   apply metis
  apply metis
  done

lemma permutation_bijective:
  assumes "permutation p"
  shows "bij p"
proof -
  from assms obtain n where n: "swapidseq n p"
    unfolding permutation_def by blast
  from swapidseq_inverse_exists[OF n] obtain q where q: "swapidseq n q" "p \ q = id" "q \ p = id"
    by blast
  then show ?thesis
    unfolding bij_iff
    apply (auto simp add: fun_eq_iff)
    apply metis
    done
qed

lemma permutation_finite_support:
  assumes "permutation p"
  shows "finite {x. p x \ x}"
proof -
  from assms obtain n where "swapidseq n p"
    unfolding permutation_def by blast
  then show ?thesis
  proof (induct n p rule: swapidseq.induct)
    case id
    then show ?case by simp
  next
    case (comp_Suc n p a b)
    let ?S = "insert a (insert b {x. p x \ x})"
    from comp_Suc.hyps(2) have *: "finite ?S"
      by simp
    from \<open>a \<noteq> b\<close> have **: "{x. (Fun.swap a b id \<circ> p) x \<noteq> x} \<subseteq> ?S"
      by (auto simp: Fun.swap_def)
    show ?case
      by (rule finite_subset[OF ** *])
  qed
qed

lemma permutation_lemma:
  assumes "finite S"
    and "bij p"
    and "\x. x\ S \ p x = x"
  shows "permutation p"
  using assms
proof (induct S arbitrary: p rule: finite_induct)
  case empty
  then show ?case
    by simp
next
  case (insert a F p)
  let ?r = "Fun.swap a (p a) id \ p"
  let ?q = "Fun.swap a (p a) id \ ?r"
  have *: "?r a = a"
    by (simp add: Fun.swap_def)
  from insert * have **: "\x. x \ F \ ?r x = x"
    by (metis bij_pointE comp_apply id_apply insert_iff swap_apply(3))
  have "bij ?r"
    by (rule bij_swap_compose_bij[OF insert(4)])
  have "permutation ?r"
    by (rule insert(3)[OF \<open>bij ?r\<close> **])
  then have "permutation ?q"
    by (simp add: permutation_compose permutation_swap_id)
  then show ?case
    by (simp add: o_assoc)
qed

lemma permutation: "permutation p \ bij p \ finite {x. p x \ x}"
  (is "?lhs \ ?b \ ?f")
proof
  assume ?lhs
  with permutation_bijective permutation_finite_support show "?b \ ?f"
    by auto
next
  assume "?b \ ?f"
  then have "?f" "?b" by blast+
  from permutation_lemma[OF this] show ?lhs
    by blast
qed

lemma permutation_inverse_works:
  assumes "permutation p"
  shows "inv p \ p = id"
    and "p \ inv p = id"
  using permutation_bijective [OF assms] by (auto simp: bij_def inj_iff surj_iff)

lemma permutation_inverse_compose:
  assumes p: "permutation p"
    and q: "permutation q"
  shows "inv (p \ q) = inv q \ inv p"
proof -
  note ps = permutation_inverse_works[OF p]
  note qs = permutation_inverse_works[OF q]
  have "p \ q \ (inv q \ inv p) = p \ (q \ inv q) \ inv p"
    by (simp add: o_assoc)
  also have "\ = id"
    by (simp add: ps qs)
  finally have *: "p \ q \ (inv q \ inv p) = id" .
  have "inv q \ inv p \ (p \ q) = inv q \ (inv p \ p) \ q"
    by (simp add: o_assoc)
  also have "\ = id"
    by (simp add: ps qs)
  finally have **: "inv q \ inv p \ (p \ q) = id" .
  show ?thesis
    by (rule inv_unique_comp[OF * **])
qed


subsection \<open>Relation to \<open>permutes\<close>\<close>

lemma permutation_permutes: "permutation p \ (\S. finite S \ p permutes S)"
  unfolding permutation permutes_def bij_iff[symmetric]
  apply (rule iffI, clarify)
   apply (rule exI[where x="{x. p x \ x}"])
   apply simp
  apply clarsimp
  apply (rule_tac B="S" in finite_subset)
   apply auto
  done


subsection \<open>Hence a sort of induction principle composing by swaps\<close>

lemma permutes_induct: "finite S \ P id \
  (\<And>a b p. a \<in> S \<Longrightarrow> b \<in> S \<Longrightarrow> P p \<Longrightarrow> P p \<Longrightarrow> permutation p \<Longrightarrow> P (Fun.swap a b id \<circ> p)) \<Longrightarrow>
  (\<And>p. p permutes S \<Longrightarrow> P p)"
proof (induct S rule: finite_induct)
  case empty
  then show ?case by auto
next
  case (insert x F p)
  let ?r = "Fun.swap x (p x) id \ p"
  let ?q = "Fun.swap x (p x) id \ ?r"
  have qp: "?q = p"
    by (simp add: o_assoc)
  from permutes_insert_lemma[OF insert.prems(3)] insert have Pr"P ?r"
    by blast
  from permutes_in_image[OF insert.prems(3), of x]
  have pxF: "p x \ insert x F"
    by simp
  have xF: "x \ insert x F"
    by simp
  have rp: "permutation ?r"
    unfolding permutation_permutes
    using insert.hyps(1) permutes_insert_lemma[OF insert.prems(3)]
    by blast
  from insert.prems(2)[OF xF pxF Pr Pr rp] qp show ?case
    by (simp only:)
qed


subsection \<open>Sign of a permutation as a real number\<close>

definition "sign p = (if evenperm p then (1::int) else -1)"

lemma sign_nz: "sign p \ 0"
  by (simp add: sign_def)

lemma sign_id: "sign id = 1"
  by (simp add: sign_def)

lemma sign_inverse: "permutation p \ sign (inv p) = sign p"
  by (simp add: sign_def evenperm_inv)

lemma sign_compose: "permutation p \ permutation q \ sign (p \ q) = sign p * sign q"
  by (simp add: sign_def evenperm_comp)

lemma sign_swap_id: "sign (Fun.swap a b id) = (if a = b then 1 else -1)"
  by (simp add: sign_def evenperm_swap)

lemma sign_idempotent: "sign p * sign p = 1"
  by (simp add: sign_def)


subsection \<open>Permuting a list\<close>

text \<open>This function permutes a list by applying a permutation to the indices.\<close>

definition permute_list :: "(nat \ nat) \ 'a list \ 'a list"
  where "permute_list f xs = map (\i. xs ! (f i)) [0..

lemma permute_list_map:
  assumes "f permutes {..
  shows "permute_list f (map g xs) = map g (permute_list f xs)"
  using permutes_in_image[OF assms] by (auto simp: permute_list_def)

lemma permute_list_nth:
  assumes "f permutes {.. "i < length xs"
  shows "permute_list f xs ! i = xs ! f i"
  using permutes_in_image[OF assms(1)] assms(2)
  by (simp add: permute_list_def)

lemma permute_list_Nil [simp]: "permute_list f [] = []"
  by (simp add: permute_list_def)

lemma length_permute_list [simp]: "length (permute_list f xs) = length xs"
  by (simp add: permute_list_def)

lemma permute_list_compose:
  assumes "g permutes {..
  shows "permute_list (f \ g) xs = permute_list g (permute_list f xs)"
  using assms[THEN permutes_in_image] by (auto simp add: permute_list_def)

lemma permute_list_ident [simp]: "permute_list (\x. x) xs = xs"
  by (simp add: permute_list_def map_nth)

lemma permute_list_id [simp]: "permute_list id xs = xs"
  by (simp add: id_def)

lemma mset_permute_list [simp]:
  fixes xs :: "'a list"
  assumes "f permutes {..
  shows "mset (permute_list f xs) = mset xs"
proof (rule multiset_eqI)
  fix y :: 'a
  from assms have [simp]: "f x < length xs \ x < length xs" for x
    using permutes_in_image[OF assms] by auto
  have "count (mset (permute_list f xs)) y = card ((\i. xs ! f i) -` {y} \ {..
    by (simp add: permute_list_def count_image_mset atLeast0LessThan)
  also have "(\i. xs ! f i) -` {y} \ {.. y = xs ! i}"
    by auto
  also from assms have "card \ = card {i. i < length xs \ y = xs ! i}"
    by (intro card_vimage_inj) (auto simp: permutes_inj permutes_surj)
  also have "\ = count (mset xs) y"
    by (simp add: count_mset length_filter_conv_card)
  finally show "count (mset (permute_list f xs)) y = count (mset xs) y"
    by simp
qed

lemma set_permute_list [simp]:
  assumes "f permutes {..
  shows "set (permute_list f xs) = set xs"
  by (rule mset_eq_setD[OF mset_permute_list]) fact

lemma distinct_permute_list [simp]:
  assumes "f permutes {..
  shows "distinct (permute_list f xs) = distinct xs"
  by (simp add: distinct_count_atmost_1 assms)

lemma permute_list_zip:
  assumes "f permutes A" "A = {..
  assumes [simp]: "length xs = length ys"
  shows "permute_list f (zip xs ys) = zip (permute_list f xs) (permute_list f ys)"
proof -
  from permutes_in_image[OF assms(1)] assms(2) have *: "f i < length ys \ i < length ys" for i
    by simp
  have "permute_list f (zip xs ys) = map (\i. zip xs ys ! f i) [0..
    by (simp_all add: permute_list_def zip_map_map)
  also have "\ = map (\(x, y). (xs ! f x, ys ! f y)) (zip [0..
    by (intro nth_equalityI) (simp_all add: *)
  also have "\ = zip (permute_list f xs) (permute_list f ys)"
    by (simp_all add: permute_list_def zip_map_map)
  finally show ?thesis .
qed

lemma map_of_permute:
  assumes "\ permutes fst ` set xs"
  shows "map_of xs \ \ = map_of (map (\(x,y). (inv \ x, y)) xs)"
    (is "_ = map_of (map ?f _)")
proof
  from assms have "inj \" "surj \"
    by (simp_all add: permutes_inj permutes_surj)
  then show "(map_of xs \ \) x = map_of (map ?f xs) x" for x
    by (induct xs) (auto simp: inv_f_f surj_f_inv_f)
qed


subsection \<open>More lemmas about permutations\<close>

text \<open>The following few lemmas were contributed by Lukas Bulwahn.\<close>

lemma count_image_mset_eq_card_vimage:
  assumes "finite A"
  shows "count (image_mset f (mset_set A)) b = card {a \ A. f a = b}"
  using assms
proof (induct A)
  case empty
  show ?case by simp
next
  case (insert x F)
  show ?case
  proof (cases "f x = b")
    case True
    with insert.hyps
    have "count (image_mset f (mset_set (insert x F))) b = Suc (card {a \ F. f a = f x})"
      by auto
    also from insert.hyps(1,2) have "\ = card (insert x {a \ F. f a = f x})"
      by simp
    also from \<open>f x = b\<close> have "card (insert x {a \<in> F. f a = f x}) = card {a \<in> insert x F. f a = b}"
      by (auto intro: arg_cong[where f="card"])
    finally show ?thesis
      using insert by auto
  next
    case False
    then have "{a \ F. f a = b} = {a \ insert x F. f a = b}"
      by auto
    with insert False show ?thesis
      by simp
  qed
qed

\<comment> \<open>Prove \<open>image_mset_eq_implies_permutes\<close> ...\<close>
lemma image_mset_eq_implies_permutes:
  fixes f :: "'a \ 'b"
  assumes "finite A"
    and mset_eq: "image_mset f (mset_set A) = image_mset f' (mset_set A)"
  obtains p where "p permutes A" and "\x\A. f x = f' (p x)"
proof -
  from \<open>finite A\<close> have [simp]: "finite {a \<in> A. f a = (b::'b)}" for f b by auto
  have "f ` A = f' ` A"
  proof -
    from \<open>finite A\<close> have "f ` A = f ` (set_mset (mset_set A))"
      by simp
    also have "\ = f' ` set_mset (mset_set A)"
      by (metis mset_eq multiset.set_map)
    also from \<open>finite A\<close> have "\<dots> = f' ` A"
      by simp
    finally show ?thesis .
  qed
  have "\b\(f ` A). \p. bij_betw p {a \ A. f a = b} {a \ A. f' a = b}"
  proof
    fix b
    from mset_eq have "count (image_mset f (mset_set A)) b = count (image_mset f' (mset_set A)) b"
      by simp
    with \<open>finite A\<close> have "card {a \<in> A. f a = b} = card {a \<in> A. f' a = b}"
      by (simp add: count_image_mset_eq_card_vimage)
    then show "\p. bij_betw p {a\A. f a = b} {a \ A. f' a = b}"
      by (intro finite_same_card_bij) simp_all
  qed
  then have "\p. \b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}"
    by (rule bchoice)
  then obtain p where p: "\b\f ` A. bij_betw (p b) {a \ A. f a = b} {a \ A. f' a = b}" ..
  define p' where "p' = (\<lambda>a. if a \<in> A then p (f a) a else a)"
  have "p' permutes A"
  proof (rule bij_imp_permutes)
    have "disjoint_family_on (\i. {a \ A. f' a = i}) (f ` A)"
      by (auto simp: disjoint_family_on_def)
    moreover
    have "bij_betw (\a. p (f a) a) {a \ A. f a = b} {a \ A. f' a = b}" if "b \ f ` A" for b
      using p that by (subst bij_betw_cong[where g="p b"]) auto
    ultimately
    have "bij_betw (\a. p (f a) a) (\b\f ` A. {a \ A. f a = b}) (\b\f ` A. {a \ A. f' a = b})"
      by (rule bij_betw_UNION_disjoint)
    moreover have "(\b\f ` A. {a \ A. f a = b}) = A"
      by auto
    moreover from \<open>f ` A = f' ` A\<close> have "(\<Union>b\<in>f ` A. {a \<in> A. f' a = b}) = A"
      by auto
    ultimately show "bij_betw p' A A"
      unfolding p'_def by (subst bij_betw_cong[where g="(\a. p (f a) a)"]) auto
  next
    show "\x. x \ A \ p' x = x"
      by (simp add: p'_def)
  qed
  moreover from p have "\x\A. f x = f' (p' x)"
    unfolding p'_def using bij_betwE by fastforce
  ultimately show ?thesis ..
qed

lemma mset_set_upto_eq_mset_upto: "mset_set {..
  by (induct n) (auto simp: add.commute lessThan_Suc)

\<comment> \<open>... and derive the existing property:\<close>
lemma mset_eq_permutation:
  fixes xs ys :: "'a list"
  assumes mset_eq: "mset xs = mset ys"
  obtains p where "p permutes {.. "permute_list p ys = xs"
proof -
  from mset_eq have length_eq: "length xs = length ys"
    by (rule mset_eq_length)
  have "mset_set {..
    by (rule mset_set_upto_eq_mset_upto)
  with mset_eq length_eq have "image_mset (\i. xs ! i) (mset_set {..
    image_mset (\<lambda>i. ys ! i) (mset_set {..<length ys})"
    by (metis map_nth mset_map)
  from image_mset_eq_implies_permutes[OF _ this]
  obtain p where p: "p permutes {.. and "\i\{..
    by auto
  with length_eq have "permute_list p ys = xs"
    by (auto intro!: nth_equalityI simp: permute_list_nth)
  with p show thesis ..
qed

lemma permutes_natset_le:
  fixes S :: "'a::wellorder set"
  assumes "p permutes S"
    and "\i \ S. p i \ i"
  shows "p = id"
proof -
  have "p n = n" for n
    using assms
  proof (induct n arbitrary: S rule: less_induct)
    case (less n)
    show ?case
    proof (cases "n \ S")
      case False
      with less(2) show ?thesis
        unfolding permutes_def by metis
    next
      case True
      with less(3) have "p n < n \ p n = n"
        by auto
      then show ?thesis
      proof
        assume "p n < n"
        with less have "p (p n) = p n"
          by metis
        with permutes_inj[OF less(2)] have "p n = n"
          unfolding inj_def by blast
        with \<open>p n < n\<close> have False
          by simp
        then show ?thesis ..
      qed
    qed
  qed
  then show ?thesis by (auto simp: fun_eq_iff)
qed

lemma permutes_natset_ge:
  fixes S :: "'a::wellorder set"
  assumes p: "p permutes S"
    and le: "\i \ S. p i \ i"
  shows "p = id"
proof -
  have "i \ inv p i" if "i \ S" for i
  proof -
    from that permutes_in_image[OF permutes_inv[OF p]] have "inv p i \ S"
      by simp
    with le have "p (inv p i) \ inv p i"
      by blast
    with permutes_inverses[OF p] show ?thesis
      by simp
  qed
  then have "\i\S. inv p i \ i"
    by blast
  from permutes_natset_le[OF permutes_inv[OF p] this] have "inv p = inv id"
    by simp
  then show ?thesis
    apply (subst permutes_inv_inv[OF p, symmetric])
    apply (rule inv_unique_comp)
     apply simp_all
    done
qed

lemma image_inverse_permutations: "{inv p |p. p permutes S} = {p. p permutes S}"
  apply (rule set_eqI)
  apply auto
  using permutes_inv_inv permutes_inv
   apply auto
  apply (rule_tac x="inv x" in exI)
  apply auto
  done

lemma image_compose_permutations_left:
  assumes "q permutes S"
  shows "{q \ p |p. p permutes S} = {p. p permutes S}"
  apply (rule set_eqI)
  apply auto
   apply (rule permutes_compose)
  using assms
    apply auto
  apply (rule_tac x = "inv q \ x" in exI)
  apply (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o)
  done

lemma image_compose_permutations_right:
  assumes "q permutes S"
  shows "{p \ q | p. p permutes S} = {p . p permutes S}"
  apply (rule set_eqI)
  apply auto
   apply (rule permutes_compose)
  using assms
    apply auto
  apply (rule_tac x = "x \ inv q" in exI)
  apply (simp add: o_assoc permutes_inv permutes_compose permutes_inv_o comp_assoc)
  done

lemma permutes_in_seg: "p permutes {1 ..n} \ i \ {1..n} \ 1 \ p i \ p i \ n"
  by (simp add: permutes_def) metis

lemma sum_permutations_inverse: "sum f {p. p permutes S} = sum (\p. f(inv p)) {p. p permutes S}"
  (is "?lhs = ?rhs")
proof -
  let ?S = "{p . p permutes S}"
  have *: "inj_on inv ?S"
  proof (auto simp add: inj_on_def)
    fix q r
    assume q: "q permutes S"
      and r: "r permutes S"
      and qr: "inv q = inv r"
    then have "inv (inv q) = inv (inv r)"
      by simp
    with permutes_inv_inv[OF q] permutes_inv_inv[OF r] show "q = r"
      by metis
  qed
  have **: "inv ` ?S = ?S"
    using image_inverse_permutations by blast
  have ***: "?rhs = sum (f \ inv) ?S"
    by (simp add: o_def)
  from sum.reindex[OF *, of f] show ?thesis
    by (simp only: ** ***)
qed

lemma setum_permutations_compose_left:
  assumes q: "q permutes S"
  shows "sum f {p. p permutes S} = sum (\p. f(q \ p)) {p. p permutes S}"
  (is "?lhs = ?rhs")
proof -
  let ?S = "{p. p permutes S}"
  have *: "?rhs = sum (f \ ((\) q)) ?S"
    by (simp add: o_def)
  have **: "inj_on ((\) q) ?S"
  proof (auto simp add: inj_on_def)
    fix p r
    assume "p permutes S"
      and r: "r permutes S"
      and rp: "q \ p = q \ r"
    then have "inv q \ q \ p = inv q \ q \ r"
      by (simp add: comp_assoc)
    with permutes_inj[OF q, unfolded inj_iff] show "p = r"
      by simp
  qed
  have "((\) q) ` ?S = ?S"
    using image_compose_permutations_left[OF q] by auto
  with * sum.reindex[OF **, of f] show ?thesis
    by (simp only:)
qed

lemma sum_permutations_compose_right:
  assumes q: "q permutes S"
  shows "sum f {p. p permutes S} = sum (\p. f(p \ q)) {p. p permutes S}"
  (is "?lhs = ?rhs")
proof -
  let ?S = "{p. p permutes S}"
  have *: "?rhs = sum (f \ (\p. p \ q)) ?S"
    by (simp add: o_def)
  have **: "inj_on (\p. p \ q) ?S"
  proof (auto simp add: inj_on_def)
    fix p r
    assume "p permutes S"
      and r: "r permutes S"
      and rp: "p \ q = r \ q"
    then have "p \ (q \ inv q) = r \ (q \ inv q)"
      by (simp add: o_assoc)
    with permutes_surj[OF q, unfolded surj_iff] show "p = r"
      by simp
  qed
  from image_compose_permutations_right[OF q] have "(\p. p \ q) ` ?S = ?S"
    by auto
  with * sum.reindex[OF **, of f] show ?thesis
    by (simp only:)
qed


subsection \<open>Sum over a set of permutations (could generalize to iteration)\<close>

lemma sum_over_permutations_insert:
  assumes fS: "finite S"
    and aS: "a \ S"
  shows "sum f {p. p permutes (insert a S)} =
    sum (\<lambda>b. sum (\<lambda>q. f (Fun.swap a b id \<circ> q)) {p. p permutes S}) (insert a S)"
proof -
  have *: "\f a b. (\(b, p). f (Fun.swap a b id \ p)) = f \ (\(b,p). Fun.swap a b id \ p)"
    by (simp add: fun_eq_iff)
  have **: "\P Q. {(a, b). a \ P \ b \ Q} = P \ Q"
    by blast
  show ?thesis
    unfolding * ** sum.cartesian_product permutes_insert
  proof (rule sum.reindex)
    let ?f = "(\(b, y). Fun.swap a b id \ y)"
    let ?P = "{p. p permutes S}"
    {
      fix b c p q
      assume b: "b \ insert a S"
      assume c: "c \ insert a S"
      assume p: "p permutes S"
      assume q: "q permutes S"
      assume eq: "Fun.swap a b id \ p = Fun.swap a c id \ q"
      from p q aS have pa: "p a = a" and qa: "q a = a"
        unfolding permutes_def by metis+
      from eq have "(Fun.swap a b id \ p) a = (Fun.swap a c id \ q) a"
        by simp
      then have bc: "b = c"
        by (simp add: permutes_def pa qa o_def fun_upd_def Fun.swap_def id_def
            cong del: if_weak_cong split: if_split_asm)
      from eq[unfolded bc] have "(\p. Fun.swap a c id \ p) (Fun.swap a c id \ p) =
        (\<lambda>p. Fun.swap a c id \<circ> p) (Fun.swap a c id \<circ> q)" by simp
      then have "p = q"
        unfolding o_assoc swap_id_idempotent by simp
      with bc have "b = c \ p = q"
        by blast
    }
    then show "inj_on ?f (insert a S \ ?P)"
      unfolding inj_on_def by clarify metis
  qed
qed


subsection \<open>Constructing permutations from association lists\<close>

definition list_permutes :: "('a \ 'a) list \ 'a set \ bool"
  where "list_permutes xs A \
    set (map fst xs) \<subseteq> A \<and>
    set (map snd xs) = set (map fst xs) \<and>
    distinct (map fst xs) \<and>
    distinct (map snd xs)"

lemma list_permutesI [simp]:
  assumes "set (map fst xs) \ A" "set (map snd xs) = set (map fst xs)" "distinct (map fst xs)"
  shows "list_permutes xs A"
proof -
  from assms(2,3) have "distinct (map snd xs)"
    by (intro card_distinct) (simp_all add: distinct_card del: set_map)
  with assms show ?thesis
    by (simp add: list_permutes_def)
qed

definition permutation_of_list :: "('a \ 'a) list \ 'a \ 'a"
  where "permutation_of_list xs x = (case map_of xs x of None \ x | Some y \ y)"

lemma permutation_of_list_Cons:
  "permutation_of_list ((x, y) # xs) x' = (if x = x' then y else permutation_of_list xs x')"
  by (simp add: permutation_of_list_def)

fun inverse_permutation_of_list :: "('a \ 'a) list \ 'a \ 'a"
  where
    "inverse_permutation_of_list [] x = x"
  | "inverse_permutation_of_list ((y, x') # xs) x =
      (if x = x' then y else inverse_permutation_of_list xs x)"

declare inverse_permutation_of_list.simps [simp del]

lemma inj_on_map_of:
  assumes "distinct (map snd xs)"
  shows "inj_on (map_of xs) (set (map fst xs))"
proof (rule inj_onI)
  fix x y
  assume xy: "x \ set (map fst xs)" "y \ set (map fst xs)"
  assume eq: "map_of xs x = map_of xs y"
  from xy obtain x' y' where x'y'"map_of xs x = Some x'" "map_of xs y = Some y'"
    by (cases "map_of xs x"; cases "map_of xs y") (simp_all add: map_of_eq_None_iff)
  moreover from x'y' have *: "(x, x') \ set xs" "(y, y') \ set xs"
    by (force dest: map_of_SomeD)+
  moreover from * eq x'y' have "x' = y'"
    by simp
  ultimately show "x = y"
    using assms by (force simp: distinct_map dest: inj_onD[of _ _ "(x,x')" "(y,y')"])
qed

lemma inj_on_the: "None \ A \ inj_on the A"
  by (auto simp: inj_on_def option.the_def split: option.splits)

lemma inj_on_map_of':
  assumes "distinct (map snd xs)"
  shows "inj_on (the \ map_of xs) (set (map fst xs))"
  by (intro comp_inj_on inj_on_map_of assms inj_on_the)
    (force simp: eq_commute[of None] map_of_eq_None_iff)

lemma image_map_of:
  assumes "distinct (map fst xs)"
  shows "map_of xs ` set (map fst xs) = Some ` set (map snd xs)"
  using assms by (auto simp: rev_image_eqI)

lemma the_Some_image [simp]: "the ` Some ` A = A"
  by (subst image_image) simp

lemma image_map_of':
  assumes "distinct (map fst xs)"
  shows "(the \ map_of xs) ` set (map fst xs) = set (map snd xs)"
  by (simp only: image_comp [symmetric] image_map_of assms the_Some_image)

lemma permutation_of_list_permutes [simp]:
  assumes "list_permutes xs A"
  shows "permutation_of_list xs permutes A"
    (is "?f permutes _")
proof (rule permutes_subset[OF bij_imp_permutes])
  from assms show "set (map fst xs) \ A"
    by (simp add: list_permutes_def)
  from assms have "inj_on (the \ map_of xs) (set (map fst xs))" (is ?P)
    by (intro inj_on_map_of') (simp_all add: list_permutes_def)
  also have "?P \ inj_on ?f (set (map fst xs))"
    by (intro inj_on_cong)
      (auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits)
  finally have "bij_betw ?f (set (map fst xs)) (?f ` set (map fst xs))"
    by (rule inj_on_imp_bij_betw)
  also from assms have "?f ` set (map fst xs) = (the \ map_of xs) ` set (map fst xs)"
    by (intro image_cong refl)
      (auto simp: permutation_of_list_def map_of_eq_None_iff split: option.splits)
  also from assms have "\ = set (map fst xs)"
    by (subst image_map_of') (simp_all add: list_permutes_def)
  finally show "bij_betw ?f (set (map fst xs)) (set (map fst xs))" .
qed (force simp: permutation_of_list_def dest!: map_of_SomeD split: option.splits)+

lemma eval_permutation_of_list [simp]:
  "permutation_of_list [] x = x"
  "x = x' \ permutation_of_list ((x',y)#xs) x = y"
  "x \ x' \ permutation_of_list ((x',y')#xs) x = permutation_of_list xs x"
  by (simp_all add: permutation_of_list_def)

lemma eval_inverse_permutation_of_list [simp]:
  "inverse_permutation_of_list [] x = x"
  "x = x' \ inverse_permutation_of_list ((y,x')#xs) x = y"
  "x \ x' \ inverse_permutation_of_list ((y',x')#xs) x = inverse_permutation_of_list xs x"
  by (simp_all add: inverse_permutation_of_list.simps)

lemma permutation_of_list_id: "x \ set (map fst xs) \ permutation_of_list xs x = x"
  by (induct xs) (auto simp: permutation_of_list_Cons)

lemma permutation_of_list_unique':
  "distinct (map fst xs) \ (x, y) \ set xs \ permutation_of_list xs x = y"
  by (induct xs) (force simp: permutation_of_list_Cons)+

lemma permutation_of_list_unique:
  "list_permutes xs A \ (x, y) \ set xs \ permutation_of_list xs x = y"
  by (intro permutation_of_list_unique') (simp_all add: list_permutes_def)

lemma inverse_permutation_of_list_id:
  "x \ set (map snd xs) \ inverse_permutation_of_list xs x = x"
  by (induct xs) auto

lemma inverse_permutation_of_list_unique':
  "distinct (map snd xs) \ (x, y) \ set xs \ inverse_permutation_of_list xs y = x"
  by (induct xs) (force simp: inverse_permutation_of_list.simps)+

lemma inverse_permutation_of_list_unique:
  "list_permutes xs A \ (x,y) \ set xs \ inverse_permutation_of_list xs y = x"
  by (intro inverse_permutation_of_list_unique') (simp_all add: list_permutes_def)

lemma inverse_permutation_of_list_correct:
  fixes A :: "'a set"
  assumes "list_permutes xs A"
  shows "inverse_permutation_of_list xs = inv (permutation_of_list xs)"
proof (rule ext, rule sym, subst permutes_inv_eq)
  from assms show "permutation_of_list xs permutes A"
    by simp
  show "permutation_of_list xs (inverse_permutation_of_list xs x) = x" for x
  proof (cases "x \ set (map snd xs)")
    case True
    then obtain y where "(y, x) \ set xs" by auto
    with assms show ?thesis
      by (simp add: inverse_permutation_of_list_unique permutation_of_list_unique)
  next
    case False
    with assms show ?thesis
      by (auto simp: list_permutes_def inverse_permutation_of_list_id permutation_of_list_id)
  qed
qed

end

¤ Dauer der Verarbeitung: 0.93 Sekunden  (vorverarbeitet)  ¤





Download des
Quellennavigators
Download des
sprechenden Kalenders

in der Quellcodebibliothek suchen




Haftungshinweis

Die Informationen auf dieser Webseite wurden nach bestem Wissen sorgfältig zusammengestellt. Es wird jedoch weder Vollständigkeit, noch Richtigkeit, noch Qualität der bereit gestellten Informationen zugesichert.


Bemerkung:

Die farbliche Syntaxdarstellung ist noch experimentell.


Bot Zugriff