src/HOL/Library/Dlist.thy
author nipkow
Wed, 13 Feb 2019 07:48:42 +0100
changeset 69801 a99a0f5474c5
parent 69593 3dda49e08b9d
child 71494 cbe0b6b0bed8
permissions -rw-r--r--
too agressive
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
     1
(* Author: Florian Haftmann, TU Muenchen 
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
45990
b7b905b23b2a incorporated More_Set and More_List into the Main body -- to be consolidated later
haftmann
parents: 45927
diff changeset
     7
imports Main
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
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
    18
setup_lifting type_definition_dlist
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
    19
39380
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 38857
diff changeset
    20
lemma dlist_eq_iff:
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 38857
diff changeset
    21
  "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
    22
  by (simp add: list_of_dlist_inject)
36274
42bd879dc1b0 lemma dlist_ext
haftmann
parents: 36176
diff changeset
    23
39380
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 38857
diff changeset
    24
lemma dlist_eqI:
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 38857
diff changeset
    25
  "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
    26
  by (simp add: dlist_eq_iff)
36112
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35688
diff changeset
    27
69593
3dda49e08b9d isabelle update -u control_cartouches;
wenzelm
parents: 67399
diff changeset
    28
text \<open>Formal, totalized constructor for \<^typ>\<open>'a dlist\<close>:\<close>
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    29
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    30
definition Dlist :: "'a list \<Rightarrow> 'a dlist" where
37765
26bdfb7b680b dropped superfluous [code del]s
haftmann
parents: 37595
diff changeset
    31
  "Dlist xs = Abs_dlist (remdups xs)"
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    32
39380
5a2662c1e44a established emerging canonical names *_eqI and *_eq_iff
haftmann
parents: 38857
diff changeset
    33
lemma distinct_list_of_dlist [simp, intro]:
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    34
  "distinct (list_of_dlist dxs)"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    35
  using list_of_dlist [of dxs] by simp
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    36
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    37
lemma list_of_dlist_Dlist [simp]:
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    38
  "list_of_dlist (Dlist xs) = remdups xs"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    39
  by (simp add: Dlist_def Abs_dlist_inverse)
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    40
39727
5dab9549c80d lemma remdups_list_of_dlist
haftmann
parents: 39380
diff changeset
    41
lemma remdups_list_of_dlist [simp]:
5dab9549c80d lemma remdups_list_of_dlist
haftmann
parents: 39380
diff changeset
    42
  "remdups (list_of_dlist dxs) = list_of_dlist dxs"
5dab9549c80d lemma remdups_list_of_dlist
haftmann
parents: 39380
diff changeset
    43
  by simp
5dab9549c80d lemma remdups_list_of_dlist
haftmann
parents: 39380
diff changeset
    44
36112
7fa17a225852 user interface for abstract datatypes is an attribute, not a command
haftmann
parents: 35688
diff changeset
    45
lemma Dlist_list_of_dlist [simp, code abstype]:
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    46
  "Dlist (list_of_dlist dxs) = dxs"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    47
  by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id)
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    48
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    49
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    50
text \<open>Fundamental operations:\<close>
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    51
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    52
context
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    53
begin
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    54
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    55
qualified definition empty :: "'a dlist" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    56
  "empty = Dlist []"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    57
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    58
qualified definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    59
  "insert x dxs = Dlist (List.insert x (list_of_dlist dxs))"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    60
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    61
qualified definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    62
  "remove x dxs = Dlist (remove1 x (list_of_dlist dxs))"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    63
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    64
qualified definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    65
  "map f dxs = Dlist (remdups (List.map f (list_of_dlist dxs)))"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    66
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    67
qualified definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    68
  "filter P dxs = Dlist (List.filter P (list_of_dlist dxs))"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    69
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
    70
