author | blanchet |
Fri, 21 Feb 2014 00:09:56 +0100 | |
changeset 55642 | 63beb38e9258 |
parent 55584 | a879f14b6f95 |
child 58889 | 5b7a9633cfa8 |
permissions | -rw-r--r-- |
45576 | 1 |
(* Title: HOL/Quotient_Examples/Lift_Set.thy |
45609 | 2 |
Author: Lukas Bulwahn and Ondrej Kuncar |
45536
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
3 |
*) |
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
4 |
|
47308 | 5 |
header {* Example of lifting definitions with the lifting infrastructure *} |
45536
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
6 |
|
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
7 |
theory Lift_Set |
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
8 |
imports Main |
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
9 |
begin |
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
10 |
|
45970
b6d0cff57d96
adjusted to set/pred distinction by means of type constructor `set`
haftmann
parents:
45694
diff
changeset
|
11 |
definition set where "set = (UNIV :: ('a \<Rightarrow> bool) set)" |
45694
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
45609
diff
changeset
|
12 |
|
49834 | 13 |
typedef 'a set = "set :: ('a \<Rightarrow> bool) set" |
45694
4a8743618257
prefer typedef without extra definition and alternative name;
wenzelm
parents:
45609
diff
changeset
|
14 |
morphisms member Set |
55584 | 15 |
unfolding set_def by (rule UNIV_witness) |
45536
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
16 |
|
47097 | 17 |
setup_lifting type_definition_set[unfolded set_def] |
45536
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
18 |
|
47308 | 19 |
text {* Now, we can employ lift_definition to lift definitions. *} |
45536
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
20 |
|
55584 | 21 |
lift_definition empty :: "'a set" is "bot :: 'a \<Rightarrow> bool" . |
45536
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
22 |
|
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
23 |
term "Lift_Set.empty" |
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
24 |
thm Lift_Set.empty_def |
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
25 |
|
55584 | 26 |
lift_definition insert :: "'a => 'a set => 'a set" is "\<lambda> x P y. y = x \<or> P y" . |
45536
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
27 |
|
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
28 |
term "Lift_Set.insert" |
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
29 |
thm Lift_Set.insert_def |
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
30 |
|
47097 | 31 |
export_code empty insert in SML |
32 |
||
45536
5b0b1dc2e40f
adding a preliminary example to show how the quotient_definition package can be generalized
bulwahn
parents:
diff
changeset
|
33 |
end |