doc-src/TutorialI/Misc/cases.thy
author kleing
Wed, 12 Jul 2000 14:47:55 +0200
changeset 9287 c406d0af9368
parent 8771 026f37a86ea7
child 9458 c613cd06d5cf
permissions -rw-r--r--
about.html -> logics.html
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     1
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     2
theory "cases" = Main:;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     3
(*>*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     4
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     5
lemma "(case xs of [] \\<Rightarrow> [] | y#ys \\<Rightarrow> xs) = xs";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     6
apply(case_tac xs);
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     7
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
     8
txt{*\noindent
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
     9
results in the proof state
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    10
\begin{isabellepar}%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    11
~1.~xs~=~[]~{\isasymLongrightarrow}~(case~xs~of~[]~{\isasymRightarrow}~[]~|~y~\#~ys~{\isasymRightarrow}~xs)~=~xs\isanewline
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    12
~2.~{\isasymAnd}a~list.~xs=a\#list~{\isasymLongrightarrow}~(case~xs~of~[]~{\isasymRightarrow}~[]~|~y\#ys~{\isasymRightarrow}~xs)~=~xs%
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    13
\end{isabellepar}%
8771
026f37a86ea7 *** empty log message ***
nipkow
parents: 8745
diff changeset
    14
which is solved automatically:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    16
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    17
apply(auto).;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
(**)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    21
(*>*)