author | wenzelm |
Mon, 26 Jun 2023 23:20:32 +0200 | |
changeset 78209 | 50c5be88ad59 |
parent 70009 | 435fb018e8ee |
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:
49834
diff
changeset
|
6 |
imports Main |
47308 | 7 |
begin |
8 |
||
63167 | 9 |
subsection \<open>The type of distinct lists\<close> |
47308 | 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 |
||
63167 | 19 |
text \<open>Fundamental operations:\<close> |
47308 | 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 |
||
63167 | 36 |
text \<open>Derived operations:\<close> |
47308 | 37 |
|
55565
f663fc1e653b
simplify proofs because of the stronger reflexivity prover
kuncar
parents:
55524
diff
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:
55524
diff
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:
55524
diff
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:
55524
diff
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:
55524
diff
changeset
|
46 |
lift_definition foldr :: "('a \<Rightarrow> 'b \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b \<Rightarrow> 'b" is List.foldr . |
47308 | 47 |
|
55732 | 48 |
lift_definition concat :: "'a dlist dlist \<Rightarrow> 'a dlist" is "remdups o List.concat" by auto |
47308 | 49 |
|
63167 | 50 |
text \<open>We can export code:\<close> |
47308 | 51 |
|
52 |
export_code empty insert remove map filter null member length fold foldr concat in SML |
|
70009
435fb018e8ee
"export_code ... file_prefix ..." is the preferred way to produce output within the logical file-system within the theory context, as well as session exports;
wenzelm
parents:
63167
diff
changeset
|
53 |
file_prefix dlist |
47308 | 54 |
|
55 |
end |