author | haftmann |
Tue, 25 Sep 2007 12:16:08 +0200 | |
changeset 24699 | c6674504103f |
parent 24286 | 7619080e49f0 |
child 24844 | 98c006a30218 |
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 |
24699
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
10 |
imports Inductive |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
11 |
uses |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
12 |
("Tools/split_rule.ML") |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
13 |
("Tools/inductive_set_package.ML") |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
14 |
("Tools/inductive_realizer.ML") |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
15 |
("Tools/datatype_realizer.ML") |
15131 | 16 |
begin |
11838 | 17 |
|
24699
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
18 |
subsection {* @{typ bool} is a datatype *} |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
19 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
20 |
rep_datatype bool |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
21 |
distinct True_not_False False_not_True |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
22 |
induction bool_induct |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
23 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
24 |
declare case_split [cases type: bool] |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
25 |
-- "prefer plain propositional version" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
26 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
27 |
|
11838 | 28 |
subsection {* Unit *} |
29 |
||
30 |
typedef unit = "{True}" |
|
31 |
proof |
|
20588 | 32 |
show "True : ?unit" .. |
11838 | 33 |
qed |
34 |
||
24699
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
35 |
definition |
11838 | 36 |
Unity :: unit ("'(')") |
24699
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
37 |
where |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
38 |
"() = Abs_unit True" |
11838 | 39 |
|
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24162
diff
changeset
|
40 |
lemma unit_eq [noatp]: "u = ()" |
11838 | 41 |
by (induct u) (simp add: unit_def Unity_def) |
42 |
||
43 |
text {* |
|
44 |
Simplification procedure for @{thm [source] unit_eq}. Cannot use |
|
45 |
this rule directly --- it loops! |
|
46 |
*} |
|
47 |
||
48 |
ML_setup {* |
|
13462 | 49 |
val unit_eq_proc = |
24699
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
50 |
let val unit_meta_eq = mk_meta_eq @{thm unit_eq} in |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
51 |
Simplifier.simproc @{theory} "unit_eq" ["x::unit"] |
15531 | 52 |
(fn _ => fn _ => fn t => if HOLogic.is_unit t then NONE else SOME unit_meta_eq) |
13462 | 53 |
end; |
11838 | 54 |
|
55 |
Addsimprocs [unit_eq_proc]; |
|
56 |
*} |
|
57 |
||
24699
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
58 |
lemma unit_induct [noatp,induct type: unit]: "P () ==> P x" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
59 |
by simp |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
60 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
61 |
rep_datatype unit |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
62 |
induction unit_induct |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
63 |
|
11838 | 64 |
lemma unit_all_eq1: "(!!x::unit. PROP P x) == PROP P ()" |
65 |
by simp |
|
66 |
||
67 |
lemma unit_all_eq2: "(!!x::unit. PROP P) == PROP P" |
|
68 |
by (rule triv_forall_equality) |
|
69 |
||
70 |
text {* |
|
71 |
This rewrite counters the effect of @{text unit_eq_proc} on @{term |
|
72 |
[source] "%u::unit. f u"}, replacing it by @{term [source] |
|
73 |
f} rather than by @{term [source] "%u. f ()"}. |
|
74 |
*} |
|
75 |
||
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24162
diff
changeset
|
76 |
lemma unit_abs_eta_conv [simp,noatp]: "(%u::unit. f ()) = f" |
11838 | 77 |
by (rule ext) simp |
10213 | 78 |
|
79 |
||
11838 | 80 |
subsection {* Pairs *} |
10213 | 81 |
|
11777 | 82 |
subsubsection {* Type definition *} |
10213 | 83 |
|
84 |
constdefs |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
85 |
Pair_Rep :: "['a, 'b] => ['a, 'b] => bool" |
11032 | 86 |
"Pair_Rep == (%a b. %x y. x=a & y=b)" |
10213 | 87 |
|
88 |
global |
|
89 |
||
90 |
typedef (Prod) |
|
22838 | 91 |
('a, 'b) "*" (infixr "*" 20) |
11032 | 92 |
= "{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
|
93 |
proof |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
94 |
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
|
95 |
by blast |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
96 |
qed |
10213 | 97 |
|
12114
a8e860c86252
eliminated old "symbols" syntax, use "xsymbols" instead;
wenzelm
parents:
11966
diff
changeset
|
98 |
syntax (xsymbols) |
15422
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
99 |
"*" :: "[type, type] => type" ("(_ \<times>/ _)" [21, 20] 20) |
10213 | 100 |
syntax (HTML output) |
15422
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
101 |
"*" :: "[type, type] => type" ("(_ \<times>/ _)" [21, 20] 20) |
10213 | 102 |
|
11777 | 103 |
local |
10213 | 104 |
|
19535 | 105 |
subsubsection {* Definitions *} |
11777 | 106 |
|
107 |
global |
|
10213 | 108 |
|
109 |
consts |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
110 |
fst :: "'a * 'b => 'a" |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
111 |
snd :: "'a * 'b => 'b" |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
112 |
split :: "[['a, 'b] => 'c, 'a * 'b] => 'c" |
14189 | 113 |
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
|
114 |
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
|
115 |
Pair :: "['a, 'b] => 'a * 'b" |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
116 |
Sigma :: "['a set, 'a => 'b set] => ('a * 'b) set" |
10213 | 117 |
|
11777 | 118 |
local |
10213 | 119 |
|
19535 | 120 |
defs |
121 |
Pair_def: "Pair a b == Abs_Prod (Pair_Rep a b)" |
|
122 |
fst_def: "fst p == THE a. EX b. p = Pair a b" |
|
123 |
snd_def: "snd p == THE b. EX a. p = Pair a b" |
|
24162
8dfd5dd65d82
split off theory Option for benefit of code generator
haftmann
parents:
23247
diff
changeset
|
124 |
split_def: "split == (%c p. c (fst p) (snd p))" |
8dfd5dd65d82
split off theory Option for benefit of code generator
haftmann
parents:
23247
diff
changeset
|
125 |
curry_def: "curry == (%c x y. c (Pair x y))" |
8dfd5dd65d82
split off theory Option for benefit of code generator
haftmann
parents:
23247
diff
changeset
|
126 |
prod_fun_def: "prod_fun f g == split (%x y. Pair (f x) (g y))" |
22744
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
haftmann
parents:
22577
diff
changeset
|
127 |
Sigma_def [code func]: "Sigma A B == UN x:A. UN y:B x. {Pair x y}" |
19535 | 128 |
|
129 |
abbreviation |
|
21404
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21331
diff
changeset
|
130 |
Times :: "['a set, 'b set] => ('a * 'b) set" |
eb85850d3eb7
more robust syntax for definition/abbreviation/notation;
wenzelm
parents:
21331
diff
changeset
|
131 |
(infixr "<*>" 80) where |
19535 | 132 |
"A <*> B == Sigma A (%_. B)" |
133 |
||
21210 | 134 |
notation (xsymbols) |
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19535
diff
changeset
|
135 |
Times (infixr "\<times>" 80) |
19535 | 136 |
|
21210 | 137 |
notation (HTML output) |
19656
09be06943252
tuned concrete syntax -- abbreviation/const_syntax;
wenzelm
parents:
19535
diff
changeset
|
138 |
Times (infixr "\<times>" 80) |
19535 | 139 |
|
140 |
||
141 |
subsubsection {* Concrete syntax *} |
|
142 |
||
11777 | 143 |
text {* |
144 |
Patterns -- extends pre-defined type @{typ pttrn} used in |
|
145 |
abstractions. |
|
146 |
*} |
|
10213 | 147 |
|
148 |
nonterminals |
|
149 |
tuple_args patterns |
|
150 |
||
151 |
syntax |
|
152 |
"_tuple" :: "'a => tuple_args => 'a * 'b" ("(1'(_,/ _'))") |
|
153 |
"_tuple_arg" :: "'a => tuple_args" ("_") |
|
154 |
"_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
|
155 |
"_pattern" :: "[pttrn, patterns] => pttrn" ("'(_,/ _')") |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
156 |
"" :: "pttrn => patterns" ("_") |
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
157 |
"_patterns" :: "[pttrn, patterns] => patterns" ("_,/ _") |
22439 | 158 |
"@Sigma" ::"[pttrn, 'a set, 'b set] => ('a * 'b) set" ("(3SIGMA _:_./ _)" [0, 0, 10] 10) |
10213 | 159 |
|
160 |
translations |
|
161 |
"(x, y)" == "Pair x y" |
|
162 |
"_tuple x (_tuple_args y z)" == "_tuple x (_tuple_arg (_tuple y z))" |
|
163 |
"%(x,y,zs).b" == "split(%x (y,zs).b)" |
|
164 |
"%(x,y).b" == "split(%x y. b)" |
|
165 |
"_abs (Pair x y) t" => "%(x,y).t" |
|
166 |
(* The last rule accommodates tuples in `case C ... (x,y) ... => ...' |
|
167 |
The (x,y) is parsed as `Pair x y' because it is logic, not pttrn *) |
|
19535 | 168 |
"SIGMA x:A. B" == "Sigma A (%x. B)" |
10213 | 169 |
|
14359 | 170 |
(* reconstructs pattern from (nested) splits, avoiding eta-contraction of body*) |
171 |
(* works best with enclosing "let", if "let" does not avoid eta-contraction *) |
|
172 |
print_translation {* |
|
173 |
let fun split_tr' [Abs (x,T,t as (Abs abs))] = |
|
174 |
(* split (%x y. t) => %(x,y) t *) |
|
175 |
let val (y,t') = atomic_abs_tr' abs; |
|
176 |
val (x',t'') = atomic_abs_tr' (x,T,t'); |
|
177 |
||
178 |
in Syntax.const "_abs" $ (Syntax.const "_pattern" $x'$y) $ t'' end |
|
179 |
| split_tr' [Abs (x,T,(s as Const ("split",_)$t))] = |
|
180 |
(* split (%x. (split (%y z. t))) => %(x,y,z). t *) |
|
181 |
let val (Const ("_abs",_)$(Const ("_pattern",_)$y$z)$t') = split_tr' [t]; |
|
182 |
val (x',t'') = atomic_abs_tr' (x,T,t'); |
|
183 |
in Syntax.const "_abs"$ |
|
184 |
(Syntax.const "_pattern"$x'$(Syntax.const "_patterns"$y$z))$t'' end |
|
185 |
| split_tr' [Const ("split",_)$t] = |
|
186 |
(* split (split (%x y z. t)) => %((x,y),z). t *) |
|
187 |
split_tr' [(split_tr' [t])] (* inner split_tr' creates next pattern *) |
|
188 |
| split_tr' [Const ("_abs",_)$x_y$(Abs abs)] = |
|
189 |
(* split (%pttrn z. t) => %(pttrn,z). t *) |
|
190 |
let val (z,t) = atomic_abs_tr' abs; |
|
191 |
in Syntax.const "_abs" $ (Syntax.const "_pattern" $x_y$z) $ t end |
|
192 |
| split_tr' _ = raise Match; |
|
193 |
in [("split", split_tr')] |
|
194 |
end |
|
195 |
*} |
|
196 |
||
15422
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
197 |
(* print "split f" as "\<lambda>(x,y). f x y" and "split (\<lambda>x. f x)" as "\<lambda>(x,y). f x y" *) |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
198 |
typed_print_translation {* |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
199 |
let |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
200 |
fun split_guess_names_tr' _ T [Abs (x,_,Abs _)] = raise Match |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
201 |
| split_guess_names_tr' _ T [Abs (x,xT,t)] = |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
202 |
(case (head_of t) of |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
203 |
Const ("split",_) => raise Match |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
204 |
| _ => let |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
205 |
val (_::yT::_) = binder_types (domain_type T) handle Bind => raise Match; |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
206 |
val (y,t') = atomic_abs_tr' ("y",yT,(incr_boundvars 1 t)$Bound 0); |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
207 |
val (x',t'') = atomic_abs_tr' (x,xT,t'); |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
208 |
in Syntax.const "_abs" $ (Syntax.const "_pattern" $x'$y) $ t'' end) |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
209 |
| split_guess_names_tr' _ T [t] = |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
210 |
(case (head_of t) of |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
211 |
Const ("split",_) => raise Match |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
212 |
| _ => let |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
213 |
val (xT::yT::_) = binder_types (domain_type T) handle Bind => raise Match; |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
214 |
val (y,t') = |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
215 |
atomic_abs_tr' ("y",yT,(incr_boundvars 2 t)$Bound 1$Bound 0); |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
216 |
val (x',t'') = atomic_abs_tr' ("x",xT,t'); |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
217 |
in Syntax.const "_abs" $ (Syntax.const "_pattern" $x'$y) $ t'' end) |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
218 |
| split_guess_names_tr' _ _ _ = raise Match; |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
219 |
in [("split", split_guess_names_tr')] |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
220 |
end |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
221 |
*} |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
222 |
|
10213 | 223 |
|
11966 | 224 |
subsubsection {* Lemmas and proof tool setup *} |
11838 | 225 |
|
226 |
lemma ProdI: "Pair_Rep a b : Prod" |
|
19535 | 227 |
unfolding Prod_def by blast |
11838 | 228 |
|
229 |
lemma Pair_Rep_inject: "Pair_Rep a b = Pair_Rep a' b' ==> a = a' & b = b'" |
|
230 |
apply (unfold Pair_Rep_def) |
|
14208 | 231 |
apply (drule fun_cong [THEN fun_cong], blast) |
11838 | 232 |
done |
10213 | 233 |
|
11838 | 234 |
lemma inj_on_Abs_Prod: "inj_on Abs_Prod Prod" |
235 |
apply (rule inj_on_inverseI) |
|
236 |
apply (erule Abs_Prod_inverse) |
|
237 |
done |
|
238 |
||
239 |
lemma Pair_inject: |
|
18372 | 240 |
assumes "(a, b) = (a', b')" |
241 |
and "a = a' ==> b = b' ==> R" |
|
242 |
shows R |
|
243 |
apply (insert prems [unfolded Pair_def]) |
|
244 |
apply (rule inj_on_Abs_Prod [THEN inj_onD, THEN Pair_Rep_inject, THEN conjE]) |
|
245 |
apply (assumption | rule ProdI)+ |
|
246 |
done |
|
10213 | 247 |
|
11838 | 248 |
lemma Pair_eq [iff]: "((a, b) = (a', b')) = (a = a' & b = b')" |
249 |
by (blast elim!: Pair_inject) |
|
250 |
||
22886 | 251 |
lemma fst_conv [simp, code]: "fst (a, b) = a" |
19535 | 252 |
unfolding fst_def by blast |
11838 | 253 |
|
22886 | 254 |
lemma snd_conv [simp, code]: "snd (a, b) = b" |
19535 | 255 |
unfolding snd_def by blast |
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
256 |
|
11838 | 257 |
lemma fst_eqD: "fst (x, y) = a ==> x = a" |
258 |
by simp |
|
259 |
||
260 |
lemma snd_eqD: "snd (x, y) = a ==> y = a" |
|
261 |
by simp |
|
262 |
||
263 |
lemma PairE_lemma: "EX x y. p = (x, y)" |
|
264 |
apply (unfold Pair_def) |
|
265 |
apply (rule Rep_Prod [unfolded Prod_def, THEN CollectE]) |
|
266 |
apply (erule exE, erule exE, rule exI, rule exI) |
|
267 |
apply (rule Rep_Prod_inverse [symmetric, THEN trans]) |
|
268 |
apply (erule arg_cong) |
|
269 |
done |
|
11032 | 270 |
|
11838 | 271 |
lemma PairE [cases type: *]: "(!!x y. p = (x, y) ==> Q) ==> Q" |
19535 | 272 |
using PairE_lemma [of p] by blast |
11838 | 273 |
|
16121 | 274 |
ML {* |
11838 | 275 |
local val PairE = thm "PairE" in |
276 |
fun pair_tac s = |
|
277 |
EVERY' [res_inst_tac [("p", s)] PairE, hyp_subst_tac, K prune_params_tac]; |
|
278 |
end; |
|
279 |
*} |
|
11032 | 280 |
|
11838 | 281 |
lemma surjective_pairing: "p = (fst p, snd p)" |
282 |
-- {* Do not add as rewrite rule: invalidates some proofs in IMP *} |
|
283 |
by (cases p) simp |
|
284 |
||
17085 | 285 |
lemmas pair_collapse = surjective_pairing [symmetric] |
286 |
declare pair_collapse [simp] |
|
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
287 |
|
11838 | 288 |
lemma surj_pair [simp]: "EX x y. z = (x, y)" |
289 |
apply (rule exI) |
|
290 |
apply (rule exI) |
|
291 |
apply (rule surjective_pairing) |
|
292 |
done |
|
293 |
||
294 |
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
|
295 |
proof |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
296 |
fix a b |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
297 |
assume "!!x. PROP P x" |
19535 | 298 |
then show "PROP P (a, b)" . |
11820
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
299 |
next |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
300 |
fix x |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
301 |
assume "!!a b. PROP P (a, b)" |
19535 | 302 |
from `PROP P (fst x, snd x)` show "PROP P x" by simp |
11820
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
303 |
qed |
015a82d4ee96
proper proof of split_paired_all (presently unused);
wenzelm
parents:
11777
diff
changeset
|
304 |
|
11838 | 305 |
lemmas split_tupled_all = split_paired_all unit_all_eq2 |
306 |
||
307 |
text {* |
|
308 |
The rule @{thm [source] split_paired_all} does not work with the |
|
309 |
Simplifier because it also affects premises in congrence rules, |
|
310 |
where this can lead to premises of the form @{text "!!a b. ... = |
|
311 |
?P(a, b)"} which cannot be solved by reflexivity. |
|
312 |
*} |
|
313 |
||
16121 | 314 |
ML_setup {* |
11838 | 315 |
(* replace parameters of product type by individual component parameters *) |
316 |
val safe_full_simp_tac = generic_simp_tac true (true, false, false); |
|
317 |
local (* filtering with exists_paired_all is an essential optimization *) |
|
16121 | 318 |
fun exists_paired_all (Const ("all", _) $ Abs (_, T, t)) = |
11838 | 319 |
can HOLogic.dest_prodT T orelse exists_paired_all t |
320 |
| exists_paired_all (t $ u) = exists_paired_all t orelse exists_paired_all u |
|
321 |
| exists_paired_all (Abs (_, _, t)) = exists_paired_all t |
|
322 |
| exists_paired_all _ = false; |
|
323 |
val ss = HOL_basic_ss |
|
16121 | 324 |
addsimps [thm "split_paired_all", thm "unit_all_eq2", thm "unit_abs_eta_conv"] |
11838 | 325 |
addsimprocs [unit_eq_proc]; |
326 |
in |
|
327 |
val split_all_tac = SUBGOAL (fn (t, i) => |
|
328 |
if exists_paired_all t then safe_full_simp_tac ss i else no_tac); |
|
329 |
val unsafe_split_all_tac = SUBGOAL (fn (t, i) => |
|
330 |
if exists_paired_all t then full_simp_tac ss i else no_tac); |
|
331 |
fun split_all th = |
|
332 |
if exists_paired_all (#prop (Thm.rep_thm th)) then full_simplify ss th else th; |
|
333 |
end; |
|
334 |
||
17875 | 335 |
change_claset (fn cs => cs addSbefore ("split_all_tac", split_all_tac)); |
16121 | 336 |
*} |
11838 | 337 |
|
338 |
lemma split_paired_All [simp]: "(ALL x. P x) = (ALL a b. P (a, b))" |
|
339 |
-- {* @{text "[iff]"} is not a good idea because it makes @{text blast} loop *} |
|
340 |
by fast |
|
341 |
||
14189 | 342 |
lemma curry_split [simp]: "curry (split f) = f" |
343 |
by (simp add: curry_def split_def) |
|
344 |
||
345 |
lemma split_curry [simp]: "split (curry f) = f" |
|
346 |
by (simp add: curry_def split_def) |
|
347 |
||
348 |
lemma curryI [intro!]: "f (a,b) ==> curry f a b" |
|
349 |
by (simp add: curry_def) |
|
350 |
||
14190
609c072edf90
Fixed blunder in the setup of the classical reasoner wrt. the constant
skalberg
parents:
14189
diff
changeset
|
351 |
lemma curryD [dest!]: "curry f a b ==> f (a,b)" |
14189 | 352 |
by (simp add: curry_def) |
353 |
||
14190
609c072edf90
Fixed blunder in the setup of the classical reasoner wrt. the constant
skalberg
parents:
14189
diff
changeset
|
354 |
lemma curryE: "[| curry f a b ; f (a,b) ==> Q |] ==> Q" |
14189 | 355 |
by (simp add: curry_def) |
356 |
||
24162
8dfd5dd65d82
split off theory Option for benefit of code generator
haftmann
parents:
23247
diff
changeset
|
357 |
lemma curry_conv [simp, code func]: "curry f a b = f (a,b)" |
14189 | 358 |
by (simp add: curry_def) |
359 |
||
11838 | 360 |
lemma prod_induct [induct type: *]: "!!x. (!!a b. P (a, b)) ==> P x" |
361 |
by fast |
|
362 |
||
24699
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
363 |
rep_datatype prod |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
364 |
inject Pair_eq |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
365 |
induction prod_induct |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
366 |
|
11838 | 367 |
lemma split_paired_Ex [simp]: "(EX x. P x) = (EX a b. P (a, b))" |
368 |
by fast |
|
369 |
||
24162
8dfd5dd65d82
split off theory Option for benefit of code generator
haftmann
parents:
23247
diff
changeset
|
370 |
lemma split_conv [simp, code func]: "split c (a, b) = c a b" |
11838 | 371 |
by (simp add: split_def) |
372 |
||
373 |
lemmas split = split_conv -- {* for backwards compatibility *} |
|
374 |
||
375 |
lemmas splitI = split_conv [THEN iffD2, standard] |
|
376 |
lemmas splitD = split_conv [THEN iffD1, standard] |
|
377 |
||
378 |
lemma split_Pair_apply: "split (%x y. f (x, y)) = f" |
|
379 |
-- {* Subsumes the old @{text split_Pair} when @{term f} is the identity function. *} |
|
380 |
apply (rule ext) |
|
14208 | 381 |
apply (tactic {* pair_tac "x" 1 *}, simp) |
11838 | 382 |
done |
383 |
||
384 |
lemma split_paired_The: "(THE x. P x) = (THE (a, b). P (a, b))" |
|
385 |
-- {* Can't be added to simpset: loops! *} |
|
386 |
by (simp add: split_Pair_apply) |
|
387 |
||
388 |
lemma The_split: "The (split P) = (THE xy. P (fst xy) (snd xy))" |
|
389 |
by (simp add: split_def) |
|
390 |
||
391 |
lemma Pair_fst_snd_eq: "!!s t. (s = t) = (fst s = fst t & snd s = snd t)" |
|
14208 | 392 |
by (simp only: split_tupled_all, simp) |
11838 | 393 |
|
394 |
lemma prod_eqI [intro?]: "fst p = fst q ==> snd p = snd q ==> p = q" |
|
395 |
by (simp add: Pair_fst_snd_eq) |
|
396 |
||
397 |
lemma split_weak_cong: "p = q ==> split c p = split c q" |
|
398 |
-- {* Prevents simplification of @{term c}: much faster *} |
|
399 |
by (erule arg_cong) |
|
400 |
||
401 |
lemma split_eta: "(%(x, y). f (x, y)) = f" |
|
402 |
apply (rule ext) |
|
403 |
apply (simp only: split_tupled_all) |
|
404 |
apply (rule split_conv) |
|
405 |
done |
|
406 |
||
407 |
lemma cond_split_eta: "(!!x y. f x y = g (x, y)) ==> (%(x, y). f x y) = g" |
|
408 |
by (simp add: split_eta) |
|
409 |
||
410 |
text {* |
|
411 |
Simplification procedure for @{thm [source] cond_split_eta}. Using |
|
412 |
@{thm [source] split_eta} as a rewrite rule is not general enough, |
|
413 |
and using @{thm [source] cond_split_eta} directly would render some |
|
414 |
existing proofs very inefficient; similarly for @{text |
|
415 |
split_beta}. *} |
|
416 |
||
417 |
ML_setup {* |
|
418 |
||
419 |
local |
|
18328 | 420 |
val cond_split_eta_ss = HOL_basic_ss addsimps [thm "cond_split_eta"] |
11838 | 421 |
fun Pair_pat k 0 (Bound m) = (m = k) |
422 |
| Pair_pat k i (Const ("Pair", _) $ Bound m $ t) = i > 0 andalso |
|
423 |
m = k+i andalso Pair_pat k (i-1) t |
|
424 |
| Pair_pat _ _ _ = false; |
|
425 |
fun no_args k i (Abs (_, _, t)) = no_args (k+1) i t |
|
426 |
| no_args k i (t $ u) = no_args k i t andalso no_args k i u |
|
427 |
| no_args k i (Bound m) = m < k orelse m > k+i |
|
428 |
| no_args _ _ _ = true; |
|
15531 | 429 |
fun split_pat tp i (Abs (_,_,t)) = if tp 0 i t then SOME (i,t) else NONE |
11838 | 430 |
| split_pat tp i (Const ("split", _) $ Abs (_, _, t)) = split_pat tp (i+1) t |
15531 | 431 |
| split_pat tp i _ = NONE; |
20044
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
19656
diff
changeset
|
432 |
fun metaeq ss lhs rhs = mk_meta_eq (Goal.prove (Simplifier.the_context ss) [] [] |
13480
bb72bd43c6c3
use Tactic.prove instead of prove_goalw_cterm in internal proofs!
wenzelm
parents:
13462
diff
changeset
|
433 |
(HOLogic.mk_Trueprop (HOLogic.mk_eq (lhs,rhs))) |
18328 | 434 |
(K (simp_tac (Simplifier.inherit_context ss cond_split_eta_ss) 1))); |
11838 | 435 |
|
436 |
fun beta_term_pat k i (Abs (_, _, t)) = beta_term_pat (k+1) i t |
|
437 |
| beta_term_pat k i (t $ u) = Pair_pat k i (t $ u) orelse |
|
438 |
(beta_term_pat k i t andalso beta_term_pat k i u) |
|
439 |
| beta_term_pat k i t = no_args k i t; |
|
440 |
fun eta_term_pat k i (f $ arg) = no_args k i f andalso Pair_pat k i arg |
|
441 |
| eta_term_pat _ _ _ = false; |
|
442 |
fun subst arg k i (Abs (x, T, t)) = Abs (x, T, subst arg (k+1) i t) |
|
443 |
| subst arg k i (t $ u) = if Pair_pat k i (t $ u) then incr_boundvars k arg |
|
444 |
else (subst arg k i t $ subst arg k i u) |
|
445 |
| subst arg k i t = t; |
|
20044
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
19656
diff
changeset
|
446 |
fun beta_proc ss (s as Const ("split", _) $ Abs (_, _, t) $ arg) = |
11838 | 447 |
(case split_pat beta_term_pat 1 t of |
20044
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
19656
diff
changeset
|
448 |
SOME (i,f) => SOME (metaeq ss s (subst arg 0 i f)) |
15531 | 449 |
| NONE => NONE) |
20044
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
19656
diff
changeset
|
450 |
| beta_proc _ _ = NONE; |
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
19656
diff
changeset
|
451 |
fun eta_proc ss (s as Const ("split", _) $ Abs (_, _, t)) = |
11838 | 452 |
(case split_pat eta_term_pat 1 t of |
20044
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
19656
diff
changeset
|
453 |
SOME (_,ft) => SOME (metaeq ss s (let val (f $ arg) = ft in f end)) |
15531 | 454 |
| NONE => NONE) |
20044
92cc2f4c7335
simprocs: no theory argument -- use simpset context instead;
wenzelm
parents:
19656
diff
changeset
|
455 |
| eta_proc _ _ = NONE; |
11838 | 456 |
in |
22577 | 457 |
val split_beta_proc = Simplifier.simproc @{theory} "split_beta" ["split f z"] (K beta_proc); |
458 |
val split_eta_proc = Simplifier.simproc @{theory} "split_eta" ["split f"] (K eta_proc); |
|
11838 | 459 |
end; |
460 |
||
461 |
Addsimprocs [split_beta_proc, split_eta_proc]; |
|
462 |
*} |
|
463 |
||
464 |
lemma split_beta: "(%(x, y). P x y) z = P (fst z) (snd z)" |
|
465 |
by (subst surjective_pairing, rule split_conv) |
|
466 |
||
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24162
diff
changeset
|
467 |
lemma split_split [noatp]: "R(split c p) = (ALL x y. p = (x, y) --> R(c x y))" |
11838 | 468 |
-- {* For use with @{text split} and the Simplifier. *} |
15481 | 469 |
by (insert surj_pair [of p], clarify, simp) |
11838 | 470 |
|
471 |
text {* |
|
472 |
@{thm [source] split_split} could be declared as @{text "[split]"} |
|
473 |
done after the Splitter has been speeded up significantly; |
|
474 |
precompute the constants involved and don't do anything unless the |
|
475 |
current goal contains one of those constants. |
|
476 |
*} |
|
477 |
||
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24162
diff
changeset
|
478 |
lemma split_split_asm [noatp]: "R (split c p) = (~(EX x y. p = (x, y) & (~R (c x y))))" |
14208 | 479 |
by (subst split_split, simp) |
11838 | 480 |
|
481 |
||
482 |
text {* |
|
483 |
\medskip @{term split} used as a logical connective or set former. |
|
484 |
||
485 |
\medskip These rules are for use with @{text blast}; could instead |
|
486 |
call @{text simp} using @{thm [source] split} as rewrite. *} |
|
487 |
||
488 |
lemma splitI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> split c p" |
|
489 |
apply (simp only: split_tupled_all) |
|
490 |
apply (simp (no_asm_simp)) |
|
491 |
done |
|
492 |
||
493 |
lemma splitI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> split c p x" |
|
494 |
apply (simp only: split_tupled_all) |
|
495 |
apply (simp (no_asm_simp)) |
|
496 |
done |
|
497 |
||
498 |
lemma splitE: "split c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q" |
|
499 |
by (induct p) (auto simp add: split_def) |
|
500 |
||
501 |
lemma splitE': "split c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q" |
|
502 |
by (induct p) (auto simp add: split_def) |
|
503 |
||
504 |
lemma splitE2: |
|
505 |
"[| Q (split P z); !!x y. [|z = (x, y); Q (P x y)|] ==> R |] ==> R" |
|
506 |
proof - |
|
507 |
assume q: "Q (split P z)" |
|
508 |
assume r: "!!x y. [|z = (x, y); Q (P x y)|] ==> R" |
|
509 |
show R |
|
510 |
apply (rule r surjective_pairing)+ |
|
511 |
apply (rule split_beta [THEN subst], rule q) |
|
512 |
done |
|
513 |
qed |
|
514 |
||
515 |
lemma splitD': "split R (a,b) c ==> R a b c" |
|
516 |
by simp |
|
517 |
||
518 |
lemma mem_splitI: "z: c a b ==> z: split c (a, b)" |
|
519 |
by simp |
|
520 |
||
521 |
lemma mem_splitI2: "!!p. [| !!a b. p = (a, b) ==> z: c a b |] ==> z: split c p" |
|
14208 | 522 |
by (simp only: split_tupled_all, simp) |
11838 | 523 |
|
18372 | 524 |
lemma mem_splitE: |
525 |
assumes major: "z: split c p" |
|
526 |
and cases: "!!x y. [| p = (x,y); z: c x y |] ==> Q" |
|
527 |
shows Q |
|
528 |
by (rule major [unfolded split_def] cases surjective_pairing)+ |
|
11838 | 529 |
|
530 |
declare mem_splitI2 [intro!] mem_splitI [intro!] splitI2' [intro!] splitI2 [intro!] splitI [intro!] |
|
531 |
declare mem_splitE [elim!] splitE' [elim!] splitE [elim!] |
|
532 |
||
16121 | 533 |
ML_setup {* |
11838 | 534 |
local (* filtering with exists_p_split is an essential optimization *) |
16121 | 535 |
fun exists_p_split (Const ("split",_) $ _ $ (Const ("Pair",_)$_$_)) = true |
11838 | 536 |
| exists_p_split (t $ u) = exists_p_split t orelse exists_p_split u |
537 |
| exists_p_split (Abs (_, _, t)) = exists_p_split t |
|
538 |
| exists_p_split _ = false; |
|
16121 | 539 |
val ss = HOL_basic_ss addsimps [thm "split_conv"]; |
11838 | 540 |
in |
541 |
val split_conv_tac = SUBGOAL (fn (t, i) => |
|
542 |
if exists_p_split t then safe_full_simp_tac ss i else no_tac); |
|
543 |
end; |
|
544 |
(* This prevents applications of splitE for already splitted arguments leading |
|
545 |
to quite time-consuming computations (in particular for nested tuples) *) |
|
17875 | 546 |
change_claset (fn cs => cs addSbefore ("split_conv_tac", split_conv_tac)); |
16121 | 547 |
*} |
11838 | 548 |
|
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24162
diff
changeset
|
549 |
lemma split_eta_SetCompr [simp,noatp]: "(%u. EX x y. u = (x, y) & P (x, y)) = P" |
18372 | 550 |
by (rule ext) fast |
11838 | 551 |
|
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24162
diff
changeset
|
552 |
lemma split_eta_SetCompr2 [simp,noatp]: "(%u. EX x y. u = (x, y) & P x y) = split P" |
18372 | 553 |
by (rule ext) fast |
11838 | 554 |
|
555 |
lemma split_part [simp]: "(%(a,b). P & Q a b) = (%ab. P & split Q ab)" |
|
556 |
-- {* Allows simplifications of nested splits in case of independent predicates. *} |
|
18372 | 557 |
by (rule ext) blast |
11838 | 558 |
|
14337
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents:
14208
diff
changeset
|
559 |
(* 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
|
560 |
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
|
561 |
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
|
562 |
*) |
e13731554e50
undid split_comp_eq[simp] because it leads to nontermination together with split_def!
nipkow
parents:
14208
diff
changeset
|
563 |
lemma split_comp_eq: |
20415 | 564 |
fixes f :: "'a => 'b => 'c" and g :: "'d => 'a" |
565 |
shows "(%u. f (g (fst u)) (snd u)) = (split (%x. f (g x)))" |
|
18372 | 566 |
by (rule ext) auto |
14101 | 567 |
|
11838 | 568 |
lemma The_split_eq [simp]: "(THE (x',y'). x = x' & y = y') = (x, y)" |
569 |
by blast |
|
570 |
||
571 |
(* |
|
572 |
the following would be slightly more general, |
|
573 |
but cannot be used as rewrite rule: |
|
574 |
### Cannot add premise as rewrite rule because it contains (type) unknowns: |
|
575 |
### ?y = .x |
|
576 |
Goal "[| P y; !!x. P x ==> x = y |] ==> (@(x',y). x = x' & P y) = (x,y)" |
|
14208 | 577 |
by (rtac some_equality 1) |
578 |
by ( Simp_tac 1) |
|
579 |
by (split_all_tac 1) |
|
580 |
by (Asm_full_simp_tac 1) |
|
11838 | 581 |
qed "The_split_eq"; |
582 |
*) |
|
583 |
||
584 |
lemma injective_fst_snd: "!!x y. [|fst x = fst y; snd x = snd y|] ==> x = y" |
|
585 |
by auto |
|
586 |
||
587 |
||
588 |
text {* |
|
589 |
\bigskip @{term prod_fun} --- action of the product functor upon |
|
590 |
functions. |
|
591 |
*} |
|
592 |
||
24162
8dfd5dd65d82
split off theory Option for benefit of code generator
haftmann
parents:
23247
diff
changeset
|
593 |
lemma prod_fun [simp, code func]: "prod_fun f g (a, b) = (f a, g b)" |
11838 | 594 |
by (simp add: prod_fun_def) |
595 |
||
596 |
lemma prod_fun_compose: "prod_fun (f1 o f2) (g1 o g2) = (prod_fun f1 g1 o prod_fun f2 g2)" |
|
597 |
apply (rule ext) |
|
14208 | 598 |
apply (tactic {* pair_tac "x" 1 *}, simp) |
11838 | 599 |
done |
600 |
||
601 |
lemma prod_fun_ident [simp]: "prod_fun (%x. x) (%y. y) = (%z. z)" |
|
602 |
apply (rule ext) |
|
14208 | 603 |
apply (tactic {* pair_tac "z" 1 *}, simp) |
11838 | 604 |
done |
605 |
||
606 |
lemma prod_fun_imageI [intro]: "(a, b) : r ==> (f a, g b) : prod_fun f g ` r" |
|
607 |
apply (rule image_eqI) |
|
14208 | 608 |
apply (rule prod_fun [symmetric], assumption) |
11838 | 609 |
done |
610 |
||
611 |
lemma prod_fun_imageE [elim!]: |
|
18372 | 612 |
assumes major: "c: (prod_fun f g)`r" |
613 |
and cases: "!!x y. [| c=(f(x),g(y)); (x,y):r |] ==> P" |
|
614 |
shows P |
|
615 |
apply (rule major [THEN imageE]) |
|
616 |
apply (rule_tac p = x in PairE) |
|
617 |
apply (rule cases) |
|
618 |
apply (blast intro: prod_fun) |
|
619 |
apply blast |
|
620 |
done |
|
11838 | 621 |
|
622 |
||
22744
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
haftmann
parents:
22577
diff
changeset
|
623 |
definition |
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
haftmann
parents:
22577
diff
changeset
|
624 |
upd_fst :: "('a \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c \<times> 'b" |
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
haftmann
parents:
22577
diff
changeset
|
625 |
where |
22886 | 626 |
[code func del]: "upd_fst f = prod_fun f id" |
14101 | 627 |
|
22744
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
haftmann
parents:
22577
diff
changeset
|
628 |
definition |
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
haftmann
parents:
22577
diff
changeset
|
629 |
upd_snd :: "('b \<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'a \<times> 'c" |
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
haftmann
parents:
22577
diff
changeset
|
630 |
where |
22886 | 631 |
[code func del]: "upd_snd f = prod_fun id f" |
14101 | 632 |
|
22886 | 633 |
lemma upd_fst_conv [simp, code]: |
22744
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
haftmann
parents:
22577
diff
changeset
|
634 |
"upd_fst f (x, y) = (f x, y)" |
18372 | 635 |
by (simp add: upd_fst_def) |
14101 | 636 |
|
22886 | 637 |
lemma upd_snd_conv [simp, code]: |
22744
5cbe966d67a2
Isar definitions are now added explicitly to code theorem table
haftmann
parents:
22577
diff
changeset
|
638 |
"upd_snd f (x, y) = (x, f y)" |
18372 | 639 |
by (simp add: upd_snd_def) |
14101 | 640 |
|
11838 | 641 |
text {* |
642 |
\bigskip Disjoint union of a family of sets -- Sigma. |
|
643 |
*} |
|
644 |
||
645 |
lemma SigmaI [intro!]: "[| a:A; b:B(a) |] ==> (a,b) : Sigma A B" |
|
646 |
by (unfold Sigma_def) blast |
|
647 |
||
14952
47455995693d
removal of x-symbol syntax <Sigma> for dependent products
paulson
parents:
14565
diff
changeset
|
648 |
lemma SigmaE [elim!]: |
11838 | 649 |
"[| c: Sigma A B; |
650 |
!!x y.[| x:A; y:B(x); c=(x,y) |] ==> P |
|
651 |
|] ==> P" |
|
652 |
-- {* The general elimination rule. *} |
|
653 |
by (unfold Sigma_def) blast |
|
654 |
||
655 |
text {* |
|
15422
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
656 |
Elimination of @{term "(a, b) : A \<times> B"} -- introduces no |
11838 | 657 |
eigenvariables. |
658 |
*} |
|
659 |
||
660 |
lemma SigmaD1: "(a, b) : Sigma A B ==> a : A" |
|
18372 | 661 |
by blast |
11838 | 662 |
|
663 |
lemma SigmaD2: "(a, b) : Sigma A B ==> b : B a" |
|
18372 | 664 |
by blast |
11838 | 665 |
|
666 |
lemma SigmaE2: |
|
667 |
"[| (a, b) : Sigma A B; |
|
668 |
[| a:A; b:B(a) |] ==> P |
|
669 |
|] ==> P" |
|
14952
47455995693d
removal of x-symbol syntax <Sigma> for dependent products
paulson
parents:
14565
diff
changeset
|
670 |
by blast |
11838 | 671 |
|
14952
47455995693d
removal of x-symbol syntax <Sigma> for dependent products
paulson
parents:
14565
diff
changeset
|
672 |
lemma Sigma_cong: |
15422
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
673 |
"\<lbrakk>A = B; !!x. x \<in> B \<Longrightarrow> C x = D x\<rbrakk> |
cbdddc0efadf
added print translation for split: split f --> %(x,y). f x y
schirmer
parents:
15404
diff
changeset
|
674 |
\<Longrightarrow> (SIGMA x: A. C x) = (SIGMA x: B. D x)" |
18372 | 675 |
by auto |
11838 | 676 |
|
677 |
lemma Sigma_mono: "[| A <= C; !!x. x:A ==> B x <= D x |] ==> Sigma A B <= Sigma C D" |
|
678 |
by blast |
|
679 |
||
680 |
lemma Sigma_empty1 [simp]: "Sigma {} B = {}" |
|
681 |
by blast |
|
682 |
||
683 |
lemma Sigma_empty2 [simp]: "A <*> {} = {}" |
|
684 |
by blast |
|
685 |
||
686 |
lemma UNIV_Times_UNIV [simp]: "UNIV <*> UNIV = UNIV" |
|
687 |
by auto |
|
688 |
||
689 |
lemma Compl_Times_UNIV1 [simp]: "- (UNIV <*> A) = UNIV <*> (-A)" |
|
690 |
by auto |
|
691 |
||
692 |
lemma Compl_Times_UNIV2 [simp]: "- (A <*> UNIV) = (-A) <*> UNIV" |
|
693 |
by auto |
|
694 |
||
695 |
lemma mem_Sigma_iff [iff]: "((a,b): Sigma A B) = (a:A & b:B(a))" |
|
696 |
by blast |
|
697 |
||
698 |
lemma Times_subset_cancel2: "x:C ==> (A <*> C <= B <*> C) = (A <= B)" |
|
699 |
by blast |
|
700 |
||
701 |
lemma Times_eq_cancel2: "x:C ==> (A <*> C = B <*> C) = (A = B)" |
|
702 |
by (blast elim: equalityE) |
|
703 |
||
704 |
lemma SetCompr_Sigma_eq: |
|
705 |
"Collect (split (%x y. P x & Q x y)) = (SIGMA x:Collect P. Collect (Q x))" |
|
706 |
by blast |
|
707 |
||
708 |
text {* |
|
709 |
\bigskip Complex rules for Sigma. |
|
710 |
*} |
|
711 |
||
712 |
lemma Collect_split [simp]: "{(a,b). P a & Q b} = Collect P <*> Collect Q" |
|
713 |
by blast |
|
714 |
||
715 |
lemma UN_Times_distrib: |
|
716 |
"(UN (a,b):(A <*> B). E a <*> F b) = (UNION A E) <*> (UNION B F)" |
|
717 |
-- {* Suggested by Pierre Chartier *} |
|
718 |
by blast |
|
719 |
||
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24162
diff
changeset
|
720 |
lemma split_paired_Ball_Sigma [simp,noatp]: |
11838 | 721 |
"(ALL z: Sigma A B. P z) = (ALL x:A. ALL y: B x. P(x,y))" |
722 |
by blast |
|
723 |
||
24286
7619080e49f0
ATP blacklisting is now in theory data, attribute noatp
paulson
parents:
24162
diff
changeset
|
724 |
lemma split_paired_Bex_Sigma [simp,noatp]: |
11838 | 725 |
"(EX z: Sigma A B. P z) = (EX x:A. EX y: B x. P(x,y))" |
726 |
by blast |
|
727 |
||
728 |
lemma Sigma_Un_distrib1: "(SIGMA i:I Un J. C(i)) = (SIGMA i:I. C(i)) Un (SIGMA j:J. C(j))" |
|
729 |
by blast |
|
730 |
||
731 |
lemma Sigma_Un_distrib2: "(SIGMA i:I. A(i) Un B(i)) = (SIGMA i:I. A(i)) Un (SIGMA i:I. B(i))" |
|
732 |
by blast |
|
733 |
||
734 |
lemma Sigma_Int_distrib1: "(SIGMA i:I Int J. C(i)) = (SIGMA i:I. C(i)) Int (SIGMA j:J. C(j))" |
|
735 |
by blast |
|
736 |
||
737 |
lemma Sigma_Int_distrib2: "(SIGMA i:I. A(i) Int B(i)) = (SIGMA i:I. A(i)) Int (SIGMA i:I. B(i))" |
|
738 |
by blast |
|
739 |
||
740 |
lemma Sigma_Diff_distrib1: "(SIGMA i:I - J. C(i)) = (SIGMA i:I. C(i)) - (SIGMA j:J. C(j))" |
|
741 |
by blast |
|
742 |
||
743 |
lemma Sigma_Diff_distrib2: "(SIGMA i:I. A(i) - B(i)) = (SIGMA i:I. A(i)) - (SIGMA i:I. B(i))" |
|
744 |
by blast |
|
745 |
||
746 |
lemma Sigma_Union: "Sigma (Union X) B = (UN A:X. Sigma A B)" |
|
747 |
by blast |
|
748 |
||
749 |
text {* |
|
750 |
Non-dependent versions are needed to avoid the need for higher-order |
|
751 |
matching, especially when the rules are re-oriented. |
|
752 |
*} |
|
753 |
||
754 |
lemma Times_Un_distrib1: "(A Un B) <*> C = (A <*> C) Un (B <*> C)" |
|
755 |
by blast |
|
756 |
||
757 |
lemma Times_Int_distrib1: "(A Int B) <*> C = (A <*> C) Int (B <*> C)" |
|
758 |
by blast |
|
759 |
||
760 |
lemma Times_Diff_distrib1: "(A - B) <*> C = (A <*> C) - (B <*> C)" |
|
761 |
by blast |
|
762 |
||
763 |
||
11493 | 764 |
lemma pair_imageI [intro]: "(a, b) : A ==> f a b : (%(a, b). f a b) ` A" |
11777 | 765 |
apply (rule_tac x = "(a, b)" in image_eqI) |
766 |
apply auto |
|
767 |
done |
|
768 |
||
11493 | 769 |
|
11838 | 770 |
text {* |
771 |
Setup of internal @{text split_rule}. |
|
772 |
*} |
|
773 |
||
11032 | 774 |
constdefs |
11425 | 775 |
internal_split :: "('a => 'b => 'c) => 'a * 'b => 'c" |
11032 | 776 |
"internal_split == split" |
777 |
||
24162
8dfd5dd65d82
split off theory Option for benefit of code generator
haftmann
parents:
23247
diff
changeset
|
778 |
lemmas [code func del] = internal_split_def |
8dfd5dd65d82
split off theory Option for benefit of code generator
haftmann
parents:
23247
diff
changeset
|
779 |
|
11032 | 780 |
lemma internal_split_conv: "internal_split c (a, b) = c a b" |
781 |
by (simp only: internal_split_def split_conv) |
|
782 |
||
783 |
hide const internal_split |
|
784 |
||
11025
a70b796d9af8
converted to Isar therory, adding attributes complete_split and split_format
oheimb
parents:
10289
diff
changeset
|
785 |
use "Tools/split_rule.ML" |
11032 | 786 |
setup SplitRule.setup |
10213 | 787 |
|
15394 | 788 |
|
24699
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
789 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
790 |
lemmas prod_caseI = prod.cases [THEN iffD2, standard] |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
791 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
792 |
lemma prod_caseI2: "!!p. [| !!a b. p = (a, b) ==> c a b |] ==> prod_case c p" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
793 |
by auto |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
794 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
795 |
lemma prod_caseI2': "!!p. [| !!a b. (a, b) = p ==> c a b x |] ==> prod_case c p x" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
796 |
by (auto simp: split_tupled_all) |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
797 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
798 |
lemma prod_caseE: "prod_case c p ==> (!!x y. p = (x, y) ==> c x y ==> Q) ==> Q" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
799 |
by (induct p) auto |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
800 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
801 |
lemma prod_caseE': "prod_case c p z ==> (!!x y. p = (x, y) ==> c x y z ==> Q) ==> Q" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
802 |
by (induct p) auto |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
803 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
804 |
lemma prod_case_unfold: "prod_case = (%c p. c (fst p) (snd p))" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
805 |
by (simp add: expand_fun_eq) |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
806 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
807 |
declare prod_caseI2' [intro!] prod_caseI2 [intro!] prod_caseI [intro!] |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
808 |
declare prod_caseE' [elim!] prod_caseE [elim!] |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
809 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
810 |
lemma prod_case_split [code post]: |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
811 |
"prod_case = split" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
812 |
by (auto simp add: expand_fun_eq) |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
813 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
814 |
lemmas [code inline] = prod_case_split [symmetric] |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
815 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
816 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
817 |
subsection {* Further cases/induct rules for tuples *} |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
818 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
819 |
lemma prod_cases3 [cases type]: |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
820 |
obtains (fields) a b c where "y = (a, b, c)" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
821 |
by (cases y, case_tac b) blast |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
822 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
823 |
lemma prod_induct3 [case_names fields, induct type]: |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
824 |
"(!!a b c. P (a, b, c)) ==> P x" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
825 |
by (cases x) blast |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
826 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
827 |
lemma prod_cases4 [cases type]: |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
828 |
obtains (fields) a b c d where "y = (a, b, c, d)" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
829 |
by (cases y, case_tac c) blast |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
830 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
831 |
lemma prod_induct4 [case_names fields, induct type]: |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
832 |
"(!!a b c d. P (a, b, c, d)) ==> P x" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
833 |
by (cases x) blast |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
834 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
835 |
lemma prod_cases5 [cases type]: |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
836 |
obtains (fields) a b c d e where "y = (a, b, c, d, e)" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
837 |
by (cases y, case_tac d) blast |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
838 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
839 |
lemma prod_induct5 [case_names fields, induct type]: |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
840 |
"(!!a b c d e. P (a, b, c, d, e)) ==> P x" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
841 |
by (cases x) blast |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
842 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
843 |
lemma prod_cases6 [cases type]: |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
844 |
obtains (fields) a b c d e f where "y = (a, b, c, d, e, f)" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
845 |
by (cases y, case_tac e) blast |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
846 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
847 |
lemma prod_induct6 [case_names fields, induct type]: |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
848 |
"(!!a b c d e f. P (a, b, c, d, e, f)) ==> P x" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
849 |
by (cases x) blast |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
850 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
851 |
lemma prod_cases7 [cases type]: |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
852 |
obtains (fields) a b c d e f g where "y = (a, b, c, d, e, f, g)" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
853 |
by (cases y, case_tac f) blast |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
854 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
855 |
lemma prod_induct7 [case_names fields, induct type]: |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
856 |
"(!!a b c d e f g. P (a, b, c, d, e, f, g)) ==> P x" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
857 |
by (cases x) blast |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
858 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
859 |
|
21195 | 860 |
subsection {* Further lemmas *} |
861 |
||
862 |
lemma |
|
863 |
split_Pair: "split Pair x = x" |
|
864 |
unfolding split_def by auto |
|
865 |
||
866 |
lemma |
|
867 |
split_comp: "split (f \<circ> g) x = f (g (fst x)) (snd x)" |
|
868 |
by (cases x, simp) |
|
869 |
||
870 |
||
15394 | 871 |
subsection {* Code generator setup *} |
872 |
||
20588 | 873 |
instance unit :: eq .. |
874 |
||
875 |
lemma [code func]: |
|
21454 | 876 |
"(u\<Colon>unit) = v \<longleftrightarrow> True" unfolding unit_eq [of u] unit_eq [of v] by rule+ |
20588 | 877 |
|
21908 | 878 |
code_type unit |
879 |
(SML "unit") |
|
880 |
(OCaml "unit") |
|
881 |
(Haskell "()") |
|
882 |
||
20588 | 883 |
code_instance unit :: eq |
884 |
(Haskell -) |
|
885 |
||
21908 | 886 |
code_const "op = \<Colon> unit \<Rightarrow> unit \<Rightarrow> bool" |
887 |
(Haskell infixl 4 "==") |
|
888 |
||
889 |
code_const Unity |
|
890 |
(SML "()") |
|
891 |
(OCaml "()") |
|
892 |
(Haskell "()") |
|
893 |
||
894 |
code_reserved SML |
|
895 |
unit |
|
896 |
||
897 |
code_reserved OCaml |
|
898 |
unit |
|
899 |
||
20588 | 900 |
instance * :: (eq, eq) eq .. |
901 |
||
902 |
lemma [code func]: |
|
21454 | 903 |
"(x1\<Colon>'a\<Colon>eq, y1\<Colon>'b\<Colon>eq) = (x2, y2) \<longleftrightarrow> x1 = x2 \<and> y1 = y2" by auto |
20588 | 904 |
|
21908 | 905 |
code_type * |
906 |
(SML infix 2 "*") |
|
907 |
(OCaml infix 2 "*") |
|
908 |
(Haskell "!((_),/ (_))") |
|
909 |
||
20588 | 910 |
code_instance * :: eq |
911 |
(Haskell -) |
|
912 |
||
21908 | 913 |
code_const "op = \<Colon> 'a\<Colon>eq \<times> 'b\<Colon>eq \<Rightarrow> 'a \<times> 'b \<Rightarrow> bool" |
20588 | 914 |
(Haskell infixl 4 "==") |
915 |
||
21908 | 916 |
code_const Pair |
917 |
(SML "!((_),/ (_))") |
|
918 |
(OCaml "!((_),/ (_))") |
|
919 |
(Haskell "!((_),/ (_))") |
|
20588 | 920 |
|
22389 | 921 |
code_const fst and snd |
922 |
(Haskell "fst" and "snd") |
|
923 |
||
15394 | 924 |
types_code |
925 |
"*" ("(_ */ _)") |
|
16770
1f1b1fae30e4
Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents:
16634
diff
changeset
|
926 |
attach (term_of) {* |
1f1b1fae30e4
Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents:
16634
diff
changeset
|
927 |
fun term_of_id_42 f T g U (x, y) = HOLogic.pair_const T U $ f x $ g y; |
1f1b1fae30e4
Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents:
16634
diff
changeset
|
928 |
*} |
1f1b1fae30e4
Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents:
16634
diff
changeset
|
929 |
attach (test) {* |
1f1b1fae30e4
Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents:
16634
diff
changeset
|
930 |
fun gen_id_42 aG bG i = (aG i, bG i); |
1f1b1fae30e4
Auxiliary functions to be used in generated code are now defined using "attach".
berghofe
parents:
16634
diff
changeset
|
931 |
*} |
15394 | 932 |
|
18706
1e7562c7afe6
Re-inserted consts_code declaration accidentally deleted
berghofe
parents:
18702
diff
changeset
|
933 |
consts_code |
1e7562c7afe6
Re-inserted consts_code declaration accidentally deleted
berghofe
parents:
18702
diff
changeset
|
934 |
"Pair" ("(_,/ _)") |
1e7562c7afe6
Re-inserted consts_code declaration accidentally deleted
berghofe
parents:
18702
diff
changeset
|
935 |
|
21908 | 936 |
setup {* |
937 |
||
938 |
let |
|
18013 | 939 |
|
19039 | 940 |
fun strip_abs_split 0 t = ([], t) |
941 |
| strip_abs_split i (Abs (s, T, t)) = |
|
18013 | 942 |
let |
943 |
val s' = Codegen.new_name t s; |
|
944 |
val v = Free (s', T) |
|
19039 | 945 |
in apfst (cons v) (strip_abs_split (i-1) (subst_bound (v, t))) end |
946 |
| strip_abs_split i (u as Const ("split", _) $ t) = (case strip_abs_split (i+1) t of |
|
15394 | 947 |
(v :: v' :: vs, u) => (HOLogic.mk_prod (v, v') :: vs, u) |
948 |
| _ => ([], u)) |
|
19039 | 949 |
| strip_abs_split i t = ([], t); |
18013 | 950 |
|
16634 | 951 |
fun let_codegen thy defs gr dep thyname brack t = (case strip_comb t of |
952 |
(t1 as Const ("Let", _), t2 :: t3 :: ts) => |
|
15394 | 953 |
let |
954 |
fun dest_let (l as Const ("Let", _) $ t $ u) = |
|
19039 | 955 |
(case strip_abs_split 1 u of |
15394 | 956 |
([p], u') => apfst (cons (p, t)) (dest_let u') |
957 |
| _ => ([], l)) |
|
958 |
| dest_let t = ([], t); |
|
959 |
fun mk_code (gr, (l, r)) = |
|
960 |
let |
|
16634 | 961 |
val (gr1, pl) = Codegen.invoke_codegen thy defs dep thyname false (gr, l); |
962 |
val (gr2, pr) = Codegen.invoke_codegen thy defs dep thyname false (gr1, r); |
|
15394 | 963 |
in (gr2, (pl, pr)) end |
16634 | 964 |
in case dest_let (t1 $ t2 $ t3) of |
15531 | 965 |
([], _) => NONE |
15394 | 966 |
| (ps, u) => |
967 |
let |
|
968 |
val (gr1, qs) = foldl_map mk_code (gr, ps); |
|
16634 | 969 |
val (gr2, pu) = Codegen.invoke_codegen thy defs dep thyname false (gr1, u); |
970 |
val (gr3, pargs) = foldl_map |
|
17021
1c361a3de73d
Fixed bug in code generator for let and split leading to ill-formed code.
berghofe
parents:
17002
diff
changeset
|
971 |
(Codegen.invoke_codegen thy defs dep thyname true) (gr2, ts) |
15394 | 972 |
in |
16634 | 973 |
SOME (gr3, Codegen.mk_app brack |
974 |
(Pretty.blk (0, [Pretty.str "let ", Pretty.blk (0, List.concat |
|
975 |
(separate [Pretty.str ";", Pretty.brk 1] (map (fn (pl, pr) => |
|
976 |
[Pretty.block [Pretty.str "val ", pl, Pretty.str " =", |
|
977 |
Pretty.brk 1, pr]]) qs))), |
|
978 |
Pretty.brk 1, Pretty.str "in ", pu, |
|
979 |
Pretty.brk 1, Pretty.str "end"])) pargs) |
|
15394 | 980 |
end |
981 |
end |
|
16634 | 982 |
| _ => NONE); |
15394 | 983 |
|
16634 | 984 |
fun split_codegen thy defs gr dep thyname brack t = (case strip_comb t of |
985 |
(t1 as Const ("split", _), t2 :: ts) => |
|
19039 | 986 |
(case strip_abs_split 1 (t1 $ t2) of |
16634 | 987 |
([p], u) => |
988 |
let |
|
989 |
val (gr1, q) = Codegen.invoke_codegen thy defs dep thyname false (gr, p); |
|
990 |
val (gr2, pu) = Codegen.invoke_codegen thy defs dep thyname false (gr1, u); |
|
991 |
val (gr3, pargs) = foldl_map |
|
17021
1c361a3de73d
Fixed bug in code generator for let and split leading to ill-formed code.
berghofe
parents:
17002
diff
changeset
|
992 |
(Codegen.invoke_codegen thy defs dep thyname true) (gr2, ts) |
16634 | 993 |
in |
994 |
SOME (gr2, Codegen.mk_app brack |
|
995 |
(Pretty.block [Pretty.str "(fn ", q, Pretty.str " =>", |
|
996 |
Pretty.brk 1, pu, Pretty.str ")"]) pargs) |
|
997 |
end |
|
998 |
| _ => NONE) |
|
999 |
| _ => NONE); |
|
15394 | 1000 |
|
21908 | 1001 |
in |
1002 |
||
20105 | 1003 |
Codegen.add_codegen "let_codegen" let_codegen |
1004 |
#> Codegen.add_codegen "split_codegen" split_codegen |
|
15394 | 1005 |
|
21908 | 1006 |
end |
15394 | 1007 |
*} |
1008 |
||
24699
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1009 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1010 |
subsection {* Legacy bindings *} |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1011 |
|
21908 | 1012 |
ML {* |
15404 | 1013 |
val Collect_split = thm "Collect_split"; |
1014 |
val Compl_Times_UNIV1 = thm "Compl_Times_UNIV1"; |
|
1015 |
val Compl_Times_UNIV2 = thm "Compl_Times_UNIV2"; |
|
1016 |
val PairE = thm "PairE"; |
|
1017 |
val PairE_lemma = thm "PairE_lemma"; |
|
1018 |
val Pair_Rep_inject = thm "Pair_Rep_inject"; |
|
1019 |
val Pair_def = thm "Pair_def"; |
|
1020 |
val Pair_eq = thm "Pair_eq"; |
|
1021 |
val Pair_fst_snd_eq = thm "Pair_fst_snd_eq"; |
|
1022 |
val Pair_inject = thm "Pair_inject"; |
|
1023 |
val ProdI = thm "ProdI"; |
|
1024 |
val SetCompr_Sigma_eq = thm "SetCompr_Sigma_eq"; |
|
1025 |
val SigmaD1 = thm "SigmaD1"; |
|
1026 |
val SigmaD2 = thm "SigmaD2"; |
|
1027 |
val SigmaE = thm "SigmaE"; |
|
1028 |
val SigmaE2 = thm "SigmaE2"; |
|
1029 |
val SigmaI = thm "SigmaI"; |
|
1030 |
val Sigma_Diff_distrib1 = thm "Sigma_Diff_distrib1"; |
|
1031 |
val Sigma_Diff_distrib2 = thm "Sigma_Diff_distrib2"; |
|
1032 |
val Sigma_Int_distrib1 = thm "Sigma_Int_distrib1"; |
|
1033 |
val Sigma_Int_distrib2 = thm "Sigma_Int_distrib2"; |
|
1034 |
val Sigma_Un_distrib1 = thm "Sigma_Un_distrib1"; |
|
1035 |
val Sigma_Un_distrib2 = thm "Sigma_Un_distrib2"; |
|
1036 |
val Sigma_Union = thm "Sigma_Union"; |
|
1037 |
val Sigma_def = thm "Sigma_def"; |
|
1038 |
val Sigma_empty1 = thm "Sigma_empty1"; |
|
1039 |
val Sigma_empty2 = thm "Sigma_empty2"; |
|
1040 |
val Sigma_mono = thm "Sigma_mono"; |
|
1041 |
val The_split = thm "The_split"; |
|
1042 |
val The_split_eq = thm "The_split_eq"; |
|
1043 |
val The_split_eq = thm "The_split_eq"; |
|
1044 |
val Times_Diff_distrib1 = thm "Times_Diff_distrib1"; |
|
1045 |
val Times_Int_distrib1 = thm "Times_Int_distrib1"; |
|
1046 |
val Times_Un_distrib1 = thm "Times_Un_distrib1"; |
|
1047 |
val Times_eq_cancel2 = thm "Times_eq_cancel2"; |
|
1048 |
val Times_subset_cancel2 = thm "Times_subset_cancel2"; |
|
1049 |
val UNIV_Times_UNIV = thm "UNIV_Times_UNIV"; |
|
1050 |
val UN_Times_distrib = thm "UN_Times_distrib"; |
|
1051 |
val Unity_def = thm "Unity_def"; |
|
1052 |
val cond_split_eta = thm "cond_split_eta"; |
|
1053 |
val fst_conv = thm "fst_conv"; |
|
1054 |
val fst_def = thm "fst_def"; |
|
1055 |
val fst_eqD = thm "fst_eqD"; |
|
1056 |
val inj_on_Abs_Prod = thm "inj_on_Abs_Prod"; |
|
1057 |
val injective_fst_snd = thm "injective_fst_snd"; |
|
1058 |
val mem_Sigma_iff = thm "mem_Sigma_iff"; |
|
1059 |
val mem_splitE = thm "mem_splitE"; |
|
1060 |
val mem_splitI = thm "mem_splitI"; |
|
1061 |
val mem_splitI2 = thm "mem_splitI2"; |
|
1062 |
val prod_eqI = thm "prod_eqI"; |
|
1063 |
val prod_fun = thm "prod_fun"; |
|
1064 |
val prod_fun_compose = thm "prod_fun_compose"; |
|
1065 |
val prod_fun_def = thm "prod_fun_def"; |
|
1066 |
val prod_fun_ident = thm "prod_fun_ident"; |
|
1067 |
val prod_fun_imageE = thm "prod_fun_imageE"; |
|
1068 |
val prod_fun_imageI = thm "prod_fun_imageI"; |
|
1069 |
val prod_induct = thm "prod_induct"; |
|
1070 |
val snd_conv = thm "snd_conv"; |
|
1071 |
val snd_def = thm "snd_def"; |
|
1072 |
val snd_eqD = thm "snd_eqD"; |
|
1073 |
val split = thm "split"; |
|
1074 |
val splitD = thm "splitD"; |
|
1075 |
val splitD' = thm "splitD'"; |
|
1076 |
val splitE = thm "splitE"; |
|
1077 |
val splitE' = thm "splitE'"; |
|
1078 |
val splitE2 = thm "splitE2"; |
|
1079 |
val splitI = thm "splitI"; |
|
1080 |
val splitI2 = thm "splitI2"; |
|
1081 |
val splitI2' = thm "splitI2'"; |
|
1082 |
val split_Pair_apply = thm "split_Pair_apply"; |
|
1083 |
val split_beta = thm "split_beta"; |
|
1084 |
val split_conv = thm "split_conv"; |
|
1085 |
val split_def = thm "split_def"; |
|
1086 |
val split_eta = thm "split_eta"; |
|
1087 |
val split_eta_SetCompr = thm "split_eta_SetCompr"; |
|
1088 |
val split_eta_SetCompr2 = thm "split_eta_SetCompr2"; |
|
1089 |
val split_paired_All = thm "split_paired_All"; |
|
1090 |
val split_paired_Ball_Sigma = thm "split_paired_Ball_Sigma"; |
|
1091 |
val split_paired_Bex_Sigma = thm "split_paired_Bex_Sigma"; |
|
1092 |
val split_paired_Ex = thm "split_paired_Ex"; |
|
1093 |
val split_paired_The = thm "split_paired_The"; |
|
1094 |
val split_paired_all = thm "split_paired_all"; |
|
1095 |
val split_part = thm "split_part"; |
|
1096 |
val split_split = thm "split_split"; |
|
1097 |
val split_split_asm = thm "split_split_asm"; |
|
1098 |
val split_tupled_all = thms "split_tupled_all"; |
|
1099 |
val split_weak_cong = thm "split_weak_cong"; |
|
1100 |
val surj_pair = thm "surj_pair"; |
|
1101 |
val surjective_pairing = thm "surjective_pairing"; |
|
1102 |
val unit_abs_eta_conv = thm "unit_abs_eta_conv"; |
|
1103 |
val unit_all_eq1 = thm "unit_all_eq1"; |
|
1104 |
val unit_all_eq2 = thm "unit_all_eq2"; |
|
1105 |
val unit_eq = thm "unit_eq"; |
|
1106 |
val unit_induct = thm "unit_induct"; |
|
1107 |
*} |
|
1108 |
||
24699
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1109 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1110 |
subsection {* Further inductive packages *} |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1111 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1112 |
use "Tools/inductive_realizer.ML" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1113 |
setup InductiveRealizer.setup |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1114 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1115 |
use "Tools/inductive_set_package.ML" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1116 |
setup InductiveSetPackage.setup |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1117 |
|
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1118 |
use "Tools/datatype_realizer.ML" |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1119 |
setup DatatypeRealizer.setup |
c6674504103f
datatype interpretators for size and datatype_realizer
haftmann
parents:
24286
diff
changeset
|
1120 |
|
10213 | 1121 |
end |