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: Permutation.thy   Sprache: Isabelle

Original von: Isabelle©

(*  Title:      HOL/Library/Permutation.thy
    Author:     Lawrence C Paulson and Thomas M Rasmussen and Norbert Voelker
*)


section \<open>Permutations\<close>

theory Permutation
imports Multiset
begin

inductive perm :: "'a list \ 'a list \ bool" (infixr \<~~>\ 50)
where
  Nil [intro!]: "[] <~~> []"
| swap [intro!]: "y # x # l <~~> x # y # l"
| Cons [intro!]: "xs <~~> ys \ z # xs <~~> z # ys"
| trans [intro]: "xs <~~> ys \ ys <~~> zs \ xs <~~> zs"

proposition perm_refl [iff]: "l <~~> l"
  by (induct l) auto


subsection \<open>Some examples of rule induction on permutations\<close>

proposition perm_empty_imp: "[] <~~> ys \ ys = []"
  by (induction "[] :: 'a list" ys pred: perm) simp_all


text \<open>\medskip This more general theorem is easier to understand!\<close>

proposition perm_length: "xs <~~> ys \ length xs = length ys"
  by (induct pred: perm) simp_all

proposition perm_sym: "xs <~~> ys \ ys <~~> xs"
  by (induct pred: perm) auto


subsection \<open>Ways of making new permutations\<close>

text \<open>We can insert the head anywhere in the list.\<close>

proposition perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys"
  by (induct xs) auto

proposition perm_append_swap: "xs @ ys <~~> ys @ xs"
  by (induct xs) (auto intro: perm_append_Cons)

proposition perm_append_single: "a # xs <~~> xs @ [a]"
  by (rule perm.trans [OF _ perm_append_swap]) simp

proposition perm_rev: "rev xs <~~> xs"
  by (induct xs) (auto intro!: perm_append_single intro: perm_sym)

proposition perm_append1: "xs <~~> ys \ l @ xs <~~> l @ ys"
  by (induct l) auto

proposition perm_append2: "xs <~~> ys \ xs @ l <~~> ys @ l"
  by (blast intro!: perm_append_swap perm_append1)


subsection \<open>Further results\<close>

proposition perm_empty [iff]: "[] <~~> xs \ xs = []"
  by (blast intro: perm_empty_imp)

proposition perm_empty2 [iff]: "xs <~~> [] \ xs = []"
  using perm_sym by auto

proposition perm_sing_imp: "ys <~~> xs \ xs = [y] \ ys = [y]"
  by (induct pred: perm) auto

proposition perm_sing_eq [iff]: "ys <~~> [y] \ ys = [y]"
  by (blast intro: perm_sing_imp)

proposition perm_sing_eq2 [iff]: "[y] <~~> ys \ ys = [y]"
  by (blast dest: perm_sym)


subsection \<open>Removing elements\<close>

proposition perm_remove: "x \ set ys \ ys <~~> x # remove1 x ys"
  by (induct ys) auto


text \<open>\medskip Congruence rule\<close>

proposition perm_remove_perm: "xs <~~> ys \ remove1 z xs <~~> remove1 z ys"
  by (induct pred: perm) auto

proposition remove_hd [simp]: "remove1 z (z # xs) = xs"
  by auto

proposition cons_perm_imp_perm: "z # xs <~~> z # ys \ xs <~~> ys"
  by (drule perm_remove_perm [where z = z]) auto

proposition cons_perm_eq [iff]: "z#xs <~~> z#ys \ xs <~~> ys"
  by (meson cons_perm_imp_perm perm.Cons)

proposition append_perm_imp_perm: "zs @ xs <~~> zs @ ys \ xs <~~> ys"
  by (induct zs arbitrary: xs ys rule: rev_induct) auto

proposition perm_append1_eq [iff]: "zs @ xs <~~> zs @ ys \ xs <~~> ys"
  by (blast intro: append_perm_imp_perm perm_append1)

proposition perm_append2_eq [iff]: "xs @ zs <~~> ys @ zs \ xs <~~> ys"
  by (meson perm.trans perm_append1_eq perm_append_swap)

theorem mset_eq_perm: "mset xs = mset ys \ xs <~~> ys"
proof
  assume "mset xs = mset ys"
  then show "xs <~~> ys"
  proof (induction xs arbitrary: ys)
    case (Cons x xs)
    then have "x \ set ys"
      using mset_eq_setD by fastforce
    then show ?case
      by (metis Cons.IH Cons.prems mset_remove1 perm.Cons perm.trans perm_remove perm_sym remove_hd)
  qed auto
