| author | wenzelm | 
| Tue, 17 Sep 2024 18:49:46 +0200 | |
| changeset 80898 | 71756d95b7df | 
| parent 69597 | ff784d5a5bfb | 
| permissions | -rw-r--r-- | 
| 8745 | 1 | (*<*) | 
| 58860 | 2 | theory pairs2 imports Main begin | 
| 9541 | 3 | (*>*) | 
| 67406 | 4 | text\<open>\label{sec:pairs}\index{pairs and tuples}
 | 
| 11428 | 5 | HOL also has ordered pairs: \isa{($a@1$,$a@2$)} is of type $\tau@1$
 | 
| 10538 | 6 | \indexboldpos{\isasymtimes}{$Isatype} $\tau@2$ provided each $a@i$ is of type
 | 
| 11428 | 7 | $\tau@i$. The functions \cdx{fst} and
 | 
| 8 | \cdx{snd} extract the components of a pair:
 | |
| 10538 | 9 |  \isa{fst($x$,$y$) = $x$} and \isa{snd($x$,$y$) = $y$}. Tuples
 | 
| 9933 | 10 | are simulated by pairs nested to the right: \isa{($a@1$,$a@2$,$a@3$)} stands
 | 
| 11 | for \isa{($a@1$,($a@2$,$a@3$))} and $\tau@1 \times \tau@2 \times \tau@3$ for
 | |
| 12 | $\tau@1 \times (\tau@2 \times \tau@3)$. Therefore we have | |
| 13 | \isa{fst(snd($a@1$,$a@2$,$a@3$)) = $a@2$}.
 | |
| 9541 | 14 | |
| 10539 | 15 | Remarks: | 
| 16 | \begin{itemize}
 | |
| 17 | \item | |
| 11428 | 18 | There is also the type \tydx{unit}, which contains exactly one
 | 
| 11456 | 19 | element denoted by~\cdx{()}.  This type can be viewed
 | 
| 10539 | 20 | as a degenerate product with 0 components. | 
| 21 | \item | |
| 69597 | 22 | Products, like type \<^typ>\<open>nat\<close>, are datatypes, which means | 
| 69505 | 23 | in particular that \<open>induct_tac\<close> and \<open>case_tac\<close> are applicable to | 
| 10539 | 24 | terms of product type. | 
| 27321 
464ac1c815ec
induct_tac/case_tac: nested tuples are split as expected;
 wenzelm parents: 
16417diff
changeset | 25 | Both split the term into a number of variables corresponding to the tuple structure | 
| 
464ac1c815ec
induct_tac/case_tac: nested tuples are split as expected;
 wenzelm parents: 
16417diff
changeset | 26 | (up to 7 components). | 
| 10539 | 27 | \item | 
| 10795 | 28 | Tuples with more than two or three components become unwieldy; | 
| 29 | records are preferable. | |
| 10539 | 30 | \end{itemize}
 | 
| 31 | For more information on pairs and records see Chapter~\ref{ch:more-types}.
 | |
| 67406 | 32 | \<close> | 
| 8745 | 33 | (*<*) | 
| 34 | end | |
| 35 | (*>*) |