src/HOL/Library/Dlist.thy
author wenzelm
Fri, 23 Aug 2024 22:47:51 +0200
changeset 80753 66893c47500d
parent 73932 fd21b4a93043
child 82691 b69e4da2604b
permissions -rw-r--r--
more markup for syntax consts;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
     2
   Author: Andreas Lochbihler, ETH Zurich *)
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
     3
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
     4
section \<open>Lists with elements distinct as canonical example for datatype invariants\<close>
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
     5
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
     6
theory Dlist
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
     7
imports Confluent_Quotient
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
     8
begin
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
     9
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    10
subsection \<open>The type of distinct lists\<close>
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    11
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 48282
diff changeset
    12
typedef 'a dlist = "{xs::'a list. distinct xs}"
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    13
  morphisms list_of_dlist Abs_dlist
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    14
proof
45694
4a8743618257 prefer typedef without extra definition and alternative name;
wenzelm
parents: 43764
diff changeset
    15
  show "[] \<in> {xs. distinct xs}" by simp
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    16
qed
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    17
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    18
context begin
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    19
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    20
qualified definition dlist_eq where "dlist_eq = BNF_Def.vimage2p remdups remdups (=)"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    21
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    22
qualified lemma equivp_dlist_eq: "equivp dlist_eq"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    23
  unfolding dlist_eq_def by(rule equivp_vimage2p)(rule identity_equivp)
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    24
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    25
qualified definition abs_dlist :: "'a list \<Rightarrow> 'a dlist" where "abs_dlist = Abs_dlist o remdups"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    26
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    27
definition qcr_dlist :: "'a list \<Rightarrow> 'a dlist \<Rightarrow> bool" where "qcr_dlist x y \<longleftrightarrow> y = abs_dlist x"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    28
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    29
qualified lemma Quotient_dlist_remdups: "Quotient dlist_eq abs_dlist list_of_dlist qcr_dlist"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    30
  unfolding Quotient_def dlist_eq_def qcr_dlist_def vimage2p_def abs_dlist_def
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    31
  by (auto simp add: fun_eq_iff Abs_dlist_inject
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    32
    list_of_dlist[simplified] list_of_dlist_inverse distinct_remdups_id)
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    33
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    34
end
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    35
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    36
locale Quotient_dlist begin
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    37
setup_lifting Dlist.Quotient_dlist_remdups Dlist.equivp_dlist_eq[THEN equivp_reflp2]
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    38
end
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
    39
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
    40
setup_lifting type_definition_dlist
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
    41
39380
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 38857
diff changeset
    42
lemma dlist_eq_iff:
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 38857
diff changeset
    43
  "dxs = dys \<longleftrightarrow> list_of_dlist dxs = list_of_dlist dys"
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 38857
diff changeset
    44
  by (simp add: list_of_dlist_inject)
36274
42bd879dc1b0 lemma dlist_ext
haftmann
parents: 36176
diff changeset
    45
39380
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 38857
diff changeset
    46
lemma dlist_eqI:
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 38857
diff changeset
    47
  "list_of_dlist dxs = list_of_dlist dys \<Longrightarrow> dxs = dys"
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 38857
diff changeset
    48
  by (simp add: dlist_eq_iff)
36112
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35688
diff changeset
    49
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
    50
text \<open>Formal, totalized constructor for \<^typ>\<open>'a dlist\<close>:\<close>
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    51
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    52
definition Dlist :: "'a list \<Rightarrow> 'a dlist" where
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 37595
diff changeset
    53
  "Dlist xs = Abs_dlist (remdups xs)"
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    54
39380
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 38857
diff changeset
    55
lemma distinct_list_of_dlist [simp, intro]:
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    56
  "distinct (list_of_dlist dxs)"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    57
  using list_of_dlist [of dxs] by simp
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    58
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    59
lemma list_of_dlist_Dlist [simp]:
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    60
  "list_of_dlist (Dlist xs) = remdups xs"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    61
  by (simp add: Dlist_def Abs_dlist_inverse)
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    62
39727
5dab9549c80d lemma remdups_list_of_dlist
haftmann
parents: 39380
diff changeset
    63
lemma remdups_list_of_dlist [simp]:
5dab9549c80d lemma remdups_list_of_dlist
haftmann
parents: 39380
diff changeset
    64
  "remdups (list_of_dlist dxs) = list_of_dlist dxs"
