author | Christian Urban <urbanc@in.tum.de> |
Wed, 31 Mar 2010 22:48:35 +0200 | |
changeset 1737 | 8b6a285ad480 |
parent 1677 | ba3f6e33d647 |
permissions | -rw-r--r-- |
1599
8b5a1ad60487
Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1596
diff
changeset
|
1 |
theory ExTySch |
8b5a1ad60487
Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1596
diff
changeset
|
2 |
imports "Parser" |
1271 | 3 |
begin |
4 |
||
1599
8b5a1ad60487
Move Leroy out of Test, rename accordingly.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1596
diff
changeset
|
5 |
(* Type Schemes *) |
1271 | 6 |
atom_decl name |
7 |
||
1676
141a36535f1d
Manually proved TySch support; All properties of TySch now true.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1673
diff
changeset
|
8 |
ML {* val _ = alpha_type := AlphaRes *} |
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
9 |
nominal_datatype t = |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
10 |
Var "name" |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
11 |
| Fun "t" "t" |
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
12 |
and tyS = |
1525
bf321f16d025
Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1515
diff
changeset
|
13 |
All xs::"name fset" ty::"t" bind xs in ty |
1271 | 14 |
|
1677
ba3f6e33d647
Automatically compute support if only one type of Abs is present in the type.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1676
diff
changeset
|
15 |
lemmas t_tyS_supp = t_tyS.fv[simplified t_tyS.supp] |
1676
141a36535f1d
Manually proved TySch support; All properties of TySch now true.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1673
diff
changeset
|
16 |
|
1568
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
17 |
lemma size_eqvt_raw: |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
18 |
"size (pi \<bullet> t :: t_raw) = size t" |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
19 |
"size (pi \<bullet> ts :: tyS_raw) = size ts" |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
20 |
apply (induct rule: t_raw_tyS_raw.inducts) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
21 |
apply simp_all |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
22 |
done |
1562 | 23 |
|
24 |
instantiation t and tyS :: size begin |
|
25 |
||
26 |
quotient_definition |
|
27 |
"size_t :: t \<Rightarrow> nat" |
|
28 |
is |
|
29 |
"size :: t_raw \<Rightarrow> nat" |
|
30 |
||
31 |
quotient_definition |
|
32 |
"size_tyS :: tyS \<Rightarrow> nat" |
|
33 |
is |
|
34 |
"size :: tyS_raw \<Rightarrow> nat" |
|
35 |
||
36 |
lemma size_rsp: |
|
37 |
"alpha_t_raw x y \<Longrightarrow> size x = size y" |
|
38 |
"alpha_tyS_raw a b \<Longrightarrow> size a = size b" |
|
39 |
apply (induct rule: alpha_t_raw_alpha_tyS_raw.inducts) |
|
40 |
apply (simp_all only: t_raw_tyS_raw.size) |
|
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1670
diff
changeset
|
41 |
apply (simp_all only: alphas) |
1562 | 42 |
apply clarify |
1568
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
43 |
apply (simp_all only: size_eqvt_raw) |
1562 | 44 |
done |
45 |
||
46 |
lemma [quot_respect]: |
|
47 |
"(alpha_t_raw ===> op =) size size" |
|
48 |
"(alpha_tyS_raw ===> op =) size size" |
|
49 |
by (simp_all add: size_rsp) |
|
50 |
||
51 |
lemma [quot_preserve]: |
|
52 |
"(rep_t ---> id) size = size" |
|
53 |
"(rep_tyS ---> id) size = size" |
|
54 |
by (simp_all add: size_t_def size_tyS_def) |
|
55 |
||
56 |
instance |
|
57 |
by default |
|
58 |
||
59 |
end |
|
60 |
||
61 |
thm t_raw_tyS_raw.size(4)[quot_lifted] |
|
62 |
thm t_raw_tyS_raw.size(5)[quot_lifted] |
|
63 |
thm t_raw_tyS_raw.size(6)[quot_lifted] |
|
64 |
||
65 |
||
1515
76fa21f27f22
Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
66 |
thm t_tyS.fv |
76fa21f27f22
Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
67 |
thm t_tyS.eq_iff |
76fa21f27f22
Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
68 |
thm t_tyS.bn |
76fa21f27f22
Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
69 |
thm t_tyS.perm |
1525
bf321f16d025
Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1515
diff
changeset
|
70 |
thm t_tyS.inducts |
1515
76fa21f27f22
Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
71 |
thm t_tyS.distinct |
1534
984ea1299cd7
The nominal infrastructure for fset. 'fs' missing, but not needed so far.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1530
diff
changeset
|
72 |
ML {* Sign.of_sort @{theory} (@{typ t}, @{sort fs}) *} |
1430
ccbcebef56f3
Trying to prove atom_image_fresh_swap
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1277
diff
changeset
|
73 |
|
1525
bf321f16d025
Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1515
diff
changeset
|
74 |
lemma induct: |
1537
0b21101157b1
Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1534
diff
changeset
|
75 |
assumes a1: "\<And>name b. P b (Var name)" |
0b21101157b1
Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1534
diff
changeset
|
76 |
and a2: "\<And>t1 t2 b. \<lbrakk>\<And>c. P c t1; \<And>c. P c t2\<rbrakk> \<Longrightarrow> P b (Fun t1 t2)" |
1538
6853ce305118
Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1537
diff
changeset
|
77 |
and a3: "\<And>fset t b. \<lbrakk>\<And>c. P c t; fset_to_set (fmap atom fset) \<sharp>* b\<rbrakk> \<Longrightarrow> P' b (All fset t)" |
1568
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
78 |
shows "P (a :: 'a :: pt) t \<and> P' (d :: 'b :: {fs}) ts " |
1537
0b21101157b1
Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1534
diff
changeset
|
79 |
proof - |
1539
78d0adf8a086
TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1538
diff
changeset
|
80 |
have " (\<forall>p a. P a (p \<bullet> t)) \<and> (\<forall>p d. P' d (p \<bullet> ts))" |
1538
6853ce305118
Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1537
diff
changeset
|
81 |
apply (rule t_tyS.induct) |
1537
0b21101157b1
Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1534
diff
changeset
|
82 |
apply (simp add: a1) |
1539
78d0adf8a086
TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1538
diff
changeset
|
83 |
apply (simp) |
78d0adf8a086
TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1538
diff
changeset
|
84 |
apply (rule allI)+ |
1538
6853ce305118
Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1537
diff
changeset
|
85 |
apply (rule a2) |
1539
78d0adf8a086
TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1538
diff
changeset
|
86 |
apply simp |
78d0adf8a086
TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1538
diff
changeset
|
87 |
apply simp |
78d0adf8a086
TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1538
diff
changeset
|
88 |
apply (rule allI) |
1538
6853ce305118
Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1537
diff
changeset
|
89 |
apply (rule allI) |
1605 | 90 |
apply(subgoal_tac "\<exists>pa. ((pa \<bullet> (fset_to_set (fmap atom (p \<bullet> fset)))) \<sharp>* d \<and> supp (p \<bullet> All fset t) \<sharp>* pa)") |
1539
78d0adf8a086
TySch strong induction looks ok.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1538
diff
changeset
|
91 |
apply clarify |
1605 | 92 |
apply(rule_tac t="p \<bullet> All fset t" and |
93 |
s="pa \<bullet> (p \<bullet> All fset t)" in subst) |
|
1568
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
94 |
apply (rule supp_perm_eq) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
95 |
apply assumption |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
96 |
apply (simp only: t_tyS.perm) |
1538
6853ce305118
Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1537
diff
changeset
|
97 |
apply (rule a3) |
1568
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
98 |
apply(erule_tac x="(pa + p)" in allE) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
99 |
apply simp |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
100 |
apply (simp add: eqvts eqvts_raw) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
101 |
apply (rule at_set_avoiding2) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
102 |
apply (simp add: fin_fset_to_set) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
103 |
apply (simp add: finite_supp) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
104 |
apply (simp add: eqvts finite_supp) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
105 |
apply (subst atom_eqvt_raw[symmetric]) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
106 |
apply (subst fmap_eqvt[symmetric]) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
107 |
apply (subst fset_to_set_eqvt[symmetric]) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
108 |
apply (simp only: fresh_star_permute_iff) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
109 |
apply (simp add: fresh_star_def) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
110 |
apply clarify |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
111 |
apply (simp add: fresh_def) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
112 |
apply (simp add: t_tyS_supp) |
2311a9fc4624
Strong induction for Type Schemes.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1562
diff
changeset
|
113 |
done |
1538
6853ce305118
Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1537
diff
changeset
|
114 |
then have "P a (0 \<bullet> t) \<and> P' d (0 \<bullet> ts)" by blast |
6853ce305118
Working on TySch strong induction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1537
diff
changeset
|
115 |
then show ?thesis by simp |
1537
0b21101157b1
Something is wrong with the statement of strong induction for TySch, as the All case is trivial and Fun case unprovable...
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1534
diff
changeset
|
116 |
qed |
1525
bf321f16d025
Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1515
diff
changeset
|
117 |
|
1271 | 118 |
lemma |
1525
bf321f16d025
Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1515
diff
changeset
|
119 |
shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|b, a|} (Fun (Var a) (Var b))" |
1515
76fa21f27f22
Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
120 |
apply(simp add: t_tyS.eq_iff) |
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
121 |
apply(rule_tac x="0::perm" in exI) |
1673
e8cf0520c820
New compose lemmas. Reverted alpha_gen sym/trans changes. Equivp for alpha_res should work now.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1670
diff
changeset
|
122 |
apply(simp add: alphas) |
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
123 |
apply(simp add: fresh_star_def fresh_zero_perm) |
1271 | 124 |
done |
125 |
||
126 |
lemma |
|
1525
bf321f16d025
Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1515
diff
changeset
|
127 |
shows "All {|a, b|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var b) (Var a))" |
1515
76fa21f27f22
Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
128 |
apply(simp add: t_tyS.eq_iff) |
1271 | 129 |
apply(rule_tac x="(atom a \<rightleftharpoons> atom b)" in exI) |
1676
141a36535f1d
Manually proved TySch support; All properties of TySch now true.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1673
diff
changeset
|
130 |
apply(simp add: alphas fresh_star_def eqvts) |
1271 | 131 |
done |
132 |
||
133 |
lemma |
|
1525
bf321f16d025
Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1515
diff
changeset
|
134 |
shows "All {|a, b, c|} (Fun (Var a) (Var b)) = All {|a, b|} (Fun (Var a) (Var b))" |
1515
76fa21f27f22
Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
135 |
apply(simp add: t_tyS.eq_iff) |
1271 | 136 |
apply(rule_tac x="0::perm" in exI) |
1676
141a36535f1d
Manually proved TySch support; All properties of TySch now true.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1673
diff
changeset
|
137 |
apply(simp add: alphas fresh_star_def eqvts t_tyS.eq_iff) |
141a36535f1d
Manually proved TySch support; All properties of TySch now true.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1673
diff
changeset
|
138 |
done |
1271 | 139 |
|
140 |
lemma |
|
141 |
assumes a: "a \<noteq> b" |
|
1525
bf321f16d025
Reached strong_induction in fset-based TySch. Will not work until isabelle changes are pushed.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1515
diff
changeset
|
142 |
shows "\<not>(All {|a, b|} (Fun (Var a) (Var b)) = All {|c|} (Fun (Var c) (Var c)))" |
1271 | 143 |
using a |
1515
76fa21f27f22
Rename "_property" to ".property"
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1510
diff
changeset
|
144 |
apply(simp add: t_tyS.eq_iff) |
1271 | 145 |
apply(clarify) |
1676
141a36535f1d
Manually proved TySch support; All properties of TySch now true.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1673
diff
changeset
|
146 |
apply(simp add: alphas fresh_star_def eqvts t_tyS.eq_iff) |
1477
4ac3485899e1
Updated Type Schemes to automatic lifting. One goal is not true because of the restriction.
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1430
diff
changeset
|
147 |
apply auto |
1271 | 148 |
done |
149 |
||
1596
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
150 |
(* PROBLEM: |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
151 |
Type schemes with separate datatypes |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
152 |
|
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
153 |
nominal_datatype T = |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
154 |
TVar "name" |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
155 |
| TFun "T" "T" |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
156 |
nominal_datatype TyS = |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
157 |
TAll xs::"name list" ty::"T" bind xs in ty |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
158 |
|
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
159 |
*** exception Datatype raised |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
160 |
*** (line 218 of "/usr/local/src/Isabelle_16-Mar-2010/src/HOL/Tools/Datatype/datatype_aux.ML") |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
161 |
*** At command "nominal_datatype". |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
162 |
*) |
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
163 |
|
c69d9fb16785
Move Ex1 and Ex2 out of Test
Cezary Kaliszyk <kaliszyk@in.tum.de>
parents:
1595
diff
changeset
|
164 |
|
1271 | 165 |
end |