doc-src/TutorialI/Misc/prime_def.thy
author chaieb
Wed, 19 May 2004 11:23:59 +0200
changeset 14758 af3b71a46a1c
parent 11457 279da0358aa9
child 16417 9bc16273c2d4
permissions -rw-r--r--
A new implementation for presburger arithmetic following the one suggested in technical report Chaieb Amine and Tobias Nipkow. It is generic an smaller. the tactic has also changed and allows the abstaction over fuction occurences whose type is nat or int.
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 prime_def = Main:;
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
     3
consts prime :: "nat \<Rightarrow> bool"
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     4
(*>*)
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
     5
text{*
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
     6
\begin{warn}
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
     7
A common mistake when writing definitions is to introduce extra free
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 9844
diff changeset
     8
variables on the right-hand side.  Consider the following, flawed definition
7eb63f63e6c6 revisions and indexing
paulson
parents: 9844
diff changeset
     9
(where @{text"dvd"} means ``divides''):
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    10
@{term[display,quotes]"prime(p) == 1 < p & (m dvd p --> (m=1 | m=p))"}
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
    11
\par\noindent\hangindent=0pt
9792
bbefb6ce5cb2 *** empty log message ***
nipkow
parents: 9458
diff changeset
    12
Isabelle rejects this ``definition'' because of the extra @{term"m"} on the
11456
7eb63f63e6c6 revisions and indexing
paulson
parents: 9844
diff changeset
    13
right-hand side, which would introduce an inconsistency (why?). 
7eb63f63e6c6 revisions and indexing
paulson
parents: 9844
diff changeset
    14
The correct version is
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    15
@{term[display,quotes]"prime(p) == 1 < p & (!m. m dvd p --> (m=1 | m=p))"}
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    16
\end{warn}
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    17
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
end
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
(*>*)