author | Cezary Kaliszyk <kaliszyk@in.tum.de> |
Tue, 05 Jul 2011 09:28:16 +0900 | |
changeset 2941 | 40991ebcda12 |
parent 2868 | 2b8e387d2dfc |
child 2950 | 0911cb7bf696 |
permissions | -rw-r--r-- |
2559
add799cf0817
adapted to changes by Florian on the quotient package and removed local fix for function package
Christian Urban <urbanc@in.tum.de>
parents:
2557
diff
changeset
|
1 |
|
2485
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
2 |
theory Datatypes |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
3 |
imports "../Nominal2" |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
4 |
begin |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
5 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
6 |
(* |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
7 |
plain nominal_datatype definition without an |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
8 |
atom_decl - example by John Matthews |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
9 |
*) |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
10 |
|
2557
781fbc8c0591
fixed locally the problem with the function package; all tests work again
Christian Urban <urbanc@in.tum.de>
parents:
2556
diff
changeset
|
11 |
|
2485
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
12 |
nominal_datatype 'a Maybe = |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
13 |
Nothing |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
14 |
| Just 'a |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
15 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
16 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
17 |
thm Maybe.distinct |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
18 |
thm Maybe.induct |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
19 |
thm Maybe.exhaust |
2617
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents:
2559
diff
changeset
|
20 |
thm Maybe.strong_exhaust |
2485
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
21 |
thm Maybe.fv_defs |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
22 |
thm Maybe.bn_defs |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
23 |
thm Maybe.perm_simps |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
24 |
thm Maybe.eq_iff |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
25 |
thm Maybe.fv_bn_eqvt |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
26 |
thm Maybe.size_eqvt |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
27 |
thm Maybe.supp |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
28 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
29 |
(* |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
30 |
using a datatype inside a nominal_datatype |
2617
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents:
2559
diff
changeset
|
31 |
|
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents:
2559
diff
changeset
|
32 |
the datatype needs to be shown to be an instance |
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents:
2559
diff
changeset
|
33 |
of the pure class (this is usually trivial) |
2485
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
34 |
*) |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
35 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
36 |
atom_decl name |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
37 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
38 |
datatype foo = Foo | Bar |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
39 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
40 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
41 |
instantiation foo :: pure |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
42 |
begin |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
43 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
44 |
definition |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
45 |
"p \<bullet> (f::foo) = f" |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
46 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
47 |
instance |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
48 |
apply(default) |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
49 |
apply(simp_all add: permute_foo_def) |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
50 |
done |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
51 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
52 |
end |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
53 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
54 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
55 |
nominal_datatype baz = |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
56 |
Baz x::name t::foo bind x in t |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
57 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
58 |
|
2868
2b8e387d2dfc
got rid of the boolean flag in the raw_equivariance function
Christian Urban <urbanc@in.tum.de>
parents:
2781
diff
changeset
|
59 |
|
2485
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
60 |
thm baz.distinct |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
61 |
thm baz.induct |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
62 |
thm baz.exhaust |
2617
e44551d067e6
properly exported strong exhaust theorem; cleaned up some examples
Christian Urban <urbanc@in.tum.de>
parents:
2559
diff
changeset
|
63 |
thm baz.strong_exhaust |
2485
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
64 |
thm baz.fv_defs |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
65 |
thm baz.bn_defs |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
66 |
thm baz.perm_simps |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
67 |
thm baz.eq_iff |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
68 |
thm baz.fv_bn_eqvt |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
69 |
thm baz.size_eqvt |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
70 |
thm baz.supp |
2777
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
71 |
|
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
72 |
(* |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
73 |
a nominal datatype with a set argument of |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
74 |
pure type |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
75 |
*) |
2485
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
76 |
|
2777
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
77 |
nominal_datatype set_ty = |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
78 |
Set_ty "nat set" |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
79 |
| Fun "nat \<Rightarrow> nat" |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
80 |
| Var "name" |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
81 |
| Lam x::"name" t::"set_ty" bind x in t |
2868
2b8e387d2dfc
got rid of the boolean flag in the raw_equivariance function
Christian Urban <urbanc@in.tum.de>
parents:
2781
diff
changeset
|
82 |
thm meta_eq_to_obj_eq |
2777
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
83 |
thm set_ty.distinct |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
84 |
thm set_ty.induct |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
85 |
thm set_ty.exhaust |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
86 |
thm set_ty.strong_exhaust |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
87 |
thm set_ty.fv_defs |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
88 |
thm set_ty.bn_defs |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
89 |
thm set_ty.perm_simps |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
90 |
thm set_ty.eq_iff |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
91 |
thm set_ty.fv_bn_eqvt |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
92 |
thm set_ty.size_eqvt |
75a95431cd8b
proved that lfp is equivariant (that simplifies equivariance proofs of inductively defined predicates)
Christian Urban <urbanc@in.tum.de>
parents:
2617
diff
changeset
|
93 |
thm set_ty.supp |
2485
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
94 |
|
2781
542ff50555f5
updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents:
2777
diff
changeset
|
95 |
|
542ff50555f5
updated to new Isabelle (> 9 May)
Christian Urban <urbanc@in.tum.de>
parents:
2777
diff
changeset
|
96 |
|
2485
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
97 |
end |
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
98 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
99 |
|
6bab47906dbe
added example about datatypes
Christian Urban <urbanc@in.tum.de>
parents:
diff
changeset
|
100 |