qualified definition rotate :: "nat \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
    71
  "rotate n dxs = Dlist (List.rotate n (list_of_dlist dxs))"
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
    72
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    73
end
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    74
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    75
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    76
text \<open>Derived operations:\<close>
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    77
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    78
context
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    79
begin
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    80
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    81
qualified definition null :: "'a dlist \<Rightarrow> bool" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    82
  "null dxs = List.null (list_of_dlist dxs)"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    83
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    84
qualified definition member :: "'a dlist \<Rightarrow> 'a \<Rightarrow> bool" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    85
  "member dxs = List.member (list_of_dlist dxs)"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    86
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    87
qualified definition length :: "'a dlist \<Rightarrow> nat" where
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    88
  "length dxs = List.length (list_of_dlist dxs)"
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    89
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    90
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
    91
  "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
    92
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    93
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
    94
  "foldr f dxs = List.foldr f (list_of_dlist dxs)"
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    95
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    96
end
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
    97
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
    98
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
    99
subsection \<open>Executable version obeying invariant\<close>
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   100
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   101
lemma list_of_dlist_empty [simp, code abstract]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   102
  "list_of_dlist Dlist.empty = []"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   103
  by (simp add: Dlist.empty_def)
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   104
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   105
lemma list_of_dlist_insert [simp, code abstract]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   106
  "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
   107
  by (simp add: Dlist.insert_def)
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   108
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   109
lemma list_of_dlist_remove [simp, code abstract]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   110
  "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
   111
  by (simp add: Dlist.remove_def)
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   112
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   113
lemma list_of_dlist_map [simp, code abstract]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   114
  "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
   115
  by (simp add: Dlist.map_def)
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   116
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   117
lemma list_of_dlist_filter [simp, code abstract]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   118
  "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
   119
  by (simp add: Dlist.filter_def)
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   120
63375
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
   121
lemma list_of_dlist_rotate [simp, code abstract]:
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
   122
  "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
   123
  by (simp add: Dlist.rotate_def)
59803048b0e8 basic facts about almost everywhere fix bijections
haftmann
parents: 62390
diff changeset
   124
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   125
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   126
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
   127
1a4cc941171a added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
haftmann
parents: 36274
diff changeset
   128
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
   129
  "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
   130
1a4cc941171a added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
haftmann
parents: 36274
diff changeset
   131
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
   132
  "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
   133
  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
   134
1a4cc941171a added implementations of Fset.Set, Fset.Coset; do not delete code equations for relational operators on fsets
haftmann
parents: 36274
diff changeset
   135
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   136
text \<open>Equality\<close>
38512
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   137
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38512
diff changeset
   138
instantiation dlist :: (equal) equal
38512
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   139
begin
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   140
38857
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38512
diff changeset
   141
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
   142
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   143
instance
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   144
  by standard (simp add: equal_dlist_def equal list_of_dlist_inject)
38512
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   145
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   146
end
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   147
43764
366d5726de09 explicit code equation for equality
haftmann
parents: 43146
diff changeset
   148
declare equal_dlist_def [code]
366d5726de09 explicit code equation for equality
haftmann
parents: 43146
diff changeset
   149
60679
ade12ef2773c tuned proofs;
wenzelm
parents: 60500
diff changeset
   150
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
   151
  by (fact equal_refl)
97775f3e8722 renamed class/constant eq to equal; tuned some instantiations
haftmann
parents: 38512
diff changeset
   152
38512
ed4703b416ed added equality instantiation
haftmann
parents: 37765
diff changeset
   153
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   154
subsection \<open>Induction principle and case distinction\<close>
37106
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   155
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   156
lemma dlist_induct [case_names empty insert, induct type: dlist]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   157
  assumes empty: "P Dlist.empty"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   158
  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
   159
  shows "P dxs"
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   160
proof (cases dxs)
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   161
  case (Abs_dlist xs)
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   162
  then have "distinct xs" and dxs: "dxs = Dlist xs"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   163
    by (simp_all add: Dlist_def distinct_remdups_id)
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   164
  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
   165
  proof (induct xs)
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   166
    case Nil from empty show ?case by (simp add: Dlist.empty_def)
37106
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   167
  next
40122
1d8ad2ff3e01 dropped (almost) redundant distinct.induct rule; distinct_simps again named distinct.simps
haftmann
parents: 39915
diff changeset
   168
    case (Cons x xs)
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   169
    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
   170
      by (simp_all add: Dlist.member_def List.member_def)
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   171
    with insrt have "P (Dlist.insert x (Dlist xs))" .
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   172
    with Cons show ?case by (simp add: Dlist.insert_def distinct_remdups_id)
