src/HOL/Library/List_Permutation.thy
author wenzelm
Sun, 14 Mar 2021 22:55:52 +0100
changeset 73439 cb127ce2c092
parent 73329 2aef2de6b17c
permissions -rw-r--r--
tuned --- following hints by IntelliJ;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
73297
beaff25452d2 more specific name
haftmann
parents: 70680
diff changeset
     1
(*  Title:      HOL/Library/List_Permutation.thy
15005
546c8e7e28d4 Norbert Voelker
paulson
parents: 14706
diff changeset
     2
    Author:     Lawrence C Paulson and Thomas M Rasmussen and Norbert Voelker
11054
a5404c70982f moved from Induct/ to Library/
wenzelm
parents:
diff changeset
     3
*)
a5404c70982f moved from Induct/ to Library/
wenzelm
parents:
diff changeset
     4
73297
beaff25452d2 more specific name
haftmann
parents: 70680
diff changeset
     5
section \<open>Permuted Lists\<close>
11054
a5404c70982f moved from Induct/ to Library/
wenzelm
parents:
diff changeset
     6
73297
beaff25452d2 more specific name
haftmann
parents: 70680
diff changeset
     7
theory List_Permutation
73327
fd32f08f4fb5 more connections between mset _ = mset _ and permutations
haftmann
parents: 73305
diff changeset
     8
imports Permutations
15131
c69542757a4d New theory header syntax.
nipkow
parents: 15072
diff changeset
     9
begin
11054
a5404c70982f moved from Induct/ to Library/
wenzelm
parents:
diff changeset
    10
73329
haftmann
parents: 73327
diff changeset
    11
text \<open>
haftmann
parents: 73327
diff changeset
    12
  Note that multisets already provide the notion of permutated list and hence
haftmann
parents: 73327
diff changeset
    13
  this theory mostly echoes material already logically present in theory
haftmann
parents: 73327
diff changeset
    14
  \<^text>\<open>Permutations\<close>; it should be seldom needed.
haftmann
parents: 73327
diff changeset
    15
\<close>
haftmann
parents: 73327
diff changeset
    16
73305
47616dc81488 repaired document
haftmann
parents: 73301
diff changeset
    17
subsection \<open>An inductive definition \ldots\<close>
73300
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    18
70680
b8cd7ea34e33 tidied up some massive ugliness
paulson <lp15@cam.ac.uk>
parents: 69597
diff changeset
    19
inductive perm :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool"  (infixr \<open><~~>\<close> 50)
53238
01ef0a103fc9 tuned proofs;
wenzelm
parents: 51542
diff changeset
    20
where
01ef0a103fc9 tuned proofs;
wenzelm
parents: 51542
diff changeset
    21
  Nil [intro!]: "[] <~~> []"
01ef0a103fc9 tuned proofs;
wenzelm
parents: 51542
diff changeset
    22
| swap [intro!]: "y # x # l <~~> x # y # l"
01ef0a103fc9 tuned proofs;
wenzelm
parents: 51542
diff changeset
    23
| Cons [intro!]: "xs <~~> ys \<Longrightarrow> z # xs <~~> z # ys"
01ef0a103fc9 tuned proofs;
wenzelm
parents: 51542
diff changeset
    24
| trans [intro]: "xs <~~> ys \<Longrightarrow> ys <~~> zs \<Longrightarrow> xs <~~> zs"
11054
a5404c70982f moved from Induct/ to Library/
wenzelm
parents:
diff changeset
    25
61699
a81dc5c4d6a9 New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents: 61585
diff changeset
    26
proposition perm_refl [iff]: "l <~~> l"
17200
3a4d03d1a31b tuned presentation;
wenzelm
parents: 15140
diff changeset
    27
  by (induct l) auto
11054
a5404c70982f moved from Induct/ to Library/
wenzelm
parents:
diff changeset
    28
73305
47616dc81488 repaired document
haftmann
parents: 73301
diff changeset
    29
text \<open>\ldots that is equivalent to an already existing notion:\<close>
11054
a5404c70982f moved from Induct/ to Library/
wenzelm
parents:
diff changeset
    30
