src/HOL/IMP/Examples.ML
author wenzelm
Tue, 23 Oct 2001 22:56:55 +0200
changeset 11912 d1b341c3aabc
parent 11704 3c50a2cd6f00
permissions -rw-r--r--
eliminated Numeral0;
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
9243
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
     1
(*  Title:      HOL/IMP/Examples.ML
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
     2
    ID:         $Id$
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
     3
    Author:     David von Oheimb, TUM
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
     4
    Copyright   2000 TUM
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
     5
*)
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
     6
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
     7
Addsimps[update_def];
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
     8
9556
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
     9
section "An example due to Tony Hoare";
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    10
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    11
Goal "[| !x. P x --> Q x; <w,s> -c-> t |] ==> \
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    12
\ !c. w = While P c --> <While Q c,t> -c-> u --> <While Q c,s> -c-> u";
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    13
by (etac evalc.induct 1);
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    14
by (Auto_tac);
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    15
val lemma1 = result() RS spec RS mp RS mp;
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    16
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    17
Goal "[| !x. P x --> Q x; <w,s> -c-> u |] ==> \
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    18
\ !c. w = While Q c --> <While P c; While Q c,s> -c-> u";
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    19
by (etac evalc.induct 1);
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    20
by (ALLGOALS Asm_simp_tac);
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    21
by (Blast_tac 1);
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    22
by (case_tac "P s" 1);
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    23
by (Auto_tac);
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    24
val lemma2 = result() RS spec RS mp;
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    25
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    26
Goal "!x. P x --> Q x ==> \
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    27
\ (<While P c; While Q c, s> -c-> t) = (<While Q c, s> -c-> t)";
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    28
by (blast_tac (claset() addIs [lemma1,lemma2]) 1);
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    29
qed "Hoare_example";
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    30
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    31
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    32
section "Factorial";
9243
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    33
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    34
val step = resolve_tac evalc.intrs 1;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    35
val simp = Asm_simp_tac 1;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    36
Goalw [factorial_def] "a~=b ==> \
11912
d1b341c3aabc eliminated Numeral0;
wenzelm
parents: 11704
diff changeset
    37
\ <factorial a b, Mem(a:=3)>  -c-> Mem(b:=6, a:=0)";
9243
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    38
by (ftac not_sym 1);
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    39
by step;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    40
by  step;
9556
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    41
by simp;
9243
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    42
by step;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    43
by   simp;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    44
by  step;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    45
by   step;
9556
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    46
by  simp;
9243
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    47
by  step;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    48
by simp;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    49
by step;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    50
by   simp;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    51
by  step;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    52
by   step;
9556
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    53
by  simp;
9243
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    54
by  step;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    55
by simp;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    56
by step;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    57
by   simp;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    58
by  step;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    59
by   step;
9556
dcdcfb0545e0 moved Hoare_example to Examples; other minor improvements
oheimb
parents: 9339
diff changeset
    60
by  simp;
9243
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    61
by  step;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    62
by simp;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    63
by step;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    64
by simp;
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    65
qed "factorial_3";