src/HOL/Quotient_Examples/Lift_DList.thy
author wenzelm
Thu, 22 Dec 2016 11:08:58 +0100
changeset 64655 ea34f36ff6a5
parent 63167 0909deb8059b
child 70009 435fb018e8ee
permissions -rw-r--r--
clarified message;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     1
(*  Title:      HOL/Quotient_Examples/Lift_DList.thy
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     2
    Author:     Ondrej Kuncar
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     3
*)
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     4
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     5
theory Lift_DList
53013
3fbcfa911863 remove unnecessary dependencies on Library/Quotient_*
kuncar
parents: 49834
diff changeset
     6
imports Main
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     7
begin
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
     8
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 55732
diff changeset
     9
subsection \<open>The type of distinct lists\<close>
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    10
49834
b27bbb021df1 discontinued obsolete typedef (open) syntax;
wenzelm
parents: 47361
diff changeset
    11
typedef 'a dlist = "{xs::'a list. distinct xs}"
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    12
  morphisms list_of_dlist Abs_dlist
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    13
proof
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    14
  show "[] \<in> {xs. distinct xs}" by simp
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    15
qed
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    16
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    17
setup_lifting type_definition_dlist
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    18
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 55732
diff changeset
    19
text \<open>Fundamental operations:\<close>
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    20
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    21
lift_definition empty :: "'a dlist" is "[]"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    22
by simp
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    23
  
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    24
lift_definition insert :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" is List.insert
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    25
by simp
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    26
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    27
lift_definition remove :: "'a \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" is List.remove1
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    28
by simp
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    29
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    30
lift_definition map :: "('a \<Rightarrow> 'b) \<Rightarrow> 'a dlist \<Rightarrow> 'b dlist" is "\<lambda>f. remdups o List.map f"
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    31
by simp
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    32
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    33
lift_definition filter :: "('a \<Rightarrow> bool) \<Rightarrow> 'a dlist \<Rightarrow> 'a dlist" is List.filter
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    34
by simp
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    35
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 55732
diff changeset
    36
text \<open>Derived operations:\<close>
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    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
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    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
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    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
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    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
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    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
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    47
55732
07906fc6af7a simplify and repair proofs due to df0fda378813
kuncar
parents: 55565
diff changeset
    48
lift_definition concat :: "'a dlist dlist \<Rightarrow> 'a dlist" is "remdups o List.concat" by auto
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    49
63167
0909deb8059b isabelle update_cartouches -c -t;
wenzelm
parents: 55732
diff changeset
    50
text \<open>We can export code:\<close>
47308
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    51
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    52
export_code empty insert remove map filter null member length fold foldr concat in SML
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    53
9caab698dbe4 new package Lifting - initial commit
kuncar
parents:
diff changeset
    54
end