src/Doc/Tutorial/Misc/prime_def.thy
author nipkow
Tue, 17 Jun 2025 14:11:40 +0200
changeset 82733 8b537e1af2ec
parent 69597 ff784d5a5bfb
permissions -rw-r--r--
reinstated intersection of lists as inter_list_set
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     1
(*<*)
58860
fee7cfa69c50 eliminated spurious semicolons;
wenzelm
parents: 48985
diff changeset
     2
theory prime_def imports Main begin
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
(*>*)
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58860
diff changeset
     5
text\<open>
9844
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
69505
cc2d676d5395 isabelle update_cartouches -t;
wenzelm
parents: 67613
diff changeset
     9
(where \<open>dvd\<close> means ``divides''):
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67406
diff changeset
    10
@{term[display,quotes]"prime(p) \<equiv> 1 < p \<and> (m dvd p \<longrightarrow> (m=1 \<or> m=p))"}
11457
279da0358aa9 additional revisions to chapters 1, 2
paulson
parents: 11456
diff changeset
    11
\par\noindent\hangindent=0pt
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 69505
diff changeset
    12
Isabelle rejects this ``definition'' because of the extra \<^term>\<open>m\<close> 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
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67406
diff changeset
    15
@{term[display,quotes]"prime(p) \<equiv> 1 < p \<and> (\<forall>m. m dvd p \<longrightarrow> (m=1 \<or> m=p))"}
9844
8016321c7de1 *** empty log message ***
nipkow
parents: 9792
diff changeset
    16
\end{warn}
67406
23307fd33906 isabelle update_cartouches -c;
wenzelm
parents: 58860
diff changeset
    17
\<close>
8745
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
(*>*)