| 8745 |      1 | (*<*)
 | 
|  |      2 | theory natsum = Main:;
 | 
|  |      3 | (*>*)
 | 
|  |      4 | text{*\noindent
 | 
| 9792 |      5 | In particular, there are @{text"case"}-expressions, for example
 | 
| 9541 |      6 | @{term[display]"case n of 0 => 0 | Suc m => m"}
 | 
| 8745 |      7 | primitive recursion, for example
 | 
|  |      8 | *}
 | 
|  |      9 | 
 | 
| 10538 |     10 | consts sum :: "nat \<Rightarrow> nat";
 | 
| 8745 |     11 | primrec "sum 0 = 0"
 | 
|  |     12 |         "sum (Suc n) = Suc n + sum n";
 | 
|  |     13 | 
 | 
|  |     14 | text{*\noindent
 | 
|  |     15 | and induction, for example
 | 
|  |     16 | *}
 | 
|  |     17 | 
 | 
|  |     18 | lemma "sum n + sum n = n*(Suc n)";
 | 
|  |     19 | apply(induct_tac n);
 | 
| 10171 |     20 | apply(auto);
 | 
|  |     21 | done
 | 
| 8745 |     22 | 
 | 
| 10538 |     23 | text{*\newcommand{\mystar}{*%
 | 
|  |     24 | }
 | 
|  |     25 | The usual arithmetic operations \ttindexboldpos{+}{$HOL2arithfun},
 | 
|  |     26 | \ttindexboldpos{-}{$HOL2arithfun}, \ttindexboldpos{\mystar}{$HOL2arithfun},
 | 
|  |     27 | \isaindexbold{div}, \isaindexbold{mod}, \isaindexbold{min} and
 | 
|  |     28 | \isaindexbold{max} are predefined, as are the relations
 | 
|  |     29 | \indexboldpos{\isasymle}{$HOL2arithrel} and
 | 
| 10654 |     30 | \ttindexboldpos{<}{$HOL2arithrel}. As usual, @{prop"m-n = 0"} if
 | 
|  |     31 | @{prop"m<n"}. There is even a least number operation
 | 
| 10538 |     32 | \isaindexbold{LEAST}. For example, @{prop"(LEAST n. 1 < n) = 2"}, although
 | 
|  |     33 | Isabelle does not prove this completely automatically. Note that @{term 1}
 | 
|  |     34 | and @{term 2} are available as abbreviations for the corresponding
 | 
|  |     35 | @{term Suc}-expressions. If you need the full set of numerals,
 | 
| 10608 |     36 | see~\S\ref{sec:numerals}.
 | 
| 10538 |     37 | 
 | 
|  |     38 | \begin{warn}
 | 
|  |     39 |   The constant \ttindexbold{0} and the operations
 | 
|  |     40 |   \ttindexboldpos{+}{$HOL2arithfun}, \ttindexboldpos{-}{$HOL2arithfun},
 | 
|  |     41 |   \ttindexboldpos{\mystar}{$HOL2arithfun}, \isaindexbold{min},
 | 
|  |     42 |   \isaindexbold{max}, \indexboldpos{\isasymle}{$HOL2arithrel} and
 | 
|  |     43 |   \ttindexboldpos{<}{$HOL2arithrel} are overloaded, i.e.\ they are available
 | 
| 10978 |     44 |   not just for natural numbers but at other types as well.
 | 
|  |     45 |   For example, given the goal @{prop"x+0 = x"},
 | 
| 10538 |     46 |   there is nothing to indicate that you are talking about natural numbers.
 | 
|  |     47 |   Hence Isabelle can only infer that @{term x} is of some arbitrary type where
 | 
|  |     48 |   @{term 0} and @{text"+"} are declared. As a consequence, you will be unable
 | 
|  |     49 |   to prove the goal (although it may take you some time to realize what has
 | 
|  |     50 |   happened if @{text show_types} is not set).  In this particular example,
 | 
|  |     51 |   you need to include an explicit type constraint, for example
 | 
| 10788 |     52 |   @{text"x+0 = (x::nat)"}. If there is enough contextual information this
 | 
| 10538 |     53 |   may not be necessary: @{prop"Suc x = x"} automatically implies
 | 
|  |     54 |   @{text"x::nat"} because @{term Suc} is not overloaded.
 | 
| 10978 |     55 | 
 | 
|  |     56 |   For details see \S\ref{sec:numbers} and \S\ref{sec:overloading};
 | 
|  |     57 |   Table~\ref{tab:overloading} in the appendix shows the most important overloaded
 | 
|  |     58 |   operations.
 | 
| 10538 |     59 | \end{warn}
 | 
|  |     60 | 
 | 
|  |     61 | Simple arithmetic goals are proved automatically by both @{term auto} and the
 | 
| 11215 |     62 | simplification method introduced in \S\ref{sec:Simplification}.  For
 | 
| 10538 |     63 | example,
 | 
|  |     64 | *}
 | 
