doc-src/TutorialI/Types/Pairs.thy
author nipkow
Tue, 01 May 2001 22:26:55 +0200
changeset 11277 a2bff98d6e5d
parent 11161 166f7d87b37f
child 11428 332347b9b942
permissions -rw-r--r--
*** empty log message ***
nipkow@10560
     1
(*<*)theory Pairs = Main:(*>*)
nipkow@10560
     2
nipkow@10560
     3
section{*Pairs*}
nipkow@10560
     4
nipkow@10560
     5
text{*\label{sec:products}
nipkow@10560
     6
Pairs were already introduced in \S\ref{sec:pairs}, but only with a minimal
nipkow@10560
     7
repertoire of operations: pairing and the two projections @{term fst} and
nipkow@11149
     8
@{term snd}. In any non-trivial application of pairs you will find that this
nipkow@11161
     9
quickly leads to unreadable formulae involving nests of projections. This
nipkow@10560
    10
section is concerned with introducing some syntactic sugar to overcome this
nipkow@10560
    11
problem: pattern matching with tuples.
nipkow@10560
    12
*}
nipkow@10560
    13
paulson@10885
    14
subsection{*Pattern Matching with Tuples*}
nipkow@10560
    15
nipkow@10560
    16
text{*
paulson@10885
    17
Tuples may be used as patterns in $\lambda$-abstractions,
nipkow@10560
    18
for example @{text"\<lambda>(x,y,z).x+y+z"} and @{text"\<lambda>((x,y),z).x+y+z"}. In fact,
paulson@10885
    19
tuple patterns can be used in most variable binding constructs,
paulson@10885
    20
and they can be nested. Here are
nipkow@10560
    21
some typical examples:
nipkow@10560
    22
\begin{quote}
nipkow@10560
    23
@{term"let (x,y) = f z in (y,x)"}\\
nipkow@10560
    24
@{term"case xs of [] => 0 | (x,y)#zs => x+y"}\\
nipkow@10560
    25
@{text"\<forall>(x,y)\<in>A. x=y"}\\
paulson@10885
    26
@{text"{(x,y,z). x=z}"}\\
nipkow@10560
    27
@{term"\<Union>(x,y)\<in>A. {x+y}"}
nipkow@10560
    28
\end{quote}
paulson@10885
    29
The intuitive meanings of these expressions should be obvious.
nipkow@10560
    30
Unfortunately, we need to know in more detail what the notation really stands
paulson@10885
    31
for once we have to reason about it.  Abstraction
nipkow@10560
    32
over pairs and tuples is merely a convenient shorthand for a more complex
nipkow@10560
    33
internal representation.  Thus the internal and external form of a term may
nipkow@10560
    34
differ, which can affect proofs. If you want to avoid this complication,
nipkow@10560
    35
stick to @{term fst} and @{term snd} and write @{term"%p. fst p + snd p"}
nipkow@10560
    36
instead of @{text"\<lambda>(x,y). x+y"} (which denote the same function but are quite
nipkow@10560
    37
different terms).
nipkow@10560
    38
nipkow@10560
    39
Internally, @{term"%(x,y). t"} becomes @{text"split (\<lambda>x y. t)"}, where
nipkow@10654
    40
@{term split}\indexbold{*split (constant)}
nipkow@10654
    41
is the uncurrying function of type @{text"('a \<Rightarrow> 'b
nipkow@10560
    42
\<Rightarrow> 'c) \<Rightarrow> 'a \<times> 'b \<Rightarrow> 'c"} defined as
nipkow@10560
    43
\begin{center}
nipkow@10560
    44
@{thm split_def}
nipkow@10560
    45
\hfill(@{thm[source]split_def})
nipkow@10560
    46
\end{center}
nipkow@10560
    47
Pattern matching in
nipkow@10560
    48
other variable binding constructs is translated similarly. Thus we need to
nipkow@10560
    49
understand how to reason about such constructs.
nipkow@10560
    50
*}
nipkow@10560
    51
paulson@10885
    52
subsection{*Theorem Proving*}
nipkow@10560
    53
nipkow@10560
    54
text{*
nipkow@10560
    55
The most obvious approach is the brute force expansion of @{term split}:
nipkow@10560
    56
*}
nipkow@10560
    57
nipkow@10560
    58
lemma "(\<lambda>(x,y).x) p = fst p"
nipkow@10560
    59
by(simp add:split_def)
nipkow@10560
    60
nipkow@10560
    61
text{* This works well if rewriting with @{thm[source]split_def} finishes the
nipkow@11149
    62
proof, as it does above.  But if it does not, you end up with exactly what
nipkow@10560
    63
we are trying to avoid: nests of @{term fst} and @{term snd}. Thus this
nipkow@10560
    64
approach is neither elegant nor very practical in large examples, although it
nipkow@10560
    65
can be effective in small ones.
nipkow@10560
    66
nipkow@10560
    67
If we step back and ponder why the above lemma presented a problem in the
nipkow@10560
    68
first place, we quickly realize that what we would like is to replace @{term
nipkow@10560
    69
p} with some concrete pair @{term"(a,b)"}, in which case both sides of the
nipkow@10560
    70
equation would simplify to @{term a} because of the simplification rules
wenzelm@10902
    71
@{thm split_conv[no_vars]} and @{thm fst_conv[no_vars]}.  This is the
nipkow@10560
    72
key problem one faces when reasoning about pattern matching with pairs: how to
nipkow@10560
    73
convert some atomic term into a pair.
nipkow@10560
    74
nipkow@10560
    75
In case of a subterm of the form @{term"split f p"} this is easy: the split
nipkow@10560
    76
rule @{thm[source]split_split} replaces @{term p} by a pair:
nipkow@10560
    77
*}
nipkow@10560
    78
nipkow@10560
    79
lemma "(\<lambda>(x,y).y) p = snd p"
nipkow@10654
    80
apply(split split_split);
nipkow@10560
    81
nipkow@10560
    82
txt{*
nipkow@10560
    83
@{subgoals[display,indent=0]}
nipkow@10824
    84
This subgoal is easily proved by simplification. Thus we could have combined
nipkow@10824
    85
simplification and splitting in one command that proves the goal outright:
nipkow@10824
    86
*}
nipkow@10824
    87
(*<*)
nipkow@10824
    88
by simp
nipkow@10824
    89
lemma "(\<lambda>(x,y).y) p = snd p"(*>*)
nipkow@10824
    90
by(simp split: split_split)
nipkow@10560
    91
nipkow@10824
    92
text{*
nipkow@10560
    93
Let us look at a second example:
nipkow@10560
    94
*}
nipkow@10560
    95
nipkow@10560
    96
lemma "let (x,y) = p in fst p = x";
nipkow@10824
    97
apply(simp only: Let_def)
nipkow@10560
    98
nipkow@10560
    99
txt{*
nipkow@10560
   100
@{subgoals[display,indent=0]}
nipkow@10560
   101
A paired @{text let} reduces to a paired $\lambda$-abstraction, which
nipkow@10560
   102
can be split as above. The same is true for paired set comprehension:
nipkow@10560
   103
*}
nipkow@10560
   104
nipkow@10560
   105
(*<*)by(simp split:split_split)(*>*)
nipkow@10560
   106
lemma "p \<in> {(x,y). x=y} \<longrightarrow> fst p = snd p"
nipkow@10560
   107
apply simp
nipkow@10560
   108
nipkow@10560
   109
txt{*
nipkow@10560
   110
@{subgoals[display,indent=0]}
nipkow@10560
   111
Again, simplification produces a term suitable for @{thm[source]split_split}
nipkow@11277
   112
as above. If you are worried about the strange form of the premise:
nipkow@11277
   113
@{term"split (op =)"} is short for @{text"\<lambda>(x,y). x=y"}.
nipkow@11277
   114
The same proof procedure works for
nipkow@10560
   115
*}
nipkow@10560
   116
nipkow@10560
   117
(*<*)by(simp split:split_split)(*>*)
nipkow@10560
   118
lemma "p \<in> {(x,y). x=y} \<Longrightarrow> fst p = snd p"
nipkow@10560
   119
nipkow@10560
   120
txt{*\noindent
nipkow@10560
   121
except that we now have to use @{thm[source]split_split_asm}, because
nipkow@10560
   122
@{term split} occurs in the assumptions.
nipkow@10560
   123
nipkow@10560
   124
However, splitting @{term split} is not always a solution, as no @{term split}
nipkow@10560
   125
may be present in the goal. Consider the following function:
nipkow@10560
   126
*}
nipkow@10560
   127
nipkow@10560
   128
(*<*)by(simp split:split_split_asm)(*>*)
nipkow@10560
   129
consts swap :: "'a \<times> 'b \<Rightarrow> 'b \<times> 'a"
nipkow@10560
   130
primrec
nipkow@10560
   131
  "swap (x,y) = (y,x)"
nipkow@10560
   132
nipkow@10560
   133
text{*\noindent
nipkow@10560
   134
Note that the above \isacommand{primrec} definition is admissible
nipkow@10560
   135
because @{text"\<times>"} is a datatype. When we now try to prove
nipkow@10560
   136
*}
nipkow@10560
   137
nipkow@10560
   138
lemma "swap(swap p) = p"
nipkow@10560
   139
nipkow@10560
   140
txt{*\noindent
nipkow@10560
   141
simplification will do nothing, because the defining equation for @{term swap}
nipkow@10560
   142
expects a pair. Again, we need to turn @{term p} into a pair first, but this
nipkow@10560
   143
time there is no @{term split} in sight. In this case the only thing we can do
nipkow@10560
   144
is to split the term by hand:
nipkow@10560
   145
*}
nipkow@10560
   146
apply(case_tac p)
nipkow@10560
   147
nipkow@10560
   148
txt{*\noindent
nipkow@10560
   149
@{subgoals[display,indent=0]}
nipkow@10560
   150
Again, @{text case_tac} is applicable because @{text"\<times>"} is a datatype.
nipkow@10560
   151
The subgoal is easily proved by @{text simp}.
nipkow@10560
   152
nipkow@10824
   153
Splitting by @{text case_tac} also solves the previous examples and may thus
nipkow@10824
   154
appear preferable to the more arcane methods introduced first. However, see
nipkow@10824
   155
the warning about @{text case_tac} in \S\ref{sec:struct-ind-case}.
nipkow@10824
   156
nipkow@10560
   157
In case the term to be split is a quantified variable, there are more options.
nipkow@10560
   158
You can split \emph{all} @{text"\<And>"}-quantified variables in a goal
nipkow@10560
   159
with the rewrite rule @{thm[source]split_paired_all}:
nipkow@10560
   160
*}
nipkow@10560
   161
nipkow@10560
   162
(*<*)by simp(*>*)
nipkow@10560
   163
lemma "\<And>p q. swap(swap p) = q \<longrightarrow> p = q"
nipkow@10560
   164
apply(simp only:split_paired_all)
nipkow@10560
   165
nipkow@10560
   166
txt{*\noindent
nipkow@10560
   167
@{subgoals[display,indent=0]}
nipkow@10560
   168
*}
nipkow@10560
   169
nipkow@10560
   170
apply simp
nipkow@10560
   171
done
nipkow@10560
   172
nipkow@10560
   173
text{*\noindent
nipkow@10560
   174
Note that we have intentionally included only @{thm[source]split_paired_all}
nipkow@10560
   175
in the first simplification step. This time the reason was not merely
nipkow@10560
   176
pedagogical:
nipkow@10560
   177
@{thm[source]split_paired_all} may interfere with certain congruence
nipkow@10560
   178
rules of the simplifier, i.e.
nipkow@10560
   179
*}
nipkow@10560
   180
nipkow@10560
   181
(*<*)
nipkow@10560
   182
lemma "\<And>p q. swap(swap p) = q \<longrightarrow> p = q"
nipkow@10560
   183
(*>*)
nipkow@10560
   184
apply(simp add:split_paired_all)
nipkow@10560
   185
(*<*)done(*>*)
nipkow@10560
   186
text{*\noindent
nipkow@10560
   187
may fail (here it does not) where the above two stages succeed.
nipkow@10560
   188
nipkow@10560
   189
Finally, all @{text"\<forall>"} and @{text"\<exists>"}-quantified variables are split
nipkow@10560
   190
automatically by the simplifier:
nipkow@10560
   191
*}
nipkow@10560
   192
nipkow@10560
   193
lemma "\<forall>p. \<exists>q. swap p = swap q"
nipkow@10560
   194
apply simp;
nipkow@10560
   195
done
nipkow@10560
   196
nipkow@10560
   197
text{*\noindent
nipkow@10560
   198
In case you would like to turn off this automatic splitting, just disable the
nipkow@10560
   199
responsible simplification rules:
nipkow@10560
   200
\begin{center}
nipkow@10654
   201
@{thm split_paired_All[no_vars]}
nipkow@10560
   202
\hfill
nipkow@10560
   203
(@{thm[source]split_paired_All})\\
nipkow@10654
   204
@{thm split_paired_Ex[no_vars]}
nipkow@10560
   205
\hfill
nipkow@10560
   206
(@{thm[source]split_paired_Ex})
nipkow@10560
   207
\end{center}
nipkow@10560
   208
*}
nipkow@10560
   209
(*<*)
nipkow@10560
   210
end
nipkow@10560
   211
(*>*)