author | wenzelm |
Sat, 13 Jun 2015 14:37:50 +0200 | |
changeset 60451 | 1f2b29f78439 |
parent 59970 | e9f73d87d904 |
child 60642 | 48dd1cefb4ae |
permissions | -rw-r--r-- |
37744 | 1 |
(* Title: HOL/Tools/hologic.ML |
923 | 2 |
Author: Lawrence C Paulson and Markus Wenzel |
3 |
||
4 |
Abstract syntax operations for HOL. |
|
5 |
*) |
|
6 |
||
7 |
signature HOLOGIC = |
|
8 |
sig |
|
41363 | 9 |
val id_const: typ -> term |
40845 | 10 |
val mk_comp: term * term -> term |
8275 | 11 |
val boolN: string |
923 | 12 |
val boolT: typ |
30450 | 13 |
val Trueprop: term |
14 |
val mk_Trueprop: term -> term |
|
15 |
val dest_Trueprop: term -> term |
|
51314
eac4bb5adbf9
just one HOLogic.Trueprop_conv, with regular exception CTERM;
wenzelm
parents:
51143
diff
changeset
|
16 |
val Trueprop_conv: conv -> conv |
44868 | 17 |
val mk_induct_forall: typ -> term |
923 | 18 |
val mk_setT: typ -> typ |
19 |
val dest_setT: typ -> typ |
|
30450 | 20 |
val Collect_const: typ -> term |
21 |
val mk_Collect: string * typ * term -> term |
|
22 |
val mk_mem: term * term -> term |
|
23 |
val dest_mem: term -> term * term |
|
24 |
val mk_set: typ -> term list -> term |
|
25 |
val dest_set: term -> term list |
|
26 |
val mk_UNIV: typ -> term |
|
59970 | 27 |
val conj_intr: Proof.context -> thm -> thm -> thm |
28 |
val conj_elim: Proof.context -> thm -> thm * thm |
|
29 |
val conj_elims: Proof.context -> thm -> thm list |
|
923 | 30 |
val conj: term |
31 |
val disj: term |
|
32 |
val imp: term |
|
8429 | 33 |
val Not: term |
7690 | 34 |
val mk_conj: term * term -> term |
35 |
val mk_disj: term * term -> term |
|
36 |
val mk_imp: term * term -> term |
|
16835 | 37 |
val mk_not: term -> term |
15151 | 38 |
val dest_conj: term -> term list |
53887 | 39 |
val conjuncts: term -> term list |
15945 | 40 |
val dest_disj: term -> term list |
24958 | 41 |
val disjuncts: term -> term list |
4571 | 42 |
val dest_imp: term -> term * term |
15151 | 43 |
val dest_not: term -> term |
51315
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
44 |
val conj_conv: conv -> conv -> conv |
923 | 45 |
val eq_const: typ -> term |
21829
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
wenzelm
parents:
21820
diff
changeset
|
46 |
val mk_eq: term * term -> term |
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
wenzelm
parents:
21820
diff
changeset
|
47 |
val dest_eq: term -> term * term |
51315
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
48 |
val eq_conv: conv -> conv -> conv |
923 | 49 |
val all_const: typ -> term |
21829
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
wenzelm
parents:
21820
diff
changeset
|
50 |
val mk_all: string * typ * term -> term |
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
wenzelm
parents:
21820
diff
changeset
|
51 |
val list_all: (string * typ) list * term -> term |
923 | 52 |
val exists_const: typ -> term |
21829
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
wenzelm
parents:
21820
diff
changeset
|
53 |
val mk_exists: string * typ * term -> term |
15945 | 54 |
val choice_const: typ -> term |
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
55 |
val class_equal: string |
2510 | 56 |
val mk_binop: string -> term * term -> term |
57 |
val mk_binrel: string -> term * term -> term |
|
58 |
val dest_bin: string -> typ -> term -> term * term |
|
4571 | 59 |
val unitT: typ |
9362 | 60 |
val is_unitT: typ -> bool |
4571 | 61 |
val unit: term |
62 |
val is_unit: term -> bool |
|
63 |
val mk_prodT: typ * typ -> typ |
|
64 |
val dest_prodT: typ -> typ * typ |
|
14048 | 65 |
val pair_const: typ -> typ -> term |
4571 | 66 |
val mk_prod: term * term -> term |
67 |
val dest_prod: term -> term * term |
|
68 |
val mk_fst: term -> term |
|
69 |
val mk_snd: term -> term |
|
21829
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
wenzelm
parents:
21820
diff
changeset
|
70 |
val split_const: typ * typ * typ -> term |
18285 | 71 |
val mk_split: term -> term |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
72 |
val flatten_tupleT: typ -> typ list |
39756
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
krauss
parents:
39250
diff
changeset
|
73 |
val tupled_lambda: term -> term -> term |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
74 |
val mk_tupleT: typ list -> typ |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
75 |
val strip_tupleT: typ -> typ list |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
76 |
val mk_tuple: term list -> term |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
77 |
val strip_tuple: term -> term list |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
78 |
val mk_ptupleT: int list list -> typ list -> typ |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
79 |
val strip_ptupleT: int list list -> typ -> typ list |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
80 |
val flat_tupleT_paths: typ -> int list list |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
81 |
val mk_ptuple: int list list -> typ -> term list -> term |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
82 |
val strip_ptuple: int list list -> term -> term list |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
83 |
val flat_tuple_paths: term -> int list list |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
84 |
val mk_psplits: int list list -> typ -> typ -> term -> term |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
85 |
val strip_psplits: term -> term * typ list * int list list |
5207 | 86 |
val natT: typ |
87 |
val zero: term |
|
88 |
val is_zero: term -> bool |
|
89 |
val mk_Suc: term -> term |
|
90 |
val dest_Suc: term -> term |
|
21621 | 91 |
val Suc_zero: term |
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23745
diff
changeset
|
92 |
val mk_nat: int -> term |
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23745
diff
changeset
|
93 |
val dest_nat: term -> int |
22994 | 94 |
val class_size: string |
95 |
val size_const: typ -> term |
|
21829
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
wenzelm
parents:
21820
diff
changeset
|
96 |
val intT: typ |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
97 |
val one_const: term |
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26036
diff
changeset
|
98 |
val bit0_const: term |
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26036
diff
changeset
|
99 |
val bit1_const: term |
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26036
diff
changeset
|
100 |
val mk_bit: int -> term |
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26036
diff
changeset
|
101 |
val dest_bit: term -> int |
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23745
diff
changeset
|
102 |
val mk_numeral: int -> term |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
103 |
val dest_num: term -> int |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
104 |
val numeral_const: typ -> term |
23269 | 105 |
val add_numerals: term -> (term * typ) list -> (term * typ) list |
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23745
diff
changeset
|
106 |
val mk_number: typ -> int -> term |
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23745
diff
changeset
|
107 |
val dest_number: term -> typ * int |
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
108 |
val code_integerT: typ |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
109 |
val code_naturalT: typ |
21778
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
wenzelm
parents:
21755
diff
changeset
|
110 |
val realT: typ |
21755 | 111 |
val nibbleT: typ |
112 |
val mk_nibble: int -> term |
|
113 |
val dest_nibble: term -> int |
|
114 |
val charT: typ |
|
115 |
val mk_char: int -> term |
|
116 |
val dest_char: term -> int |
|
21829
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
wenzelm
parents:
21820
diff
changeset
|
117 |
val listT: typ -> typ |
25887 | 118 |
val nil_const: typ -> term |
119 |
val cons_const: typ -> term |
|
21829
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
wenzelm
parents:
21820
diff
changeset
|
120 |
val mk_list: typ -> term list -> term |
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
wenzelm
parents:
21820
diff
changeset
|
121 |
val dest_list: term -> term list |
21778
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
wenzelm
parents:
21755
diff
changeset
|
122 |
val stringT: typ |
21820
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents:
21788
diff
changeset
|
123 |
val mk_string: string -> term |
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents:
21788
diff
changeset
|
124 |
val dest_string: term -> string |
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31183
diff
changeset
|
125 |
val literalT: typ |
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31183
diff
changeset
|
126 |
val mk_literal: string -> term |
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31183
diff
changeset
|
127 |
val dest_literal: term -> string |
31135 | 128 |
val mk_typerep: typ -> term |
32339 | 129 |
val termT: typ |
130 |
val term_of_const: typ -> term |
|
31135 | 131 |
val mk_term_of: typ -> term -> term |
132 |
val reflect_term: term -> term |
|
31463 | 133 |
val mk_valtermify_app: string -> (string * typ) list -> typ -> term |
31183 | 134 |
val mk_return: typ -> typ -> term -> term |
135 |
val mk_ST: ((term * typ) * (string * typ) option) list -> term -> typ -> typ option * typ -> term |
|
31463 | 136 |
val mk_random: typ -> term -> term |
923 | 137 |
end; |
138 |
||
139 |
structure HOLogic: HOLOGIC = |
|
140 |
struct |
|
141 |
||
40845 | 142 |
(* functions *) |
143 |
||
41363 | 144 |
fun id_const T = Const ("Fun.id", T --> T); |
41339 | 145 |
|
40845 | 146 |
fun mk_comp (f, g) = |
147 |
let |
|
148 |
val fT = fastype_of f; |
|
149 |
val gT = fastype_of g; |
|
150 |
val compT = fT --> gT --> domain_type gT --> range_type fT; |
|
151 |
in Const ("Fun.comp", compT) $ f $ g end; |
|
152 |
||
153 |
||
2510 | 154 |
(* bool and set *) |
923 | 155 |
|
38555 | 156 |
val boolN = "HOL.bool"; |
8275 | 157 |
val boolT = Type (boolN, []); |
923 | 158 |
|
44868 | 159 |
fun mk_induct_forall T = Const ("HOL.induct_forall", (T --> boolT) --> boolT); |
160 |
||
45977 | 161 |
fun mk_setT T = Type ("Set.set", [T]); |
923 | 162 |
|
45977 | 163 |
fun dest_setT (Type ("Set.set", [T])) = T |
3794 | 164 |
| dest_setT T = raise TYPE ("dest_setT: set type expected", [T], []); |
923 | 165 |
|
30450 | 166 |
fun mk_set T ts = |
167 |
let |
|
168 |
val sT = mk_setT T; |
|
32264
0be31453f698
Set.UNIV and Set.empty are mere abbreviations for top and bot
haftmann
parents:
31736
diff
changeset
|
169 |
val empty = Const ("Orderings.bot_class.bot", sT); |
31456 | 170 |
fun insert t u = Const ("Set.insert", T --> sT --> sT) $ t $ u; |
30450 | 171 |
in fold_rev insert ts empty end; |
172 |
||
32264
0be31453f698
Set.UNIV and Set.empty are mere abbreviations for top and bot
haftmann
parents:
31736
diff
changeset
|
173 |
fun mk_UNIV T = Const ("Orderings.top_class.top", mk_setT T); |
30450 | 174 |
|
32264
0be31453f698
Set.UNIV and Set.empty are mere abbreviations for top and bot
haftmann
parents:
31736
diff
changeset
|
175 |
fun dest_set (Const ("Orderings.bot_class.bot", _)) = [] |
31456 | 176 |
| dest_set (Const ("Set.insert", _) $ t $ u) = t :: dest_set u |
30450 | 177 |
| dest_set t = raise TERM ("dest_set", [t]); |
178 |
||
37676
f433cb9caea4
"prod" and "sum" replace "*" and "+" respectively; qualified constants Set.member and Set.Collect
haftmann
parents:
37591
diff
changeset
|
179 |
fun Collect_const T = Const ("Set.Collect", (T --> boolT) --> mk_setT T); |
44241 | 180 |
fun mk_Collect (a, T, t) = Collect_const T $ absfree (a, T) t; |
30450 | 181 |
|
182 |
fun mk_mem (x, A) = |
|
183 |
let val setT = fastype_of A in |
|
37676
f433cb9caea4
"prod" and "sum" replace "*" and "+" respectively; qualified constants Set.member and Set.Collect
haftmann
parents:
37591
diff
changeset
|
184 |
Const ("Set.member", dest_setT setT --> setT --> boolT) $ x $ A |
30450 | 185 |
end; |
186 |
||
37676
f433cb9caea4
"prod" and "sum" replace "*" and "+" respectively; qualified constants Set.member and Set.Collect
haftmann
parents:
37591
diff
changeset
|
187 |
fun dest_mem (Const ("Set.member", _) $ x $ A) = (x, A) |
30450 | 188 |
| dest_mem t = raise TERM ("dest_mem", [t]); |
189 |
||
8275 | 190 |
|
7073 | 191 |
(* logic *) |
923 | 192 |
|
51314
eac4bb5adbf9
just one HOLogic.Trueprop_conv, with regular exception CTERM;
wenzelm
parents:
51143
diff
changeset
|
193 |
val Trueprop = Const (@{const_name Trueprop}, boolT --> propT); |
923 | 194 |
|
195 |
fun mk_Trueprop P = Trueprop $ P; |
|
196 |
||
51314
eac4bb5adbf9
just one HOLogic.Trueprop_conv, with regular exception CTERM;
wenzelm
parents:
51143
diff
changeset
|
197 |
fun dest_Trueprop (Const (@{const_name Trueprop}, _) $ P) = P |
3794 | 198 |
| dest_Trueprop t = raise TERM ("dest_Trueprop", [t]); |
923 | 199 |
|
51314
eac4bb5adbf9
just one HOLogic.Trueprop_conv, with regular exception CTERM;
wenzelm
parents:
51143
diff
changeset
|
200 |
fun Trueprop_conv cv ct = |
eac4bb5adbf9
just one HOLogic.Trueprop_conv, with regular exception CTERM;
wenzelm
parents:
51143
diff
changeset
|
201 |
(case Thm.term_of ct of |
eac4bb5adbf9
just one HOLogic.Trueprop_conv, with regular exception CTERM;
wenzelm
parents:
51143
diff
changeset
|
202 |
Const (@{const_name Trueprop}, _) $ _ => Conv.arg_conv cv ct |
eac4bb5adbf9
just one HOLogic.Trueprop_conv, with regular exception CTERM;
wenzelm
parents:
51143
diff
changeset
|
203 |
| _ => raise CTERM ("Trueprop_conv", [ct])) |
eac4bb5adbf9
just one HOLogic.Trueprop_conv, with regular exception CTERM;
wenzelm
parents:
51143
diff
changeset
|
204 |
|
eac4bb5adbf9
just one HOLogic.Trueprop_conv, with regular exception CTERM;
wenzelm
parents:
51143
diff
changeset
|
205 |
|
59970 | 206 |
fun conj_intr ctxt thP thQ = |
23555
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
wenzelm
parents:
23297
diff
changeset
|
207 |
let |
59970 | 208 |
val (P, Q) = apply2 (Object_Logic.dest_judgment ctxt o Thm.cprop_of) (thP, thQ) |
23555
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
wenzelm
parents:
23297
diff
changeset
|
209 |
handle CTERM (msg, _) => raise THM (msg, 0, [thP, thQ]); |
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
wenzelm
parents:
23297
diff
changeset
|
210 |
val inst = Thm.instantiate ([], [(@{cpat "?P::bool"}, P), (@{cpat "?Q::bool"}, Q)]); |
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
wenzelm
parents:
23297
diff
changeset
|
211 |
in Drule.implies_elim_list (inst @{thm conjI}) [thP, thQ] end; |
923 | 212 |
|
59970 | 213 |
fun conj_elim ctxt thPQ = |
23555
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
wenzelm
parents:
23297
diff
changeset
|
214 |
let |
59970 | 215 |
val (P, Q) = Thm.dest_binop (Object_Logic.dest_judgment ctxt (Thm.cprop_of thPQ)) |
23555
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
wenzelm
parents:
23297
diff
changeset
|
216 |
handle CTERM (msg, _) => raise THM (msg, 0, [thPQ]); |
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
wenzelm
parents:
23297
diff
changeset
|
217 |
val inst = Thm.instantiate ([], [(@{cpat "?P::bool"}, P), (@{cpat "?Q::bool"}, Q)]); |
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
wenzelm
parents:
23297
diff
changeset
|
218 |
val thP = Thm.implies_elim (inst @{thm conjunct1}) thPQ; |
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
wenzelm
parents:
23297
diff
changeset
|
219 |
val thQ = Thm.implies_elim (inst @{thm conjunct2}) thPQ; |
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
wenzelm
parents:
23297
diff
changeset
|
220 |
in (thP, thQ) end; |
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
wenzelm
parents:
23297
diff
changeset
|
221 |
|
59970 | 222 |
fun conj_elims ctxt th = |
223 |
let val (th1, th2) = conj_elim ctxt th |
|
224 |
in conj_elims ctxt th1 @ conj_elims ctxt th2 end handle THM _ => [th]; |
|
23555
16e5fd18905c
assume basic HOL context for compilation (antiquotations);
wenzelm
parents:
23297
diff
changeset
|
225 |
|
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
226 |
val conj = @{term HOL.conj} |
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
227 |
and disj = @{term HOL.disj} |
38786
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
haftmann
parents:
38555
diff
changeset
|
228 |
and imp = @{term implies} |
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
haftmann
parents:
38555
diff
changeset
|
229 |
and Not = @{term Not}; |
923 | 230 |
|
7690 | 231 |
fun mk_conj (t1, t2) = conj $ t1 $ t2 |
232 |
and mk_disj (t1, t2) = disj $ t1 $ t2 |
|
16835 | 233 |
and mk_imp (t1, t2) = imp $ t1 $ t2 |
234 |
and mk_not t = Not $ t; |
|
7690 | 235 |
|
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
236 |
fun dest_conj (Const ("HOL.conj", _) $ t $ t') = t :: dest_conj t' |
15151 | 237 |
| dest_conj t = [t]; |
238 |
||
53887 | 239 |
(*Like dest_conj, but flattens conjunctions however nested*) |
240 |
fun conjuncts_aux (Const ("HOL.conj", _) $ t $ t') conjs = conjuncts_aux t (conjuncts_aux t' conjs) |
|
241 |
| conjuncts_aux t conjs = t::conjs; |
|
242 |
||
243 |
fun conjuncts t = conjuncts_aux t []; |
|
244 |
||
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
245 |
fun dest_disj (Const ("HOL.disj", _) $ t $ t') = t :: dest_disj t' |
15945 | 246 |
| dest_disj t = [t]; |
247 |
||
24958 | 248 |
(*Like dest_disj, but flattens disjunctions however nested*) |
38795
848be46708dc
formerly unnamed infix conjunction and disjunction now named HOL.conj and HOL.disj
haftmann
parents:
38786
diff
changeset
|
249 |
fun disjuncts_aux (Const ("HOL.disj", _) $ t $ t') disjs = disjuncts_aux t (disjuncts_aux t' disjs) |
24958 | 250 |
| disjuncts_aux t disjs = t::disjs; |
251 |
||
252 |
fun disjuncts t = disjuncts_aux t []; |
|
253 |
||
38786
e46e7a9cb622
formerly unnamed infix impliciation now named HOL.implies
haftmann
parents:
38555
diff
changeset
|
254 |
fun dest_imp (Const ("HOL.implies", _) $ A $ B) = (A, B) |
4466
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
paulson
parents:
4294
diff
changeset
|
255 |
| dest_imp t = raise TERM ("dest_imp", [t]); |
305390f23734
Better equality handling in Blast_tac, usingd a new variant of hyp_subst_tac
paulson
parents:
4294
diff
changeset
|
256 |
|
38555 | 257 |
fun dest_not (Const ("HOL.Not", _) $ t) = t |
15151 | 258 |
| dest_not t = raise TERM ("dest_not", [t]); |
8302 | 259 |
|
51315
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
260 |
|
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
261 |
fun conj_conv cv1 cv2 ct = |
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
262 |
(case Thm.term_of ct of |
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
263 |
Const (@{const_name HOL.conj}, _) $ _ $ _ => |
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
264 |
Conv.combination_conv (Conv.arg_conv cv1) cv2 ct |
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
265 |
| _ => raise CTERM ("conj_conv", [ct])); |
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
266 |
|
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
267 |
|
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
268 |
fun eq_const T = Const (@{const_name HOL.eq}, T --> T --> boolT); |
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
269 |
|
923 | 270 |
fun mk_eq (t, u) = eq_const (fastype_of t) $ t $ u; |
271 |
||
51315
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
272 |
fun dest_eq (Const (@{const_name HOL.eq}, _) $ lhs $ rhs) = (lhs, rhs) |
6031 | 273 |
| dest_eq t = raise TERM ("dest_eq", [t]) |
274 |
||
51315
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
275 |
fun eq_conv cv1 cv2 ct = |
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
276 |
(case Thm.term_of ct of |
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
277 |
Const (@{const_name HOL.eq}, _) $ _ $ _ => Conv.combination_conv (Conv.arg_conv cv1) cv2 ct |
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
278 |
| _ => raise CTERM ("eq_conv", [ct])); |
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
279 |
|
536a5603a138
provide common HOLogic.conj_conv and HOLogic.eq_conv;
wenzelm
parents:
51314
diff
changeset
|
280 |
|
46216 | 281 |
fun all_const T = Const ("HOL.All", (T --> boolT) --> boolT); |
44241 | 282 |
fun mk_all (x, T, P) = all_const T $ absfree (x, T) P; |
21173
663a7b39894c
removed is_Trueprop (use can dest_Trueprop'' instead);
wenzelm
parents:
21078
diff
changeset
|
283 |
fun list_all (xs, t) = fold_rev (fn (x, T) => fn P => all_const T $ Abs (x, T, P)) xs t; |
923 | 284 |
|
46216 | 285 |
fun exists_const T = Const ("HOL.Ex", (T --> boolT) --> boolT); |
44241 | 286 |
fun mk_exists (x, T, P) = exists_const T $ absfree (x, T) P; |
923 | 287 |
|
21755 | 288 |
fun choice_const T = Const("Hilbert_Choice.Eps", (T --> boolT) --> T); |
15945 | 289 |
|
38857
97775f3e8722
renamed class/constant eq to equal; tuned some instantiations
haftmann
parents:
38795
diff
changeset
|
290 |
val class_equal = "HOL.equal"; |
21045 | 291 |
|
923 | 292 |
|
13743
f8f9393be64c
Fixed bug in simpdata.ML that prevented the use of congruence rules from a
ballarin
parents:
13640
diff
changeset
|
293 |
(* binary operations and relations *) |
2510 | 294 |
|
295 |
fun mk_binop c (t, u) = |
|
46216 | 296 |
let val T = fastype_of t |
297 |
in Const (c, T --> T --> T) $ t $ u end; |
|
2510 | 298 |
|
299 |
fun mk_binrel c (t, u) = |
|
46216 | 300 |
let val T = fastype_of t |
301 |
in Const (c, T --> T --> boolT) $ t $ u end; |
|
2510 | 302 |
|
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14103
diff
changeset
|
303 |
(*destruct the application of a binary operator. The dummyT case is a crude |
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14103
diff
changeset
|
304 |
way of handling polymorphic operators.*) |
2510 | 305 |
fun dest_bin c T (tm as Const (c', Type ("fun", [T', _])) $ t $ u) = |
14387
e96d5c42c4b0
Polymorphic treatment of binary arithmetic using axclasses
paulson
parents:
14103
diff
changeset
|
306 |
if c = c' andalso (T=T' orelse T=dummyT) then (t, u) |
3794 | 307 |
else raise TERM ("dest_bin " ^ c, [tm]) |
308 |
| dest_bin c _ tm = raise TERM ("dest_bin " ^ c, [tm]); |
|
2510 | 309 |
|
310 |
||
4571 | 311 |
(* unit *) |
312 |
||
11604 | 313 |
val unitT = Type ("Product_Type.unit", []); |
4571 | 314 |
|
11604 | 315 |
fun is_unitT (Type ("Product_Type.unit", [])) = true |
9362 | 316 |
| is_unitT _ = false; |
317 |
||
11604 | 318 |
val unit = Const ("Product_Type.Unity", unitT); |
4571 | 319 |
|
11604 | 320 |
fun is_unit (Const ("Product_Type.Unity", _)) = true |
4571 | 321 |
| is_unit _ = false; |
322 |
||
323 |
||
324 |
(* prod *) |
|
325 |
||
37676
f433cb9caea4
"prod" and "sum" replace "*" and "+" respectively; qualified constants Set.member and Set.Collect
haftmann
parents:
37591
diff
changeset
|
326 |
fun mk_prodT (T1, T2) = Type ("Product_Type.prod", [T1, T2]); |
4571 | 327 |
|
37676
f433cb9caea4
"prod" and "sum" replace "*" and "+" respectively; qualified constants Set.member and Set.Collect
haftmann
parents:
37591
diff
changeset
|
328 |
fun dest_prodT (Type ("Product_Type.prod", [T1, T2])) = (T1, T2) |
4571 | 329 |
| dest_prodT T = raise TYPE ("dest_prodT", [T], []); |
330 |
||
46216 | 331 |
fun pair_const T1 T2 = Const ("Product_Type.Pair", T1 --> T2 --> mk_prodT (T1, T2)); |
14048 | 332 |
|
4571 | 333 |
fun mk_prod (t1, t2) = |
334 |
let val T1 = fastype_of t1 and T2 = fastype_of t2 in |
|
14048 | 335 |
pair_const T1 T2 $ t1 $ t2 |
4571 | 336 |
end; |
337 |
||
37389
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
haftmann
parents:
37387
diff
changeset
|
338 |
fun dest_prod (Const ("Product_Type.Pair", _) $ t1 $ t2) = (t1, t2) |
4571 | 339 |
| dest_prod t = raise TERM ("dest_prod", [t]); |
340 |
||
341 |
fun mk_fst p = |
|
342 |
let val pT = fastype_of p in |
|
55411 | 343 |
Const ("Product_Type.prod.fst", pT --> fst (dest_prodT pT)) $ p |
4571 | 344 |
end; |
345 |
||
346 |
fun mk_snd p = |
|
347 |
let val pT = fastype_of p in |
|
55411 | 348 |
Const ("Product_Type.prod.snd", pT --> snd (dest_prodT pT)) $ p |
4571 | 349 |
end; |
350 |
||
18285 | 351 |
fun split_const (A, B, C) = |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
55411
diff
changeset
|
352 |
Const ("Product_Type.prod.case_prod", (A --> B --> C) --> mk_prodT (A, B) --> C); |
18285 | 353 |
|
354 |
fun mk_split t = |
|
355 |
(case Term.fastype_of t of |
|
356 |
T as (Type ("fun", [A, Type ("fun", [B, C])])) => |
|
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
55411
diff
changeset
|
357 |
Const ("Product_Type.prod.case_prod", T --> mk_prodT (A, B) --> C) $ t |
18285 | 358 |
| _ => raise TERM ("mk_split: bad body type", [t])); |
359 |
||
5096
84b00be693b4
Moved most of the Prod_Syntax - stuff to HOLogic.
berghofe
parents:
4571
diff
changeset
|
360 |
(*Maps the type T1 * ... * Tn to [T1, ..., Tn], however nested*) |
37676
f433cb9caea4
"prod" and "sum" replace "*" and "+" respectively; qualified constants Set.member and Set.Collect
haftmann
parents:
37591
diff
changeset
|
361 |
fun flatten_tupleT (Type ("Product_Type.prod", [T1, T2])) = flatten_tupleT T1 @ flatten_tupleT T2 |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
362 |
| flatten_tupleT T = [T]; |
23745 | 363 |
|
39756
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
krauss
parents:
39250
diff
changeset
|
364 |
(*abstraction over nested tuples*) |
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
krauss
parents:
39250
diff
changeset
|
365 |
fun tupled_lambda (x as Free _) b = lambda x b |
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
krauss
parents:
39250
diff
changeset
|
366 |
| tupled_lambda (x as Var _) b = lambda x b |
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
krauss
parents:
39250
diff
changeset
|
367 |
| tupled_lambda (Const ("Product_Type.Pair", _) $ u $ v) b = |
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
krauss
parents:
39250
diff
changeset
|
368 |
mk_split (tupled_lambda u (tupled_lambda v b)) |
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
krauss
parents:
39250
diff
changeset
|
369 |
| tupled_lambda (Const ("Product_Type.Unity", _)) b = |
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
krauss
parents:
39250
diff
changeset
|
370 |
Abs ("x", unitT, b) |
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
krauss
parents:
39250
diff
changeset
|
371 |
| tupled_lambda t _ = raise TERM ("tupled_lambda: bad tuple", [t]); |
6c8e83d94536
consolidated tupled_lambda; moved to structure HOLogic
krauss
parents:
39250
diff
changeset
|
372 |
|
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
373 |
|
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
374 |
(* tuples with right-fold structure *) |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
375 |
|
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
376 |
fun mk_tupleT [] = unitT |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
377 |
| mk_tupleT Ts = foldr1 mk_prodT Ts; |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
378 |
|
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
379 |
fun strip_tupleT (Type ("Product_Type.unit", [])) = [] |
37676
f433cb9caea4
"prod" and "sum" replace "*" and "+" respectively; qualified constants Set.member and Set.Collect
haftmann
parents:
37591
diff
changeset
|
380 |
| strip_tupleT (Type ("Product_Type.prod", [T1, T2])) = T1 :: strip_tupleT T2 |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
381 |
| strip_tupleT T = [T]; |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
382 |
|
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
383 |
fun mk_tuple [] = unit |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
384 |
| mk_tuple ts = foldr1 mk_prod ts; |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
385 |
|
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
386 |
fun strip_tuple (Const ("Product_Type.Unity", _)) = [] |
37389
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
haftmann
parents:
37387
diff
changeset
|
387 |
| strip_tuple (Const ("Product_Type.Pair", _) $ t1 $ t2) = t1 :: strip_tuple t2 |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
388 |
| strip_tuple t = [t]; |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
389 |
|
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
390 |
|
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
391 |
(* tuples with specific arities |
23745 | 392 |
|
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
393 |
an "arity" of a tuple is a list of lists of integers, |
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
394 |
denoting paths to subterms that are pairs |
25172 | 395 |
*) |
23745 | 396 |
|
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
397 |
fun ptuple_err s = raise TERM (s ^ ": inconsistent use of nested products", []); |
23745 | 398 |
|
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
399 |
fun mk_ptupleT ps = |
23745 | 400 |
let |
401 |
fun mk p Ts = |
|
36692
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents:
35625
diff
changeset
|
402 |
if member (op =) ps p then |
23745 | 403 |
let |
404 |
val (T, Ts') = mk (1::p) Ts; |
|
405 |
val (U, Ts'') = mk (2::p) Ts' |
|
406 |
in (mk_prodT (T, U), Ts'') end |
|
407 |
else (hd Ts, tl Ts) |
|
408 |
in fst o mk [] end; |
|
409 |
||
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
410 |
fun strip_ptupleT ps = |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
411 |
let |
36692
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents:
35625
diff
changeset
|
412 |
fun factors p T = if member (op =) ps p then (case T of |
37676
f433cb9caea4
"prod" and "sum" replace "*" and "+" respectively; qualified constants Set.member and Set.Collect
haftmann
parents:
37591
diff
changeset
|
413 |
Type ("Product_Type.prod", [T1, T2]) => |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
414 |
factors (1::p) T1 @ factors (2::p) T2 |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
415 |
| _ => ptuple_err "strip_ptupleT") else [T] |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
416 |
in factors [] end; |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
417 |
|
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
418 |
val flat_tupleT_paths = |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
419 |
let |
37676
f433cb9caea4
"prod" and "sum" replace "*" and "+" respectively; qualified constants Set.member and Set.Collect
haftmann
parents:
37591
diff
changeset
|
420 |
fun factors p (Type ("Product_Type.prod", [T1, T2])) = |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
421 |
p :: factors (1::p) T1 @ factors (2::p) T2 |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
422 |
| factors p _ = [] |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
423 |
in factors [] end; |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
424 |
|
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
425 |
fun mk_ptuple ps = |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
426 |
let |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
427 |
fun mk p T ts = |
36692
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents:
35625
diff
changeset
|
428 |
if member (op =) ps p then (case T of |
37676
f433cb9caea4
"prod" and "sum" replace "*" and "+" respectively; qualified constants Set.member and Set.Collect
haftmann
parents:
37591
diff
changeset
|
429 |
Type ("Product_Type.prod", [T1, T2]) => |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
430 |
let |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
431 |
val (t, ts') = mk (1::p) T1 ts; |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
432 |
val (u, ts'') = mk (2::p) T2 ts' |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
433 |
in (pair_const T1 T2 $ t $ u, ts'') end |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
434 |
| _ => ptuple_err "mk_ptuple") |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
435 |
else (hd ts, tl ts) |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
436 |
in fst oo mk [] end; |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
437 |
|
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
438 |
fun strip_ptuple ps = |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
439 |
let |
36692
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents:
35625
diff
changeset
|
440 |
fun dest p t = if member (op =) ps p then (case t of |
37389
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
haftmann
parents:
37387
diff
changeset
|
441 |
Const ("Product_Type.Pair", _) $ t $ u => |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
442 |
dest (1::p) t @ dest (2::p) u |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
443 |
| _ => ptuple_err "strip_ptuple") else [t] |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
444 |
in dest [] end; |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
445 |
|
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
446 |
val flat_tuple_paths = |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
447 |
let |
37389
09467cdfa198
qualified type "*"; qualified constants Pair, fst, snd, split
haftmann
parents:
37387
diff
changeset
|
448 |
fun factors p (Const ("Product_Type.Pair", _) $ t $ u) = |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
449 |
p :: factors (1::p) t @ factors (2::p) u |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
450 |
| factors p _ = [] |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
451 |
in factors [] end; |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
452 |
|
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
453 |
(*In mk_psplits ps S T u, term u expects separate arguments for the factors of S, |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
454 |
with result type T. The call creates a new term expecting one argument |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
455 |
of type S.*) |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
456 |
fun mk_psplits ps T T3 u = |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
457 |
let |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
458 |
fun ap ((p, T) :: pTs) = |
36692
54b64d4ad524
farewell to old-style mem infixes -- type inference in situations with mem_int and mem_string should provide enough information to resolve the type of (op =)
haftmann
parents:
35625
diff
changeset
|
459 |
if member (op =) ps p then (case T of |
37676
f433cb9caea4
"prod" and "sum" replace "*" and "+" respectively; qualified constants Set.member and Set.Collect
haftmann
parents:
37591
diff
changeset
|
460 |
Type ("Product_Type.prod", [T1, T2]) => |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
461 |
split_const (T1, T2, map snd pTs ---> T3) $ |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
462 |
ap ((1::p, T1) :: (2::p, T2) :: pTs) |
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
463 |
| _ => ptuple_err "mk_psplits") |
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
464 |
else Abs ("x", T, ap pTs) |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
465 |
| ap [] = |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
466 |
let val k = length ps |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
467 |
in list_comb (incr_boundvars (k + 1) u, map Bound (k downto 0)) end |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
468 |
in ap [([], T)] end; |
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
469 |
|
32342
3fabf5b5fc83
path-sensitive tuple combinators carry a "p"(ath) prefix; combinators for standard right-fold tuples
haftmann
parents:
32339
diff
changeset
|
470 |
val strip_psplits = |
23745 | 471 |
let |
34962
807f6ce0273d
HOLogic.strip_psplits: types are ordered after syntactic appearance, not after corresponding de-Bruin index (closer correspondence to similar strip operations)
haftmann
parents:
33245
diff
changeset
|
472 |
fun strip [] qs Ts t = (t, rev Ts, qs) |
55414
eab03e9cee8a
renamed '{prod,sum,bool,unit}_case' to 'case_...'
blanchet
parents:
55411
diff
changeset
|
473 |
| strip (p :: ps) qs Ts (Const ("Product_Type.prod.case_prod", _) $ t) = |
23745 | 474 |
strip ((1 :: p) :: (2 :: p) :: ps) (p :: qs) Ts t |
475 |
| strip (p :: ps) qs Ts (Abs (s, T, t)) = strip ps qs (T :: Ts) t |
|
476 |
| strip (p :: ps) qs Ts t = strip ps qs |
|
477 |
(hd (binder_types (fastype_of1 (Ts, t))) :: Ts) |
|
478 |
(incr_boundvars 1 t $ Bound 0) |
|
32287
65d5c5b30747
cleaned up abstract tuple operations and named them consistently
haftmann
parents:
32264
diff
changeset
|
479 |
in strip [[]] [] [] end; |
23745 | 480 |
|
5207 | 481 |
|
482 |
(* nat *) |
|
483 |
||
37387
3581483cca6c
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
haftmann
parents:
37145
diff
changeset
|
484 |
val natT = Type ("Nat.nat", []); |
5207 | 485 |
|
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
486 |
val zero = Const ("Groups.zero_class.zero", natT); |
5207 | 487 |
|
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
488 |
fun is_zero (Const ("Groups.zero_class.zero", _)) = true |
5207 | 489 |
| is_zero _ = false; |
490 |
||
37387
3581483cca6c
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
haftmann
parents:
37145
diff
changeset
|
491 |
fun mk_Suc t = Const ("Nat.Suc", natT --> natT) $ t; |
5207 | 492 |
|
37387
3581483cca6c
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
haftmann
parents:
37145
diff
changeset
|
493 |
fun dest_Suc (Const ("Nat.Suc", _) $ t) = t |
5207 | 494 |
| dest_Suc t = raise TERM ("dest_Suc", [t]); |
495 |
||
21621 | 496 |
val Suc_zero = mk_Suc zero; |
497 |
||
24630
351a308ab58d
simplified type int (eliminated IntInf.int, integer);
wenzelm
parents:
23745
diff
changeset
|
498 |
fun mk_nat n = |
22994 | 499 |
let |
500 |
fun mk 0 = zero |
|
23297 | 501 |
| mk n = mk_Suc (mk (n - 1)); |
23576
1ffe739e5ee0
removed dest_cTrueprop (cf. ObjectLogic.dest_judgment);
wenzelm
parents:
23566
diff
changeset
|
502 |
in if n < 0 then raise TERM ("mk_nat: negative number", []) else mk n end; |
5207 | 503 |
|
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
504 |
fun dest_nat (Const ("Groups.zero_class.zero", _)) = 0 |
37387
3581483cca6c
qualified types "+" and nat; qualified constants Ball, Bex, Suc, curry; modernized some specifications
haftmann
parents:
37145
diff
changeset
|
505 |
| dest_nat (Const ("Nat.Suc", _) $ t) = dest_nat t + 1 |
5207 | 506 |
| dest_nat t = raise TERM ("dest_nat", [t]); |
507 |
||
22994 | 508 |
val class_size = "Nat.size"; |
509 |
||
510 |
fun size_const T = Const ("Nat.size_class.size", T --> natT); |
|
511 |
||
5207 | 512 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
513 |
(* binary numerals and int *) |
21778
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
wenzelm
parents:
21755
diff
changeset
|
514 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
515 |
val numT = Type ("Num.num", []); |
25919
8b1c0d434824
joined theories IntDef, Numeral, IntArith to theory Int
haftmann
parents:
25887
diff
changeset
|
516 |
val intT = Type ("Int.int", []); |
21778
66440bf72cdc
binary numerals: restricted to actual abstract syntax;
wenzelm
parents:
21755
diff
changeset
|
517 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
518 |
val one_const = Const ("Num.num.One", numT) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
519 |
and bit0_const = Const ("Num.num.Bit0", numT --> numT) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
520 |
and bit1_const = Const ("Num.num.Bit1", numT --> numT); |
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26036
diff
changeset
|
521 |
|
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26036
diff
changeset
|
522 |
fun mk_bit 0 = bit0_const |
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26036
diff
changeset
|
523 |
| mk_bit 1 = bit1_const |
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26036
diff
changeset
|
524 |
| mk_bit _ = raise TERM ("mk_bit", []); |
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26036
diff
changeset
|
525 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
526 |
fun dest_bit (Const ("Num.num.Bit0", _)) = 0 |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
527 |
| dest_bit (Const ("Num.num.Bit1", _)) = 1 |
26086
3c243098b64a
New simpler representation of numerals, using Bit0 and Bit1 instead of BIT, B0, and B1
huffman
parents:
26036
diff
changeset
|
528 |
| dest_bit t = raise TERM ("dest_bit", [t]); |
8768 | 529 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
530 |
fun mk_numeral i = |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
531 |
let fun mk 1 = one_const |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
532 |
| mk i = let val (q, r) = Integer.div_mod i 2 in mk_bit r $ mk q end |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
533 |
in if i > 0 then mk i else raise TERM ("mk_numeral: " ^ string_of_int i, []) |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
534 |
end |
13755 | 535 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
536 |
fun dest_num (Const ("Num.num.One", _)) = 1 |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
537 |
| dest_num (Const ("Num.num.Bit0", _) $ bs) = 2 * dest_num bs |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
538 |
| dest_num (Const ("Num.num.Bit1", _) $ bs) = 2 * dest_num bs + 1 |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
539 |
| dest_num t = raise TERM ("dest_num", [t]); |
13755 | 540 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
541 |
fun numeral_const T = Const ("Num.numeral_class.numeral", numT --> T); |
21820
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents:
21788
diff
changeset
|
542 |
|
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
543 |
fun add_numerals (Const ("Num.numeral_class.numeral", Type (_, [_, T])) $ t) = cons (t, T) |
23269 | 544 |
| add_numerals (t $ u) = add_numerals t #> add_numerals u |
545 |
| add_numerals (Abs (_, _, t)) = add_numerals t |
|
546 |
| add_numerals _ = I; |
|
22391 | 547 |
|
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
548 |
fun mk_number T 0 = Const ("Groups.zero_class.zero", T) |
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
549 |
| mk_number T 1 = Const ("Groups.one_class.one", T) |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
550 |
| mk_number T i = |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
551 |
if i > 0 then numeral_const T $ mk_numeral i |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
53887
diff
changeset
|
552 |
else Const ("Groups.uminus_class.uminus", T --> T) $ mk_number T (~ i); |
21820
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents:
21788
diff
changeset
|
553 |
|
35267
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
554 |
fun dest_number (Const ("Groups.zero_class.zero", T)) = (T, 0) |
8dfd816713c6
moved remaning class operations from Algebras.thy to Groups.thy
haftmann
parents:
34974
diff
changeset
|
555 |
| dest_number (Const ("Groups.one_class.one", T)) = (T, 1) |
47108
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
556 |
| dest_number (Const ("Num.numeral_class.numeral", Type ("fun", [_, T])) $ t) = |
2a1953f0d20d
merged fork with new numeral representation (see NEWS)
huffman
parents:
46216
diff
changeset
|
557 |
(T, dest_num t) |
54489
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
53887
diff
changeset
|
558 |
| dest_number (Const ("Groups.uminus_class.uminus", Type ("fun", [_, T])) $ t) = |
03ff4d1e6784
eliminiated neg_numeral in favour of - (numeral _)
haftmann
parents:
53887
diff
changeset
|
559 |
apsnd (op ~) (dest_number t) |
21820
2f2b6a965ccc
introduced mk/dest_numeral/number for mk/dest_binum etc.
haftmann
parents:
21788
diff
changeset
|
560 |
| dest_number t = raise TERM ("dest_number", [t]); |
13755 | 561 |
|
21829
016eff9c5699
simplified mk/dest_numeral (cf. mk/dest_binum of 1.65);
wenzelm
parents:
21820
diff
changeset
|
562 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
563 |
(* code target numerals *) |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
564 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
565 |
val code_integerT = Type ("Code_Numeral.integer", []); |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
566 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
567 |
val code_naturalT = Type ("Code_Numeral.natural", []); |
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
568 |
|
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
569 |
|
13755 | 570 |
(* real *) |
571 |
||
51523 | 572 |
val realT = Type ("Real.real", []); |
13755 | 573 |
|
574 |
||
21755 | 575 |
(* list *) |
21455 | 576 |
|
21755 | 577 |
fun listT T = Type ("List.list", [T]); |
21455 | 578 |
|
25887 | 579 |
fun nil_const T = Const ("List.list.Nil", listT T); |
580 |
||
581 |
fun cons_const T = |
|
582 |
let val lT = listT T |
|
583 |
in Const ("List.list.Cons", T --> lT --> lT) end; |
|
584 |
||
21755 | 585 |
fun mk_list T ts = |
21455 | 586 |
let |
21755 | 587 |
val lT = listT T; |
588 |
val Nil = Const ("List.list.Nil", lT); |
|
589 |
fun Cons t u = Const ("List.list.Cons", T --> lT --> lT) $ t $ u; |
|
590 |
in fold_rev Cons ts Nil end; |
|
591 |
||
592 |
fun dest_list (Const ("List.list.Nil", _)) = [] |
|
593 |
| dest_list (Const ("List.list.Cons", _) $ t $ u) = t :: dest_list u |
|
594 |
| dest_list t = raise TERM ("dest_list", [t]); |
|
21455 | 595 |
|
596 |
||
31048
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
597 |
(* nibble *) |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
598 |
|
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
599 |
val nibbleT = Type ("String.nibble", []); |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
600 |
|
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
601 |
fun mk_nibble n = |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
602 |
let val s = |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
603 |
if 0 <= n andalso n <= 9 then chr (n + ord "0") |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
604 |
else if 10 <= n andalso n <= 15 then chr (n + ord "A" - 10) |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
605 |
else raise TERM ("mk_nibble", []) |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
606 |
in Const ("String.nibble.Nibble" ^ s, nibbleT) end; |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
607 |
|
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
608 |
fun dest_nibble t = |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
609 |
let fun err () = raise TERM ("dest_nibble", [t]) in |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
610 |
(case try (unprefix "String.nibble.Nibble" o fst o Term.dest_Const) t of |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
611 |
NONE => err () |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
612 |
| SOME c => |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
613 |
if size c <> 1 then err () |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
614 |
else if "0" <= c andalso c <= "9" then ord c - ord "0" |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
615 |
else if "A" <= c andalso c <= "F" then ord c - ord "A" + 10 |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
616 |
else err ()) |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
617 |
end; |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
618 |
|
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
619 |
|
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
620 |
(* char *) |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
621 |
|
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
622 |
val charT = Type ("String.char", []); |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
623 |
|
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
624 |
fun mk_char n = |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
625 |
if 0 <= n andalso n <= 255 then |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
626 |
Const ("String.char.Char", nibbleT --> nibbleT --> charT) $ |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
627 |
mk_nibble (n div 16) $ mk_nibble (n mod 16) |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
628 |
else raise TERM ("mk_char", []); |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
629 |
|
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
630 |
fun dest_char (Const ("String.char.Char", _) $ t $ u) = |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
631 |
dest_nibble t * 16 + dest_nibble u |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
632 |
| dest_char t = raise TERM ("dest_char", [t]); |
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
633 |
|
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
634 |
|
21455 | 635 |
(* string *) |
636 |
||
32446 | 637 |
val stringT = listT charT; |
21455 | 638 |
|
40627
becf5d5187cc
renamed raw "explode" function to "raw_explode" to emphasize its meaning;
wenzelm
parents:
39756
diff
changeset
|
639 |
val mk_string = mk_list charT o map (mk_char o ord) o raw_explode; |
21755 | 640 |
val dest_string = implode o map (chr o dest_char) o dest_list; |
21455 | 641 |
|
31048
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
642 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31183
diff
changeset
|
643 |
(* literal *) |
31048
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
644 |
|
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31183
diff
changeset
|
645 |
val literalT = Type ("String.literal", []); |
31048
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
646 |
|
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
38864
diff
changeset
|
647 |
fun mk_literal s = Const ("String.STR", stringT --> literalT) |
31048
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
648 |
$ mk_string s; |
39250
548a3e5521ab
changing String.literal to a type instead of a datatype
bulwahn
parents:
38864
diff
changeset
|
649 |
fun dest_literal (Const ("String.STR", _) $ t) = |
31048
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
650 |
dest_string t |
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31183
diff
changeset
|
651 |
| dest_literal t = raise TERM ("dest_literal", [t]); |
31048
ac146fc38b51
refined HOL string theories and corresponding ML fragments
haftmann
parents:
30453
diff
changeset
|
652 |
|
31135 | 653 |
|
654 |
(* typerep and term *) |
|
655 |
||
656 |
val typerepT = Type ("Typerep.typerep", []); |
|
657 |
||
31736 | 658 |
fun mk_typerep (Type (tyco, Ts)) = Const ("Typerep.typerep.Typerep", |
659 |
literalT --> listT typerepT --> typerepT) $ mk_literal tyco |
|
660 |
$ mk_list typerepT (map mk_typerep Ts) |
|
661 |
| mk_typerep (T as TFree _) = Const ("Typerep.typerep_class.typerep", |
|
662 |
Term.itselfT T --> typerepT) $ Logic.mk_type T; |
|
31135 | 663 |
|
32657 | 664 |
val termT = Type ("Code_Evaluation.term", []); |
31135 | 665 |
|
32657 | 666 |
fun term_of_const T = Const ("Code_Evaluation.term_of_class.term_of", T --> termT); |
32339 | 667 |
|
668 |
fun mk_term_of T t = term_of_const T $ t; |
|
31135 | 669 |
|
670 |
fun reflect_term (Const (c, T)) = |
|
32657 | 671 |
Const ("Code_Evaluation.Const", literalT --> typerepT --> termT) |
31205
98370b26c2ce
String.literal replaces message_string, code_numeral replaces (code_)index
haftmann
parents:
31183
diff
changeset
|
672 |
$ mk_literal c $ mk_typerep T |
31135 | 673 |
| reflect_term (t1 $ t2) = |
32657 | 674 |
Const ("Code_Evaluation.App", termT --> termT --> termT) |
31135 | 675 |
$ reflect_term t1 $ reflect_term t2 |
31183 | 676 |
| reflect_term (Abs (v, _, t)) = Abs (v, termT, reflect_term t) |
677 |
| reflect_term t = t; |
|
678 |
||
31463 | 679 |
fun mk_valtermify_app c vs T = |
680 |
let |
|
681 |
fun termifyT T = mk_prodT (T, unitT --> termT); |
|
32657 | 682 |
fun valapp T T' = Const ("Code_Evaluation.valapp", |
31463 | 683 |
termifyT (T --> T') --> termifyT T --> termifyT T'); |
684 |
fun mk_fTs [] _ = [] |
|
685 |
| mk_fTs (_ :: Ts) T = (Ts ---> T) :: mk_fTs Ts T; |
|
686 |
val Ts = map snd vs; |
|
687 |
val t = Const (c, Ts ---> T); |
|
688 |
val tt = mk_prod (t, Abs ("u", unitT, reflect_term t)); |
|
33245 | 689 |
fun app (fT, (v, T)) t = valapp T fT $ t $ Free (v, termifyT T); |
690 |
in fold app (mk_fTs Ts T ~~ vs) tt end; |
|
31463 | 691 |
|
31183 | 692 |
|
693 |
(* open state monads *) |
|
694 |
||
695 |
fun mk_return T U x = pair_const T U $ x; |
|
696 |
||
697 |
fun mk_ST clauses t U (someT, V) = |
|
698 |
let |
|
699 |
val R = case someT of SOME T => mk_prodT (T, V) | NONE => V |
|
700 |
fun mk_clause ((t, U), SOME (v, T)) (t', U') = |
|
701 |
(Const ("Product_Type.scomp", (U --> mk_prodT (T, U')) --> (T --> U' --> R) --> U --> R) |
|
702 |
$ t $ lambda (Free (v, T)) t', U) |
|
703 |
| mk_clause ((t, U), NONE) (t', U') = |
|
704 |
(Const ("Product_Type.fcomp", (U --> U') --> (U' --> R) --> U --> R) |
|
705 |
$ t $ t', U) |
|
706 |
in fold_rev mk_clause clauses (t, U) |> fst end; |
|
31135 | 707 |
|
38550 | 708 |
|
709 |
(* random seeds *) |
|
710 |
||
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
711 |
val random_seedT = mk_prodT (code_naturalT, code_naturalT); |
31463 | 712 |
|
51143
0a2371e7ced3
two target language numeral types: integer and natural, as replacement for code_numeral;
haftmann
parents:
51126
diff
changeset
|
713 |
fun mk_random T t = Const ("Quickcheck_Random.random_class.random", code_naturalT |
31463 | 714 |
--> random_seedT --> mk_prodT (mk_prodT (T, unitT --> termT), random_seedT)) $ t; |
715 |
||
923 | 716 |
end; |