| 8745 |      1 | (*<*)
 | 
| 16523 |      2 | theory natsum imports Main begin
 | 
| 8745 |      3 | (*>*)
 | 
| 67406 |      4 | text\<open>\noindent
 | 
| 69505 |      5 | In particular, there are \<open>case\<close>-expressions, for example
 | 
| 9541 |      6 | @{term[display]"case n of 0 => 0 | Suc m => m"}
 | 
| 8745 |      7 | primitive recursion, for example
 | 
| 67406 |      8 | \<close>
 | 
| 8745 |      9 | 
 | 
| 27015 |     10 | primrec sum :: "nat \<Rightarrow> nat" where
 | 
|  |     11 | "sum 0 = 0" |
 | 
|  |     12 | "sum (Suc n) = Suc n + sum n"
 | 
| 8745 |     13 | 
 | 
| 67406 |     14 | text\<open>\noindent
 | 
| 8745 |     15 | and induction, for example
 | 
| 67406 |     16 | \<close>
 | 
| 8745 |     17 | 
 | 
| 16523 |     18 | lemma "sum n + sum n = n*(Suc n)"
 | 
|  |     19 | apply(induct_tac n)
 | 
|  |     20 | apply(auto)
 | 
| 10171 |     21 | done
 | 
| 8745 |     22 | 
 | 
| 67406 |     23 | text\<open>\newcommand{\mystar}{*%
 | 
| 10538 |     24 | }
 | 
| 11456 |     25 | \index{arithmetic operations!for \protect\isa{nat}}%
 | 
| 15364 |     26 | The arithmetic operations \isadxboldpos{+}{$HOL2arithfun},
 | 
|  |     27 | \isadxboldpos{-}{$HOL2arithfun}, \isadxboldpos{\mystar}{$HOL2arithfun},
 | 
| 11428 |     28 | \sdx{div}, \sdx{mod}, \cdx{min} and
 | 
|  |     29 | \cdx{max} are predefined, as are the relations
 | 
| 15364 |     30 | \isadxboldpos{\isasymle}{$HOL2arithrel} and
 | 
| 69597 |     31 | \isadxboldpos{<}{$HOL2arithrel}. As usual, \<^prop>\<open>m-n = (0::nat)\<close> if
 | 
|  |     32 | \<^prop>\<open>m<n\<close>. There is even a least number operation
 | 
|  |     33 | \sdx{LEAST}\@.  For example, \<^prop>\<open>(LEAST n. 0 < n) = Suc 0\<close>.
 | 
| 11456 |     34 | \begin{warn}\index{overloading}
 | 
| 12327 |     35 |   The constants \cdx{0} and \cdx{1} and the operations
 | 
| 15364 |     36 |   \isadxboldpos{+}{$HOL2arithfun}, \isadxboldpos{-}{$HOL2arithfun},
 | 
|  |     37 |   \isadxboldpos{\mystar}{$HOL2arithfun}, \cdx{min},
 | 
|  |     38 |   \cdx{max}, \isadxboldpos{\isasymle}{$HOL2arithrel} and
 | 
|  |     39 |   \isadxboldpos{<}{$HOL2arithrel} are overloaded: they are available
 | 
| 12329 |     40 |   not just for natural numbers but for other types as well.
 | 
| 69505 |     41 |   For example, given the goal \<open>x + 0 = x\<close>, there is nothing to indicate
 | 
| 12327 |     42 |   that you are talking about natural numbers. Hence Isabelle can only infer
 | 
| 69597 |     43 |   that \<^term>\<open>x\<close> is of some arbitrary type where \<open>0\<close> and \<open>+\<close> are
 | 
| 12327 |     44 |   declared. As a consequence, you will be unable to prove the
 | 
|  |     45 |   goal. To alert you to such pitfalls, Isabelle flags numerals without a
 | 
| 69597 |     46 |   fixed type in its output: \<^prop>\<open>x+0 = x\<close>. (In the absence of a numeral,
 | 
| 16523 |     47 |   it may take you some time to realize what has happened if \pgmenu{Show
 | 
|  |     48 |   Types} is not set).  In this particular example, you need to include
 | 
| 69505 |     49 |   an explicit type constraint, for example \<open>x+0 = (x::nat)\<close>. If there
 | 
| 69597 |     50 |   is enough contextual information this may not be necessary: \<^prop>\<open>Suc x =
 | 
|  |     51 |   x\<close> automatically implies \<open>x::nat\<close> because \<^term>\<open>Suc\<close> is not
 | 
| 12327 |     52 |   overloaded.
 | 
| 10978 |     53 | 
 | 
| 12327 |     54 |   For details on overloading see \S\ref{sec:overloading}.
 | 
|  |     55 |   Table~\ref{tab:overloading} in the appendix shows the most important
 | 
|  |     56 |   overloaded operations.
 | 
|  |     57 | \end{warn}
 | 
|  |     58 | \begin{warn}
 | 
| 15364 |     59 |   The symbols \isadxboldpos{>}{$HOL2arithrel} and
 | 
| 69505 |     60 |   \isadxboldpos{\isasymge}{$HOL2arithrel} are merely syntax: \<open>x > y\<close>
 | 
| 69597 |     61 |   stands for \<^prop>\<open>y < x\<close> and similary for \<open>\<ge>\<close> and
 | 
| 69505 |     62 |   \<open>\<le>\<close>.
 | 
| 15364 |     63 | \end{warn}
 | 
|  |     64 | \begin{warn}
 | 
