| author | traytel |
| Thu, 25 Jul 2013 12:25:07 +0200 | |
| changeset 52730 | 6bf02eb4ddf7 |
| parent 52323 | a11bbb5fef56 |
| permissions | -rw-r--r-- |
|
49591
91b228e26348
generate high-level "coinduct" and "strong_coinduct" properties
blanchet
parents:
49510
diff
changeset
|
1 |
(* Title: HOL/BNF/Examples/Misc_Codata.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 |
| 52323 | 4 |
Author: Jasmin Blanchette, TU Muenchen |
5 |
Copyright 2012, 2013 |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
6 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
7 |
Miscellaneous codatatype declarations. |
|
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 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
10 |
header {* Miscellaneous Codatatype Declarations *}
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
11 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
12 |
theory Misc_Codata |
|
49509
163914705f8d
renamed top-level theory from "Codatatype" to "BNF"
blanchet
parents:
49463
diff
changeset
|
13 |
imports "../BNF" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
14 |
begin |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
15 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
16 |
codatatype simple = X1 | X2 | X3 | X4 |
| 49158 | 17 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
18 |
codatatype 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
|
19 |
|
| 52323 | 20 |
codatatype simple'' = X1'' nat int | X2'' |
21 |
||
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
22 |
codatatype 'a stream = Stream 'a "'a stream" |
| 49158 | 23 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
24 |
codatatype 'a mylist = MyNil | MyCons 'a "'a mylist" |
| 49158 | 25 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
26 |
codatatype ('b, 'c, 'd, 'e) some_passive =
|
| 49158 | 27 |
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
|
28 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
29 |
codatatype lambda = |
| 49158 | 30 |
Var string | |
31 |
App lambda lambda | |
|
32 |
Abs string lambda | |
|
33 |
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
|
34 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
35 |
codatatype 'a par_lambda = |
| 49158 | 36 |
PVar 'a | |
37 |
PApp "'a par_lambda" "'a par_lambda" | |
|
38 |
PAbs 'a "'a par_lambda" | |
|
39 |
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
|
40 |
|
|
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 |
('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
|
43 |
('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
|
44 |
*) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
45 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
46 |
codatatype 'a J1 = J11 'a "'a J1" | J12 'a "'a J2" |
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
47 |
and 'a J2 = J21 | J22 "'a J1" "'a J2" |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
48 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
49 |
codatatype 'a tree = TEmpty | TNode 'a "'a forest" |
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
50 |
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
|
51 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
52 |
codatatype 'a tree' = TEmpty' | TNode' "'a branch" "'a branch" |
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
53 |
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
|
54 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
55 |
codatatype ('a, 'b) exp = Term "('a, 'b) trm" | Sum "('a, 'b) trm" "('a, 'b) exp"
|
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
56 |
and ('a, 'b) trm = Factor "('a, 'b) factor" | Prod "('a, 'b) factor" "('a, 'b) trm"
|
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
57 |
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
|
58 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
59 |
codatatype ('a, 'b, 'c) some_killing =
|
| 49158 | 60 |
SK "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing + ('a, 'b, 'c) in_here"
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
61 |
and ('a, 'b, 'c) in_here =
|
| 49158 | 62 |
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
|
63 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
64 |
codatatype ('a, 'b, 'c) some_killing' =
|
| 49617 | 65 |
SK' "'a \<Rightarrow> 'b \<Rightarrow> ('a, 'b, 'c) some_killing' + ('a, 'b, 'c) in_here'"
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
66 |
and ('a, 'b, 'c) in_here' =
|
| 49617 | 67 |
IH1' 'b | IH2' 'c |
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
68 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
69 |
codatatype ('a, 'b, 'c) some_killing'' =
|
| 49617 | 70 |
SK'' "'a \<Rightarrow> ('a, 'b, 'c) in_here''"
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
71 |
and ('a, 'b, 'c) in_here'' =
|
| 49617 | 72 |
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
|
73 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
74 |
codatatype ('b, 'c) less_killing = LK "'b \<Rightarrow> 'c"
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
75 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
76 |
codatatype 'b cps = CPS1 'b | CPS2 "'b \<Rightarrow> 'b cps" |
| 49158 | 77 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
78 |
codatatype ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9) fun_rhs =
|
| 49158 | 79 |
FR "'b1 \<Rightarrow> 'b2 \<Rightarrow> 'b3 \<Rightarrow> 'b4 \<Rightarrow> 'b5 \<Rightarrow> 'b6 \<Rightarrow> 'b7 \<Rightarrow> 'b8 \<Rightarrow> 'b9 \<Rightarrow> |
80 |
('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9) fun_rhs"
|
|
|
49102
ce2ef34eb828
added examples for testing of coinductive witnesses
traytel
parents:
49074
diff
changeset
|
81 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
82 |
codatatype ('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9, 'b10, 'b11, 'b12, 'b13, 'b14, 'b15, 'b16, 'b17,
|
| 49158 | 83 |
'b18, 'b19, 'b20) fun_rhs' = |
84 |
FR' "'b1 \<Rightarrow> 'b2 \<Rightarrow> 'b3 \<Rightarrow> 'b4 \<Rightarrow> 'b5 \<Rightarrow> 'b6 \<Rightarrow> 'b7 \<Rightarrow> 'b8 \<Rightarrow> 'b9 \<Rightarrow> 'b10 \<Rightarrow> 'b11 \<Rightarrow> 'b12 \<Rightarrow> 'b13 \<Rightarrow> 'b14 \<Rightarrow> |
|
85 |
'b15 \<Rightarrow> 'b16 \<Rightarrow> 'b17 \<Rightarrow> 'b18 \<Rightarrow> 'b19 \<Rightarrow> 'b20 \<Rightarrow> |
|
86 |
('b1, 'b2, 'b3, 'b4, 'b5, 'b6, 'b7, 'b8, 'b9, 'b10, 'b11, 'b12, 'b13, 'b14, 'b15, 'b16, 'b17,
|
|
87 |
'b18, 'b19, 'b20) fun_rhs'" |
|
88 |
||
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
89 |
codatatype ('a, 'b, 'c) wit3_F1 = W1 'a "('a, 'b, 'c) wit3_F1" "('a, 'b, 'c) wit3_F2"
|
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
90 |
and ('a, 'b, 'c) wit3_F2 = W2 'b "('a, 'b, 'c) wit3_F2"
|
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
91 |
and ('a, 'b, 'c) wit3_F3 = W31 'a 'b "('a, 'b, 'c) wit3_F1" | W32 'c 'a 'b "('a, 'b, 'c) wit3_F1"
|
| 49158 | 92 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
93 |
codatatype ('c, 'e, 'g) coind_wit1 =
|
| 49158 | 94 |
CW1 'c "('c, 'e, 'g) coind_wit1" "('c, 'e, 'g) ind_wit" "('c, 'e, 'g) coind_wit2"
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
95 |
and ('c, 'e, 'g) coind_wit2 =
|
| 49158 | 96 |
CW21 "('c, 'e, 'g) coind_wit2" 'e | CW22 'c 'g
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
97 |
and ('c, 'e, 'g) ind_wit =
|
| 49158 | 98 |
IW1 | IW2 'c |
|
49102
ce2ef34eb828
added examples for testing of coinductive witnesses
traytel
parents:
49074
diff
changeset
|
99 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
100 |
codatatype ('b, 'a) bar = BAR "'a \<Rightarrow> 'b"
|
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
101 |
codatatype ('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:
49158
diff
changeset
|
102 |
|
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
103 |
codatatype 'a dead_foo = A |
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
104 |
codatatype ('a, 'b) use_dead_foo = Y "'a" "'b dead_foo"
|
|
49185
073d7d1b7488
respect order of/additional type variables supplied by the user in fixed point constructions;
traytel
parents:
49158
diff
changeset
|
105 |
|
|
48975
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
106 |
(* SLOW, MEMORY-HUNGRY |
|
51804
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
107 |
codatatype ('a, 'c) D1 = A1 "('a, 'c) D2" | B1 "'a list"
|
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
108 |
and ('a, 'c) D2 = A2 "('a, 'c) D3" | B2 "'c list"
|
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
109 |
and ('a, 'c) D3 = A3 "('a, 'c) D3" | B3 "('a, 'c) D4" | C3 "('a, 'c) D4" "('a, 'c) D5"
|
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
110 |
and ('a, 'c) D4 = A4 "('a, 'c) D5" | B4 "'a list list list"
|
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
111 |
and ('a, 'c) D5 = A5 "('a, 'c) D6"
|
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
112 |
and ('a, 'c) D6 = A6 "('a, 'c) D7"
|
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
113 |
and ('a, 'c) D7 = A7 "('a, 'c) D8"
|
|
be6e703908f4
renamed BNF "(co)data" commands to names that are closer to their final names
blanchet
parents:
49617
diff
changeset
|
114 |
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
|
115 |
*) |
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
116 |
|
|
7f79f94a432c
added new (co)datatype package + theories of ordinals and cardinals (with Dmitriy and Andrei)
blanchet
parents:
diff
changeset
|
117 |
end |