|  |     65 | 
 | 
|  |     66 | lemma "\<lbrakk> \<not> m < n; m < n+1 \<rbrakk> \<Longrightarrow> m = n"
 | 
|  |     67 | (*<*)by(auto)(*>*)
 | 
|  |     68 | 
 | 
|  |     69 | text{*\noindent
 | 
|  |     70 | is proved automatically. The main restriction is that only addition is taken
 | 
|  |     71 | into account; other arithmetic operations and quantified formulae are ignored.
 | 
|  |     72 | 
 | 
|  |     73 | For more complex goals, there is the special method \isaindexbold{arith}
 | 
|  |     74 | (which attacks the first subgoal). It proves arithmetic goals involving the
 | 
|  |     75 | usual logical connectives (@{text"\<not>"}, @{text"\<and>"}, @{text"\<or>"},
 | 
| 10971 |     76 | @{text"\<longrightarrow>"}), the relations @{text"="}, @{text"\<le>"} and @{text"<"},
 | 
|  |     77 | and the operations
 | 
| 10654 |     78 | @{text"+"}, @{text"-"}, @{term min} and @{term max}. Technically, this is
 | 
|  |     79 | known as the class of (quantifier free) \bfindex{linear arithmetic} formulae.
 | 
|  |     80 | For example,
 | 
| 10538 |     81 | *}
 | 
|  |     82 | 
 | 
|  |     83 | lemma "min i (max j (k*k)) = max (min (k*k) i) (min i (j::nat))";
 | 
|  |     84 | apply(arith)
 | 
|  |     85 | (*<*)done(*>*)
 | 
|  |     86 | 
 | 
|  |     87 | text{*\noindent
 | 
|  |     88 | succeeds because @{term"k*k"} can be treated as atomic. In contrast,
 | 
|  |     89 | *}
 | 
|  |     90 | 
 | 
|  |     91 | lemma "n*n = n \<Longrightarrow> n=0 \<or> n=1"
 | 
|  |     92 | (*<*)oops(*>*)
 | 
|  |     93 | 
 | 
|  |     94 | text{*\noindent
 | 
|  |     95 | is not even proved by @{text arith} because the proof relies essentially
 | 
|  |     96 | on properties of multiplication.
 | 
|  |     97 | 
 | 
|  |     98 | \begin{warn}
 | 
|  |     99 |   The running time of @{text arith} is exponential in the number of occurrences
 | 
|  |    100 |   of \ttindexboldpos{-}{$HOL2arithfun}, \isaindexbold{min} and
 | 
|  |    101 |   \isaindexbold{max} because they are first eliminated by case distinctions.
 | 
|  |    102 | 
 | 
| 10654 |    103 |   \isa{arith} is incomplete even for the restricted class of
 | 
|  |    104 |   linear arithmetic formulae. If divisibility plays a
 | 
| 10538 |    105 |   role, it may fail to prove a valid formula, for example
 | 
|  |    106 |   @{prop"m+m \<noteq> n+n+1"}. Fortunately, such examples are rare in practice.
 | 
|  |    107 | \end{warn}
 | 
|  |    108 | *}
 | 
|  |    109 | 
 | 
| 8745 |    110 | (*<*)
 | 
|  |    111 | end
 | 
|  |    112 | (*>*)
 |