73300
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    31
lemma perm_iff_eq_mset:
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    32
  \<open>xs <~~> ys \<longleftrightarrow> mset xs = mset ys\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    33
proof
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    34
  assume \<open>mset xs = mset ys\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    35
  then show \<open>xs <~~> ys\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    36
  proof (induction xs arbitrary: ys)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    37
    case Nil
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    38
    then show ?case
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    39
      by simp
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    40
  next
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    41
    case (Cons x xs)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    42
    from Cons.prems [symmetric] have \<open>mset xs = mset (remove1 x ys)\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    43
      by simp
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    44
    then have \<open>xs <~~> remove1 x ys\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    45
      by (rule Cons.IH)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    46
    then have \<open>x # xs <~~> x # remove1 x ys\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    47
      by (rule perm.Cons)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    48
    moreover from Cons.prems have \<open>x \<in> set ys\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    49
      by (auto dest: union_single_eq_member)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    50
    then have \<open>x # remove1 x ys <~~> ys\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    51
      by (induction ys) auto
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    52
    ultimately show \<open>x # xs <~~> ys\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    53
      by (rule perm.trans)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    54
  qed
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    55
next
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    56
  assume \<open>xs <~~> ys\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    57
  then show \<open>mset xs = mset ys\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    58
    by induction simp_all
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    59
qed
11054
a5404c70982f moved from Induct/ to Library/
wenzelm
parents:
diff changeset
    60
73300
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    61
theorem mset_eq_perm: \<open>mset xs = mset ys \<longleftrightarrow> xs <~~> ys\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    62
  by (simp add: perm_iff_eq_mset)
11054
a5404c70982f moved from Induct/ to Library/
wenzelm
parents:
diff changeset
    63
a5404c70982f moved from Induct/ to Library/
wenzelm
parents:
diff changeset
    64
73300
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    65
subsection \<open>Nontrivial conclusions\<close>
11054
a5404c70982f moved from Induct/ to Library/
wenzelm
parents:
diff changeset
    66
73300
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    67
proposition perm_swap:
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    68
  \<open>xs[i := xs ! j, j := xs ! i] <~~> xs\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    69
  if \<open>i < length xs\<close> \<open>j < length xs\<close>
73301
bfe92e4f6ea4 multiset as equivalence class of permuted lists
haftmann
parents: 73300
diff changeset
    70
  using that by (simp add: perm_iff_eq_mset mset_swap)
15005
546c8e7e28d4 Norbert Voelker
paulson
parents: 14706
diff changeset
    71
64587
8355a6e2df79 standardized notation
haftmann
parents: 63649
diff changeset
    72
proposition mset_le_perm_append: "mset xs \<subseteq># mset ys \<longleftrightarrow> (\<exists>zs. xs @ zs <~~> ys)"
73300
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    73
  by (auto simp add: perm_iff_eq_mset mset_subset_eq_exists_conv ex_mset dest: sym)
15005
546c8e7e28d4 Norbert Voelker
paulson
parents: 14706
diff changeset
    74
61699
a81dc5c4d6a9 New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents: 61585
diff changeset
    75
proposition perm_set_eq: "xs <~~> ys \<Longrightarrow> set xs = set ys"
73300
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    76
  by (rule mset_eq_setD) (simp add: perm_iff_eq_mset)
25277
95128fcdd7e8 added lemmas
nipkow
parents: 23755
diff changeset
    77
73301
bfe92e4f6ea4 multiset as equivalence class of permuted lists
haftmann
parents: 73300
diff changeset
    78
proposition perm_distinct_iff: "xs <~~> ys \<Longrightarrow> distinct xs \<longleftrightarrow> distinct ys"
bfe92e4f6ea4 multiset as equivalence class of permuted lists
haftmann
parents: 73300
diff changeset
    79
  by (rule mset_eq_imp_distinct_iff) (simp add: perm_iff_eq_mset)
25277
95128fcdd7e8 added lemmas
nipkow
parents: 23755
diff changeset
    80
