author | paulson |
Mon, 13 Dec 2004 15:06:59 +0100 | |
changeset 15404 | a9a762f586b5 |
parent 15394 | a2c34e6ca4f8 |
child 15422 | cbdddc0efadf |
permissions | -rw-r--r-- |
10213 | 1 |
(* Title: HOL/Product_Type.thy |
2 |
ID: $Id$ |
|
3 |
Author: Lawrence C Paulson, Cambridge University Computer Laboratory |
|
4 |
Copyright 1992 University of Cambridge |
|
11777 | 5 |
*) |
10213 | 6 |
|
11838 | 7 |
header {* Cartesian products *} |
10213 | 8 |
|
15131 | 9 |
theory Product_Type |
15140 | 10 |
imports Fun |
15131 | 11 |
files ("Tools/split_rule.ML") |
12 |
begin |
|
11838 | 13 |
|
14 |
subsection {* Unit *} |
|
15 |
||
16 |
typedef unit = "{True}" |
|
17 |
proof |
|
18 |
show "True : ?unit" by blast |
|
19 |
qed |
|
20 |
||
21 |
constdefs |
|
22 |
Unity :: unit ("'(')") |
|
23 |
"() == Abs_unit True" |
|
24 |
||
25 |
lemma unit_eq: "u = ()" |
|
26 |
by (induct u) (simp add: unit_def Unity_def) |
|
27 |
||
28 |
text {* |
|
29 |
Simplification procedure for @{thm [source] unit_eq}. Cannot use |
|
30 |
this rule directly --- it loops! |
|
31 |
*} |
|
32 |
||
33 |
ML_setup {* |
|
13462 | 34 |
val unit_eq_proc = |
35 |
let val unit_meta_eq = mk_meta_eq (thm "unit_eq") in |
|
36 |
Simplifier.simproc (Theory.sign_of (the_context ())) "unit_eq" ["x::unit"] |
|
37 |
(fn _ => fn _ => fn t => if HOLogic.is_unit t then None else Some unit_meta_eq) |
|
38 |
end; |
|
11838 | 39 |
|
40 |
Addsimprocs [unit_eq_proc]; |
|
41 |
*} |
|
42 |
||
43 |
lemma unit_all_eq1: "(!!x::unit. PROP P x) == PROP P ()" |
|
44 |
by simp |
|
45 |
||
46 |
lemma unit_all_eq2: "(!!x::unit. PROP P) == PROP P" |
|
47 |
by (rule triv_forall_equality) |
|
48 |
||
49 |
lemma unit_induct [induct type: unit]: "P () ==> P x" |
|
50 |
by simp |
|
51 |
||
52 |
text {* |
|
53 |
This rewrite counters the effect of @{text unit_eq_proc} on @{term |
|
54 |
[source] "%u::unit. f u"}, replacing it by @{term [source] |
|
55 |
f} rather than by @{term [source] "%u. f ()"}. |
|
56 |
*} |
|
57 |
||
58 |
lemma unit_abs_eta_conv [simp]: "(%u::unit. f ()) = f" |
|
59 |
by (rule ext) simp |
|
10213 | 60 |
|
61 |
||
11838 | 62 |
subsection {* Pairs *} |
10213 | 63 |
|
11777 | 64 |
subsubsection {* Type definition *} |
10213 | 65 |
|
66 |
constdefs |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
67 |
Pair_Rep :: "['a, 'b] => ['a, 'b] => bool" |
11032 | 68 |
"Pair_Rep == (%a b. %x y. x=a & y=b)" |
10213 | 69 |
|
70 |
global |
|
71 |
||
72 |
typedef (Prod) |
|
11838 | 73 |
('a, 'b) "*" (infixr 20) |
11032 | 74 |
= "{f. EX a b. f = Pair_Rep (a::'a) (b::'b)}" |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
75 |
proof |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
76 |
fix a b show "Pair_Rep a b : ?Prod" |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
77 |
by blast |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
78 |
qed |
10213 | 79 |
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
11966
diff
changeset
|
80 |
syntax (xsymbols) |
15404 | 81 |
"*" :: "[type, type] => type" ("(_ \\<times>/ _)" [21, 20] 20) |
10213 | 82 |
syntax (HTML output) |
15404 | 83 |
"*" :: "[type, type] => type" ("(_ \\<times>/ _)" [21, 20] 20) |
10213 | 84 |
|
11777 | 85 |
local |
10213 | 86 |
|
11777 | 87 |
|
88 |
subsubsection {* Abstract constants and syntax *} |
|
89 |
||
90 |
global |
|
10213 | 91 |
|
92 |
consts |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
93 |
fst :: "'a * 'b => 'a" |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
94 |
snd :: "'a * 'b => 'b" |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
95 |
split :: "[['a, 'b] => 'c, 'a * 'b] => 'c" |
14189 | 96 |
curry :: "['a * 'b => 'c, 'a, 'b] => 'c" |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
97 |
prod_fun :: "['a => 'b, 'c => 'd, 'a * 'c] => 'b * 'd" |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
98 |
Pair :: "['a, 'b] => 'a * 'b" |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
99 |
Sigma :: "['a set, 'a => 'b set] => ('a * 'b) set" |
10213 | 100 |
|
11777 | 101 |
local |
10213 | 102 |
|
11777 | 103 |
text {* |
104 |
Patterns -- extends pre-defined type @{typ pttrn} used in |
|
105 |
abstractions. |
|
106 |
*} |
|
10213 | 107 |
|
108 |
nonterminals |
|
109 |
tuple_args patterns |
|
110 |
||
111 |
syntax |
|
112 |
"_tuple" :: "'a => tuple_args => 'a * 'b" ("(1'(_,/ _'))") |
|
113 |
"_tuple_arg" :: "'a => tuple_args" ("_") |
|
114 |
"_tuple_args" :: "'a => tuple_args => tuple_args" ("_,/ _") |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
115 |
"_pattern" :: "[pttrn, patterns] => pttrn" ("'(_,/ _')") |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
116 |
"" :: "pttrn => patterns" ("_") |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
117 |
"_patterns" :: "[pttrn, patterns] => patterns" ("_,/ _") |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
118 |
"@Sigma" ::"[pttrn, 'a set, 'b set] => ('a * 'b) set" ("(3SIGMA _:_./ _)" 10) |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
119 |
"@Times" ::"['a set, 'a => 'b set] => ('a * 'b) set" (infixr "<*>" 80) |
10213 | 120 |
|
121 |
translations |
|
122 |
"(x, y)" == "Pair x y" |
|
123 |
"_tuple x (_tuple_args y z)" == "_tuple x (_tuple_arg (_tuple y z))" |
|
124 |
"%(x,y,zs).b" == "split(%x (y,zs).b)" |
|
125 |
"%(x,y).b" == "split(%x y. b)" |
|
126 |
"_abs (Pair x y) t" => "%(x,y).t" |
|
127 |
(* The last rule accommodates tuples in `case C ... (x,y) ... => ...' |
|
128 |
The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *) |
|
129 |
||
130 |
"SIGMA x:A. B" => "Sigma A (%x. B)" |
|
131 |
"A <*> B" => "Sigma A (_K B)" |
|
132 |
||
14359 | 133 |
(* reconstructs pattern from (nested) splits, avoiding eta-contraction of body*) |
134 |
(* works best with enclosing "let", if "let" does not avoid eta-contraction *) |
|
135 |
print_translation {* |
|
136 |
let fun split_tr' [Abs (x,T,t as (Abs abs))] = |
|
137 |
(* split (%x y. t) => %(x,y) t *) |
|
138 |
let val (y,t') = atomic_abs_tr' abs; |
|
139 |
val (x',t'') = atomic_abs_tr' (x,T,t'); |
|
140 |
||
141 |
in Syntax.const "_abs" $ (Syntax.const "_pattern" $x'$y) $ t'' end |
|
142 |
| split_tr' [Abs (x,T,(s as Const ("split",_)$t))] = |
|
143 |
(* split (%x. (split (%y z. t))) => %(x,y,z). t *) |
|
144 |
let val (Const ("_abs",_)$(Const ("_pattern",_)$y$z)$t') = split_tr' [t]; |
|
145 |
val (x',t'') = atomic_abs_tr' (x,T,t'); |
|
146 |
in Syntax.const "_abs"$ |
|
147 |
(Syntax.const "_pattern"$x'$(Syntax.const "_patterns"$y$z))$t'' end |
|
148 |
| split_tr' [Const ("split",_)$t] = |
|
149 |
(* split (split (%x y z. t)) => %((x,y),z). t *) |
|
150 |
split_tr' [(split_tr' [t])] (* inner split_tr' creates next pattern *) |
|
151 |
| split_tr' [Const ("_abs",_)$x_y$(Abs abs)] = |
|
152 |
(* split (%pttrn z. t) => %(pttrn,z). t *) |
|
153 |
let val (z,t) = atomic_abs_tr' abs; |
|
154 |
in Syntax.const "_abs" $ (Syntax.const "_pattern" $x_y$z) $ t end |
|
155 |
| split_tr' _ = raise Match; |
|
156 |
in [("split", split_tr')] |
|
157 |
end |
|
158 |
*} |
|
159 |
||
15404 | 160 |
text{*Deleted x-symbol and html support using @{text"\\<Sigma>"} (Sigma) because of the danger of confusion with Sum.*} |
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
11966
diff
changeset
|
161 |
syntax (xsymbols) |
15404 | 162 |
"@Times" :: "['a set, 'a => 'b set] => ('a * 'b) set" ("_ \\<times> _" [81, 80] 80) |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
163 |
|
14565 | 164 |
syntax (HTML output) |
15404 | 165 |
"@Times" :: "['a set, 'a => 'b set] => ('a * 'b) set" ("_ \\<times> _" [81, 80] 80) |
14565 | 166 |
|
11032 | 167 |
print_translation {* [("Sigma", dependent_tr' ("@Sigma", "@Times"))] *} |
10213 | 168 |
|
169 |
||
11777 | 170 |
subsubsection {* Definitions *} |
10213 | 171 |
|
172 |
defs |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
173 |
Pair_def: "Pair a b == Abs_Prod(Pair_Rep a b)" |
11451
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11425
diff
changeset
|
174 |
fst_def: "fst p == THE a. EX b. p = (a, b)" |
8abfb4f7bd02
partial restructuring to reduce dependence on Axiom of Choice
paulson
parents:
11425
diff
changeset
|
175 |
snd_def: "snd p == THE b. EX a. p = (a, b)" |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
176 |
split_def: "split == (%c p. c (fst p) (snd p))" |
14189 | 177 |
curry_def: "curry == (%c x y. c (x,y))" |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
178 |
prod_fun_def: "prod_fun f g == split(%x y.(f(x), g(y)))" |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
179 |
Sigma_def: "Sigma A B == UN x:A. UN y:B(x). {(x, y)}" |
10213 | 180 |
|
181 |
||
11966 | 182 |
subsubsection {* Lemmas and proof tool setup *} |
11838 | 183 |
|
184 |
lemma ProdI: "Pair_Rep a b : Prod" |
|
185 |
by (unfold Prod_def) blast |
|
186 |
||
187 |
lemma Pair_Rep_inject: "Pair_Rep a b = Pair_Rep a' b' ==> a = a' & b = b'" |
|
188 |
apply (unfold Pair_Rep_def) |
|
14208 | 189 |
apply (drule fun_cong [THEN fun_cong], blast) |
11838 | 190 |
done |
10213 | 191 |
|
11838 | 192 |
lemma inj_on_Abs_Prod: "inj_on Abs_Prod Prod" |
193 |
apply (rule inj_on_inverseI) |
|
194 |
apply (erule Abs_Prod_inverse) |
|
195 |
done |
|
196 |
||
197 |
lemma Pair_inject: |
|
198 |
"(a, b) = (a', b') ==> (a = a' ==> b = b' ==> R) ==> R" |
|
199 |
proof - |
|
200 |
case rule_context [unfolded Pair_def] |
|
201 |
show ?thesis |
|
202 |
apply (rule inj_on_Abs_Prod [THEN inj_onD, THEN Pair_Rep_inject, THEN conjE]) |
|
203 |
apply (rule rule_context ProdI)+ |
|
204 |
. |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
205 |
qed |
10213 | 206 |
|
11838 | 207 |
lemma Pair_eq [iff]: "((a, b) = (a', b')) = (a = a' & b = b')" |
208 |
by (blast elim!: Pair_inject) |
|
209 |
||
210 |
lemma fst_conv [simp]: "fst (a, b) = a" |
|
211 |
by (unfold fst_def) blast |
|
212 |
||
213 |
lemma snd_conv [simp]: "snd (a, b) = b" |
|
214 |
by (unfold snd_def) blast |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
215 |
|
11838 | 216 |
lemma fst_eqD: "fst (x, y) = a ==> x = a" |
217 |
by simp |
|
218 |
||
219 |
lemma snd_eqD: "snd (x, y) = a ==> y = a" |
|
220 |
by simp |
|
221 |
||
222 |
lemma PairE_lemma: "EX x y. p = (x, y)" |
|
223 |
apply (unfold Pair_def) |
|
224 |
apply (rule Rep_Prod [unfolded Prod_def, THEN CollectE]) |
|
225 |
apply (erule exE, erule exE, rule exI, rule exI) |
|
226 |
apply (rule Rep_Prod_inverse [symmetric, THEN trans]) |
|
227 |
apply (erule arg_cong) |
|
228 |
done |
|
11032 | 229 |
|
11838 | 230 |
lemma PairE [cases type: *]: "(!!x y. p = (x, y) ==> Q) ==> Q" |
231 |
by (insert PairE_lemma [of p]) blast |
|
232 |
||
233 |
ML_setup {* |
|
234 |
local val PairE = thm "PairE" in |
|
235 |
fun pair_tac s = |
|
236 |
EVERY' [res_inst_tac [("p", s)] PairE, hyp_subst_tac, K prune_params_tac]; |
|
237 |
end; |
|
238 |
*} |
|
11032 | 239 |
|
11838 | 240 |
lemma surjective_pairing: "p = (fst p, snd p)" |
241 |
-- {* Do not add as rewrite rule: invalidates some proofs in IMP *} |
|
242 |
by (cases p) simp |
|
243 |
||
244 |
declare surjective_pairing [symmetric, simp] |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
245 |
|
11838 | 246 |
lemma surj_pair [simp]: "EX x y. z = (x, y)" |
247 |
apply (rule exI) |
|
248 |
apply (rule exI) |
|
249 |
apply (rule surjective_pairing) |
|
250 |
done |
|
251 |
||
252 |
lemma split_paired_all: "(!!x. PROP P x) == (!!a b. PROP P (a, b))" |
|
11820
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
253 |
proof |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
254 |
fix a b |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
255 |
assume "!!x. PROP P x" |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
256 |
thus "PROP P (a, b)" . |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
257 |
next |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
258 |
fix x |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
259 |
assume "!!a b. PROP P (a, b)" |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
260 |
hence "PROP P (fst x, snd x)" . |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
261 |
thus "PROP P x" by simp |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
262 |
qed |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
263 |
|
11838 | 264 |
lemmas split_tupled_all = split_paired_all unit_all_eq2 |
265 |
||
266 |
text {* |
|
267 |
The rule @{thm [source] split_paired_all} does not work with the |
|
268 |
Simplifier because it also affects premises in congrence rules, |
|
269 |
where this can lead to premises of the form @{text "!!a b. ... = |
|
270 |
?P(a, b)"} which cannot be solved by reflexivity. |
|
271 |
*} |
|
272 |
||
273 |
ML_setup " |
|
274 |
(* replace parameters of product type by individual component parameters *) |
|
275 |
val safe_full_simp_tac = generic_simp_tac true (true, false, false); |
|
276 |
local (* filtering with exists_paired_all is an essential optimization *) |
|
277 |
fun exists_paired_all (Const (\"all\", _) $ Abs (_, T, t)) = |
|
278 |
can HOLogic.dest_prodT T orelse exists_paired_all t |
|
279 |
| exists_paired_all (t $ u) = exists_paired_all t orelse exists_paired_all u |
|
280 |
| exists_paired_all (Abs (_, _, t)) = exists_paired_all t |
|
281 |
| exists_paired_all _ = false; |
|
282 |
val ss = HOL_basic_ss |
|
283 |
addsimps [thm \"split_paired_all\", thm \"unit_all_eq2\", thm \"unit_abs_eta_conv\"] |
|
284 |
addsimprocs [unit_eq_proc]; |
|
285 |
in |
|
286 |
val split_all_tac = SUBGOAL (fn (t, i) => |
|
287 |
if exists_paired_all t then safe_full_simp_tac ss i else no_tac); |
|
288 |
val unsafe_split_all_tac = SUBGOAL (fn (t, i) => |
|
289 |
if exists_paired_all t then full_simp_tac ss i else no_tac); |
|
290 |
fun split_all th = |
|
291 |
if exists_paired_all (#prop (Thm.rep_thm th)) then full_simplify ss th else th; |
|
292 |
end; |
|
293 |
||
294 |
claset_ref() := claset() addSbefore (\"split_all_tac\", split_all_tac); |
|
295 |
" |
|
296 |
||
297 |
lemma split_paired_All [simp]: "(ALL x. P x) = (ALL a b. P (a, b))" |
|
298 |
-- {* @{text "[iff]"} is not a good idea because it makes @{text blast} loop *} |
|
299 |
by fast |
|
300 |
||
14189 | 301 |
lemma curry_split [simp]: "curry (split f) = f" |
302 |
by (simp add: curry_def split_def) |
|
303 |
||
304 |
lemma split_curry [simp]: "split (curry f) = f" |
|
305 |
by (simp add: curry_def split_def) |
|
306 |
||
307 |
lemma curryI [intro!]: "f (a,b) ==> curry f a b" |
|
308 |
by (simp add: curry_def) |
|
309 |
||
14190
609c072edf90
Fixed blunder in the setup of the classical reasoner wrt. the constant
skalberg
parents:
14189
diff
changeset
|
310 |
lemma curryD [dest!]: "curry f a b ==> f (a,b)" |
14189 | 311 |
by (simp add: curry_def) |
312 |
||
14190
609c072edf90
Fixed blunder in the setup of the classical reasoner wrt. the constant
skalberg
parents:
14189
diff
changeset
|
313 |
lemma curryE: "[| curry f a b ; f (a,b) ==> Q |] ==> Q" |
14189 | 314 |
by (simp add: curry_def) |
315 |
||
316 |
lemma curry_conv [simp]: "curry f a b = f (a,b)" |
|
317 |
by (simp add: curry_def) |
|
318 |
||
11838 | 319 |
lemma prod_induct [induct type: *]: "!!x. (!!a b. P (a, b)) ==> P x" |
320 |
by fast |
|
321 |
||
322 |
lemma split_paired_Ex [simp]: "(EX x. P x) = (EX a b. P (a, b))" |
|
323 |
by fast |
|
324 |
||
325 |
lemma split_conv [simp]: "split c (a, b) = c a b" |
|
326 |
by (simp add: split_def) |
|
327 |
||
328 |
lemmas split = split_conv -- {* for backwards compatibility *} |
|
329 |
||
330 |
lemmas splitI = split_conv [THEN iffD2, standard] |
|
331 |
lemmas splitD = split_conv [THEN iffD1, standard] |
|
332 |
||
333 |
lemma split_Pair_apply: "split (%x y. f (x, y)) = f" |
|
334 |
-- {* Subsumes the old @{text split_Pair} when @{term f} is the identity function. *} |
|
335 |
apply (rule ext) |
|
14208 | 336 |
apply (tactic {* pair_tac "x" 1 *}, simp) |
11838 | 337 |
done |
338 |
||
339 |
lemma split_paired_The: "(THE x. P x) = (THE (a, b). P (a, b))" |
|
340 |
-- {* Can't be added to simpset: loops! *} |
|
341 |
by (simp add: split_Pair_apply) |
|
342 |
||
343 |
lemma The_split: "The (split P) = (THE xy. P (fst xy) (snd xy))" |
|
344 |
by (simp add: split_def) |
|
345 |
||
346 |
lemma Pair_fst_snd_eq: "!!s t. (s = t) = (fst s = fst t & snd s = snd t)" |
|
14208 | 347 |
by (simp only: split_tupled_all, simp) |
11838 | 348 |
|
349 |
lemma prod_eqI [intro?]: "fst p = fst q ==> snd p = snd q ==> p = q" |
|
350 |
by (simp add: Pair_fst_snd_eq) |
|
351 |
||
352 |
lemma split_weak_cong: "p = q ==> split c p = split c q" |
|
353 |
-- {* Prevents simplification of @{term c}: much faster *} |
|
354 |
by (erule arg_cong) |
|
355 |
||
356 |
lemma split_eta: "(%(x, y). f (x, y)) = f" |
|
357 |
apply (rule ext) |
|
358 |
apply (simp only: split_tupled_all) |
|
359 |
apply (rule split_conv) |
|
360 |
done |
|
361 |
||
362 |
lemma cond_split_eta: "(!!x y. f x y = g (x, y)) ==> (%(x, y). f x y) = g" |
|
363 |
by (simp add: split_eta) |
|
364 |
||
365 |
text {* |
|
366 |
Simplification procedure for @{thm [source] cond_split_eta}. Using |
|
367 |
@{thm [source] split_eta} as a rewrite rule is not general enough, |
|
368 |
and using @{thm [source] cond_split_eta} directly would render some |
|
369 |
existing proofs very inefficient; similarly for @{text |
|
370 |
split_beta}. *} |
|
371 |
||
372 |
ML_setup {* |
|
373 |
||
374 |
local |
|
375 |
val cond_split_eta = thm "cond_split_eta"; |
|
376 |
fun Pair_pat k 0 (Bound m) = (m = k) |
|
377 |
| Pair_pat k i (Const ("Pair", _) $ Bound m $ t) = i > 0 andalso |
|
378 |
m = k+i andalso Pair_pat k (i-1) t |
|
379 |
| Pair_pat _ _ _ = false; |
|
380 |
fun no_args k i (Abs (_, _, t)) = no_args (k+1) i t |
|
381 |
| no_args k i (t $ u) = no_args k i t andalso no_args k i u |
|
382 |
| no_args k i (Bound m) = m < k orelse m > k+i |
|
383 |
| no_args _ _ _ = true; |
|
384 |
fun split_pat tp i (Abs (_,_,t)) = if tp 0 i t then Some (i,t) else None |
|
385 |
| split_pat tp i (Const ("split", _) $ Abs (_, _, t)) = split_pat tp (i+1) t |
|
386 |
| split_pat tp i _ = None; |
|
13480
bb72bd43c6c3
use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents:
13462
diff
changeset
|
387 |
fun metaeq sg lhs rhs = mk_meta_eq (Tactic.prove sg [] [] |
bb72bd43c6c3
use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents:
13462
diff
changeset
|
388 |
(HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs,rhs))) |
bb72bd43c6c3
use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents:
13462
diff
changeset
|
389 |
(K (simp_tac (HOL_basic_ss addsimps [cond_split_eta]) 1))); |
11838 | 390 |
|
391 |
fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k+1) i t |
|
392 |
| beta_term_pat k i (t $ u) = Pair_pat k i (t $ u) orelse |
|
393 |
(beta_term_pat k i t andalso beta_term_pat k i u) |
|
394 |
| beta_term_pat k i t = no_args k i t; |
|
395 |
fun eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg |
|
396 |
| eta_term_pat _ _ _ = false; |
|
397 |
fun subst arg k i (Abs (x, T, t)) = Abs (x, T, subst arg (k+1) i t) |
|
398 |
| subst arg k i (t $ u) = if Pair_pat k i (t $ u) then incr_boundvars k arg |
|
399 |
else (subst arg k i t $ subst arg k i u) |
|
400 |
| subst arg k i t = t; |
|
401 |
fun beta_proc sg _ (s as Const ("split", _) $ Abs (_, _, t) $ arg) = |
|
402 |
(case split_pat beta_term_pat 1 t of |
|
403 |
Some (i,f) => Some (metaeq sg s (subst arg 0 i f)) |
|
404 |
| None => None) |
|
405 |
| beta_proc _ _ _ = None; |
|
406 |
fun eta_proc sg _ (s as Const ("split", _) $ Abs (_, _, t)) = |
|
407 |
(case split_pat eta_term_pat 1 t of |
|
408 |
Some (_,ft) => Some (metaeq sg s (let val (f $ arg) = ft in f end)) |
|
409 |
| None => None) |
|
410 |
| eta_proc _ _ _ = None; |
|
411 |
in |
|
13462 | 412 |
val split_beta_proc = Simplifier.simproc (Theory.sign_of (the_context ())) |
413 |
"split_beta" ["split f z"] beta_proc; |
|
414 |
val split_eta_proc = Simplifier.simproc (Theory.sign_of (the_context ())) |
|
415 |
"split_eta" ["split f"] eta_proc; |
|
11838 | 416 |
end; |
417 |
||
418 |
Addsimprocs [split_beta_proc, split_eta_proc]; |
|
419 |
*} |
|
420 |
||
421 |
lemma split_beta: "(%(x, y). P x y) z = P (fst z) (snd z)" |
|
422 |
by (subst surjective_pairing, rule split_conv) |
|
423 |
||
424 |
lemma split_split: "R (split c p) = (ALL x y. p = (x, y) --> R (c x y))" |
|
425 |
-- {* For use with @{text split} and the Simplifier. *} |
|
426 |
apply (subst surjective_pairing) |
|
14208 | 427 |
apply (subst split_conv, blast) |
11838 | 428 |
done |
429 |
||
430 |
text {* |
|
431 |
@{thm [source] split_split} could be declared as @{text "[split]"} |
|
432 |
done after the Splitter has been speeded up significantly; |
|
433 |
precompute the constants involved and don't do anything unless the |
|
434 |
current goal contains one of those constants. |
|
435 |
*} |
|
436 |
||
437 |
lemma split_split_asm: "R (split c p) = (~(EX x y. p = (x, y) & (~R (c x y))))" |
|
14208 | 438 |
by (subst split_split, simp) |
11838 | 439 |
|
440 |
||
441 |
text {* |
|
442 |
\medskip @{term split} used as a logical connective or set former. |
|
443 |
||
444 |
\medskip These rules are for use with @{text blast}; could instead |
|
445 |
call @{text simp} using @{thm [source] split} as rewrite. *} |
|
446 |
||
447 |
lemma splitI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> split c p" |
|
448 |
apply (simp only: split_tupled_all) |
|
449 |
apply (simp (no_asm_simp)) |
|
450 |
done |
|
451 |
||
452 |
lemma splitI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> split c p x" |
|
453 |
apply (simp only: split_tupled_all) |
|
454 |
apply (simp (no_asm_simp)) |
|
455 |
done |
|
456 |
||
457 |
lemma splitE: "split c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q" |
|
458 |
by (induct p) (auto simp add: split_def) |
|
459 |
||
460 |
lemma splitE': "split c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q" |
|
461 |
by (induct p) (auto simp add: split_def) |
|
462 |
||
463 |
lemma splitE2: |
|
464 |
"[| Q (split P z); !!x y. [|z = (x, y); Q (P x y)|] ==> R |] ==> R" |
|
465 |
proof - |
|
466 |
assume q: "Q (split P z)" |
|
467 |
assume r: "!!x y. [|z = (x, y); Q (P x y)|] ==> R" |
|
468 |
show R |
|
469 |
apply (rule r surjective_pairing)+ |
|
470 |
apply (rule split_beta [THEN subst], rule q) |
|
471 |
done |
|
472 |
qed |
|
473 |
||
474 |
lemma splitD': "split R (a,b) c ==> R a b c" |
|
475 |
by simp |
|
476 |
||
477 |
lemma mem_splitI: "z: c a b ==> z: split c (a, b)" |
|
478 |
by simp |
|
479 |
||
480 |
lemma mem_splitI2: "!!p. [| !!a b. p = (a, b) ==> z: c a b |] ==> z: split c p" |
|
14208 | 481 |
by (simp only: split_tupled_all, simp) |
11838 | 482 |
|
483 |
lemma mem_splitE: "[| z: split c p; !!x y. [| p = (x,y); z: c x y |] ==> Q |] ==> Q" |
|
484 |
proof - |
|
485 |
case rule_context [unfolded split_def] |
|
486 |
show ?thesis by (rule rule_context surjective_pairing)+ |
|
487 |
qed |
|
488 |
||
489 |
declare mem_splitI2 [intro!] mem_splitI [intro!] splitI2' [intro!] splitI2 [intro!] splitI [intro!] |
|
490 |
declare mem_splitE [elim!] splitE' [elim!] splitE [elim!] |
|
491 |
||
492 |
ML_setup " |
|
493 |
local (* filtering with exists_p_split is an essential optimization *) |
|
494 |
fun exists_p_split (Const (\"split\",_) $ _ $ (Const (\"Pair\",_)$_$_)) = true |
|
495 |
| exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u |
|
496 |
| exists_p_split (Abs (_, _, t)) = exists_p_split t |
|
497 |
| exists_p_split _ = false; |
|
498 |
val ss = HOL_basic_ss addsimps [thm \"split_conv\"]; |
|
499 |
in |
|
500 |
val split_conv_tac = SUBGOAL (fn (t, i) => |
|
501 |
if exists_p_split t then safe_full_simp_tac ss i else no_tac); |
|
502 |
end; |
|
503 |
(* This prevents applications of splitE for already splitted arguments leading |
|
504 |
to quite time-consuming computations (in particular for nested tuples) *) |
|
505 |
claset_ref() := claset() addSbefore (\"split_conv_tac\", split_conv_tac); |
|
506 |
" |
|
507 |
||
508 |
lemma split_eta_SetCompr [simp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P" |
|
14208 | 509 |
by (rule ext, fast) |
11838 | 510 |
|
511 |
lemma split_eta_SetCompr2 [simp]: "(%u. EX x y. u = (x, y) & P x y) = split P" |
|
14208 | 512 |
by (rule ext, fast) |
11838 | 513 |
|
514 |
lemma split_part [simp]: "(%(a,b). P & Q a b) = (%ab. P & split Q ab)" |
|
515 |
-- {* Allows simplifications of nested splits in case of independent predicates. *} |
|
14208 | 516 |
apply (rule ext, blast) |
11838 | 517 |
done |
518 |
||
14337
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents:
14208
diff
changeset
|
519 |
(* Do NOT make this a simp rule as it |
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents:
14208
diff
changeset
|
520 |
a) only helps in special situations |
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents:
14208
diff
changeset
|
521 |
b) can lead to nontermination in the presence of split_def |
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents:
14208
diff
changeset
|
522 |
*) |
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents:
14208
diff
changeset
|
523 |
lemma split_comp_eq: |
14101 | 524 |
"(%u. f (g (fst u)) (snd u)) = (split (%x. f (g x)))" |
525 |
by (rule ext, auto) |
|
526 |
||
11838 | 527 |
lemma The_split_eq [simp]: "(THE (x',y'). x = x' & y = y') = (x, y)" |
528 |
by blast |
|
529 |
||
530 |
(* |
|
531 |
the following would be slightly more general, |
|
532 |
but cannot be used as rewrite rule: |
|
533 |
### Cannot add premise as rewrite rule because it contains (type) unknowns: |
|
534 |
### ?y = .x |
|
535 |
Goal "[| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)" |
|
14208 | 536 |
by (rtac some_equality 1) |
537 |
by ( Simp_tac 1) |
|
538 |
by (split_all_tac 1) |
|
539 |
by (Asm_full_simp_tac 1) |
|
11838 | 540 |
qed "The_split_eq"; |
541 |
*) |
|
542 |
||
543 |
lemma injective_fst_snd: "!!x y. [|fst x = fst y; snd x = snd y|] ==> x = y" |
|
544 |
by auto |
|
545 |
||
546 |
||
547 |
text {* |
|
548 |
\bigskip @{term prod_fun} --- action of the product functor upon |
|
549 |
functions. |
|
550 |
*} |
|
551 |
||
552 |
lemma prod_fun [simp]: "prod_fun f g (a, b) = (f a, g b)" |
|
553 |
by (simp add: prod_fun_def) |
|
554 |
||
555 |
lemma prod_fun_compose: "prod_fun (f1 o f2) (g1 o g2) = (prod_fun f1 g1 o prod_fun f2 g2)" |
|
556 |
apply (rule ext) |
|
14208 | 557 |
apply (tactic {* pair_tac "x" 1 *}, simp) |
11838 | 558 |
done |
559 |
||
560 |
lemma prod_fun_ident [simp]: "prod_fun (%x. x) (%y. y) = (%z. z)" |
|
561 |
apply (rule ext) |
|
14208 | 562 |
apply (tactic {* pair_tac "z" 1 *}, simp) |
11838 | 563 |
done |
564 |
||
565 |
lemma prod_fun_imageI [intro]: "(a, b) : r ==> (f a, g b) : prod_fun f g ` r" |
|
566 |
apply (rule image_eqI) |
|
14208 | 567 |
apply (rule prod_fun [symmetric], assumption) |
11838 | 568 |
done |
569 |
||
570 |
lemma prod_fun_imageE [elim!]: |
|
571 |
"[| c: (prod_fun f g)`r; !!x y. [| c=(f(x),g(y)); (x,y):r |] ==> P |
|
572 |
|] ==> P" |
|
573 |
proof - |
|
574 |
case rule_context |
|
575 |
assume major: "c: (prod_fun f g)`r" |
|
576 |
show ?thesis |
|
577 |
apply (rule major [THEN imageE]) |
|
578 |
apply (rule_tac p = x in PairE) |
|
579 |
apply (rule rule_context) |
|
580 |
prefer 2 |
|
581 |
apply blast |
|
582 |
apply (blast intro: prod_fun) |
|
583 |
done |
|
584 |
qed |
|
585 |
||
586 |
||
14101 | 587 |
constdefs |
588 |
upd_fst :: "('a => 'c) => 'a * 'b => 'c * 'b" |
|
589 |
"upd_fst f == prod_fun f id" |
|
590 |
||
591 |
upd_snd :: "('b => 'c) => 'a * 'b => 'a * 'c" |
|
592 |
"upd_snd f == prod_fun id f" |
|
593 |
||
594 |
lemma upd_fst_conv [simp]: "upd_fst f (x,y) = (f x,y)" |
|
595 |
by (simp add: upd_fst_def) |
|
596 |
||
597 |
lemma upd_snd_conv [simp]: "upd_snd f (x,y) = (x,f y)" |
|
598 |
by (simp add: upd_snd_def) |
|
599 |
||
11838 | 600 |
text {* |
601 |
\bigskip Disjoint union of a family of sets -- Sigma. |
|
602 |
*} |
|
603 |
||
604 |
lemma SigmaI [intro!]: "[| a:A; b:B(a) |] ==> (a,b) : Sigma A B" |
|
605 |
by (unfold Sigma_def) blast |
|
606 |
||
14952
47455995693d
removal of x-symbol syntax <Sigma> for dependent products
paulson
parents:
14565
diff
changeset
|
607 |
lemma SigmaE [elim!]: |
11838 | 608 |
"[| c: Sigma A B; |
609 |
!!x y.[| x:A; y:B(x); c=(x,y) |] ==> P |
|
610 |
|] ==> P" |
|
611 |
-- {* The general elimination rule. *} |
|
612 |
by (unfold Sigma_def) blast |
|
613 |
||
614 |
text {* |
|
15404 | 615 |
Elimination of @{term "(a, b) : A \\<times> B"} -- introduces no |
11838 | 616 |
eigenvariables. |
617 |
*} |
|
618 |
||
619 |
lemma SigmaD1: "(a, b) : Sigma A B ==> a : A" |
|
14952
47455995693d
removal of x-symbol syntax <Sigma> for dependent products
paulson
parents:
14565
diff
changeset
|
620 |
by blast |
11838 | 621 |
|
622 |
lemma SigmaD2: "(a, b) : Sigma A B ==> b : B a" |
|
14952
47455995693d
removal of x-symbol syntax <Sigma> for dependent products
paulson
parents:
14565
diff
changeset
|
623 |
by blast |
11838 | 624 |
|
625 |
lemma SigmaE2: |
|
626 |
"[| (a, b) : Sigma A B; |
|
627 |
[| a:A; b:B(a) |] ==> P |
|
628 |
|] ==> P" |
|
14952
47455995693d
removal of x-symbol syntax <Sigma> for dependent products
paulson
parents:
14565
diff
changeset
|
629 |
by blast |
11838 | 630 |
|
14952
47455995693d
removal of x-symbol syntax <Sigma> for dependent products
paulson
parents:
14565
diff
changeset
|
631 |
lemma Sigma_cong: |
15404 | 632 |
"\\<lbrakk>A = B; !!x. x \\<in> B \\<Longrightarrow> C x = D x\\<rbrakk> |
633 |
\\<Longrightarrow> (SIGMA x: A. C x) = (SIGMA x: B. D x)" |
|
14952
47455995693d
removal of x-symbol syntax <Sigma> for dependent products
paulson
parents:
14565
diff
changeset
|
634 |
by auto |
11838 | 635 |
|
636 |
lemma Sigma_mono: "[| A <= C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D" |
|
637 |
by blast |
|
638 |
||
639 |
lemma Sigma_empty1 [simp]: "Sigma {} B = {}" |
|
640 |
by blast |
|
641 |
||
642 |
lemma Sigma_empty2 [simp]: "A <*> {} = {}" |
|
643 |
by blast |
|
644 |
||
645 |
lemma UNIV_Times_UNIV [simp]: "UNIV <*> UNIV = UNIV" |
|
646 |
by auto |
|
647 |
||
648 |
lemma Compl_Times_UNIV1 [simp]: "- (UNIV <*> A) = UNIV <*> (-A)" |
|
649 |
by auto |
|
650 |
||
651 |
lemma Compl_Times_UNIV2 [simp]: "- (A <*> UNIV) = (-A) <*> UNIV" |
|
652 |
by auto |
|
653 |
||
654 |
lemma mem_Sigma_iff [iff]: "((a,b): Sigma A B) = (a:A & b:B(a))" |
|
655 |
by blast |
|
656 |
||
657 |
lemma Times_subset_cancel2: "x:C ==> (A <*> C <= B <*> C) = (A <= B)" |
|
658 |
by blast |
|
659 |
||
660 |
lemma Times_eq_cancel2: "x:C ==> (A <*> C = B <*> C) = (A = B)" |
|
661 |
by (blast elim: equalityE) |
|
662 |
||
663 |
lemma SetCompr_Sigma_eq: |
|
664 |
"Collect (split (%x y. P x & Q x y)) = (SIGMA x:Collect P. Collect (Q x))" |
|
665 |
by blast |
|
666 |
||
667 |
text {* |
|
668 |
\bigskip Complex rules for Sigma. |
|
669 |
*} |
|
670 |
||
671 |
lemma Collect_split [simp]: "{(a,b). P a & Q b} = Collect P <*> Collect Q" |
|
672 |
by blast |
|
673 |
||
674 |
lemma UN_Times_distrib: |
|
675 |
"(UN (a,b):(A <*> B). E a <*> F b) = (UNION A E) <*> (UNION B F)" |
|
676 |
-- {* Suggested by Pierre Chartier *} |
|
677 |
by blast |
|
678 |
||
679 |
lemma split_paired_Ball_Sigma [simp]: |
|
680 |
"(ALL z: Sigma A B. P z) = (ALL x:A. ALL y: B x. P(x,y))" |
|
681 |
by blast |
|
682 |
||
683 |
lemma split_paired_Bex_Sigma [simp]: |
|
684 |
"(EX z: Sigma A B. P z) = (EX x:A. EX y: B x. P(x,y))" |
|
685 |
by blast |
|
686 |
||
687 |
lemma Sigma_Un_distrib1: "(SIGMA i:I Un J. C(i)) = (SIGMA i:I. C(i)) Un (SIGMA j:J. C(j))" |
|
688 |
by blast |
|
689 |
||
690 |
lemma Sigma_Un_distrib2: "(SIGMA i:I. A(i) Un B(i)) = (SIGMA i:I. A(i)) Un (SIGMA i:I. B(i))" |
|
691 |
by blast |
|
692 |
||
693 |
lemma Sigma_Int_distrib1: "(SIGMA i:I Int J. C(i)) = (SIGMA i:I. C(i)) Int (SIGMA j:J. C(j))" |
|
694 |
by blast |
|
695 |
||
696 |
lemma Sigma_Int_distrib2: "(SIGMA i:I. A(i) Int B(i)) = (SIGMA i:I. A(i)) Int (SIGMA i:I. B(i))" |
|
697 |
by blast |
|
698 |
||
699 |
lemma Sigma_Diff_distrib1: "(SIGMA i:I - J. C(i)) = (SIGMA i:I. C(i)) - (SIGMA j:J. C(j))" |
|
700 |
by blast |
|
701 |
||
702 |
lemma Sigma_Diff_distrib2: "(SIGMA i:I. A(i) - B(i)) = (SIGMA i:I. A(i)) - (SIGMA i:I. B(i))" |
|
703 |
by blast |
|
704 |
||
705 |
lemma Sigma_Union: "Sigma (Union X) B = (UN A:X. Sigma A B)" |
|
706 |
by blast |
|
707 |
||
708 |
text {* |
|
709 |
Non-dependent versions are needed to avoid the need for higher-order |
|
710 |
matching, especially when the rules are re-oriented. |
|
711 |
*} |
|
712 |
||
713 |
lemma Times_Un_distrib1: "(A Un B) <*> C = (A <*> C) Un (B <*> C)" |
|
714 |
by blast |
|
715 |
||
716 |
lemma Times_Int_distrib1: "(A Int B) <*> C = (A <*> C) Int (B <*> C)" |
|
717 |
by blast |
|
718 |
||
719 |
lemma Times_Diff_distrib1: "(A - B) <*> C = (A <*> C) - (B <*> C)" |
|
720 |
by blast |
|
721 |
||
722 |
||
11493 | 723 |
lemma pair_imageI [intro]: "(a, b) : A ==> f a b : (%(a, b). f a b) ` A" |
11777 | 724 |
apply (rule_tac x = "(a, b)" in image_eqI) |
725 |
apply auto |
|
726 |
done |
|
727 |
||
11493 | 728 |
|
11838 | 729 |
text {* |
730 |
Setup of internal @{text split_rule}. |
|
731 |
*} |
|
732 |
||
11032 | 733 |
constdefs |
11425 | 734 |
internal_split :: "('a => 'b => 'c) => 'a * 'b => 'c" |
11032 | 735 |
"internal_split == split" |
736 |
||
737 |
lemma internal_split_conv: "internal_split c (a, b) = c a b" |
|
738 |
by (simp only: internal_split_def split_conv) |
|
739 |
||
740 |
hide const internal_split |
|
741 |
||
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
742 |
use "Tools/split_rule.ML" |
11032 | 743 |
setup SplitRule.setup |
10213 | 744 |
|
15394 | 745 |
|
746 |
subsection {* Code generator setup *} |
|
747 |
||
748 |
types_code |
|
749 |
"*" ("(_ */ _)") |
|
750 |
||
751 |
consts_code |
|
752 |
"Pair" ("(_,/ _)") |
|
753 |
"fst" ("fst") |
|
754 |
"snd" ("snd") |
|
755 |
||
756 |
ML {* |
|
757 |
fun term_of_id_42 f T g U (x, y) = HOLogic.pair_const T U $ f x $ g y; |
|
758 |
||
759 |
fun gen_id_42 aG bG i = (aG i, bG i); |
|
760 |
||
761 |
local |
|
762 |
||
763 |
fun strip_abs 0 t = ([], t) |
|
764 |
| strip_abs i (Abs (s, T, t)) = |
|
765 |
let |
|
766 |
val s' = Codegen.new_name t s; |
|
767 |
val v = Free (s', T) |
|
768 |
in apfst (cons v) (strip_abs (i-1) (subst_bound (v, t))) end |
|
769 |
| strip_abs i (u as Const ("split", _) $ t) = (case strip_abs (i+1) t of |
|
770 |
(v :: v' :: vs, u) => (HOLogic.mk_prod (v, v') :: vs, u) |
|
771 |
| _ => ([], u)) |
|
772 |
| strip_abs i t = ([], t); |
|
773 |
||
774 |
fun let_codegen thy gr dep brack (t as Const ("Let", _) $ _ $ _) = |
|
775 |
let |
|
776 |
fun dest_let (l as Const ("Let", _) $ t $ u) = |
|
777 |
(case strip_abs 1 u of |
|
778 |
([p], u') => apfst (cons (p, t)) (dest_let u') |
|
779 |
| _ => ([], l)) |
|
780 |
| dest_let t = ([], t); |
|
781 |
fun mk_code (gr, (l, r)) = |
|
782 |
let |
|
783 |
val (gr1, pl) = Codegen.invoke_codegen thy dep false (gr, l); |
|
784 |
val (gr2, pr) = Codegen.invoke_codegen thy dep false (gr1, r); |
|
785 |
in (gr2, (pl, pr)) end |
|
786 |
in case dest_let t of |
|
787 |
([], _) => None |
|
788 |
| (ps, u) => |
|
789 |
let |
|
790 |
val (gr1, qs) = foldl_map mk_code (gr, ps); |
|
791 |
val (gr2, pu) = Codegen.invoke_codegen thy dep false (gr1, u) |
|
792 |
in |
|
793 |
Some (gr2, Pretty.blk (0, [Pretty.str "let ", Pretty.blk (0, flat |
|
794 |
(separate [Pretty.str ";", Pretty.brk 1] (map (fn (pl, pr) => |
|
795 |
[Pretty.block [Pretty.str "val ", pl, Pretty.str " =", |
|
796 |
Pretty.brk 1, pr]]) qs))), |
|
797 |
Pretty.brk 1, Pretty.str "in ", pu, |
|
798 |
Pretty.brk 1, Pretty.str "end"])) |
|
799 |
end |
|
800 |
end |
|
801 |
| let_codegen thy gr dep brack t = None; |
|
802 |
||
803 |
fun split_codegen thy gr dep brack (t as Const ("split", _) $ _) = |
|
804 |
(case strip_abs 1 t of |
|
805 |
([p], u) => |
|
806 |
let |
|
807 |
val (gr1, q) = Codegen.invoke_codegen thy dep false (gr, p); |
|
808 |
val (gr2, pu) = Codegen.invoke_codegen thy dep false (gr1, u) |
|
809 |
in |
|
810 |
Some (gr2, Pretty.block [Pretty.str "(fn ", q, Pretty.str " =>", |
|
811 |
Pretty.brk 1, pu, Pretty.str ")"]) |
|
812 |
end |
|
813 |
| _ => None) |
|
814 |
| split_codegen thy gr dep brack t = None; |
|
815 |
||
816 |
in |
|
817 |
||
818 |
val prod_codegen_setup = |
|
819 |
[Codegen.add_codegen "let_codegen" let_codegen, |
|
820 |
Codegen.add_codegen "split_codegen" split_codegen]; |
|
821 |
||
822 |
end; |
|
823 |
*} |
|
824 |
||
825 |
setup prod_codegen_setup |
|
826 |
||
15404 | 827 |
ML |
828 |
{* |
|
829 |
val Collect_split = thm "Collect_split"; |
|
830 |
val Compl_Times_UNIV1 = thm "Compl_Times_UNIV1"; |
|
831 |
val Compl_Times_UNIV2 = thm "Compl_Times_UNIV2"; |
|
832 |
val PairE = thm "PairE"; |
|
833 |
val PairE_lemma = thm "PairE_lemma"; |
|
834 |
val Pair_Rep_inject = thm "Pair_Rep_inject"; |
|
835 |
val Pair_def = thm "Pair_def"; |
|
836 |
val Pair_eq = thm "Pair_eq"; |
|
837 |
val Pair_fst_snd_eq = thm "Pair_fst_snd_eq"; |
|
838 |
val Pair_inject = thm "Pair_inject"; |
|
839 |
val ProdI = thm "ProdI"; |
|
840 |
val SetCompr_Sigma_eq = thm "SetCompr_Sigma_eq"; |
|
841 |
val SigmaD1 = thm "SigmaD1"; |
|
842 |
val SigmaD2 = thm "SigmaD2"; |
|
843 |
val SigmaE = thm "SigmaE"; |
|
844 |
val SigmaE2 = thm "SigmaE2"; |
|
845 |
val SigmaI = thm "SigmaI"; |
|
846 |
val Sigma_Diff_distrib1 = thm "Sigma_Diff_distrib1"; |
|
847 |
val Sigma_Diff_distrib2 = thm "Sigma_Diff_distrib2"; |
|
848 |
val Sigma_Int_distrib1 = thm "Sigma_Int_distrib1"; |
|
849 |
val Sigma_Int_distrib2 = thm "Sigma_Int_distrib2"; |
|
850 |
val Sigma_Un_distrib1 = thm "Sigma_Un_distrib1"; |
|
851 |
val Sigma_Un_distrib2 = thm "Sigma_Un_distrib2"; |
|
852 |
val Sigma_Union = thm "Sigma_Union"; |
|
853 |
val Sigma_def = thm "Sigma_def"; |
|
854 |
val Sigma_empty1 = thm "Sigma_empty1"; |
|
855 |
val Sigma_empty2 = thm "Sigma_empty2"; |
|
856 |
val Sigma_mono = thm "Sigma_mono"; |
|
857 |
val The_split = thm "The_split"; |
|
858 |
val The_split_eq = thm "The_split_eq"; |
|
859 |
val The_split_eq = thm "The_split_eq"; |
|
860 |
val Times_Diff_distrib1 = thm "Times_Diff_distrib1"; |
|
861 |
val Times_Int_distrib1 = thm "Times_Int_distrib1"; |
|
862 |
val Times_Un_distrib1 = thm "Times_Un_distrib1"; |
|
863 |
val Times_eq_cancel2 = thm "Times_eq_cancel2"; |
|
864 |
val Times_subset_cancel2 = thm "Times_subset_cancel2"; |
|
865 |
val UNIV_Times_UNIV = thm "UNIV_Times_UNIV"; |
|
866 |
val UN_Times_distrib = thm "UN_Times_distrib"; |
|
867 |
val Unity_def = thm "Unity_def"; |
|
868 |
val cond_split_eta = thm "cond_split_eta"; |
|
869 |
val fst_conv = thm "fst_conv"; |
|
870 |
val fst_def = thm "fst_def"; |
|
871 |
val fst_eqD = thm "fst_eqD"; |
|
872 |
val inj_on_Abs_Prod = thm "inj_on_Abs_Prod"; |
|
873 |
val injective_fst_snd = thm "injective_fst_snd"; |
|
874 |
val mem_Sigma_iff = thm "mem_Sigma_iff"; |
|
875 |
val mem_splitE = thm "mem_splitE"; |
|
876 |
val mem_splitI = thm "mem_splitI"; |
|
877 |
val mem_splitI2 = thm "mem_splitI2"; |
|
878 |
val prod_eqI = thm "prod_eqI"; |
|
879 |
val prod_fun = thm "prod_fun"; |
|
880 |
val prod_fun_compose = thm "prod_fun_compose"; |
|
881 |
val prod_fun_def = thm "prod_fun_def"; |
|
882 |
val prod_fun_ident = thm "prod_fun_ident"; |
|
883 |
val prod_fun_imageE = thm "prod_fun_imageE"; |
|
884 |
val prod_fun_imageI = thm "prod_fun_imageI"; |
|
885 |
val prod_induct = thm "prod_induct"; |
|
886 |
val snd_conv = thm "snd_conv"; |
|
887 |
val snd_def = thm "snd_def"; |
|
888 |
val snd_eqD = thm "snd_eqD"; |
|
889 |
val split = thm "split"; |
|
890 |
val splitD = thm "splitD"; |
|
891 |
val splitD' = thm "splitD'"; |
|
892 |
val splitE = thm "splitE"; |
|
893 |
val splitE' = thm "splitE'"; |
|
894 |
val splitE2 = thm "splitE2"; |
|
895 |
val splitI = thm "splitI"; |
|
896 |
val splitI2 = thm "splitI2"; |
|
897 |
val splitI2' = thm "splitI2'"; |
|
898 |
val split_Pair_apply = thm "split_Pair_apply"; |
|
899 |
val split_beta = thm "split_beta"; |
|
900 |
val split_conv = thm "split_conv"; |
|
901 |
val split_def = thm "split_def"; |
|
902 |
val split_eta = thm "split_eta"; |
|
903 |
val split_eta_SetCompr = thm "split_eta_SetCompr"; |
|
904 |
val split_eta_SetCompr2 = thm "split_eta_SetCompr2"; |
|
905 |
val split_paired_All = thm "split_paired_All"; |
|
906 |
val split_paired_Ball_Sigma = thm "split_paired_Ball_Sigma"; |
|
907 |
val split_paired_Bex_Sigma = thm "split_paired_Bex_Sigma"; |
|
908 |
val split_paired_Ex = thm "split_paired_Ex"; |
|
909 |
val split_paired_The = thm "split_paired_The"; |
|
910 |
val split_paired_all = thm "split_paired_all"; |
|
911 |
val split_part = thm "split_part"; |
|
912 |
val split_split = thm "split_split"; |
|
913 |
val split_split_asm = thm "split_split_asm"; |
|
914 |
val split_tupled_all = thms "split_tupled_all"; |
|
915 |
val split_weak_cong = thm "split_weak_cong"; |
|
916 |
val surj_pair = thm "surj_pair"; |
|
917 |
val surjective_pairing = thm "surjective_pairing"; |
|
918 |
val unit_abs_eta_conv = thm "unit_abs_eta_conv"; |
|
919 |
val unit_all_eq1 = thm "unit_all_eq1"; |
|
920 |
val unit_all_eq2 = thm "unit_all_eq2"; |
|
921 |
val unit_eq = thm "unit_eq"; |
|
922 |
val unit_induct = thm "unit_induct"; |
|
923 |
*} |
|
924 |
||
10213 | 925 |
end |