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