| 69597 |     65 |   Constant \<open>1::nat\<close> is defined to equal \<^term>\<open>Suc 0\<close>. This definition
 | 
| 12327 |     66 |   (see \S\ref{sec:ConstDefinitions}) is unfolded automatically by some
 | 
| 69505 |     67 |   tactics (like \<open>auto\<close>, \<open>simp\<close> and \<open>arith\<close>) but not by
 | 
| 12327 |     68 |   others (especially the single step tactics in Chapter~\ref{chap:rules}).
 | 
|  |     69 |   If you need the full set of numerals, see~\S\ref{sec:numerals}.
 | 
| 69597 |     70 |   \emph{Novices are advised to stick to \<^term>\<open>0::nat\<close> and \<^term>\<open>Suc\<close>.}
 | 
| 10538 |     71 | \end{warn}
 | 
|  |     72 | 
 | 
| 69505 |     73 | Both \<open>auto\<close> and \<open>simp\<close>
 | 
| 11456 |     74 | (a method introduced below, \S\ref{sec:Simplification}) prove 
 | 
|  |     75 | simple arithmetic goals automatically:
 | 
| 67406 |     76 | \<close>
 | 
| 10538 |     77 | 
 | 
| 11711 |     78 | lemma "\<lbrakk> \<not> m < n; m < n + (1::nat) \<rbrakk> \<Longrightarrow> m = n"
 | 
| 10538 |     79 | (*<*)by(auto)(*>*)
 | 
|  |     80 | 
 | 
| 67406 |     81 | text\<open>\noindent
 | 
| 11458 |     82 | For efficiency's sake, this built-in prover ignores quantified formulae,
 | 
| 16768 |     83 | many logical connectives, and all arithmetic operations apart from addition.
 | 
| 69505 |     84 | In consequence, \<open>auto\<close> and \<open>simp\<close> cannot prove this slightly more complex goal:
 | 
| 67406 |     85 | \<close>
 | 
| 10538 |     86 | 
 | 
| 13181 |     87 | lemma "m \<noteq> (n::nat) \<Longrightarrow> m < n \<or> n < m"
 | 
| 11458 |     88 | (*<*)by(arith)(*>*)
 | 
|  |     89 | 
 | 
| 67406 |     90 | text\<open>\noindent The method \methdx{arith} is more general.  It attempts to
 | 
| 13996 |     91 | prove the first subgoal provided it is a \textbf{linear arithmetic} formula.
 | 
| 69505 |     92 | Such formulas may involve the usual logical connectives (\<open>\<not>\<close>,
 | 
|  |     93 | \<open>\<and>\<close>, \<open>\<or>\<close>, \<open>\<longrightarrow>\<close>, \<open>=\<close>,
 | 
|  |     94 | \<open>\<forall>\<close>, \<open>\<exists>\<close>), the relations \<open>=\<close>,
 | 
|  |     95 | \<open>\<le>\<close> and \<open><\<close>, and the operations \<open>+\<close>, \<open>-\<close>,
 | 
| 69597 |     96 | \<^term>\<open>min\<close> and \<^term>\<open>max\<close>.  For example,\<close>
 | 
| 10538 |     97 | 
 | 
| 16523 |     98 | lemma "min i (max j (k*k)) = max (min (k*k) i) (min i (j::nat))"
 | 
| 10538 |     99 | apply(arith)
 | 
|  |    100 | (*<*)done(*>*)
 | 
|  |    101 | 
 | 
| 67406 |    102 | text\<open>\noindent
 | 
| 69597 |    103 | succeeds because \<^term>\<open>k*k\<close> can be treated as atomic. In contrast,
 | 
| 67406 |    104 | \<close>
 | 
| 10538 |    105 | 
 | 
| 27168 |    106 | lemma "n*n = n+1 \<Longrightarrow> n=0"
 | 
| 10538 |    107 | (*<*)oops(*>*)
 | 
|  |    108 | 
 | 
| 67406 |    109 | text\<open>\noindent
 | 
| 69505 |    110 | is not proved by \<open>arith\<close> because the proof relies 
 | 
| 13996 |    111 | on properties of multiplication. Only multiplication by numerals (which is
 | 
| 27168 |    112 | the same as iterated addition) is taken into account.
 | 
| 10538 |    113 | 
 | 
| 69505 |    114 | \begin{warn} The running time of \<open>arith\<close> is exponential in the number
 | 
| 13996 |    115 |   of occurrences of \ttindexboldpos{-}{$HOL2arithfun}, \cdx{min} and
 | 
| 11428 |    116 |   \cdx{max} because they are first eliminated by case distinctions.
 | 
| 10538 |    117 | 
 | 
| 69505 |    118 | If \<open>k\<close> is a numeral, \sdx{div}~\<open>k\<close>, \sdx{mod}~\<open>k\<close> and
 | 
|  |    119 | \<open>k\<close>~\sdx{dvd} are also supported, where the former two are eliminated
 | 
| 13996 |    120 | by case distinctions, again blowing up the running time.
 | 
|  |    121 | 
 | 
| 69505 |    122 | If the formula involves quantifiers, \<open>arith\<close> may take
 | 
| 13996 |    123 | super-exponential time and space.
 | 
| 10538 |    124 | \end{warn}
 | 
| 67406 |    125 | \<close>
 | 
| 10538 |    126 | 
 | 
| 8745 |    127 | (*<*)
 | 
|  |    128 | end
 | 
|  |    129 | (*>*)
 |