| author | Fabian Huch <huch@in.tum.de> | 
| Tue, 20 Aug 2024 17:28:51 +0200 | |
| changeset 80730 | be4c1fbccfe8 | 
| parent 67312 | 0d25e02759b7 | 
| child 81583 | b6df83045178 | 
| permissions | -rw-r--r-- | 
| 42151 | 1  | 
(* Title: HOL/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  | 
|
| 62175 | 5  | 
section \<open>The cpo of cartesian products\<close>  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
6  | 
|
| 15577 | 7  | 
theory Cprod  | 
| 67312 | 8  | 
imports Cfun  | 
| 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  | 
|
| 67312 | 13  | 
|
| 62175 | 14  | 
subsection \<open>Continuous case function for unit type\<close>  | 
| 16008 | 15  | 
|
| 67312 | 16  | 
definition unit_when :: "'a \<rightarrow> unit \<rightarrow> 'a"  | 
17  | 
where "unit_when = (\<Lambda> a _. a)"  | 
|
| 16008 | 18  | 
|
| 18289 | 19  | 
translations  | 
| 67312 | 20  | 
"\<Lambda>(). t" \<rightleftharpoons> "CONST unit_when\<cdot>t"  | 
| 18289 | 21  | 
|
22  | 
lemma unit_when [simp]: "unit_when\<cdot>a\<cdot>u = a"  | 
|
| 67312 | 23  | 
by (simp add: unit_when_def)  | 
24  | 
||
| 18289 | 25  | 
|
| 62175 | 26  | 
subsection \<open>Continuous version of split function\<close>  | 
| 17834 | 27  | 
|
| 67312 | 28  | 
definition csplit :: "('a \<rightarrow> 'b \<rightarrow> 'c) \<rightarrow> ('a \<times> 'b) \<rightarrow> 'c"
 | 
29  | 
where "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  | 
| 67312 | 32  | 
"\<Lambda>(CONST Pair x y). t" \<rightleftharpoons> "CONST csplit\<cdot>(\<Lambda> x y. t)"  | 
| 
17816
 
9942c5ed866a
new syntax translations for continuous lambda abstraction
 
huffman 
parents: 
16916 
diff
changeset
 | 
33  | 
|
| 67312 | 34  | 
abbreviation cfst :: "'a \<times> 'b \<rightarrow> 'a"  | 
35  | 
where "cfst \<equiv> Abs_cfun fst"  | 
|
| 
45786
 
3f07a7a91180
reinstate old functions cfst and csnd as abbreviations
 
huffman 
parents: 
42151 
diff
changeset
 | 
36  | 
|
| 67312 | 37  | 
abbreviation csnd :: "'a \<times> 'b \<rightarrow> 'b"  | 
38  | 
where "csnd \<equiv> Abs_cfun snd"  | 
|
39  | 
||
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
40  | 
|
| 62175 | 41  | 
subsection \<open>Convert all lemmas to the continuous versions\<close>  | 
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
42  | 
|
| 
18077
 
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
 
huffman 
parents: 
17837 
diff
changeset
 | 
43  | 
lemma csplit1 [simp]: "csplit\<cdot>f\<cdot>\<bottom> = f\<cdot>\<bottom>\<cdot>\<bottom>"  | 
| 67312 | 44  | 
by (simp add: csplit_def)  | 
| 
18077
 
f1f4f951ec8d
make cpair_less, cpair_defined_iff into iff rules; add lemma csplit1
 
huffman 
parents: 
17837 
diff
changeset
 | 
45  | 
|
| 33399 | 46  | 
lemma csplit_Pair [simp]: "csplit\<cdot>f\<cdot>(x, y) = f\<cdot>x\<cdot>y"  | 
| 67312 | 47  | 
by (simp add: csplit_def)  | 
| 33399 | 48  | 
|
| 
15576
 
efb95d0d01f7
converted to new-style theories, and combined numbered files
 
huffman 
parents:  
diff
changeset
 | 
49  | 
end  |