61699
a81dc5c4d6a9 New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents: 61585
diff changeset
    81
theorem eq_set_perm_remdups: "set xs = set ys \<Longrightarrow> remdups xs <~~> remdups ys"
73300
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    82
  by (simp add: perm_iff_eq_mset set_eq_iff_mset_remdups_eq)
25287
094dab519ff5 added lemmas
nipkow
parents: 25277
diff changeset
    83
61699
a81dc5c4d6a9 New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents: 61585
diff changeset
    84
proposition perm_remdups_iff_eq_set: "remdups x <~~> remdups y \<longleftrightarrow> set x = set y"
73300
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
    85
  by (simp add: perm_iff_eq_mset set_eq_iff_mset_remdups_eq)
25287
094dab519ff5 added lemmas
nipkow
parents: 25277
diff changeset
    86
61699
a81dc5c4d6a9 New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents: 61585
diff changeset
    87
theorem permutation_Ex_bij:
39075
a18e5946d63c Permutation implies bij function
hoelzl
parents: 36903
diff changeset
    88
  assumes "xs <~~> ys"
a18e5946d63c Permutation implies bij function
hoelzl
parents: 36903
diff changeset
    89
  shows "\<exists>f. bij_betw f {..<length xs} {..<length ys} \<and> (\<forall>i<length xs. xs ! i = ys ! (f i))"
73301
bfe92e4f6ea4 multiset as equivalence class of permuted lists
haftmann
parents: 73300
diff changeset
    90
proof -
73327
fd32f08f4fb5 more connections between mset _ = mset _ and permutations
haftmann
parents: 73305
diff changeset
    91
  from assms have \<open>mset xs = mset ys\<close> \<open>length xs = length ys\<close>
fd32f08f4fb5 more connections between mset _ = mset _ and permutations
haftmann
parents: 73305
diff changeset
    92
    by (auto simp add: perm_iff_eq_mset dest: mset_eq_length)
fd32f08f4fb5 more connections between mset _ = mset _ and permutations
haftmann
parents: 73305
diff changeset
    93
  from \<open>mset xs = mset ys\<close> obtain p where \<open>p permutes {..<length ys}\<close> \<open>permute_list p ys = xs\<close>
73301
bfe92e4f6ea4 multiset as equivalence class of permuted lists
haftmann
parents: 73300
diff changeset
    94
    by (rule mset_eq_permutation)
73327
fd32f08f4fb5 more connections between mset _ = mset _ and permutations
haftmann
parents: 73305
diff changeset
    95
  then have \<open>bij_betw p {..<length xs} {..<length ys}\<close>
fd32f08f4fb5 more connections between mset _ = mset _ and permutations
haftmann
parents: 73305
diff changeset
    96
    by (simp add: \<open>length xs = length ys\<close> permutes_imp_bij)
fd32f08f4fb5 more connections between mset _ = mset _ and permutations
haftmann
parents: 73305
diff changeset
    97
  moreover have \<open>\<forall>i<length xs. xs ! i = ys ! (p i)\<close>
fd32f08f4fb5 more connections between mset _ = mset _ and permutations
haftmann
parents: 73305
diff changeset
    98
    using \<open>permute_list p ys = xs\<close> \<open>length xs = length ys\<close> \<open>p permutes {..<length ys}\<close> permute_list_nth
fd32f08f4fb5 more connections between mset _ = mset _ and permutations
haftmann
parents: 73305
diff changeset
    99
    by auto
fd32f08f4fb5 more connections between mset _ = mset _ and permutations
haftmann
parents: 73305
diff changeset
   100
  ultimately show ?thesis
fd32f08f4fb5 more connections between mset _ = mset _ and permutations
haftmann
parents: 73305
diff changeset
   101
    by blast
39075
a18e5946d63c Permutation implies bij function
hoelzl
parents: 36903
diff changeset
   102
qed
a18e5946d63c Permutation implies bij function
hoelzl
parents: 36903
diff changeset
   103
61699
a81dc5c4d6a9 New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents: 61585
diff changeset
   104
