author | wenzelm |
Mon, 11 Feb 2013 14:39:04 +0100 | |
changeset 51085 | d90218288d51 |
parent 49601 | ba31032887db |
child 51744 | 0468af6546ff |
permissions | -rw-r--r-- |
49509
163914705f8d
renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents:
49463
diff
changeset
|
1 |
(* Title: HOL/BNF/Examples/Misc_Data.thy |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
2 |
Author: Dmitriy Traytel, TU Muenchen |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
3 |
Author: Andrei Popescu, TU Muenchen |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
4 |
Copyright 2012 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
5 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
6 |
Miscellaneous datatype declarations. |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
7 |
*) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
8 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
9 |
header {* Miscellaneous Datatype Declarations *} |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
10 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
11 |
theory Misc_Data |
49509
163914705f8d
renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents:
49463
diff
changeset
|
12 |
imports "../BNF" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
13 |
begin |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
14 |
|
49157 | 15 |
data simple = X1 | X2 | X3 | X4 |
16 |
||
17 |
data simple' = X1' unit | X2' unit | X3' unit | X4' unit |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
18 |
|
49157 | 19 |
data 'a mylist = MyNil | MyCons 'a "'a mylist" |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
20 |
|
49157 | 21 |
data ('b, 'c, 'd, 'e) some_passive = |
22 |
SP1 "('b, 'c, 'd, 'e) some_passive" | SP2 'b | SP3 'c | SP4 'd | SP5 'e |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
23 |
|
49601 | 24 |
data hfset = HFset "hfset fset" |
25 |
||
49157 | 26 |
data lambda = |
27 |
Var string | |
|
28 |
App lambda lambda | |
|
29 |
Abs string lambda | |
|
30 |
Let "(string \<times> lambda) fset" lambda |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
31 |
|
49157 | 32 |
data 'a par_lambda = |
33 |
PVar 'a | |
|
34 |
PApp "'a par_lambda" "'a par_lambda" | |
|
35 |
PAbs 'a "'a par_lambda" | |
|
36 |
PLet "('a \<times> 'a par_lambda) fset" "'a par_lambda" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
37 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
38 |
(* |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
39 |
('a, 'b1, 'b2) F1 = 'a * 'b1 + 'a * 'b2 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
40 |
('a, 'b1, 'b2) F2 = unit + 'b1 * 'b2 |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
41 |
*) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
42 |
|
49157 | 43 |
data 'a I1 = I11 'a "'a I1" | I12 'a "'a I2" |
44 |
and 'a I2 = I21 | I22 "'a I1" "'a I2" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
45 |
|
49157 | 46 |
data 'a tree = TEmpty | TNode 'a "'a forest" |
47 |
and 'a forest = FNil | FCons "'a tree" "'a forest" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
48 |
|
49157 | 49 |
data 'a tree' = TEmpty' | TNode' "'a branch" "'a branch" |
50 |
and 'a branch = Branch 'a "'a tree'" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
51 |
|
49157 | 52 |
data ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp" |
53 |
and ('a, 'b) trm = Factor "('a, 'b) factor" | Prod "('a, 'b) factor" "('a, 'b) trm" |
|
54 |
and ('a, 'b) factor = C 'a | V 'b | Paren "('a, 'b) exp" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
55 |
|
49157 | 56 |
data ('a, 'b, 'c) some_killing = |
57 |
SK "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing + ('a, 'b, 'c) in_here" |
|
58 |
and ('a, 'b, 'c) in_here = |
|
59 |
IH1 'b 'a | IH2 'c |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
60 |
|
49219 | 61 |
data 'b nofail1 = NF11 "'b nofail1" 'b | NF12 'b |
49157 | 62 |
data 'b nofail2 = NF2 "('b nofail2 \<times> 'b \<times> 'b nofail2 \<times> 'b) list" |
49219 | 63 |
data 'b nofail3 = NF3 'b "('b nofail3 \<times> 'b \<times> 'b nofail3 \<times> 'b) fset" |
64 |
data 'b nofail4 = NF4 "('b nofail4 \<times> ('b nofail4 \<times> 'b \<times> 'b nofail4 \<times> 'b) fset) list" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
65 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
66 |
(* |
49219 | 67 |
data 'b fail = F "'b fail" 'b "'b fail" "'b list" |
68 |
data 'b fail = F "'b fail" 'b "'b fail" 'b |
|
69 |
data 'b fail = F1 "'b fail" 'b | F2 "'b fail" |
|
70 |
data 'b fail = F "'b fail" 'b |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
71 |
*) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
72 |
|
49157 | 73 |
data l1 = L1 "l2 list" |
74 |
and l2 = L21 "l1 fset" | L22 l2 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
75 |
|
49157 | 76 |
data kk1 = KK1 kk2 |
77 |
and kk2 = KK2 kk3 |
|
78 |
and kk3 = KK3 "kk1 list" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
79 |
|
49157 | 80 |
data t1 = T11 t3 | T12 t2 |
81 |
and t2 = T2 t1 |
|
82 |
and t3 = T3 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
83 |
|
49157 | 84 |
data t1' = T11' t2' | T12' t3' |
85 |
and t2' = T2' t1' |
|
86 |
and t3' = T3' |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
87 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
88 |
(* |
49157 | 89 |
data fail1 = F1 fail2 |
90 |
and fail2 = F2 fail3 |
|
91 |
and fail3 = F3 fail1 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
92 |
|
49157 | 93 |
data fail1 = F1 "fail2 list" fail2 |
94 |
and fail2 = F2 "fail2 fset" fail3 |
|
95 |
and fail3 = F3 fail1 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
96 |
|
49157 | 97 |
data fail1 = F1 "fail2 list" fail2 |
98 |
and fail2 = F2 "fail1 fset" fail1 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
99 |
*) |
49157 | 100 |
|
49162 | 101 |
(* SLOW |
49157 | 102 |
data ('a, 'c) D1 = A1 "('a, 'c) D2" | B1 "'a list" |
103 |
and ('a, 'c) D2 = A2 "('a, 'c) D3" | B2 "'c list" |
|
104 |
and ('a, 'c) D3 = A3 "('a, 'c) D3" | B3 "('a, 'c) D4" | C3 "('a, 'c) D4" "('a, 'c) D5" |
|
105 |
and ('a, 'c) D4 = A4 "('a, 'c) D5" | B4 "'a list list list" |
|
106 |
and ('a, 'c) D5 = A5 "('a, 'c) D6" |
|
107 |
and ('a, 'c) D6 = A6 "('a, 'c) D7" |
|
108 |
and ('a, 'c) D7 = A7 "('a, 'c) D8" |
|
109 |
and ('a, 'c) D8 = A8 "('a, 'c) D1 list" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
110 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
111 |
(*time comparison*) |
49157 | 112 |
datatype ('a, 'c) D1' = A1' "('a, 'c) D2'" | B1' "'a list" |
113 |
and ('a, 'c) D2' = A2' "('a, 'c) D3'" | B2' "'c list" |
|
114 |
and ('a, 'c) D3' = A3' "('a, 'c) D3'" | B3' "('a, 'c) D4'" | C3' "('a, 'c) D4'" "('a, 'c) D5'" |
|
115 |
and ('a, 'c) D4' = A4' "('a, 'c) D5'" | B4' "'a list list list" |
|
116 |
and ('a, 'c) D5' = A5' "('a, 'c) D6'" |
|
117 |
and ('a, 'c) D6' = A6' "('a, 'c) D7'" |
|
118 |
and ('a, 'c) D7' = A7' "('a, 'c) D8'" |
|
119 |
and ('a, 'c) D8' = A8' "('a, 'c) D1' list" |
|
49162 | 120 |
*) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
121 |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
122 |
(* fail: |
49157 | 123 |
data tt1 = TT11 tt2 tt3 | TT12 tt2 tt4 |
124 |
and tt2 = TT2 |
|
125 |
and tt3 = TT3 tt4 |
|
126 |
and tt4 = TT4 tt1 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
127 |
*) |
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
128 |
|
49157 | 129 |
data k1 = K11 k2 k3 | K12 k2 k4 |
130 |
and k2 = K2 |
|
131 |
and k3 = K3 k4 |
|
132 |
and k4 = K4 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
133 |
|
49157 | 134 |
data tt1 = TT11 tt3 tt2 | TT12 tt2 tt4 |
135 |
and tt2 = TT2 |
|
136 |
and tt3 = TT3 tt1 |
|
137 |
and tt4 = TT4 |
|
138 |
||
49162 | 139 |
(* SLOW |
49157 | 140 |
data s1 = S11 s2 s3 s4 | S12 s3 | S13 s2 s6 | S14 s4 s2 | S15 s2 s2 |
141 |
and s2 = S21 s7 s5 | S22 s5 s4 s6 |
|
142 |
and s3 = S31 s1 s7 s2 | S32 s3 s3 | S33 s4 s5 |
|
143 |
and s4 = S4 s5 |
|
144 |
and s5 = S5 |
|
145 |
and s6 = S61 s6 | S62 s1 s2 | S63 s6 |
|
146 |
and s7 = S71 s8 | S72 s5 |
|
147 |
and s8 = S8 nat |
|
49162 | 148 |
*) |
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
149 |
|
49436 | 150 |
data ('a, 'b) bar = Bar "'b \<Rightarrow> 'a" |
151 |
data ('a, 'b, 'c, 'd) foo = Foo "'d + 'b \<Rightarrow> 'c + 'a" |
|
49185
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents:
49166
diff
changeset
|
152 |
|
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents:
49166
diff
changeset
|
153 |
data 'a dead_foo = A |
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents:
49166
diff
changeset
|
154 |
data ('a, 'b) use_dead_foo = Y "'a" "'b dead_foo" |
49186
4b5fa9d5e330
handle type constructors not known to be a BNF using the DEADID BNF
traytel
parents:
49185
diff
changeset
|
155 |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
156 |
end |