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