| author | blanchet | 
| Wed, 19 Feb 2014 16:32:37 +0100 | |
| changeset 55584 | a879f14b6f95 | 
| parent 55565 | f663fc1e653b | 
| child 55732 | 07906fc6af7a | 
| permissions | -rw-r--r-- | 
| 47308 | 1 | (* Title: HOL/Quotient_Examples/Lift_DList.thy | 
| 2 | Author: Ondrej Kuncar | |
| 3 | *) | |
| 4 | ||
| 5 | theory Lift_DList | |
| 53013 
3fbcfa911863
remove unnecessary dependencies on Library/Quotient_*
 kuncar parents: 
49834diff
changeset | 6 | imports Main | 
| 47308 | 7 | begin | 
| 8 | ||
| 9 | subsection {* The type of distinct lists *}
 | |
| 10 | ||
| 49834 | 11 | typedef 'a dlist = "{xs::'a list. distinct xs}"
 | 
| 47308 | 12 | morphisms list_of_dlist Abs_dlist | 
| 13 | proof | |
| 14 |   show "[] \<in> {xs. distinct xs}" by simp
 | |
| 15 | qed | |
| 16 | ||
| 17 | setup_lifting type_definition_dlist | |
| 18 | ||
| 19 | text {* Fundamental operations: *}
 | |
| 20 | ||
| 21 | lift_definition empty :: "'a dlist" is "[]" | |
| 22 | by simp | |
| 23 | ||
| 24 | lift_definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" is List.insert | |
| 25 | by simp | |
| 26 | ||
| 27 | lift_definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" is List.remove1 | |
| 28 | by simp | |
| 29 | ||
| 30 | lift_definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" is "\<lambda>f. remdups o List.map f"
 | |
| 31 | by simp | |
| 32 | ||
| 33 | lift_definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" is List.filter
 | |
| 34 | by simp | |
| 35 | ||
| 36 | text {* Derived operations: *}
 | |
| 37 | ||
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55524diff
changeset | 38 | lift_definition null :: "'a dlist \<Rightarrow> bool" is List.null . | 
| 47308 | 39 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55524diff
changeset | 40 | lift_definition member :: "'a dlist \<Rightarrow> 'a \<Rightarrow> bool" is List.member . | 
| 47308 | 41 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55524diff
changeset | 42 | lift_definition length :: "'a dlist \<Rightarrow> nat" is List.length . | 
| 47308 | 43 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55524diff
changeset | 44 | lift_definition fold :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" is List.fold .
 | 
| 47308 | 45 | |
| 55565 
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
 kuncar parents: 
55524diff
changeset | 46 | lift_definition foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" is List.foldr .
 | 
| 47308 | 47 | |
| 48 | lift_definition concat :: "'a dlist dlist \<Rightarrow> 'a dlist" is "remdups o List.concat" | |
| 49 | proof - | |
| 50 |   {
 | |
| 51 | fix x y | |
| 47361 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 kuncar parents: 
47308diff
changeset | 52 | have "list_all2 cr_dlist x y \<Longrightarrow> x = List.map list_of_dlist y" | 
| 55524 
f41ef840f09d
folded 'list_all2' with the relator generated by 'datatype_new'
 blanchet parents: 
53013diff
changeset | 53 | unfolding list_all2_iff cr_dlist_def by (induction x y rule: list_induct2') auto | 
| 47308 | 54 | } | 
| 55 | note cr = this | |
| 56 | fix x :: "'a list list" and y :: "'a list list" | |
| 47361 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 kuncar parents: 
47308diff
changeset | 57 | assume "(list_all2 cr_dlist OO Lifting.invariant distinct OO (list_all2 cr_dlist)\<inverse>\<inverse>) x y" | 
| 
87c0eaf04bad
support non-open typedefs; define cr_rel in terms of a rep function for typedefs
 kuncar parents: 
47308diff
changeset | 58 | then have "x = y" by (auto dest: cr simp add: Lifting.invariant_def) | 
| 47308 | 59 | then show "?thesis x y" unfolding Lifting.invariant_def by auto | 
| 60 | qed | |
| 61 | ||
| 62 | text {* We can export code: *}
 | |
| 63 | ||
| 64 | export_code empty insert remove map filter null member length fold foldr concat in SML | |
| 65 | ||
| 66 | end |