src/HOL/Library/Permutation.thy
changeset 11054 a5404c70982f
child 11153 950ede59c05a
equal deleted inserted replaced
11053:026007eb2ccc 11054:a5404c70982f
       
     1 (*  Title:      HOL/Library/Permutation.thy
       
     2     ID:         $Id$
       
     3     Author:     Lawrence C Paulson, Cambridge University Computer Laboratory
       
     4     Copyright   1995  University of Cambridge
       
     5 
       
     6 TODO: it would be nice to prove (for "multiset", defined on
       
     7 HOL/ex/Sorting.thy) xs <~~> ys = (\<forall>x. multiset xs x = multiset ys x)
       
     8 *)
       
     9 
       
    10 header {*
       
    11  \title{Permutations}
       
    12  \author{Lawrence C Paulson and Thomas M Rasmussen}
       
    13 *}
       
    14 
       
    15 theory Permutation = Main:
       
    16 
       
    17 consts
       
    18   perm :: "('a list * 'a list) set"
       
    19 
       
    20 syntax
       
    21   "_perm" :: "'a list => 'a list => bool"    ("_ <~~> _"  [50, 50] 50)
       
    22 translations
       
    23   "x <~~> y" == "(x, y) \<in> perm"
       
    24 
       
    25 inductive perm
       
    26   intros [intro]
       
    27     Nil: "[] <~~> []"
       
    28     swap: "y # x # l <~~> x # y # l"
       
    29     Cons: "xs <~~> ys ==> z # xs <~~> z # ys"
       
    30     trans: "xs <~~> ys ==> ys <~~> zs ==> xs <~~> zs"
       
    31 
       
    32 lemma perm_refl [iff]: "l <~~> l"
       
    33   apply (induct l)
       
    34    apply auto
       
    35   done
       
    36 
       
    37 
       
    38 subsection {* Some examples of rule induction on permutations *}
       
    39 
       
    40 lemma xperm_empty_imp_aux: "xs <~~> ys ==> xs = [] --> ys = []"
       
    41     -- {* the form of the premise lets the induction bind @{term xs} and @{term ys} *}
       
    42   apply (erule perm.induct)
       
    43      apply (simp_all (no_asm_simp))
       
    44   done
       
    45 
       
    46 lemma xperm_empty_imp: "[] <~~> ys ==> ys = []"
       
    47   apply (insert xperm_empty_imp_aux)
       
    48   apply blast
       
    49   done
       
    50 
       
    51 
       
    52 text {*
       
    53   \medskip This more general theorem is easier to understand!
       
    54   *}
       
    55 
       
    56 lemma perm_length: "xs <~~> ys ==> length xs = length ys"
       
    57   apply (erule perm.induct)
       
    58      apply simp_all
       
    59   done
       
    60 
       
    61 lemma perm_empty_imp: "[] <~~> xs ==> xs = []"
       
    62   apply (drule perm_length)
       
    63   apply auto
       
    64   done
       
    65 
       
    66 lemma perm_sym: "xs <~~> ys ==> ys <~~> xs"
       
    67   apply (erule perm.induct)
       
    68      apply auto
       
    69   done
       
    70 
       
    71 lemma perm_mem [rule_format]: "xs <~~> ys ==> x mem xs --> x mem ys"
       
    72   apply (erule perm.induct)
       
    73      apply auto
       
    74   done
       
    75 
       
    76 
       
    77 subsection {* Ways of making new permutations *}
       
    78 
       
    79 text {*
       
    80   We can insert the head anywhere in the list.
       
    81 *}
       
    82 
       
    83 lemma perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys"
       
    84   apply (induct xs)
       
    85    apply auto
       
    86   done
       
    87 
       
    88 lemma perm_append_swap: "xs @ ys <~~> ys @ xs"
       
    89   apply (induct xs)
       
    90     apply simp_all
       
    91   apply (blast intro: perm_append_Cons)
       
    92   done
       
    93 
       
    94 lemma perm_append_single: "a # xs <~~> xs @ [a]"
       
    95   apply (rule perm.trans)
       
    96    prefer 2
       
    97    apply (rule perm_append_swap)
       
    98   apply simp
       
    99   done
       
   100 
       
   101 lemma perm_rev: "rev xs <~~> xs"
       
   102   apply (induct xs)
       
   103    apply simp_all
       
   104   apply (blast intro: perm_sym perm_append_single)
       
   105   done
       
   106 
       
   107 lemma perm_append1: "xs <~~> ys ==> l @ xs <~~> l @ ys"
       
   108   apply (induct l)
       
   109    apply auto
       
   110   done
       
   111 
       
   112 lemma perm_append2: "xs <~~> ys ==> xs @ l <~~> ys @ l"
       
   113   apply (blast intro!: perm_append_swap perm_append1)
       
   114   done
       
   115 
       
   116 
       
   117 subsection {* Further results *}
       
   118 
       
   119 lemma perm_empty [iff]: "([] <~~> xs) = (xs = [])"
       
   120   apply (blast intro: perm_empty_imp)
       
   121   done
       
   122 
       
   123 lemma perm_empty2 [iff]: "(xs <~~> []) = (xs = [])"
       
   124   apply auto
       
   125   apply (erule perm_sym [THEN perm_empty_imp])
       
   126   done
       
   127 
       
   128 lemma perm_sing_imp [rule_format]: "ys <~~> xs ==> xs = [y] --> ys = [y]"
       
   129   apply (erule perm.induct)
       
   130      apply auto
       
   131   done
       
   132 
       
   133 lemma perm_sing_eq [iff]: "(ys <~~> [y]) = (ys = [y])"
       
   134   apply (blast intro: perm_sing_imp)
       
   135   done
       
   136 
       
   137 lemma perm_sing_eq2 [iff]: "([y] <~~> ys) = (ys = [y])"
       
   138   apply (blast dest: perm_sym)
       
   139   done
       
   140 
       
   141 
       
   142 subsection {* Removing elements *}
       
   143 
       
   144 consts
       
   145   remove :: "'a => 'a list => 'a list"
       
   146 primrec
       
   147   "remove x [] = []"
       
   148   "remove x (y # ys) = (if x = y then ys else y # remove x ys)"
       
   149 
       
   150 lemma perm_remove: "x \<in> set ys ==> ys <~~> x # remove x ys"
       
   151   apply (induct ys)
       
   152    apply auto
       
   153   done
       
   154 
       
   155 lemma remove_commute: "remove x (remove y l) = remove y (remove x l)"
       
   156   apply (induct l)
       
   157    apply auto
       
   158   done
       
   159 
       
   160 
       
   161 text {* \medskip Congruence rule *}
       
   162 
       
   163 lemma perm_remove_perm: "xs <~~> ys ==> remove z xs <~~> remove z ys"
       
   164   apply (erule perm.induct)
       
   165      apply auto
       
   166   done
       
   167 
       
   168 lemma remove_hd [simp]: "remove z (z # xs) = xs"
       
   169   apply auto
       
   170   done
       
   171 
       
   172 lemma cons_perm_imp_perm: "z # xs <~~> z # ys ==> xs <~~> ys"
       
   173   apply (drule_tac z = z in perm_remove_perm)
       
   174   apply auto
       
   175   done
       
   176 
       
   177 lemma cons_perm_eq [iff]: "(z#xs <~~> z#ys) = (xs <~~> ys)"
       
   178   apply (blast intro: cons_perm_imp_perm)
       
   179   done
       
   180 
       
   181 lemma append_perm_imp_perm: "!!xs ys. zs @ xs <~~> zs @ ys ==> xs <~~> ys"
       
   182   apply (induct zs rule: rev_induct)
       
   183    apply (simp_all (no_asm_use))
       
   184   apply blast
       
   185   done
       
   186 
       
   187 lemma perm_append1_eq [iff]: "(zs @ xs <~~> zs @ ys) = (xs <~~> ys)"
       
   188   apply (blast intro: append_perm_imp_perm perm_append1)
       
   189   done
       
   190 
       
   191 lemma perm_append2_eq [iff]: "(xs @ zs <~~> ys @ zs) = (xs <~~> ys)"
       
   192   apply (safe intro!: perm_append2)
       
   193   apply (rule append_perm_imp_perm)
       
   194   apply (rule perm_append_swap [THEN perm.trans])
       
   195     -- {* the previous step helps this @{text blast} call succeed quickly *}
       
   196   apply (blast intro: perm_append_swap)
       
   197   done
       
   198 
       
   199 end