5dab9549c80d lemma remdups_list_of_dlist
haftmann
parents: 39380
diff changeset
    65
  by simp
5dab9549c80d lemma remdups_list_of_dlist
haftmann
parents: 39380
diff changeset
    66
36112
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35688
diff changeset
    67
lemma Dlist_list_of_dlist [simp, code abstype]:
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    68
  "Dlist (list_of_dlist dxs) = dxs"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    69
  by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id)
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    70
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    71
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    72
text \<open>Fundamental operations:\<close>
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    73
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    74
context
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    75
begin
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    76
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    77
qualified definition empty :: "'a dlist" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    78
  "empty = Dlist []"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    79
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    80
qualified definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    81
  "insert x dxs = Dlist (List.insert x (list_of_dlist dxs))"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    82
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    83
qualified definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    84
  "remove x dxs = Dlist (remove1 x (list_of_dlist dxs))"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    85
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    86
qualified definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    87
  "map f dxs = Dlist (remdups (List.map f (list_of_dlist dxs)))"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    88
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    89
qualified definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    90
  "filter P dxs = Dlist (List.filter P (list_of_dlist dxs))"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    91
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
    92
qualified definition rotate :: "nat \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
    93
  "rotate n dxs = Dlist (List.rotate n (list_of_dlist dxs))"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
    94
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    95
end
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    96
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    97
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    98
text \<open>Derived operations:\<close>
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    99
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   100
context
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   101
begin
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   102
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   103
qualified definition null :: "'a dlist \<Rightarrow> bool" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   104
  "null dxs = List.null (list_of_dlist dxs)"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   105
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   106
qualified definition member :: "'a dlist \<Rightarrow> 'a \<Rightarrow> bool" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   107
  "member dxs = List.member (list_of_dlist dxs)"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   108
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   109
qualified definition length :: "'a dlist \<Rightarrow> nat" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   110
  "length dxs = List.length (list_of_dlist dxs)"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   111
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   112
qualified definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
46133
d9fe85d3d2cd incorporated canonical fold combinator on lists into body of List theory; refactored passages on List.fold(l/r)
haftmann
parents: 45990
diff changeset
   113
  "fold f dxs = List.fold f (list_of_dlist dxs)"
37022
f9681d9d1d56 moved generic List operations to theory More_List
haftmann
parents: 36980
diff changeset
   114
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   115
qualified definition foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
37022
f9681d9d1d56 moved generic List operations to theory More_List
haftmann
parents: 36980
diff changeset
   116
  "foldr f dxs = List.foldr f (list_of_dlist dxs)"
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   117
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   118
end
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   119
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   120
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   121
subsection \<open>Executable version obeying invariant\<close>
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   122
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   123
lemma list_of_dlist_empty [simp, code abstract]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   124
  "list_of_dlist Dlist.empty = []"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   125
  by (simp add: Dlist.empty_def)
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   126
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   127
lemma list_of_dlist_insert [simp, code abstract]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   128
  "list_of_dlist (Dlist.insert x dxs) = List.insert x (list_of_dlist dxs)"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   129
  by (simp add: Dlist.insert_def)
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   130
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   131
lemma list_of_dlist_remove [simp, code abstract]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   132
  "list_of_dlist (Dlist.remove x dxs) = remove1 x (list_of_dlist dxs)"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   133
  by (simp add: Dlist.remove_def)
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   134
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   135
lemma list_of_dlist_map [simp, code abstract]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   136
  "list_of_dlist (Dlist.map f dxs) = remdups (List.map f (list_of_dlist dxs))"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   137
  by (simp add: Dlist.map_def)
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   138
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   139
lemma list_of_dlist_filter [simp, code abstract]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   140
  "list_of_dlist (Dlist.filter P dxs) = List.filter P (list_of_dlist dxs)"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   141
  by (simp add: Dlist.filter_def)
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   142
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
   143