37106
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   173
  qed
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   174
  with dxs show "P dxs" by simp
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   175
qed
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   176
55913
c1409c103b77 proper UTF-8;
wenzelm
parents: 55467
diff changeset
   177
lemma dlist_case [cases type: dlist]:
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   178
  obtains (empty) "dxs = Dlist.empty"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   179
    | (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
   180
proof (cases dxs)
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   181
  case (Abs_dlist xs)
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   182
  then have dxs: "dxs = Dlist xs" and distinct: "distinct xs"
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   183
    by (simp_all add: Dlist_def distinct_remdups_id)
55913
c1409c103b77 proper UTF-8;
wenzelm
parents: 55467
diff changeset
   184
  show thesis
c1409c103b77 proper UTF-8;
wenzelm
parents: 55467
diff changeset
   185
  proof (cases xs)
c1409c103b77 proper UTF-8;
wenzelm
parents: 55467
diff changeset
   186
    case Nil with dxs
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   187
    have "dxs = Dlist.empty" by (simp add: Dlist.empty_def) 
55913
c1409c103b77 proper UTF-8;
wenzelm
parents: 55467
diff changeset
   188
    with empty show ?thesis .
37106
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   189
  next
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   190
    case (Cons x xs)
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   191
    with dxs distinct have "\<not> Dlist.member (Dlist xs) x"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   192
      and "dxs = Dlist.insert x (Dlist xs)"
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   193
      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
   194
    with insert show ?thesis .
37106
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   195
  qed
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   196
qed
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   197
d56e0b30ce5a induction and case rules
haftmann
parents: 37029
diff changeset
   198
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   199
subsection \<open>Functorial structure\<close>
40603
963ee2331d20 mapper for dlist type
haftmann
parents: 40122
diff changeset
   200
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
   201
functor map: map
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   202
  by (simp_all add: remdups_map_remdups fun_eq_iff dlist_eq_iff)
40603
963ee2331d20 mapper for dlist type
haftmann
parents: 40122
diff changeset
   203
48282
39bfb2844b9e tuned whitespace
haftmann
parents: 46565
diff changeset
   204
60500
903bb1495239 isabelle update_cartouches;
wenzelm
parents: 58881
diff changeset
   205
subsection \<open>Quickcheck generators\<close>
45927
e0305e4f02c9 adding quickcheck generator for distinct lists; adding examples
bulwahn
parents: 45694
diff changeset
   206
61115
3a4400985780 modernized name space management -- more uniform qualification;
wenzelm
parents: 60679
diff changeset
   207
quickcheck_generator dlist predicate: distinct constructors: Dlist.empty, Dlist.insert
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   208
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   209
subsection \<open>BNF instance\<close>
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   210
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   211
context begin
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   212
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   213
qualified fun wpull :: "('a \<times> 'b) list \<Rightarrow> ('b \<times> 'c) list \<Rightarrow> ('a \<times> 'c) list"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   214
where
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   215
  "wpull [] ys = []"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   216
| "wpull xs [] = []"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   217
| "wpull ((a, b) # xs) ((b', c) # ys) =
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   218
  (if b \<in> snd ` set xs then
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   219
     (a, the (map_of (rev ((b', c) # ys)) b)) # wpull xs ((b', c) # ys)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   220
   else if b' \<in> fst ` set ys then
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   221
     (the (map_of (map prod.swap (rev ((a, b) # xs))) b'), c) # wpull ((a, b) # xs) ys
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   222
   else (a, c) # wpull xs ys)"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   223
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   224
qualified lemma wpull_eq_Nil_iff [simp]: "wpull xs ys = [] \<longleftrightarrow> xs = [] \<or> ys = []"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   225
by(cases "(xs, ys)" rule: wpull.cases) simp_all
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   226
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   227
qualified lemma wpull_induct
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   228
  [consumes 1, 
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   229
   case_names Nil left[xs eq in_set IH] right[xs ys eq in_set IH] step[xs ys eq IH] ]:
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   230
  assumes eq: "remdups (map snd xs) = remdups (map fst ys)"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   231
  and Nil: "P [] []"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   232
  and left: "\<And>a b xs b' c ys.
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   233
    \<lbrakk> b \<in> snd ` set xs; remdups (map snd xs) = remdups (map fst ((b', c) # ys)); 
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   234
      (b, the (map_of (rev ((b', c) # ys)) b)) \<in> set ((b', c) # ys); P xs ((b', c) # ys) \<rbrakk>
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   235
    \<Longrightarrow> P ((a, b) # xs) ((b', c) # ys)"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   236
  and right: "\<And>a b xs b' c ys.
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   237
    \<lbrakk> b \<notin> snd ` set xs; b' \<in> fst ` set ys;
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   238
      remdups (map snd ((a, b) # xs)) = remdups (map fst ys);
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   239
      (the (map_of (map prod.swap (rev ((a, b) #xs))) b'), b') \<in> set ((a, b) # xs);
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   240
      P ((a, b) # xs) ys \<rbrakk>
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   241
    \<Longrightarrow> P ((a, b) # xs) ((b', c) # ys)"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   242
  and step: "\<And>a b xs c ys.
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   243
    \<lbrakk> b \<notin> snd ` set xs; b \<notin> fst ` set ys; remdups (map snd xs) = remdups (map fst ys); 
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   244
      P xs ys \<rbrakk>
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   245
    \<Longrightarrow> P ((a, b) # xs) ((b, c) # ys)"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   246
  shows "P xs ys"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   247
using eq
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   248
proof(induction xs ys rule: wpull.induct)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   249
  case 1 thus ?case by(simp add: Nil)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   250
next
62390
842917225d56 more canonical names
nipkow
parents: 62324
diff changeset
   251
  case 2 thus ?case by(simp split: if_split_asm)
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   252
next
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   253
  case Cons: (3 a b xs b' c ys)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   254
  let ?xs = "(a, b) # xs" and ?ys = "(b', c) # ys"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   255
  consider (xs) "b \<in> snd ` set xs" | (ys) "b \<notin> snd ` set xs" "b' \<in> fst ` set ys"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   256
    | (step) "b \<notin> snd ` set xs" "b' \<notin> fst ` set ys" by auto
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   257
  thus ?case
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   258
  proof cases
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   259
    case xs
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   260
    with Cons.prems have eq: "remdups (map snd xs) = remdups (map fst ?ys)" by auto
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   261
    from xs eq have "b \<in> fst ` set ?ys" by (metis list.set_map set_remdups)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   262
    hence "map_of (rev ?ys) b \<noteq> None" unfolding map_of_eq_None_iff by auto
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   263
    then obtain c' where "map_of (rev ?ys) b = Some c'" by blast
62390
842917225d56 more canonical names
nipkow
parents: 62324
diff changeset
   264
    then have "(b, the (map_of (rev ?ys) b)) \<in> set ?ys" by(auto dest: map_of_SomeD split: if_split_asm)
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   265
    from xs eq this Cons.IH(1)[OF xs eq] show ?thesis by(rule left)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   266
  next
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   267
    case ys
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   268
    from ys Cons.prems have eq: "remdups (map snd ?xs) = remdups (map fst ys)" by auto
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   269
    from ys eq have "b' \<in> snd ` set ?xs" by (metis list.set_map set_remdups)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   270
    hence "map_of (map prod.swap (rev ?xs)) b' \<noteq> None"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   271
      unfolding map_of_eq_None_iff by(auto simp add: image_image)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   272
    then obtain a' where "map_of (map prod.swap (rev ?xs)) b' = Some a'" by blast
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   273
    then have "(the (map_of (map prod.swap (rev ?xs)) b'), b') \<in> set ?xs"
62390
842917225d56 more canonical names
nipkow
parents: 62324
diff changeset
   274
      by(auto dest: map_of_SomeD split: if_split_asm)
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   275
    from ys eq this Cons.IH(2)[OF ys eq] show ?thesis by(rule right)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   276
  next
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   277
    case *: step
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   278
    hence "remdups (map snd xs) = remdups (map fst ys)" "b = b'" using Cons.prems by auto
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   279
    from * this(1) Cons.IH(3)[OF * this(1)] show ?thesis unfolding \<open>b = b'\<close> by(rule step)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   280
  qed
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   281
qed
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   282
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   283
qualified lemma set_wpull_subset:
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   284
  assumes "remdups (map snd xs) = remdups (map fst ys)"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   285
  shows "set (wpull xs ys) \<subseteq> set xs O set ys"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   286
using assms by(induction xs ys rule: wpull_induct) auto
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   287
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   288
qualified lemma set_fst_wpull:
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   289
  assumes "remdups (map snd xs) = remdups (map fst ys)"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   290
  shows "fst ` set (wpull xs ys) = fst ` set xs"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   291
using assms by(induction xs ys rule: wpull_induct)(auto intro: rev_image_eqI)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   292
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   293
qualified lemma set_snd_wpull:
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   294
  assumes "remdups (map snd xs) = remdups (map fst ys)"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   295
  shows "snd ` set (wpull xs ys) = snd ` set ys"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   296
using assms by(induction xs ys rule: wpull_induct)(auto intro: rev_image_eqI)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   297
  
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   298
qualified lemma wpull:
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   299
  assumes "distinct xs"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   300
  and "distinct ys"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   301
  and "set xs \<subseteq> {(x, y). R x y}"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   302
  and "set ys \<subseteq> {(x, y). S x y}"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   303
  and eq: "remdups (map snd xs) = remdups (map fst ys)"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   304
  shows "\<exists>zs. distinct zs \<and> set zs \<subseteq> {(x, y). (R OO S) x y} \<and>
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   305
         remdups (map fst zs) = remdups (map fst xs) \<and> remdups (map snd zs) = remdups (map snd ys)"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   306
proof(intro exI conjI)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   307
  let ?zs = "remdups (wpull xs ys)"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   308
  show "distinct ?zs" by simp
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   309
  show "set ?zs \<subseteq> {(x, y). (R OO S) x y}" using assms(3-4) set_wpull_subset[OF eq] by fastforce
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   310
  show "remdups (map fst ?zs) = remdups (map fst xs)" unfolding remdups_map_remdups using eq
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   311
    by(induction xs ys rule: wpull_induct)(auto simp add: set_fst_wpull intro: rev_image_eqI)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   312
  show "remdups (map snd ?zs) = remdups (map snd ys)" unfolding remdups_map_remdups using eq
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   313
    by(induction xs ys rule: wpull_induct)(auto simp add: set_snd_wpull intro: rev_image_eqI)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   314
qed
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   315
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   316
qualified lift_definition set :: "'a dlist \<Rightarrow> 'a set" is List.set .
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   317
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   318
qualified lemma map_transfer [transfer_rule]:
67399
eab6ce8368fa ran isabelle update_op on all sources
nipkow
parents: 63375
diff changeset
   319
  "(rel_fun (=) (rel_fun (pcr_dlist (=)) (pcr_dlist (=)))) (\<lambda>f x. remdups (List.map f x)) Dlist.map"
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   320
by(simp add: rel_fun_def dlist.pcr_cr_eq cr_dlist_def Dlist.map_def remdups_remdups)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   321
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   322
bnf "'a dlist"
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   323
  map: Dlist.map
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   324
  sets: set
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   325
  bd: natLeq
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   326
  wits: Dlist.empty
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   327
unfolding OO_Grp_alt mem_Collect_eq
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   328
subgoal by(rule ext)(simp add: dlist_eq_iff)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   329
subgoal by(rule ext)(simp add: dlist_eq_iff remdups_map_remdups)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   330
subgoal by(simp add: dlist_eq_iff set_def cong: list.map_cong)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   331
subgoal by(simp add: set_def fun_eq_iff)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   332
subgoal by(simp add: natLeq_card_order)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   333
subgoal by(simp add: natLeq_cinfinite)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   334
subgoal by(rule ordLess_imp_ordLeq)(simp add: finite_iff_ordLess_natLeq[symmetric] set_def)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   335
subgoal by(rule predicate2I)(transfer; auto simp add: wpull)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   336
subgoal by(simp add: set_def)
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   337
done
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   338
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   339
lifting_update dlist.lifting
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   340
lifting_forget dlist.lifting
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   341
35303
816e48d60b13 added Dlist
haftmann
parents:
diff changeset
   342
end
62139
519362f817c7 add BNF instance for Dlist
Andreas Lochbihler
parents: 61115
diff changeset
   343
62390
842917225d56 more canonical names
nipkow
parents: 62324
diff changeset
   344
end