src/HOL/Library/Dlist.thy
author haftmann
Tue, 17 Jul 2012 23:11:24 +0200
changeset 49297 39bfb2844b9e
parent 47436 ad21900e0ee9
child 50849 b27bbb021df1
permissions -rw-r--r--
tuned whitespace
     1 (* Author: Florian Haftmann, TU Muenchen *)
     2 
     3 header {* Lists with elements distinct as canonical example for datatype invariants *}
     4 
     5 theory Dlist
     6 imports Main
     7 begin
     8 
     9 subsection {* The type of distinct lists *}
    10 
    11 typedef (open) 'a dlist = "{xs::'a list. distinct xs}"
    12   morphisms list_of_dlist Abs_dlist
    13 proof
    14   show "[] \<in> {xs. distinct xs}" by simp
    15 qed
    16 
    17 lemma dlist_eq_iff:
    18   "dxs = dys \<longleftrightarrow> list_of_dlist dxs = list_of_dlist dys"
    19   by (simp add: list_of_dlist_inject)
    20 
    21 lemma dlist_eqI:
    22   "list_of_dlist dxs = list_of_dlist dys \<Longrightarrow> dxs = dys"
    23   by (simp add: dlist_eq_iff)
    24 
    25 text {* Formal, totalized constructor for @{typ "'a dlist"}: *}
    26 
    27 definition Dlist :: "'a list \<Rightarrow> 'a dlist" where
    28   "Dlist xs = Abs_dlist (remdups xs)"
    29 
    30 lemma distinct_list_of_dlist [simp, intro]:
    31   "distinct (list_of_dlist dxs)"
    32   using list_of_dlist [of dxs] by simp
    33 
    34 lemma list_of_dlist_Dlist [simp]:
    35   "list_of_dlist (Dlist xs) = remdups xs"
    36   by (simp add: Dlist_def Abs_dlist_inverse)
    37 
    38 lemma remdups_list_of_dlist [simp]:
    39   "remdups (list_of_dlist dxs) = list_of_dlist dxs"
    40   by simp
    41 
    42 lemma Dlist_list_of_dlist [simp, code abstype]:
    43   "Dlist (list_of_dlist dxs) = dxs"
    44   by (simp add: Dlist_def list_of_dlist_inverse distinct_remdups_id)
    45 
    46 
    47 text {* Fundamental operations: *}
    48 
    49 definition empty :: "'a dlist" where
    50   "empty = Dlist []"
    51 
    52 definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
    53   "insert x dxs = Dlist (List.insert x (list_of_dlist dxs))"
    54 
    55 definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
    56   "remove x dxs = Dlist (remove1 x (list_of_dlist dxs))"
    57 
    58 definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" where
    59   "map f dxs = Dlist (remdups (List.map f (list_of_dlist dxs)))"
    60 
    61 definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" where
    62   "filter P dxs = Dlist (List.filter P (list_of_dlist dxs))"
    63 
    64 
    65 text {* Derived operations: *}
    66 
    67 definition null :: "'a dlist \<Rightarrow> bool" where
    68   "null dxs = List.null (list_of_dlist dxs)"
    69 
    70 definition member :: "'a dlist \<Rightarrow> 'a \<Rightarrow> bool" where
    71   "member dxs = List.member (list_of_dlist dxs)"
    72 
    73 definition length :: "'a dlist \<Rightarrow> nat" where
    74   "length dxs = List.length (list_of_dlist dxs)"
    75 
    76 definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
    77   "fold f dxs = List.fold f (list_of_dlist dxs)"
    78 
    79 definition foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" where
    80   "foldr f dxs = List.foldr f (list_of_dlist dxs)"
    81 
    82 
    83 subsection {* Executable version obeying invariant *}
    84 
    85 lemma list_of_dlist_empty [simp, code abstract]:
    86   "list_of_dlist empty = []"
    87   by (simp add: empty_def)
    88 
    89 lemma list_of_dlist_insert [simp, code abstract]:
    90   "list_of_dlist (insert x dxs) = List.insert x (list_of_dlist dxs)"
    91   by (simp add: insert_def)
    92 
    93 lemma list_of_dlist_remove [simp, code abstract]:
    94   "list_of_dlist (remove x dxs) = remove1 x (list_of_dlist dxs)"
    95   by (simp add: remove_def)
    96 
    97 lemma list_of_dlist_map [simp, code abstract]:
    98   "list_of_dlist (map f dxs) = remdups (List.map f (list_of_dlist dxs))"
    99   by (simp add: map_def)
   100 
   101 lemma list_of_dlist_filter [simp, code abstract]:
   102   "list_of_dlist (filter P dxs) = List.filter P (list_of_dlist dxs)"
   103   by (simp add: filter_def)
   104 
   105 
   106 text {* Explicit executable conversion *}
   107 
   108 definition dlist_of_list [simp]:
   109   "dlist_of_list = Dlist"
   110 
   111 lemma [code abstract]:
   112   "list_of_dlist (dlist_of_list xs) = remdups xs"
   113   by simp
   114 
   115 
   116 text {* Equality *}
   117 
   118 instantiation dlist :: (equal) equal
   119 begin
   120 
   121 definition "HOL.equal dxs dys \<longleftrightarrow> HOL.equal (list_of_dlist dxs) (list_of_dlist dys)"
   122 
   123 instance proof
   124 qed (simp add: equal_dlist_def equal list_of_dlist_inject)
   125 
   126 end
   127 
   128 declare equal_dlist_def [code]
   129 
   130 lemma [code nbe]:
   131   "HOL.equal (dxs :: 'a::equal dlist) dxs \<longleftrightarrow> True"
   132   by (fact equal_refl)
   133 
   134 
   135 subsection {* Induction principle and case distinction *}
   136 
   137 lemma dlist_induct [case_names empty insert, induct type: dlist]:
   138   assumes empty: "P empty"
   139   assumes insrt: "\<And>x dxs. \<not> member dxs x \<Longrightarrow> P dxs \<Longrightarrow> P (insert x dxs)"
   140   shows "P dxs"
   141 proof (cases dxs)
   142   case (Abs_dlist xs)
   143   then have "distinct xs" and dxs: "dxs = Dlist xs" by (simp_all add: Dlist_def distinct_remdups_id)
   144   from `distinct xs` have "P (Dlist xs)"
   145   proof (induct xs)
   146     case Nil from empty show ?case by (simp add: empty_def)
   147   next
   148     case (Cons x xs)
   149     then have "\<not> member (Dlist xs) x" and "P (Dlist xs)"
   150       by (simp_all add: member_def List.member_def)
   151     with insrt have "P (insert x (Dlist xs))" .
   152     with Cons show ?case by (simp add: insert_def distinct_remdups_id)
   153   qed
   154   with dxs show "P dxs" by simp
   155 qed
   156 
   157 lemma dlist_case [case_names empty insert, cases type: dlist]:
   158   assumes empty: "dxs = empty \<Longrightarrow> P"
   159   assumes insert: "\<And>x dys. \<not> member dys x \<Longrightarrow> dxs = insert x dys \<Longrightarrow> P"
   160   shows P
   161 proof (cases dxs)
   162   case (Abs_dlist xs)
   163   then have dxs: "dxs = Dlist xs" and distinct: "distinct xs"
   164     by (simp_all add: Dlist_def distinct_remdups_id)
   165   show P proof (cases xs)
   166     case Nil with dxs have "dxs = empty" by (simp add: empty_def) 
   167     with empty show P .
   168   next
   169     case (Cons x xs)
   170     with dxs distinct have "\<not> member (Dlist xs) x"
   171       and "dxs = insert x (Dlist xs)"
   172       by (simp_all add: member_def List.member_def insert_def distinct_remdups_id)
   173     with insert show P .
   174   qed
   175 qed
   176 
   177 
   178 subsection {* Functorial structure *}
   179 
   180 enriched_type map: map
   181   by (simp_all add: List.map.id remdups_map_remdups fun_eq_iff dlist_eq_iff)
   182 
   183 
   184 subsection {* Quickcheck generators *}
   185 
   186 quickcheck_generator dlist predicate: distinct constructors: empty, insert
   187 
   188 
   189 hide_const (open) member fold foldr empty insert remove map filter null member length fold
   190 
   191 end