| author | blanchet | 
| Thu, 26 Jun 2014 13:35:30 +0200 | |
| changeset 57366 | d01d1befe4a3 | 
| parent 47308 | 9caab698dbe4 | 
| child 58889 | 5b7a9633cfa8 | 
| permissions | -rw-r--r-- | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 1 | (* Title: HOL/Quotient_Examples/DList.thy | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 2 | Author: Cezary Kaliszyk, University of Tsukuba | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 3 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 4 | Based on typedef version "Library/Dlist" by Florian Haftmann | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 5 | and theory morphism version by Maksym Bortin | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 6 | *) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 7 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 8 | header {* Lists with distinct elements *}
 | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 9 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 10 | theory DList | 
| 45990 
b7b905b23b2a
incorporated More_Set and More_List into the Main body -- to be consolidated later
 haftmann parents: 
45129diff
changeset | 11 | imports "~~/src/HOL/Library/Quotient_List" | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 12 | begin | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 13 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 14 | text {* Some facts about lists *}
 | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 15 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 16 | lemma remdups_removeAll_commute[simp]: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 17 | "remdups (removeAll x l) = removeAll x (remdups l)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 18 | by (induct l) auto | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 19 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 20 | lemma removeAll_distinct[simp]: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 21 | assumes "distinct l" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 22 | shows "distinct (removeAll x l)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 23 | using assms by (induct l) simp_all | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 24 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 25 | lemma removeAll_commute: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 26 | "removeAll x (removeAll y l) = removeAll y (removeAll x l)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 27 | by (induct l) auto | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 28 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 29 | lemma remdups_hd_notin_tl: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 30 | "remdups dl = h # t \<Longrightarrow> h \<notin> set t" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 31 | by (induct dl arbitrary: h t) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 32 | (case_tac [!] "a \<in> set dl", auto) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 33 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 34 | lemma remdups_repeat: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 35 | "remdups dl = h # t \<Longrightarrow> t = remdups t" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 36 | by (induct dl arbitrary: h t, case_tac [!] "a \<in> set dl") | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 37 | (simp_all, metis remdups_remdups) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 38 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 39 | lemma remdups_nil_noteq_cons: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 40 | "remdups (h # t) \<noteq> remdups []" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 41 | "remdups [] \<noteq> remdups (h # t)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 42 | by auto | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 43 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 44 | lemma remdups_eq_map_eq: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 45 | assumes "remdups xa = remdups ya" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 46 | shows "remdups (map f xa) = remdups (map f ya)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 47 | using assms | 
| 45129 
1fce03e3e8ad
tuned proofs -- eliminated vacuous "induct arbitrary: ..." situations;
 wenzelm parents: 
