src/HOL/IMP/Examples.thy
author wenzelm
Sun, 16 Jan 2011 15:53:03 +0100
changeset 41589 bbd861837ebc
parent 27362 a6dc1769fdda
permissions -rw-r--r--
tuned headers;
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.thy
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
     2
    Author:     David von Oheimb, TUM
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
     3
*)
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
     4
12431
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
     5
header "Examples"
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
     6
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 12431
diff changeset
     7
theory Examples imports Natural begin
12431
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
     8
27362
a6dc1769fdda modernized specifications;
wenzelm
parents: 18372
diff changeset
     9
definition
a6dc1769fdda modernized specifications;
wenzelm
parents: 18372
diff changeset
    10
  factorial :: "loc => loc => com" where
a6dc1769fdda modernized specifications;
wenzelm
parents: 18372
diff changeset
    11
  "factorial a b = (b :== (%s. 1);
12431
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    12
                    \<WHILE> (%s. s a ~= 0) \<DO>
27362
a6dc1769fdda modernized specifications;
wenzelm
parents: 18372
diff changeset
    13
                    (b :== (%s. s b * s a); a :== (%s. s a - 1)))"
12431
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    14
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    15
declare update_def [simp]
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    16
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    17
subsection "An example due to Tony Hoare"
9243
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    18
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    19
lemma lemma1:
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    20
  assumes 1: "!x. P x \<longrightarrow> Q x"
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    21
    and 2: "\<langle>w,s\<rangle> \<longrightarrow>\<^sub>c t"
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    22
  shows "w = While P c \<Longrightarrow> \<langle>While Q c,t\<rangle> \<longrightarrow>\<^sub>c u \<Longrightarrow> \<langle>While Q c,s\<rangle> \<longrightarrow>\<^sub>c u"
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    23
  using 2 apply induct
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    24
  using 1 apply auto
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    25
  done
12431
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    26
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    27
lemma lemma2 [rule_format (no_asm)]:
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    28
  "[| !x. P x \<longrightarrow> Q x; \<langle>w,s\<rangle> \<longrightarrow>\<^sub>c u |] ==>
12431
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    29
  !c. w = While Q c \<longrightarrow> \<langle>While P c; While Q c,s\<rangle> \<longrightarrow>\<^sub>c u"
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    30
apply (erule evalc.induct)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    31
apply (simp_all (no_asm_simp))
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    32
apply blast
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    33
apply (case_tac "P s")
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    34
apply auto
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    35
done
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    36
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    37
lemma Hoare_example: "!x. P x \<longrightarrow> Q x ==>
12431
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    38
  (\<langle>While P c; While Q c, s\<rangle> \<longrightarrow>\<^sub>c t) = (\<langle>While Q c, s\<rangle> \<longrightarrow>\<^sub>c t)"
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    39
  by (blast intro: lemma1 lemma2 dest: semi [THEN iffD1])
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    40
9243
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    41
12431
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    42
subsection "Factorial"
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    43
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    44
lemma factorial_3: "a~=b ==>
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    45
    \<langle>factorial a b, Mem(a:=3)\<rangle> \<longrightarrow>\<^sub>c Mem(b:=6, a:=0)"
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    46
  by (simp add: factorial_def)
12431
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    47
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    48
text {* the same in single step mode: *}
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    49
lemmas [simp del] = evalc_cases
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    50
lemma  "a~=b \<Longrightarrow> \<langle>factorial a b, Mem(a:=3)\<rangle> \<longrightarrow>\<^sub>c Mem(b:=6, a:=0)"
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    51
apply (unfold factorial_def)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    52
apply (frule not_sym)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    53
apply (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    54
apply  (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    55
apply simp
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    56
apply (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    57
apply   simp
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    58
apply  (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    59
apply   (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    60
apply  simp
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    61
apply  (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    62
apply simp
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    63
apply (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    64
apply   simp
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    65
apply  (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    66
apply   (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    67
apply  simp
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    68
apply  (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    69
apply simp
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    70
apply (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    71
apply   simp
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    72
apply  (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    73
apply   (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    74
apply  simp
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    75
apply  (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    76
apply simp
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    77
apply (rule evalc.intros)
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    78
apply simp
07ec657249e5 converted to Isar
kleing
parents: 9243
diff changeset
    79
done
18372
2bffdf62fe7f tuned proofs;
wenzelm
parents: 16417
diff changeset
    80
9243
22b460a0b676 disambiguated := ; added Examples (factorial)
oheimb
parents:
diff changeset
    81
end