next
  assume "xs <~~> ys"
  then show "mset xs = mset ys"
    by induction (simp_all add: union_ac)
qed

proposition mset_le_perm_append: "mset xs \# mset ys \ (\zs. xs @ zs <~~> ys)"
  apply (rule iffI)
  apply (metis mset_append mset_eq_perm mset_subset_eq_exists_conv surjD surj_mset)
  by (metis mset_append mset_eq_perm mset_subset_eq_exists_conv)

proposition perm_set_eq: "xs <~~> ys \ set xs = set ys"
  by (metis mset_eq_perm mset_eq_setD)

proposition perm_distinct_iff: "xs <~~> ys \ distinct xs = distinct ys"
  by (metis card_distinct distinct_card perm_length perm_set_eq)

theorem eq_set_perm_remdups: "set xs = set ys \ remdups xs <~~> remdups ys"
proof (induction xs arbitrary: ys rule: length_induct)
  case (1 xs)
  show ?case
  proof (cases "remdups xs")
    case Nil
    with "1.prems" show ?thesis
      using "1.prems" by auto
  next
    case (Cons x us)
    then have "x \ set (remdups ys)"
      using "1.prems" set_remdups by fastforce
    then show ?thesis
      using "1.prems" mset_eq_perm set_eq_iff_mset_remdups_eq by blast
  qed
qed 

proposition perm_remdups_iff_eq_set: "remdups x <~~> remdups y \ set x = set y"
  by (metis List.set_remdups perm_set_eq eq_set_perm_remdups)

theorem permutation_Ex_bij:
  assumes "xs <~~> ys"
  shows "\f. bij_betw f {.. (\i
  using assms
proof induct
  case Nil
  then show ?case
    unfolding bij_betw_def by simp
next
  case (swap y x l)
  show ?case
  proof (intro exI[of _ "Fun.swap 0 1 id"] conjI allI impI)
    show "bij_betw (Fun.swap 0 1 id) {..
      by (auto simp: bij_betw_def)
    fix i
    assume "i < length (y # x # l)"
    show "(y # x # l) ! i = (x # y # l) ! (Fun.swap 0 1 id) i"
      by (cases i) (auto simp: Fun.swap_def gr0_conv_Suc)
  qed
next
  case (Cons xs ys z)
  then obtain f where bij: "bij_betw f {..
    and perm: "\i
    by blast
  let ?f = "\i. case i of Suc n \ Suc (f n) | 0 \ 0"
  show ?case
  proof (intro exI[of _ ?f] allI conjI impI)
    have *: "{.. Suc ` {..
            "{.. Suc ` {..
      by (simp_all add: lessThan_Suc_eq_insert_0)
    show "bij_betw ?f {..
      unfolding *
    proof (rule bij_betw_combine)
      show "bij_betw ?f (Suc ` {..
        using bij unfolding bij_betw_def
        by (auto intro!: inj_onI imageI dest: inj_onD simp: image_comp comp_def)
    qed (auto simp: bij_betw_def)
    fix i
    assume "i < length (z # xs)"
    then show "(z # xs) ! i = (z # ys) ! (?f i)"
      using perm by (cases i) auto
  qed
next
  case (trans xs ys zs)
  then obtain f g
    where bij: "bij_betw f {.. "bij_betw g {..
    and perm: "\ii
    by blast
  show ?case
  proof (intro exI[of _ "g \ f"] conjI allI impI)
    show "bij_betw (g \ f) {..
      using bij by (rule bij_betw_trans)
    fix i
    assume "i < length xs"
    with bij have "f i < length ys"
      unfolding bij_betw_def by force
    with \<open>i < length xs\<close> show "xs ! i = zs ! (g \<circ> f) i"
      using trans(1,3)[THEN perm_length] perm by auto
  qed
qed

proposition perm_finite: "finite {B. B <~~> A}"
proof (rule finite_subset[where B="{xs. set xs \ set A \ length xs \ length A}"])
 show "finite {xs. set xs \ set A \ length xs \ length A}"
   using finite_lists_length_le by blast
next
 show "{B. B <~~> A} \ {xs. set xs \ set A \ length xs \ length A}"
   by (clarsimp simp add: perm_length perm_set_eq)
qed

proposition perm_swap:
    assumes "i < length xs" "j < length xs"
    shows "xs[i := xs ! j, j := xs ! i] <~~> xs"
  using assms by (simp add: mset_eq_perm[symmetric] mset_swap)

end

¤ Dauer der Verarbeitung: 0.14 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