44263diff
changeset | 48 | by (induct xa ya rule: list_induct2') | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 49 | (metis (full_types) remdups_nil_noteq_cons(2) remdups_map_remdups)+ | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 50 | |
| 44263 
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43766diff
changeset | 51 | lemma remdups_eq_member_eq: | 
| 
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43766diff
changeset | 52 | assumes "remdups xa = remdups ya" | 
| 
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43766diff
changeset | 53 | shows "List.member xa = List.member ya" | 
| 
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43766diff
changeset | 54 | using assms | 
| 
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43766diff
changeset | 55 | unfolding fun_eq_iff List.member_def | 
| 
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43766diff
changeset | 56 | by (induct xa ya rule: list_induct2') | 
| 
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43766diff
changeset | 57 | (metis remdups_nil_noteq_cons set_remdups)+ | 
| 
971d1be5d5ce
Quotient_Examples/DList: explicit proof of remdups_eq_member_eq needed for explicit set type.
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: 
43766diff
changeset | 58 | |
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 59 | text {* Setting up the quotient type *}
 | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 60 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 61 | definition | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 62 | dlist_eq :: "'a list \<Rightarrow> 'a list \<Rightarrow> bool" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 63 | where | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 64 | [simp]: "dlist_eq xs ys \<longleftrightarrow> remdups xs = remdups ys" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 65 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 66 | lemma dlist_eq_reflp: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 67 | "reflp dlist_eq" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 68 | by (auto intro: reflpI) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 69 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 70 | lemma dlist_eq_symp: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 71 | "symp dlist_eq" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 72 | by (auto intro: sympI) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 73 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 74 | lemma dlist_eq_transp: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 75 | "transp dlist_eq" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 76 | by (auto intro: transpI) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 77 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 78 | lemma dlist_eq_equivp: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 79 | "equivp dlist_eq" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 80 | by (auto intro: equivpI dlist_eq_reflp dlist_eq_symp dlist_eq_transp) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 81 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 82 | quotient_type | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 83 | 'a dlist = "'a list" / "dlist_eq" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 84 | by (rule dlist_eq_equivp) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 85 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 86 | text {* respectfulness and constant definitions *}
 | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 87 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 88 | definition [simp]: "card_remdups = length \<circ> remdups" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 89 | definition [simp]: "foldr_remdups f xs e = foldr f (remdups xs) e" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 90 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 91 | quotient_definition empty where "empty :: 'a dlist" | 
| 47092 | 92 | is "Nil" done | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 93 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 94 | quotient_definition insert where "insert :: 'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" | 
| 47092 | 95 | is "Cons" by (metis (mono_tags) List.insert_def dlist_eq_def remdups.simps(2) set_remdups) | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 96 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 97 | quotient_definition "member :: 'a dlist \<Rightarrow> 'a \<Rightarrow> bool" | 
| 47092 | 98 | is "List.member" by (metis dlist_eq_def remdups_eq_member_eq) | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 99 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 100 | quotient_definition foldr where "foldr :: ('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b"
 | 
| 47092 | 101 | is "foldr_remdups" by auto | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 102 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 103 | quotient_definition "remove :: 'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" | 
| 47092 | 104 | is "removeAll" by force | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 105 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 106 | quotient_definition card where "card :: 'a dlist \<Rightarrow> nat" | 
| 47092 | 107 | is "card_remdups" by fastforce | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 108 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 109 | quotient_definition map where "map :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist"
 | 
| 47092 | 110 |   is "List.map :: ('a \<Rightarrow> 'b) \<Rightarrow> 'a list \<Rightarrow> 'b list" by (metis dlist_eq_def remdups_eq_map_eq)
 | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 111 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 112 | quotient_definition filter where "filter :: ('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist"
 | 
| 47092 | 113 | is "List.filter" by (metis dlist_eq_def remdups_filter) | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 114 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 115 | quotient_definition "list_of_dlist :: 'a dlist \<Rightarrow> 'a list" | 
| 47092 | 116 | is "remdups" by simp | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 117 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 118 | text {* lifted theorems *}
 | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 119 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 120 | lemma dlist_member_insert: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 121 | "member dl x \<Longrightarrow> insert x dl = dl" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 122 | by descending (simp add: List.member_def) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 123 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 124 | lemma dlist_member_insert_eq: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 125 | "member (insert y dl) x = (x = y \<or> member dl x)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 126 | by descending (simp add: List.member_def) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 127 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 128 | lemma dlist_insert_idem: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 129 | "insert x (insert x dl) = insert x dl" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 130 | by descending simp | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 131 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 132 | lemma dlist_insert_not_empty: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 133 | "insert x dl \<noteq> empty" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 134 | by descending auto | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 135 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 136 | lemma not_dlist_member_empty: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 137 | "\<not> member empty x" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 138 | by descending (simp add: List.member_def) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 139 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 140 | lemma not_dlist_member_remove: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 141 | "\<not> member (remove x dl) x" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 142 | by descending (simp add: List.member_def) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 143 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 144 | lemma dlist_in_remove: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 145 | "a \<noteq> b \<Longrightarrow> member (remove b dl) a = member dl a" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 146 | by descending (simp add: List.member_def) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 147 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 148 | lemma dlist_not_memb_remove: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 149 | "\<not> member dl x \<Longrightarrow> remove x dl = dl" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 150 | by descending (simp add: List.member_def) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 151 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 152 | lemma dlist_no_memb_remove_insert: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 153 | "\<not> member dl x \<Longrightarrow> remove x (insert x dl) = dl" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 154 | by descending (simp add: List.member_def) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 155 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 156 | lemma dlist_remove_empty: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 157 | "remove x empty = empty" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 158 | by descending simp | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 159 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 160 | lemma dlist_remove_insert_commute: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 161 | "a \<noteq> b \<Longrightarrow> remove a (insert b dl) = insert b (remove a dl)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 162 | by descending simp | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 163 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 164 | lemma dlist_remove_commute: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 165 | "remove a (remove b dl) = remove b (remove a dl)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 166 | by (lifting removeAll_commute) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 167 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 168 | lemma dlist_foldr_empty: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 169 | "foldr f empty e = e" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 170 | by descending simp | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 171 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 172 | lemma dlist_no_memb_foldr: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 173 | assumes "\<not> member dl x" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 174 | shows "foldr f (insert x dl) e = f x (foldr f dl e)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 175 | using assms by descending (simp add: List.member_def) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 176 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 177 | lemma dlist_foldr_insert_not_memb: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 178 | assumes "\<not>member t h" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 179 | shows "foldr f (insert h t) e = f h (foldr f t e)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 180 | using assms by descending (simp add: List.member_def) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 181 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 182 | lemma list_of_dlist_empty[simp]: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 183 | "list_of_dlist empty = []" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 184 | by descending simp | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 185 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 186 | lemma list_of_dlist_insert[simp]: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 187 | "list_of_dlist (insert x xs) = | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 188 | (if member xs x then (remdups (list_of_dlist xs)) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 189 | else x # (remdups (list_of_dlist xs)))" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 190 | by descending (simp add: List.member_def remdups_remdups) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 191 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 192 | lemma list_of_dlist_remove[simp]: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 193 | "list_of_dlist (remove x dxs) = remove1 x (list_of_dlist dxs)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 194 | by descending (simp add: distinct_remove1_removeAll) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 195 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 196 | lemma list_of_dlist_map[simp]: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 197 | "list_of_dlist (map f dxs) = remdups (List.map f (list_of_dlist dxs))" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 198 | by descending (simp add: remdups_map_remdups) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 199 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 200 | lemma list_of_dlist_filter [simp]: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 201 | "list_of_dlist (filter P dxs) = List.filter P (list_of_dlist dxs)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 202 | by descending (simp add: remdups_filter) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 203 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 204 | lemma dlist_map_empty: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 205 | "map f empty = empty" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 206 | by descending simp | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 207 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 208 | lemma dlist_map_insert: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 209 | "map f (insert x xs) = insert (f x) (map f xs)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 210 | by descending simp | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 211 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 212 | lemma dlist_eq_iff: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 213 | "dxs = dys \<longleftrightarrow> list_of_dlist dxs = list_of_dlist dys" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 214 | by descending simp | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 215 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 216 | lemma dlist_eqI: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 217 | "list_of_dlist dxs = list_of_dlist dys \<Longrightarrow> dxs = dys" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 218 | by (simp add: dlist_eq_iff) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 219 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 220 | abbreviation | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 221 | "dlist xs \<equiv> abs_dlist xs" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 222 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 223 | lemma distinct_list_of_dlist [simp, intro]: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 224 | "distinct (list_of_dlist dxs)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 225 | by descending simp | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 226 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 227 | lemma list_of_dlist_dlist [simp]: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 228 | "list_of_dlist (dlist xs) = remdups xs" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 229 | unfolding list_of_dlist_def map_fun_apply id_def | 
| 47308 | 230 | by (metis Quotient3_rep_abs[OF Quotient3_dlist] dlist_eq_def) | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 231 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 232 | lemma remdups_list_of_dlist [simp]: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 233 | "remdups (list_of_dlist dxs) = list_of_dlist dxs" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 234 | by simp | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 235 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 236 | lemma dlist_list_of_dlist [simp, code abstype]: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 237 | "dlist (list_of_dlist dxs) = dxs" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 238 | by (simp add: list_of_dlist_def) | 
| 47308 | 239 | (metis Quotient3_def Quotient3_dlist dlist_eqI list_of_dlist_dlist remdups_list_of_dlist) | 
| 43766 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 240 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 241 | lemma dlist_filter_simps: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 242 | "filter P empty = empty" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 243 | "filter P (insert x xs) = (if P x then insert x (filter P xs) else filter P xs)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 244 | by (lifting filter.simps) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 245 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 246 | lemma dlist_induct: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 247 | assumes "P empty" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 248 | and "\<And>a dl. P dl \<Longrightarrow> P (insert a dl)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 249 | shows "P dl" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 250 | using assms by descending (drule list.induct, simp) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 251 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 252 | lemma dlist_induct_stronger: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 253 | assumes a1: "P empty" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 254 | and a2: "\<And>x dl. \<lbrakk>\<not>member dl x; P dl\<rbrakk> \<Longrightarrow> P (insert x dl)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 255 | shows "P dl" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 256 | proof(induct dl rule: dlist_induct) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 257 | show "P empty" using a1 by simp | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 258 | next | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 259 | fix x dl | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 260 | assume "P dl" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 261 | then show "P (insert x dl)" using a2 | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 262 | by (cases "member dl x") (simp_all add: dlist_member_insert) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 263 | qed | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 264 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 265 | lemma dlist_card_induct: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 266 | assumes "\<And>xs. (\<And>ys. card ys < card xs \<Longrightarrow> P ys) \<Longrightarrow> P xs" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 267 | shows "P xs" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 268 | using assms | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 269 | by descending (rule measure_induct [of card_remdups], blast) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 270 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 271 | lemma dlist_cases: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 272 | "dl = empty \<or> (\<exists>h t. dl = insert h t \<and> \<not> member t h)" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 273 | apply (descending, simp add: List.member_def) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 274 | by (metis list.exhaust remdups_eq_nil_iff remdups_hd_notin_tl remdups_repeat) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 275 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 276 | lemma dlist_exhaust: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 277 | assumes "y = empty \<Longrightarrow> P" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 278 | and "\<And>a dl. y = insert a dl \<Longrightarrow> P" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 279 | shows "P" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 280 | using assms by (lifting list.exhaust) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 281 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 282 | lemma dlist_exhaust_stronger: | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 283 | assumes "y = empty \<Longrightarrow> P" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 284 | and "\<And>a dl. y = insert a dl \<Longrightarrow> \<not> member dl a \<Longrightarrow> P" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 285 | shows "P" | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 286 | using assms by (metis dlist_cases) | 
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 287 | |
| 
9545bb3cefac
Quotient example: Lists with distinct elements
 Cezary Kaliszyk <kaliszyk@in.tum.de> parents: diff
changeset | 288 | end |