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