proposition perm_finite: "finite {B. B <~~> A}"
73301
bfe92e4f6ea4 multiset as equivalence class of permuted lists
haftmann
parents: 73300
diff changeset
   105
  using mset_eq_finite by (auto simp add: perm_iff_eq_mset)
61699
a81dc5c4d6a9 New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents: 61585
diff changeset
   106
73300
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   107
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   108
subsection \<open>Trivial conclusions:\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   109
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   110
proposition perm_empty_imp: "[] <~~> ys \<Longrightarrow> ys = []"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   111
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   112
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   113
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   114
text \<open>\medskip This more general theorem is easier to understand!\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   115
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   116
proposition perm_length: "xs <~~> ys \<Longrightarrow> length xs = length ys"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   117
  by (rule mset_eq_length) (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   118
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   119
proposition perm_sym: "xs <~~> ys \<Longrightarrow> ys <~~> xs"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   120
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   121
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   122
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   123
text \<open>We can insert the head anywhere in the list.\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   124
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   125
proposition perm_append_Cons: "a # xs @ ys <~~> xs @ a # ys"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   126
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   127
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   128
proposition perm_append_swap: "xs @ ys <~~> ys @ xs"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   129
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   130
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   131
proposition perm_append_single: "a # xs <~~> xs @ [a]"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   132
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   133
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   134
proposition perm_rev: "rev xs <~~> xs"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   135
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   136
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   137
proposition perm_append1: "xs <~~> ys \<Longrightarrow> l @ xs <~~> l @ ys"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   138
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   139
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   140
proposition perm_append2: "xs <~~> ys \<Longrightarrow> xs @ l <~~> ys @ l"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   141
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   142
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   143
proposition perm_empty [iff]: "[] <~~> xs \<longleftrightarrow> xs = []"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   144
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   145
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   146
proposition perm_empty2 [iff]: "xs <~~> [] \<longleftrightarrow> xs = []"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   147
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   148
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   149
proposition perm_sing_imp: "ys <~~> xs \<Longrightarrow> xs = [y] \<Longrightarrow> ys = [y]"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   150
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   151
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   152
proposition perm_sing_eq [iff]: "ys <~~> [y] \<longleftrightarrow> ys = [y]"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   153
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   154
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   155
proposition perm_sing_eq2 [iff]: "[y] <~~> ys \<longleftrightarrow> ys = [y]"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   156
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   157
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   158
proposition perm_remove: "x \<in> set ys \<Longrightarrow> ys <~~> x # remove1 x ys"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   159
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   160
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   161
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   162
text \<open>\medskip Congruence rule\<close>
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   163
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   164
proposition perm_remove_perm: "xs <~~> ys \<Longrightarrow> remove1 z xs <~~> remove1 z ys"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   165
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   166
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   167
proposition remove_hd [simp]: "remove1 z (z # xs) = xs"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   168
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   169
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   170
proposition cons_perm_imp_perm: "z # xs <~~> z # ys \<Longrightarrow> xs <~~> ys"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   171
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   172
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   173
proposition cons_perm_eq [simp]: "z#xs <~~> z#ys \<longleftrightarrow> xs <~~> ys"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   174
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   175
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   176
proposition append_perm_imp_perm: "zs @ xs <~~> zs @ ys \<Longrightarrow> xs <~~> ys"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   177
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   178
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   179
proposition perm_append1_eq [iff]: "zs @ xs <~~> zs @ ys \<longleftrightarrow> xs <~~> ys"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   180
  by (simp add: perm_iff_eq_mset)
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   181
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   182
proposition perm_append2_eq [iff]: "xs @ zs <~~> ys @ zs \<longleftrightarrow> xs <~~> ys"
c52c5a5bf4e6 emphasize connection to multisets
haftmann
parents: 73297
diff changeset
   183
  by (simp add: perm_iff_eq_mset)
61699
a81dc5c4d6a9 New theorems mostly from Peter Gammie
paulson <lp15@cam.ac.uk>
parents: 61585
diff changeset
   184
11054
a5404c70982f moved from Induct/ to Library/
wenzelm
parents:
diff changeset
   185
end