author | huffman |
Tue, 26 Oct 2010 14:19:59 -0700 | |
changeset 40216 | 366309dfaf60 |
parent 40002 | c5b5f7a3a3b1 |
child 40502 | 8e92772bc0e8 |
permissions | -rw-r--r-- |
15600 | 1 |
(* Title: HOLCF/Cprod.thy |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
2 |
Author: Franz Regensburger |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
3 |
*) |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
4 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
5 |
header {* The cpo of cartesian products *} |
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
6 |
|
15577 | 7 |
theory Cprod |
39987
8c2f449af35a
move all bifinite class instances to Bifinite.thy
huffman
parents:
39986
diff
changeset
|
8 |
imports Deflation |
15577 | 9 |
begin |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
10 |
|
36452 | 11 |
default_sort cpo |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
12 |
|
35900
aa5dfb03eb1e
remove LaTeX hyperref warnings by avoiding antiquotations within section headings
huffman
parents:
33399
diff
changeset
|
13 |
subsection {* Continuous case function for unit type *} |
16008 | 14 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
15 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
16 |
unit_when :: "'a \<rightarrow> unit \<rightarrow> 'a" where |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
17 |
"unit_when = (\<Lambda> a _. a)" |
16008 | 18 |
|
18289 | 19 |
translations |
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
20 |
"\<Lambda>(). t" == "CONST unit_when\<cdot>t" |
18289 | 21 |
|
22 |
lemma unit_when [simp]: "unit_when\<cdot>a\<cdot>u = a" |
|
23 |
by (simp add: unit_when_def) |
|
24 |
||
35926 | 25 |
subsection {* Continuous version of split function *} |
17834 | 26 |
|
25131
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
27 |
definition |
2c8caac48ade
modernized specifications ('definition', 'abbreviation', 'notation');
wenzelm
parents:
18289
diff
changeset
|
28 |
csplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a * 'b) \<rightarrow> 'c" where |
35922 | 29 |
"csplit = (\<Lambda> f p. f\<cdot>(fst p)\<cdot>(snd p))" |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
30 |
|
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
31 |
translations |
35926 | 32 |
"\<Lambda>(CONST Pair x y). t" == "CONST csplit\<cdot>(\<Lambda> x y. t)" |
17816
9942c5ed866a
new syntax translations for continuous lambda abstraction
huffman
parents:
16916
diff
changeset
|
33 |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
34 |
|
15593
24d770bbc44a
reordered and arranged for document generation, cleaned up some proofs
huffman
parents:
15577
diff
changeset
|
35 |
subsection {* Convert all lemmas to the continuous versions *} |
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
36 |
|
18077
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
37 |
lemma csplit1 [simp]: "csplit\<cdot>f\<cdot>\<bottom> = f\<cdot>\<bottom>\<cdot>\<bottom>" |
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
38 |
by (simp add: csplit_def) |
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
huffman
parents:
17837
diff
changeset
|
39 |
|
33399 | 40 |
lemma csplit_Pair [simp]: "csplit\<cdot>f\<cdot>(x, y) = f\<cdot>x\<cdot>y" |
35922 | 41 |
by (simp add: csplit_def) |
33399 | 42 |
|
39985
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
43 |
subsection {* Map operator for product type *} |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
44 |
|
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
45 |
definition |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
46 |
cprod_map :: "('a \<rightarrow> 'b) \<rightarrow> ('c \<rightarrow> 'd) \<rightarrow> 'a \<times> 'c \<rightarrow> 'b \<times> 'd" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
47 |
where |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
48 |
"cprod_map = (\<Lambda> f g p. (f\<cdot>(fst p), g\<cdot>(snd p)))" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
49 |
|
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
50 |
lemma cprod_map_Pair [simp]: "cprod_map\<cdot>f\<cdot>g\<cdot>(x, y) = (f\<cdot>x, g\<cdot>y)" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
51 |
unfolding cprod_map_def by simp |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
52 |
|
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
53 |
lemma cprod_map_ID: "cprod_map\<cdot>ID\<cdot>ID = ID" |
40002
c5b5f7a3a3b1
new theorem names: fun_below_iff, fun_belowI, cfun_eq_iff, cfun_eqI, cfun_below_iff, cfun_belowI
huffman
parents:
39987
diff
changeset
|
54 |
unfolding cfun_eq_iff by auto |
39985
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
55 |
|
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
56 |
lemma cprod_map_map: |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
57 |
"cprod_map\<cdot>f1\<cdot>g1\<cdot>(cprod_map\<cdot>f2\<cdot>g2\<cdot>p) = |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
58 |
cprod_map\<cdot>(\<Lambda> x. f1\<cdot>(f2\<cdot>x))\<cdot>(\<Lambda> x. g1\<cdot>(g2\<cdot>x))\<cdot>p" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
59 |
by (induct p) simp |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
60 |
|
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
61 |
lemma ep_pair_cprod_map: |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
62 |
assumes "ep_pair e1 p1" and "ep_pair e2 p2" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
63 |
shows "ep_pair (cprod_map\<cdot>e1\<cdot>e2) (cprod_map\<cdot>p1\<cdot>p2)" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
64 |
proof |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
65 |
interpret e1p1: ep_pair e1 p1 by fact |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
66 |
interpret e2p2: ep_pair e2 p2 by fact |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
67 |
fix x show "cprod_map\<cdot>p1\<cdot>p2\<cdot>(cprod_map\<cdot>e1\<cdot>e2\<cdot>x) = x" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
68 |
by (induct x) simp |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
69 |
fix y show "cprod_map\<cdot>e1\<cdot>e2\<cdot>(cprod_map\<cdot>p1\<cdot>p2\<cdot>y) \<sqsubseteq> y" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
70 |
by (induct y) (simp add: e1p1.e_p_below e2p2.e_p_below) |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
71 |
qed |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
72 |
|
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
73 |
lemma deflation_cprod_map: |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
74 |
assumes "deflation d1" and "deflation d2" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
75 |
shows "deflation (cprod_map\<cdot>d1\<cdot>d2)" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
76 |
proof |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
77 |
interpret d1: deflation d1 by fact |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
78 |
interpret d2: deflation d2 by fact |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
79 |
fix x |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
80 |
show "cprod_map\<cdot>d1\<cdot>d2\<cdot>(cprod_map\<cdot>d1\<cdot>d2\<cdot>x) = cprod_map\<cdot>d1\<cdot>d2\<cdot>x" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
81 |
by (induct x) (simp add: d1.idem d2.idem) |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
82 |
show "cprod_map\<cdot>d1\<cdot>d2\<cdot>x \<sqsubseteq> x" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
83 |
by (induct x) (simp add: d1.below d2.below) |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
84 |
qed |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
85 |
|
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
86 |
lemma finite_deflation_cprod_map: |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
87 |
assumes "finite_deflation d1" and "finite_deflation d2" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
88 |
shows "finite_deflation (cprod_map\<cdot>d1\<cdot>d2)" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
89 |
proof (rule finite_deflation_intro) |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
90 |
interpret d1: finite_deflation d1 by fact |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
91 |
interpret d2: finite_deflation d2 by fact |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
92 |
have "deflation d1" and "deflation d2" by fact+ |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
93 |
thus "deflation (cprod_map\<cdot>d1\<cdot>d2)" by (rule deflation_cprod_map) |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
94 |
have "{p. cprod_map\<cdot>d1\<cdot>d2\<cdot>p = p} \<subseteq> {x. d1\<cdot>x = x} \<times> {y. d2\<cdot>y = y}" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
95 |
by clarsimp |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
96 |
thus "finite {p. cprod_map\<cdot>d1\<cdot>d2\<cdot>p = p}" |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
97 |
by (rule finite_subset, simp add: d1.finite_fixes d2.finite_fixes) |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
98 |
qed |
310f98585107
move stuff from Algebraic.thy to Bifinite.thy and elsewhere
huffman
parents:
39974
diff
changeset
|
99 |
|
15576
efb95d0d01f7
converted to new-style theories, and combined numbered files
huffman
parents:
diff
changeset
|
100 |
end |