doc-src/TutorialI/Misc/natsum.thy
author Walther Neuper <neuper@ist.tugraz.at>
Thu, 12 Aug 2010 15:03:34 +0200
branchisac-from-Isabelle2009-2
changeset 37913 20e3616b2d9c
parent 27168 9a9cc62932d9
permissions -rw-r--r--
prepare reactivation of isac-update-Isa09-2
     1 (*<*)
     2 theory natsum imports Main begin
     3 (*>*)
     4 text{*\noindent
     5 In particular, there are @{text"case"}-expressions, for example
     6 @{term[display]"case n of 0 => 0 | Suc m => m"}
     7 primitive recursion, for example
     8 *}
     9 
    10 primrec sum :: "nat \<Rightarrow> nat" where
    11 "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)
    20 apply(auto)
    21 done
    22 
    23 text{*\newcommand{\mystar}{*%
    24 }
    25 \index{arithmetic operations!for \protect\isa{nat}}%
    26 The arithmetic operations \isadxboldpos{+}{$HOL2arithfun},
    27 \isadxboldpos{-}{$HOL2arithfun}, \isadxboldpos{\mystar}{$HOL2arithfun},
    28 \sdx{div}, \sdx{mod}, \cdx{min} and
    29 \cdx{max} are predefined, as are the relations
    30 \isadxboldpos{\isasymle}{$HOL2arithrel} and
    31 \isadxboldpos{<}{$HOL2arithrel}. As usual, @{prop"m-n = (0::nat)"} if
    32 @{prop"m<n"}. There is even a least number operation
    33 \sdx{LEAST}\@.  For example, @{prop"(LEAST n. 0 < n) = Suc 0"}.
    34 \begin{warn}\index{overloading}
    35   The constants \cdx{0} and \cdx{1} and the operations
    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
    40   not just for natural numbers but for other types as well.
    41   For example, given the goal @{text"x + 0 = x"}, there is nothing to indicate
    42   that you are talking about natural numbers. Hence Isabelle can only infer
    43   that @{term x} is of some arbitrary type where @{text 0} and @{text"+"} are
    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
    46   fixed type in its output: @{prop"x+0 = x"}. (In the absence of a numeral,
    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
    49   an explicit type constraint, for example @{text"x+0 = (x::nat)"}. If there
    50   is enough contextual information this may not be necessary: @{prop"Suc x =
    51   x"} automatically implies @{text"x::nat"} because @{term Suc} is not
    52   overloaded.
    53 
    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}
    59   The symbols \isadxboldpos{>}{$HOL2arithrel} and
    60   \isadxboldpos{\isasymge}{$HOL2arithrel} are merely syntax: @{text"x > y"}
    61   stands for @{prop"y < x"} and similary for @{text"\<ge>"} and
    62   @{text"\<le>"}.
    63 \end{warn}
    64 \begin{warn}
    65   Constant @{text"1::nat"} is defined to equal @{term"Suc 0"}. This definition
    66   (see \S\ref{sec:ConstDefinitions}) is unfolded automatically by some
    67   tactics (like @{text auto}, @{text simp} and @{text arith}) but not by
    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}.
    70   \emph{Novices are advised to stick to @{term"0::nat"} and @{term Suc}.}
    71 \end{warn}
    72 
    73 Both @{text auto} and @{text simp}
    74 (a method introduced below, \S\ref{sec:Simplification}) prove 
    75 simple arithmetic goals automatically:
    76 *}
    77 
    78 lemma "\<lbrakk> \<not> m < n; m < n + (1::nat) \<rbrakk> \<Longrightarrow> m = n"
    79 (*<*)by(auto)(*>*)
    80 
    81 text{*\noindent
    82 For efficiency's sake, this built-in prover ignores quantified formulae,
    83 many logical connectives, and all arithmetic operations apart from addition.
    84 In consequence, @{text auto} and @{text simp} cannot prove this slightly more complex goal:
    85 *}
    86 
    87 lemma "m \<noteq> (n::nat) \<Longrightarrow> m < n \<or> n < m"
    88 (*<*)by(arith)(*>*)
    89 
    90 text{*\noindent The method \methdx{arith} is more general.  It attempts to
    91 prove the first subgoal provided it is a \textbf{linear arithmetic} formula.
    92 Such formulas may involve the usual logical connectives (@{text"\<not>"},
    93 @{text"\<and>"}, @{text"\<or>"}, @{text"\<longrightarrow>"}, @{text"="},
    94 @{text"\<forall>"}, @{text"\<exists>"}), the relations @{text"="},
    95 @{text"\<le>"} and @{text"<"}, and the operations @{text"+"}, @{text"-"},
    96 @{term min} and @{term max}.  For example, *}
    97 
    98 lemma "min i (max j (k*k)) = max (min (k*k) i) (min i (j::nat))"
    99 apply(arith)
   100 (*<*)done(*>*)
   101 
   102 text{*\noindent
   103 succeeds because @{term"k*k"} can be treated as atomic. In contrast,
   104 *}
   105 
   106 lemma "n*n = n+1 \<Longrightarrow> n=0"
   107 (*<*)oops(*>*)
   108 
   109 text{*\noindent
   110 is not proved by @{text arith} because the proof relies 
   111 on properties of multiplication. Only multiplication by numerals (which is
   112 the same as iterated addition) is taken into account.
   113 
   114 \begin{warn} The running time of @{text arith} is exponential in the number
   115   of occurrences of \ttindexboldpos{-}{$HOL2arithfun}, \cdx{min} and
   116   \cdx{max} because they are first eliminated by case distinctions.
   117 
   118 If @{text k} is a numeral, \sdx{div}~@{text k}, \sdx{mod}~@{text k} and
   119 @{text k}~\sdx{dvd} are also supported, where the former two are eliminated
   120 by case distinctions, again blowing up the running time.
   121 
   122 If the formula involves quantifiers, @{text arith} may take
   123 super-exponential time and space.
   124 \end{warn}
   125 *}
   126 
   127 (*<*)
   128 end
   129 (*>*)