doc-src/TutorialI/Misc/asm_simp.thy
author wenzelm
Tue, 01 Aug 2000 13:43:22 +0200
changeset 9490 c2606af9922c
parent 9458 c613cd06d5cf
child 9494 44fefb6e9994
permissions -rw-r--r--
tuned msg;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     1
(*<*)
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     2
theory asm_simp = Main:;
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     3
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     4
(*>*)text{*
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     5
By default, assumptions are part of the simplification process: they are used
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     6
as simplification rules and are simplified themselves. For example:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     7
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     8
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
     9
lemma "\\<lbrakk> xs @ zs = ys @ xs; [] @ xs = [] @ [] \\<rbrakk> \\<Longrightarrow> ys = zs";
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8823
diff changeset
    10
by simp;
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    11
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    12
text{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    13
The second assumption simplifies to \isa{xs = []}, which in turn
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    14
simplifies the first assumption to \isa{zs = ys}, thus reducing the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    15
conclusion to \isa{ys = ys} and hence to \isa{True}.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    16
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    17
In some cases this may be too much of a good thing and may lead to
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    18
nontermination:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    19
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    20
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    21
lemma "\\<forall>x. f x = g (f (g x)) \\<Longrightarrow> f [] = f [] @ []";
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    22
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    23
txt{*\noindent
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    24
cannot be solved by an unmodified application of \isa{simp} because the
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    25
simplification rule \isa{f x = g(f(g x))} extracted from the assumption
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    26
does not terminate. Isabelle notices certain simple forms of
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    27
nontermination but not this one. The problem can be circumvented by
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    28
explicitly telling the simplifier to ignore the assumptions:
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    29
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    30
9458
c613cd06d5cf apply. -> by
nipkow
parents: 8823
diff changeset
    31
by(simp (no_asm));
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    32
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    33
text{*\noindent
8823
bd8f8dbda512 updated syntax of simp options: (no_asm) etc.;
wenzelm
parents: 8745
diff changeset
    34
There are three options that influence the treatment of assumptions:
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    35
\begin{description}
8823
bd8f8dbda512 updated syntax of simp options: (no_asm) etc.;
wenzelm
parents: 8745
diff changeset
    36
\item[\isaindexbold{(no_asm)}] means that assumptions are completely ignored.
bd8f8dbda512 updated syntax of simp options: (no_asm) etc.;
wenzelm
parents: 8745
diff changeset
    37
\item[\isaindexbold{(no_asm_simp)}] means that the assumptions are not simplified but
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    38
  are used in the simplification of the conclusion.
8823
bd8f8dbda512 updated syntax of simp options: (no_asm) etc.;
wenzelm
parents: 8745
diff changeset
    39
\item[\isaindexbold{(no_asm_use)}] means that the assumptions are simplified
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    40
but are not
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    41
  used in the simplification of each other or the conclusion.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    42
\end{description}
8823
bd8f8dbda512 updated syntax of simp options: (no_asm) etc.;
wenzelm
parents: 8745
diff changeset
    43
Neither \isa{(no_asm_simp)} nor \isa{(no_asm_use)} allow to simplify the above
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    44
problematic subgoal.
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    45
8823
bd8f8dbda512 updated syntax of simp options: (no_asm) etc.;
wenzelm
parents: 8745
diff changeset
    46
Note that only one of the above options is allowed, and it must precede all
bd8f8dbda512 updated syntax of simp options: (no_asm) etc.;
wenzelm
parents: 8745
diff changeset
    47
other arguments.
8745
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    48
*}
13b32661dde4 I wonder which files i forgot.
nipkow
parents:
diff changeset
    49
(*<*)end(*>*)