src/HOL/Hoare/Pointer_Examples.thy
author haftmann
Fri, 17 Jun 2005 16:12:49 +0200
changeset 16417 9bc16273c2d4
parent 14074 93dfce3b6f86
child 19397 524f1cb4652a
permissions -rw-r--r--
migrated theory headers to new format
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
     1
(*  Title:      HOL/Hoare/Pointers.thy
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
     2
    ID:         $Id$
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
     3
    Author:     Tobias Nipkow
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
     4
    Copyright   2002 TUM
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
     5
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
     6
Examples of verifications of pointer programs
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
     7
*)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
     8
16417
9bc16273c2d4 migrated theory headers to new format
haftmann
parents: 14074
diff changeset
     9
theory Pointer_Examples imports HeapSyntax begin
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    10
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    11
section "Verifications"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    12
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    13
subsection "List reversal"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    14
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    15
text "A short but unreadable proof:"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    16
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    17
lemma "VARS tl p q r
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    18
  {List tl p Ps \<and> List tl q Qs \<and> set Ps \<inter> set Qs = {}}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    19
  WHILE p \<noteq> Null
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    20
  INV {\<exists>ps qs. List tl p ps \<and> List tl q qs \<and> set ps \<inter> set qs = {} \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    21
                 rev ps @ qs = rev Ps @ Qs}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    22
  DO r := p; p := p^.tl; r^.tl := q; q := r OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    23
  {List tl q (rev Ps @ Qs)}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    24
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    25
  apply fastsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    26
 apply(fastsimp intro:notin_List_update[THEN iffD2])
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    27
(* explicit:
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    28
 apply clarify
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    29
 apply(rename_tac ps b qs)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    30
 apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    31
 apply(rename_tac ps')
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    32
 apply(fastsimp intro:notin_List_update[THEN iffD2])
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    33
 apply(rule_tac x = ps' in exI)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    34
 apply simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    35
 apply(rule_tac x = "b#qs" in exI)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    36
 apply simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    37
*)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    38
apply fastsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    39
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    40
14062
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    41
text{* And now with ghost variables @{term ps} and @{term qs}. Even
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    42
``more automatic''. *}
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    43
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    44
lemma "VARS next p ps q qs r
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    45
  {List next p Ps \<and> List next q Qs \<and> set Ps \<inter> set Qs = {} \<and>
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    46
   ps = Ps \<and> qs = Qs}
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    47
  WHILE p \<noteq> Null
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    48
  INV {List next p ps \<and> List next q qs \<and> set ps \<inter> set qs = {} \<and>
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    49
       rev ps @ qs = rev Ps @ Qs}
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    50
  DO r := p; p := p^.next; r^.next := q; q := r;
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    51
     qs := (hd ps) # qs; ps := tl ps OD
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    52
  {List next q (rev Ps @ Qs)}"
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    53
apply vcg_simp
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    54
 apply fastsimp
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    55
apply fastsimp
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    56
done
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    57
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    58
text "A longer readable version:"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    59
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    60
lemma "VARS tl p q r
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    61
  {List tl p Ps \<and> List tl q Qs \<and> set Ps \<inter> set Qs = {}}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    62
  WHILE p \<noteq> Null
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    63
  INV {\<exists>ps qs. List tl p ps \<and> List tl q qs \<and> set ps \<inter> set qs = {} \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    64
               rev ps @ qs = rev Ps @ Qs}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    65
  DO r := p; p := p^.tl; r^.tl := q; q := r OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    66
  {List tl q (rev Ps @ Qs)}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    67
proof vcg
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    68
  fix tl p q r
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    69
  assume "List tl p Ps \<and> List tl q Qs \<and> set Ps \<inter> set Qs = {}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    70
  thus "\<exists>ps qs. List tl p ps \<and> List tl q qs \<and> set ps \<inter> set qs = {} \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    71
                rev ps @ qs = rev Ps @ Qs" by fastsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    72
next
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    73
  fix tl p q r
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    74
  assume "(\<exists>ps qs. List tl p ps \<and> List tl q qs \<and> set ps \<inter> set qs = {} \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    75
                   rev ps @ qs = rev Ps @ Qs) \<and> p \<noteq> Null"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    76
         (is "(\<exists>ps qs. ?I ps qs) \<and> _")
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    77
  then obtain ps qs a where I: "?I ps qs \<and> p = Ref a"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    78
    by fast
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    79
  then obtain ps' where "ps = a # ps'" by fastsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    80
  hence "List (tl(p \<rightarrow> q)) (p^.tl) ps' \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    81
         List (tl(p \<rightarrow> q)) p       (a#qs) \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    82
         set ps' \<inter> set (a#qs) = {} \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    83
         rev ps' @ (a#qs) = rev Ps @ Qs"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    84
    using I by fastsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    85
  thus "\<exists>ps' qs'. List (tl(p \<rightarrow> q)) (p^.tl) ps' \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    86
                  List (tl(p \<rightarrow> q)) p       qs' \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    87
                  set ps' \<inter> set qs' = {} \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    88
                  rev ps' @ qs' = rev Ps @ Qs" by fast
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    89
next
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    90
  fix tl p q r
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    91
  assume "(\<exists>ps qs. List tl p ps \<and> List tl q qs \<and> set ps \<inter> set qs = {} \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    92
                   rev ps @ qs = rev Ps @ Qs) \<and> \<not> p \<noteq> Null"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    93
  thus "List tl q (rev Ps @ Qs)" by fastsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    94
qed
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    95
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    96
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    97
text{* Finaly, the functional version. A bit more verbose, but automatic! *}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    98
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    99
lemma "VARS tl p q r
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   100
  {islist tl p \<and> islist tl q \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   101
   Ps = list tl p \<and> Qs = list tl q \<and> set Ps \<inter> set Qs = {}}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   102
  WHILE p \<noteq> Null
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   103
  INV {islist tl p \<and> islist tl q \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   104
       set(list tl p) \<inter> set(list tl q) = {} \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   105
       rev(list tl p) @ (list tl q) = rev Ps @ Qs}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   106
  DO r := p; p := p^.tl; r^.tl := q; q := r OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   107
  {islist tl q \<and> list tl q = rev Ps @ Qs}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   108
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   109
  apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   110
 apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   111
apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   112
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   113
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   114
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   115
subsection "Searching in a list"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   116
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   117
text{*What follows is a sequence of successively more intelligent proofs that
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   118
a simple loop finds an element in a linked list.
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   119
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   120
We start with a proof based on the @{term List} predicate. This means it only
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   121
works for acyclic lists. *}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   122
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   123
lemma "VARS tl p
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   124
  {List tl p Ps \<and> X \<in> set Ps}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   125
  WHILE p \<noteq> Null \<and> p \<noteq> Ref X
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   126
  INV {\<exists>ps. List tl p ps \<and> X \<in> set ps}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   127
  DO p := p^.tl OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   128
  {p = Ref X}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   129
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   130
  apply blast
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   131
 apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   132
apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   133
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   134
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   135
text{*Using @{term Path} instead of @{term List} generalizes the correctness
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   136
statement to cyclic lists as well: *}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   137
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   138
lemma "VARS tl p
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   139
  {Path tl p Ps X}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   140
  WHILE p \<noteq> Null \<and> p \<noteq> X
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   141
  INV {\<exists>ps. Path tl p ps X}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   142
  DO p := p^.tl OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   143
  {p = X}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   144
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   145
  apply blast
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   146
 apply fastsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   147
apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   148
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   149
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   150
text{*Now it dawns on us that we do not need the list witness at all --- it
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   151
suffices to talk about reachability, i.e.\ we can use relations directly. The
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   152
first version uses a relation on @{typ"'a ref"}: *}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   153
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   154
lemma "VARS tl p
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   155
  {(p,X) \<in> {(Ref x,tl x) |x. True}^*}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   156
  WHILE p \<noteq> Null \<and> p \<noteq> X
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   157
  INV {(p,X) \<in> {(Ref x,tl x) |x. True}^*}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   158
  DO p := p^.tl OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   159
  {p = X}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   160
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   161
 apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   162
 apply(erule converse_rtranclE)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   163
  apply simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   164
 apply(clarsimp elim:converse_rtranclE)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   165
apply(fast elim:converse_rtranclE)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   166
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   167
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   168
text{*Finally, a version based on a relation on type @{typ 'a}:*}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   169
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   170
lemma "VARS tl p
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   171
  {p \<noteq> Null \<and> (addr p,X) \<in> {(x,y). tl x = Ref y}^*}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   172
  WHILE p \<noteq> Null \<and> p \<noteq> Ref X
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   173
  INV {p \<noteq> Null \<and> (addr p,X) \<in> {(x,y). tl x = Ref y}^*}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   174
  DO p := p^.tl OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   175
  {p = Ref X}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   176
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   177
 apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   178
 apply(erule converse_rtranclE)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   179
  apply simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   180
 apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   181
apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   182
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   183
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   184
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   185
subsection "Merging two lists"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   186
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   187
text"This is still a bit rough, especially the proof."
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   188
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   189
constdefs
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   190
 cor :: "bool \<Rightarrow> bool \<Rightarrow> bool"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   191
"cor P Q == if P then True else Q"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   192
 cand :: "bool \<Rightarrow> bool \<Rightarrow> bool"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   193
"cand P Q == if P then Q else False"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   194
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   195
consts merge :: "'a list * 'a list * ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   196
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   197
recdef merge "measure(%(xs,ys,f). size xs + size ys)"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   198
"merge(x#xs,y#ys,f) = (if f x y then x # merge(xs,y#ys,f)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   199
                                else y # merge(x#xs,ys,f))"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   200
"merge(x#xs,[],f) = x # merge(xs,[],f)"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   201
"merge([],y#ys,f) = y # merge([],ys,f)"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   202
"merge([],[],f) = []"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   203
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   204
text{* Simplifies the proof a little: *}
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   205
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   206
lemma [simp]: "({} = insert a A \<inter> B) = (a \<notin> B & {} = A \<inter> B)"
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   207
by blast
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   208
lemma [simp]: "({} = A \<inter> insert b B) = (b \<notin> A & {} = A \<inter> B)"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   209
by blast
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   210
lemma [simp]: "({} = A \<inter> (B \<union> C)) = ({} = A \<inter> B & {} = A \<inter> C)"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   211
by blast
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   212
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   213
lemma "VARS hd tl p q r s
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   214
 {List tl p Ps \<and> List tl q Qs \<and> set Ps \<inter> set Qs = {} \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   215
  (p \<noteq> Null \<or> q \<noteq> Null)}
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   216
 IF cor (q = Null) (cand (p \<noteq> Null) (p^.hd \<le> q^.hd))
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   217
 THEN r := p; p := p^.tl ELSE r := q; q := q^.tl FI;
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   218
 s := r;
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   219
 WHILE p \<noteq> Null \<or> q \<noteq> Null
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   220
 INV {EX rs ps qs a. Path tl r rs s \<and> List tl p ps \<and> List tl q qs \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   221
      distinct(a # ps @ qs @ rs) \<and> s = Ref a \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   222
      merge(Ps,Qs,\<lambda>x y. hd x \<le> hd y) =
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   223
      rs @ a # merge(ps,qs,\<lambda>x y. hd x \<le> hd y) \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   224
      (tl a = p \<or> tl a = q)}
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   225
 DO IF cor (q = Null) (cand (p \<noteq> Null) (p^.hd \<le> q^.hd))
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   226
    THEN s^.tl := p; p := p^.tl ELSE s^.tl := q; q := q^.tl FI;
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   227
    s := s^.tl
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   228
 OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   229
 {List tl r (merge(Ps,Qs,\<lambda>x y. hd x \<le> hd y))}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   230
apply vcg_simp
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   231
apply (simp_all add: cand_def cor_def)
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   232
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   233
apply (fastsimp)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   234
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   235
apply clarsimp
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   236
apply(rule conjI)
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   237
apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   238
apply(rule conjI)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   239
apply (fastsimp intro!:Path_snoc intro:Path_upd[THEN iffD2] notin_List_update[THEN iffD2] simp:eq_sym_conv)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   240
apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   241
apply(rule conjI)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   242
apply (clarsimp)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   243
apply(rule_tac x = "rs @ [a]" in exI)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   244
apply(clarsimp simp:eq_sym_conv)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   245
apply(rule_tac x = "bs" in exI)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   246
apply(clarsimp simp:eq_sym_conv)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   247
apply(rule_tac x = "ya#bsa" in exI)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   248
apply(simp)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   249
apply(clarsimp simp:eq_sym_conv)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   250
apply(rule_tac x = "rs @ [a]" in exI)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   251
apply(clarsimp simp:eq_sym_conv)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   252
apply(rule_tac x = "y#bs" in exI)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   253
apply(clarsimp simp:eq_sym_conv)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   254
apply(rule_tac x = "bsa" in exI)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   255
apply(simp)
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   256
apply (fastsimp intro!:Path_snoc intro:Path_upd[THEN iffD2] notin_List_update[THEN iffD2] simp:eq_sym_conv)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   257
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   258
apply(clarsimp simp add:List_app)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   259
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   260
14074
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   261
text{* And now with ghost variables: *}
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   262
14074
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   263
lemma "VARS elem next p q r s ps qs rs a
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   264
 {List next p Ps \<and> List next q Qs \<and> set Ps \<inter> set Qs = {} \<and>
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   265
  (p \<noteq> Null \<or> q \<noteq> Null) \<and> ps = Ps \<and> qs = Qs}
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   266
 IF cor (q = Null) (cand (p \<noteq> Null) (p^.elem \<le> q^.elem))
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   267
 THEN r := p; p := p^.next; ps := tl ps
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   268
 ELSE r := q; q := q^.next; qs := tl qs FI;
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   269
 s := r; rs := []; a := addr s;
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   270
 WHILE p \<noteq> Null \<or> q \<noteq> Null
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   271
 INV {Path next r rs s \<and> List next p ps \<and> List next q qs \<and>
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   272
      distinct(a # ps @ qs @ rs) \<and> s = Ref a \<and>
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   273
      merge(Ps,Qs,\<lambda>x y. elem x \<le> elem y) =
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   274
      rs @ a # merge(ps,qs,\<lambda>x y. elem x \<le> elem y) \<and>
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   275
      (next a = p \<or> next a = q)}
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   276
 DO IF cor (q = Null) (cand (p \<noteq> Null) (p^.elem \<le> q^.elem))
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   277
    THEN s^.next := p; p := p^.next; ps := tl ps
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   278
    ELSE s^.next := q; q := q^.next; qs := tl qs FI;
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   279
    rs := rs @ [a]; s := s^.next; a := addr s
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   280
 OD
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   281
 {List next r (merge(Ps,Qs,\<lambda>x y. elem x \<le> elem y))}"
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   282
apply vcg_simp
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   283
apply (simp_all add: cand_def cor_def)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   284
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   285
apply (fastsimp)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   286
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   287
apply clarsimp
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   288
apply(rule conjI)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   289
apply(clarsimp)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   290
apply(rule conjI)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   291
apply(clarsimp simp:neq_commute)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   292
apply(clarsimp simp:neq_commute)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   293
apply(clarsimp simp:neq_commute)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   294
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   295
apply(clarsimp simp add:List_app)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   296
done
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   297
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   298
text{* The proof is a LOT simpler because it does not need
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   299
instantiations anymore, but it is still not quite automatic, probably
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   300
because of this wrong orientation business. *}
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   301
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   302
text{* More of the previous proof without ghost variables can be
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   303
automated, but the runtime goes up drastically. In general it is
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   304
usually more efficient to give the witness directly than to have it
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   305
found by proof.
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   306
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   307
Now we try a functional version of the abstraction relation @{term
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   308
Path}. Since the result is not that convincing, we do not prove any of
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   309
the lemmas.*}
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   310
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   311
consts ispath:: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   312
       path:: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a ref \<Rightarrow> 'a list"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   313
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   314
ML"set quick_and_dirty"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   315
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   316
text"First some basic lemmas:"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   317
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   318
lemma [simp]: "ispath f p p"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   319
sorry
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   320
lemma [simp]: "path f p p = []"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   321
sorry
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   322
lemma [simp]: "ispath f p q \<Longrightarrow> a \<notin> set(path f p q) \<Longrightarrow> ispath (f(a := r)) p q"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   323
sorry
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   324
lemma [simp]: "ispath f p q \<Longrightarrow> a \<notin> set(path f p q) \<Longrightarrow>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   325
 path (f(a := r)) p q = path f p q"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   326
sorry
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   327
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   328
text"Some more specific lemmas needed by the example:"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   329
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   330
lemma [simp]: "ispath (f(a := q)) p (Ref a) \<Longrightarrow> ispath (f(a := q)) p q"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   331
sorry
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   332
lemma [simp]: "ispath (f(a := q)) p (Ref a) \<Longrightarrow>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   333
 path (f(a := q)) p q = path (f(a := q)) p (Ref a) @ [a]"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   334
sorry
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   335
lemma [simp]: "ispath f p (Ref a) \<Longrightarrow> f a = Ref b \<Longrightarrow>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   336
 b \<notin> set (path f p (Ref a))"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   337
sorry
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   338
lemma [simp]: "ispath f p (Ref a) \<Longrightarrow> f a = Null \<Longrightarrow> islist f p"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   339
sorry
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   340
lemma [simp]: "ispath f p (Ref a) \<Longrightarrow> f a = Null \<Longrightarrow> list f p = path f p (Ref a) @ [a]"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   341
sorry
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   342
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   343
lemma [simp]: "islist f p \<Longrightarrow> distinct (list f p)"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   344
sorry
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   345
ML"reset quick_and_dirty"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   346
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   347
lemma "VARS hd tl p q r s
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   348
 {islist tl p & Ps = list tl p \<and> islist tl q & Qs = list tl q \<and>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   349
  set Ps \<inter> set Qs = {} \<and>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   350
  (p \<noteq> Null \<or> q \<noteq> Null)}
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   351
 IF cor (q = Null) (cand (p \<noteq> Null) (p^.hd \<le> q^.hd))
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   352
 THEN r := p; p := p^.tl ELSE r := q; q := q^.tl FI;
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   353
 s := r;
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   354
 WHILE p \<noteq> Null \<or> q \<noteq> Null
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   355
 INV {EX rs ps qs a. ispath tl r s & rs = path tl r s \<and>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   356
      islist tl p & ps = list tl p \<and> islist tl q & qs = list tl q \<and>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   357
      distinct(a # ps @ qs @ rs) \<and> s = Ref a \<and>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   358
      merge(Ps,Qs,\<lambda>x y. hd x \<le> hd y) =
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   359
      rs @ a # merge(ps,qs,\<lambda>x y. hd x \<le> hd y) \<and>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   360
      (tl a = p \<or> tl a = q)}
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   361
 DO IF cor (q = Null) (cand (p \<noteq> Null) (p^.hd \<le> q^.hd))
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   362
    THEN s^.tl := p; p := p^.tl ELSE s^.tl := q; q := q^.tl FI;
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   363
    s := s^.tl
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   364
 OD
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   365
 {islist tl r & list tl r = (merge(Ps,Qs,\<lambda>x y. hd x \<le> hd y))}"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   366
apply vcg_simp
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   367
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   368
apply (simp_all add: cand_def cor_def)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   369
  apply (fastsimp)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   370
 apply (fastsimp simp: eq_sym_conv)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   371
apply(clarsimp)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   372
done
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   373
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   374
text"The proof is automatic, but requires a numbet of special lemmas."
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   375
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   376
subsection "Storage allocation"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   377
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   378
constdefs new :: "'a set \<Rightarrow> 'a"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   379
"new A == SOME a. a \<notin> A"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   380
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   381
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   382
lemma new_notin:
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   383
 "\<lbrakk> ~finite(UNIV::'a set); finite(A::'a set); B \<subseteq> A \<rbrakk> \<Longrightarrow> new (A) \<notin> B"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   384
apply(unfold new_def)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   385
apply(rule someI2_ex)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   386
 apply (fast intro:ex_new_if_finite)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   387
apply (fast)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   388
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   389
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   390
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   391
lemma "~finite(UNIV::'a set) \<Longrightarrow>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   392
  VARS xs elem next alloc p q
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   393
  {Xs = xs \<and> p = (Null::'a ref)}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   394
  WHILE xs \<noteq> []
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   395
  INV {islist next p \<and> set(list next p) \<subseteq> set alloc \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   396
       map elem (rev(list next p)) @ xs = Xs}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   397
  DO q := Ref(new(set alloc)); alloc := (addr q)#alloc;
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   398
     q^.next := p; q^.elem := hd xs; xs := tl xs; p := q
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   399
  OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   400
  {islist next p \<and> map elem (rev(list next p)) = Xs}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   401
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   402
 apply (clarsimp simp: subset_insert_iff neq_Nil_conv fun_upd_apply new_notin)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   403
apply fastsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   404
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   405
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   406
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   407
end