lemma list_of_dlist_rotate [simp, code abstract]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
   144
  "list_of_dlist (Dlist.rotate n dxs) = List.rotate n (list_of_dlist dxs)"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
   145
  by (simp add: Dlist.rotate_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
   146
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   147
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   148
text \<open>Explicit executable conversion\<close>
36980
1a4cc941171a added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
haftmann
parents: 36274
diff changeset
   149
1a4cc941171a added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
haftmann
parents: 36274
diff changeset
   150
definition dlist_of_list [simp]:
1a4cc941171a added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
haftmann
parents: 36274
diff changeset
   151
  "dlist_of_list = Dlist"
1a4cc941171a added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
haftmann
parents: 36274
diff changeset
   152
1a4cc941171a added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
haftmann
parents: 36274
diff changeset
   153
lemma [code abstract]:
1a4cc941171a added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
haftmann
parents: 36274
diff changeset
   154
  "list_of_dlist (dlist_of_list xs) = remdups xs"
1a4cc941171a added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
haftmann
parents: 36274
diff changeset
   155
  by simp
1a4cc941171a added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
haftmann
parents: 36274
diff changeset
   156
1a4cc941171a added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
haftmann
parents: 36274
diff changeset
   157
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   158
text \<open>Equality\<close>
38512
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   159
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38512
diff changeset
   160
instantiation dlist :: (equal) equal
38512
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   161
begin
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   162
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38512
diff changeset
   163
definition "HOL.equal dxs dys \<longleftrightarrow> HOL.equal (list_of_dlist dxs) (list_of_dlist dys)"
38512
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   164
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   165
instance
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   166
  by standard (simp add: equal_dlist_def equal list_of_dlist_inject)
38512
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   167
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   168
end
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   169
43764
366d5726de09 explicit code equation for equality
haftmann
parents: 43146
diff changeset
   170
declare equal_dlist_def [code]
366d5726de09 explicit code equation for equality
haftmann
parents: 43146
diff changeset
   171
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   172
lemma [code nbe]: "HOL.equal (dxs :: 'a::equal dlist) dxs \<longleftrightarrow> True"
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38512
diff changeset
   173
  by (fact equal_refl)
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38512
diff changeset
   174
38512
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   175
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   176
subsection \<open>Induction principle and case distinction\<close>
37106
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   177
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   178
lemma dlist_induct [case_names empty insert, induct type: dlist]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   179
  assumes empty: "P Dlist.empty"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   180
  assumes insrt: "\<And>x dxs. \<not> Dlist.member dxs x \<Longrightarrow> P dxs \<Longrightarrow> P (Dlist.insert x dxs)"
37106
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   181
  shows "P dxs"
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   182
proof (cases dxs)
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   183
  case (Abs_dlist xs)
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   184
  then have "distinct xs" and dxs: "dxs = Dlist xs"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   185
    by (simp_all add: Dlist_def distinct_remdups_id)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   186
  from \<open>distinct xs\<close> have "P (Dlist xs)"
39915
ecf97cf3d248 turned distinct and sorted into inductive predicates: yields nice induction principles for free; more elegant proofs
haftmann
parents: 39727
diff changeset
   187
  proof (induct xs)
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   188
    case Nil from empty show ?case by (simp add: Dlist.empty_def)
37106
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   189
  next
40122
1d8ad2ff3e01 dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
haftmann
parents: 39915
diff changeset
   190
    case (Cons x xs)
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   191
    then have "\<not> Dlist.member (Dlist xs) x" and "P (Dlist xs)"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   192
      by (simp_all add: Dlist.member_def List.member_def)
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   193
    with insrt have "P (Dlist.insert x (Dlist xs))" .
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   194
    with Cons show ?case by (simp add: Dlist.insert_def distinct_remdups_id)
37106
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   195
  qed
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   196
  with dxs show "P dxs" by simp
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   197
qed
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   198
55913
c1409c103b77 proper UTF-8;
wenzelm
parents: 55467
diff changeset
   199
lemma dlist_case [cases type: dlist]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   200
  obtains (empty) "dxs = Dlist.empty"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   201
    | (insert) x dys where "\<not> Dlist.member dys x" and "dxs = Dlist.insert x dys"
37106
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   202
proof (cases dxs)
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   203
  case (Abs_dlist xs)
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   204
  then have dxs: "dxs = Dlist xs" and distinct: "distinct xs"
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   205
    by (simp_all add: Dlist_def distinct_remdups_id)
55913
c1409c103b77 proper UTF-8;
wenzelm
parents: 55467
diff changeset
   206
  show thesis
c1409c103b77 proper UTF-8;
wenzelm
parents: 55467
diff changeset
   207
  proof (cases xs)
c1409c103b77 proper UTF-8;
wenzelm
parents: 55467
diff changeset
   208
    case Nil with dxs
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   209
    have "dxs = Dlist.empty" by (simp add: Dlist.empty_def)
55913
c1409c103b77 proper UTF-8;
wenzelm
parents: 55467
diff changeset
   210
    with empty show ?thesis .
37106
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   211
  next
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   212
    case (Cons x xs)
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   213
    with dxs distinct have "\<not> Dlist.member (Dlist xs) x"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   214
      and "dxs = Dlist.insert x (Dlist xs)"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   215
      by (simp_all add: Dlist.member_def List.member_def Dlist.insert_def distinct_remdups_id)
55913
c1409c103b77 proper UTF-8;
wenzelm
parents: 55467
diff changeset
   216
    with insert show ?thesis .
37106
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   217
  qed
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   218
qed
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   219
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   220
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   221
subsection \<open>Functorial structure\<close>
40603
963ee2331d20 mapper for dlist type
haftmann
parents: 40122
diff changeset
   222
55467
a5c9002bc54d renamed 'enriched_type' to more informative 'functor' (following the renaming of enriched type constructors to bounded natural functors)
blanchet
parents: 49834
diff changeset
   223
functor map: map
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   224
  by (simp_all add: remdups_map_remdups fun_eq_iff dlist_eq_iff)
40603
963ee2331d20 mapper for dlist type
haftmann
parents: 40122
diff changeset
   225
48282
39bfb2844b9e tuned whitespace
haftmann
parents: 46565
diff changeset
   226
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   227
subsection \<open>Quickcheck generators\<close>
45927
e0305e4f02c9 adding quickcheck generator for distinct lists; adding examples
bulwahn
parents: 45694
diff changeset
   228
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   229
quickcheck_generator dlist predicate: distinct constructors: Dlist.empty, Dlist.insert
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   230
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   231
subsection \<open>BNF instance\<close>
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   232
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   233
context begin
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   234
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   235
qualified inductive double :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" where
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   236
  "double (xs @ ys) (xs @ x # ys)" if "x \<in> set ys"
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   237
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   238
qualified lemma strong_confluentp_double: "strong_confluentp double"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   239
proof
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   240
  fix xs ys zs :: "'a list"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   241
  assume ys: "double xs ys" and zs: "double xs zs"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   242
  consider (left) as y bs z cs where "xs = as @ bs @ cs" "ys = as @ y # bs @ cs" "zs = as @ bs @ z # cs" "y \<in> set (bs @ cs)" "z \<in> set cs"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   243
    | (right) as y bs z cs where "xs = as @ bs @ cs" "ys = as @ bs @ y # cs" "zs = as @ z # bs @ cs" "y \<in> set cs" "z \<in> set (bs @ cs)"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   244
  proof -
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   245
    show thesis using ys zs
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   246
      by(clarsimp simp add: double.simps append_eq_append_conv2)(auto intro: that)
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   247
  qed
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   248
  then show "\<exists>us. double\<^sup>*\<^sup>* ys us \<and> double\<^sup>=\<^sup>= zs us"
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   249
  proof cases
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   250
    case left
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   251
    let ?us = "as @ y # bs @ z # cs"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   252
    have "double ys ?us" "double zs ?us" using left
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   253
      by(auto 4 4 simp add: double.simps)(metis append_Cons append_assoc)+
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   254
    then show ?thesis by blast
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   255
  next
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   256
    case right
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   257
    let ?us = "as @ z # bs @ y # cs"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   258
    have "double ys ?us" "double zs ?us" using right
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   259
      by(auto 4 4 simp add: double.simps)(metis append_Cons append_assoc)+
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   260
    then show ?thesis by blast
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   261
  qed
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   262
qed
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   263
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   264
qualified lemma double_Cons1 [simp]: "double xs (x # xs)" if "x \<in> set xs"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   265
  using double.intros[of x xs "[]"] that by simp
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   266
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   267
qualified lemma double_Cons_same [simp]: "double xs ys \<Longrightarrow> double (x # xs) (x # ys)"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   268
  by(auto simp add: double.simps Cons_eq_append_conv)
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   269
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   270
qualified lemma doubles_Cons_same: "double\<^sup>*\<^sup>* xs ys \<Longrightarrow> double\<^sup>*\<^sup>* (x # xs) (x # ys)"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   271
  by(induction rule: rtranclp_induct)(auto intro: rtranclp.rtrancl_into_rtrancl)
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   272
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   273
qualified lemma remdups_into_doubles: "double\<^sup>*\<^sup>* (remdups xs) xs"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   274
  by(induction xs)(auto intro: doubles_Cons_same rtranclp.rtrancl_into_rtrancl)
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   275
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   276
qualified lemma dlist_eq_into_doubles: "Dlist.dlist_eq \<le> equivclp double"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   277
  by(auto 4 4 simp add: Dlist.dlist_eq_def vimage2p_def
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   278
     intro: equivclp_trans converse_rtranclp_into_equivclp rtranclp_into_equivclp remdups_into_doubles)
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   279
73398
180981b87929 generalized confluence-based subdistributivity theorem for quotients;
traytel
parents: 71494
diff changeset
   280
qualified lemma factor_double_map: "double (map f xs) ys \<Longrightarrow> \<exists>zs. Dlist.dlist_eq xs zs \<and> ys = map f zs \<and> set zs \<subseteq> set xs"
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   281
  by(auto simp add: double.simps Dlist.dlist_eq_def vimage2p_def map_eq_append_conv)
73932
fd21b4a93043 added opaque_combs and renamed hide_lams to opaque_lifting
desharna
parents: 73398
diff changeset
   282
    (metis (no_types, opaque_lifting) list.simps(9) map_append remdups.simps(2) remdups_append2 set_append set_eq_subset set_remdups)
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   283
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   284
qualified lemma dlist_eq_set_eq: "Dlist.dlist_eq xs ys \<Longrightarrow> set xs = set ys"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   285
  by(simp add: Dlist.dlist_eq_def vimage2p_def)(metis set_remdups)
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   286
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   287
qualified lemma dlist_eq_map_respect: "Dlist.dlist_eq xs ys \<Longrightarrow> Dlist.dlist_eq (map f xs) (map f ys)"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   288
  by(clarsimp simp add: Dlist.dlist_eq_def vimage2p_def)(metis remdups_map_remdups)
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   289
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   290
qualified lemma confluent_quotient_dlist:
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   291
  "confluent_quotient double Dlist.dlist_eq Dlist.dlist_eq Dlist.dlist_eq Dlist.dlist_eq Dlist.dlist_eq
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   292
     (map fst) (map snd) (map fst) (map snd) list_all2 list_all2 list_all2 set set"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   293
  by(unfold_locales)(auto intro: strong_confluentp_imp_confluentp strong_confluentp_double
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   294
    dest: factor_double_map dlist_eq_into_doubles[THEN predicate2D] dlist_eq_set_eq
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   295
    simp add: list.in_rel list.rel_compp dlist_eq_map_respect Dlist.equivp_dlist_eq equivp_imp_transp)
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   296
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   297
lifting_update dlist.lifting
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   298
lifting_forget dlist.lifting
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   299
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   300
end
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   301
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   302
context begin
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   303
interpretation Quotient_dlist: Quotient_dlist .
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   304
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   305
lift_bnf (plugins del: code) 'a dlist
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   306
  subgoal for A B by(rule confluent_quotient.subdistributivity[OF Dlist.confluent_quotient_dlist])
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   307
  subgoal by(force dest: Dlist.dlist_eq_set_eq intro: equivp_reflp[OF Dlist.equivp_dlist_eq])
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   308
  done
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   309
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   310
qualified lemma list_of_dlist_transfer[transfer_rule]:
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   311
  "bi_unique R \<Longrightarrow> (rel_fun (Quotient_dlist.pcr_dlist R) (list_all2 R)) remdups list_of_dlist"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   312
  unfolding rel_fun_def Quotient_dlist.pcr_dlist_def qcr_dlist_def Dlist.abs_dlist_def
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   313
  by (auto simp: Abs_dlist_inverse intro!: remdups_transfer[THEN rel_funD])
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   314
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   315
lemma list_of_dlist_map_dlist[simp]:
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   316
  "list_of_dlist (map_dlist f xs) = remdups (map f (list_of_dlist xs))"
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   317
  by transfer (auto simp: remdups_map_remdups)
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   318
62390
842917225d56 more canonical names
nipkow
parents: 62324
diff changeset
   319
end
71494
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   320
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   321
cbe0b6b0bed8 tuned lift_bnf's user interface for quotients
traytel
parents: 69593
diff changeset
   322
end