8745
|
1 |
(*<*)
|
|
2 |
theory "cases" = Main:;
|
|
3 |
(*>*)
|
|
4 |
|
|
5 |
lemma "(case xs of [] \\<Rightarrow> [] | y#ys \\<Rightarrow> xs) = xs";
|
|
6 |
apply(case_tac xs);
|
|
7 |
|
8771
|
8 |
txt{*\noindent
|
|
9 |
results in the proof state
|
8745
|
10 |
\begin{isabellepar}%
|
|
11 |
~1.~xs~=~[]~{\isasymLongrightarrow}~(case~xs~of~[]~{\isasymRightarrow}~[]~|~y~\#~ys~{\isasymRightarrow}~xs)~=~xs\isanewline
|
|
12 |
~2.~{\isasymAnd}a~list.~xs=a\#list~{\isasymLongrightarrow}~(case~xs~of~[]~{\isasymRightarrow}~[]~|~y\#ys~{\isasymRightarrow}~xs)~=~xs%
|
|
13 |
\end{isabellepar}%
|
8771
|
14 |
which is solved automatically:
|
8745
|
15 |
*}
|
|
16 |
|
9458
|
17 |
by(auto);
|
8745
|
18 |
(**)
|
|
19 |
(*<*)
|
|
20 |
end
|
|
21 |
(*>*)
|