src/HOL/Hoare/Pointer_Examples.thy
author wenzelm
Tue, 19 Oct 2021 18:24:33 +0200
changeset 74551 375e8e1a2139
parent 72990 db8f94656024
permissions -rw-r--r--
merged
Ignore whitespace changes - Everywhere: Within whitespace: At end of lines:
41959
b460124855b8 tuned headers;
wenzelm
parents: 38353
diff changeset
     1
(*  Title:      HOL/Hoare/Pointer_Examples.thy
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
     2
    Author:     Tobias Nipkow
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
     3
    Copyright   2002 TUM
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
     4
*)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
     5
72990
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
     6
section \<open>Examples of verifications of pointer programs\<close>
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
     7
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
     8
theory Pointer_Examples
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
     9
  imports HeapSyntax
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
    10
begin
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    11
24499
5a3ee202e0b0 do not touch quick_and_dirty;
wenzelm
parents: 19399
diff changeset
    12
axiomatization where unproven: "PROP A"
5a3ee202e0b0 do not touch quick_and_dirty;
wenzelm
parents: 19399
diff changeset
    13
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    14
72990
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
    15
subsection "Verifications"
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
    16
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
    17
subsubsection "List reversal"
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    18
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    19
text "A short but unreadable proof:"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    20
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    21
lemma "VARS tl p q r
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    22
  {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
    23
  WHILE p \<noteq> Null
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    24
  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
    25
                 rev ps @ qs = rev Ps @ Qs}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    26
  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
    27
  {List tl q (rev Ps @ Qs)}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    28
apply vcg_simp
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
    29
  apply fastforce
67444
100247708f31 clarified comments;
wenzelm
parents: 62042
diff changeset
    30
  apply(fastforce intro:notin_List_update[THEN iffD2])
100247708f31 clarified comments;
wenzelm
parents: 62042
diff changeset
    31
  \<comment> \<open>explicit:\<close>
100247708f31 clarified comments;
wenzelm
parents: 62042
diff changeset
    32
  \<^cancel>\<open>apply clarify\<close>
100247708f31 clarified comments;
wenzelm
parents: 62042
diff changeset
    33
  \<^cancel>\<open>apply(rename_tac ps b qs)\<close>
100247708f31 clarified comments;
wenzelm
parents: 62042
diff changeset
    34
  \<^cancel>\<open>apply clarsimp\<close>
100247708f31 clarified comments;
wenzelm
parents: 62042
diff changeset
    35
  \<^cancel>\<open>apply(rename_tac ps')\<close>
100247708f31 clarified comments;
wenzelm
parents: 62042
diff changeset
    36
  \<^cancel>\<open>apply(fastforce intro:notin_List_update[THEN iffD2])\<close>
100247708f31 clarified comments;
wenzelm
parents: 62042
diff changeset
    37
  done
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    38
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
    39
text\<open>And now with ghost variables \<^term>\<open>ps\<close> and \<^term>\<open>qs\<close>. Even
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
    40
``more automatic''.\<close>
14062
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    41
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    42
lemma "VARS next p ps q qs r
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    43
  {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
    44
   ps = Ps \<and> qs = Qs}
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    45
  WHILE p \<noteq> Null
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    46
  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
    47
       rev ps @ qs = rev Ps @ Qs}
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    48
  DO r := p; p := p^.next; r^.next := q; q := r;
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    49
     qs := (hd ps) # qs; ps := tl ps OD
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    50
  {List next q (rev Ps @ Qs)}"
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    51
apply vcg_simp
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
    52
 apply fastforce
14062
7f0d5cc52615 *** empty log message ***
nipkow
parents: 13875
diff changeset
    53
done
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    54
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    55
text "A longer readable version:"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    56
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    57
lemma "VARS tl p q r
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    58
  {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
    59
  WHILE p \<noteq> Null
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    60
  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
    61
               rev ps @ qs = rev Ps @ Qs}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    62
  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
    63
  {List tl q (rev Ps @ Qs)}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    64
proof vcg
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    65
  fix tl p q r
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    66
  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
    67
  thus "\<exists>ps qs. List tl p ps \<and> List tl q qs \<and> set ps \<inter> set qs = {} \<and>
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
    68
                rev ps @ qs = rev Ps @ Qs" by fastforce
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    69
next
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    70
  fix tl p q r
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    71
  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
    72
                   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
    73
         (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
    74
  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
    75
    by fast
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
    76
  then obtain ps' where "ps = a # ps'" by fastforce
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    77
  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
    78
         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
    79
         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
    80
         rev ps' @ (a#qs) = rev Ps @ Qs"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
    81
    using I by fastforce
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    82
  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
    83
                  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
    84
                  set ps' \<inter> set qs' = {} \<and>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    85
                  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
    86
next
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    87
  fix tl p q r
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    88
  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
    89
                   rev ps @ qs = rev Ps @ Qs) \<and> \<not> p \<noteq> Null"
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
    90
  thus "List tl q (rev Ps @ Qs)" by fastforce
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    91
qed
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    92
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    93
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
    94
text\<open>Finaly, the functional version. A bit more verbose, but automatic!\<close>
13772
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
lemma "VARS tl p q r
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
    97
  {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
    98
   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
    99
  WHILE p \<noteq> Null
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   100
  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
   101
       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
   102
       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
   103
  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
   104
  {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
   105
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   106
  apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   107
 apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   108
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   109
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   110
72990
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
   111
subsubsection "Searching in a list"
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   112
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   113
text\<open>What follows is a sequence of successively more intelligent proofs that
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   114
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
   115
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   116
We start with a proof based on the \<^term>\<open>List\<close> predicate. This means it only
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   117
works for acyclic lists.\<close>
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   118
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   119
lemma "VARS tl p
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   120
  {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
   121
  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
   122
  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
   123
  DO p := p^.tl OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   124
  {p = Ref X}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   125
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   126
  apply blast
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   127
 apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   128
apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   129
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   130
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   131
text\<open>Using \<^term>\<open>Path\<close> instead of \<^term>\<open>List\<close> generalizes the correctness
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   132
statement to cyclic lists as well:\<close>
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   133
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   134
lemma "VARS tl p
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   135
  {Path tl p Ps X}
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   136
  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
   137
  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
   138
  DO p := p^.tl OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   139
  {p = X}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   140
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   141
  apply blast
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
   142
 apply fastforce
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   143
apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   144
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   145
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   146
text\<open>Now it dawns on us that we do not need the list witness at all --- it
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   147
suffices to talk about reachability, i.e.\ we can use relations directly. The
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   148
first version uses a relation on \<^typ>\<open>'a ref\<close>:\<close>
13772
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
lemma "VARS tl p
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67444
diff changeset
   151
  {(p,X) \<in> {(Ref x,tl x) |x. True}\<^sup>*}
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   152
  WHILE p \<noteq> Null \<and> p \<noteq> X
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67444
diff changeset
   153
  INV {(p,X) \<in> {(Ref x,tl x) |x. True}\<^sup>*}
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   154
  DO p := p^.tl OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   155
  {p = X}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   156
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   157
 apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   158
 apply(erule converse_rtranclE)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   159
  apply simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   160
 apply(clarsimp elim:converse_rtranclE)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   161
apply(fast elim:converse_rtranclE)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   162
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   163
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   164
text\<open>Finally, a version based on a relation on type \<^typ>\<open>'a\<close>:\<close>
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   165
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   166
lemma "VARS tl p
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67444
diff changeset
   167
  {p \<noteq> Null \<and> (addr p,X) \<in> {(x,y). tl x = Ref y}\<^sup>*}
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   168
  WHILE p \<noteq> Null \<and> p \<noteq> Ref X
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67444
diff changeset
   169
  INV {p \<noteq> Null \<and> (addr p,X) \<in> {(x,y). tl x = Ref y}\<^sup>*}
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   170
  DO p := p^.tl OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   171
  {p = Ref X}"
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   172
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   173
 apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   174
 apply(erule converse_rtranclE)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   175
  apply simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   176
 apply clarsimp
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
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   179
72990
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
   180
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
   181
subsubsection "Splicing two lists"
19397
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   182
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   183
lemma "VARS tl p q pp qq
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   184
  {List tl p Ps \<and> List tl q Qs \<and> set Ps \<inter> set Qs = {} \<and> size Qs \<le> size Ps}
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   185
  pp := p;
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   186
  WHILE q \<noteq> Null
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   187
  INV {\<exists>as bs qs.
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   188
    distinct as \<and> Path tl p as pp \<and> List tl pp bs \<and> List tl q qs \<and>
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   189
    set bs \<inter> set qs = {} \<and> set as \<inter> (set bs \<union> set qs) = {} \<and>
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   190
    size qs \<le> size bs \<and> splice Ps Qs = as @ splice bs qs}
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   191
  DO qq := q^.tl; q^.tl := pp^.tl; pp^.tl := q; pp := q^.tl; q := qq OD
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   192
  {List tl p (splice Ps Qs)}"
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   193
apply vcg_simp
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   194
  apply(rule_tac x = "[]" in exI)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
   195
  apply fastforce
19397
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   196
 apply clarsimp
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   197
 apply(rename_tac y bs qqs)
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   198
 apply(case_tac bs) apply simp
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   199
 apply clarsimp
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   200
 apply(rename_tac x bbs)
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   201
 apply(rule_tac x = "as @ [x,y]" in exI)
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   202
 apply simp
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   203
 apply(rule_tac x = "bbs" in exI)
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   204
 apply simp
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   205
 apply(rule_tac x = "qqs" in exI)
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   206
 apply simp
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
   207
apply (fastforce simp:List_app)
19397
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   208
done
524f1cb4652a Added splicing algorithm.
nipkow
parents: 16417
diff changeset
   209
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   210
72990
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
   211
subsubsection "Merging two lists"
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
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
   214
38353
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   215
definition cor :: "bool \<Rightarrow> bool \<Rightarrow> bool"
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   216
  where "cor P Q \<longleftrightarrow> (if P then True else Q)"
35416
d8d7d1b785af replaced a couple of constsdefs by definitions (also some old primrecs by modern ones)
haftmann
parents: 35316
diff changeset
   217
38353
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   218
definition cand :: "bool \<Rightarrow> bool \<Rightarrow> bool"
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   219
  where "cand P Q \<longleftrightarrow> (if P then Q else False)"
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   220
38353
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   221
fun merge :: "'a list * 'a list * ('a \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a list"
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   222
where
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   223
  "merge(x#xs,y#ys,f) = (if f x y then x # merge(xs,y#ys,f)
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   224
                                  else y # merge(x#xs,ys,f))"
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   225
| "merge(x#xs,[],f) = x # merge(xs,[],f)"
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   226
| "merge([],y#ys,f) = y # merge([],ys,f)"
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   227
| "merge([],[],f) = []"
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   228
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   229
text\<open>Simplifies the proof a little:\<close>
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   230
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   231
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
   232
by blast
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   233
lemma [simp]: "({} = A \<inter> insert b B) = (b \<notin> A & {} = A \<inter> B)"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   234
by blast
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   235
lemma [simp]: "({} = A \<inter> (B \<union> C)) = ({} = A \<inter> B & {} = A \<inter> C)"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   236
by blast
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   237
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   238
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
   239
 {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
   240
  (p \<noteq> Null \<or> q \<noteq> Null)}
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   241
 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
   242
 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
   243
 s := r;
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   244
 WHILE p \<noteq> Null \<or> q \<noteq> Null
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67444
diff changeset
   245
 INV {\<exists>rs ps qs a. Path tl r rs s \<and> List tl p ps \<and> List tl q qs \<and>
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   246
      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
   247
      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
   248
      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
   249
      (tl a = p \<or> tl a = q)}
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   250
 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
   251
    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
   252
    s := s^.tl
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   253
 OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   254
 {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
   255
apply vcg_simp
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   256
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
   257
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
   258
apply (fastforce)
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   259
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   260
apply clarsimp
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   261
apply(rule conjI)
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   262
apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   263
apply(rule conjI)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
   264
apply (fastforce intro!:Path_snoc intro:Path_upd[THEN iffD2] notin_List_update[THEN iffD2] simp:eq_sym_conv)
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   265
apply clarsimp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   266
apply(rule conjI)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   267
apply (clarsimp)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   268
apply(rule_tac x = "rs @ [a]" in exI)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   269
apply(clarsimp simp:eq_sym_conv)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   270
apply(rule_tac x = "bs" in exI)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   271
apply(clarsimp simp:eq_sym_conv)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   272
apply(rule_tac x = "ya#bsa" in exI)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   273
apply(simp)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   274
apply(clarsimp simp:eq_sym_conv)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   275
apply(rule_tac x = "rs @ [a]" in exI)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   276
apply(clarsimp simp:eq_sym_conv)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   277
apply(rule_tac x = "y#bs" in exI)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   278
apply(clarsimp simp:eq_sym_conv)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   279
apply(rule_tac x = "bsa" in exI)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   280
apply(simp)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
   281
apply (fastforce intro!:Path_snoc intro:Path_upd[THEN iffD2] notin_List_update[THEN iffD2] simp:eq_sym_conv)
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   282
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   283
apply(clarsimp simp add:List_app)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   284
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   285
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   286
text\<open>And now with ghost variables:\<close>
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   287
14074
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   288
lemma "VARS elem next p q r s ps qs rs a
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   289
 {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
   290
  (p \<noteq> Null \<or> q \<noteq> Null) \<and> ps = Ps \<and> qs = Qs}
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   291
 IF cor (q = Null) (cand (p \<noteq> Null) (p^.elem \<le> q^.elem))
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   292
 THEN r := p; p := p^.next; ps := tl ps
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   293
 ELSE r := q; q := q^.next; qs := tl qs FI;
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   294
 s := r; rs := []; a := addr s;
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   295
 WHILE p \<noteq> Null \<or> q \<noteq> Null
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   296
 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
   297
      distinct(a # ps @ qs @ rs) \<and> s = Ref a \<and>
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   298
      merge(Ps,Qs,\<lambda>x y. elem x \<le> elem y) =
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   299
      rs @ a # merge(ps,qs,\<lambda>x y. elem x \<le> elem y) \<and>
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   300
      (next a = p \<or> next a = q)}
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   301
 DO IF cor (q = Null) (cand (p \<noteq> Null) (p^.elem \<le> q^.elem))
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   302
    THEN s^.next := p; p := p^.next; ps := tl ps
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   303
    ELSE s^.next := q; q := q^.next; qs := tl qs FI;
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   304
    rs := rs @ [a]; s := s^.next; a := addr s
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   305
 OD
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   306
 {List next r (merge(Ps,Qs,\<lambda>x y. elem x \<le> elem y))}"
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   307
apply vcg_simp
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   308
apply (simp_all add: cand_def cor_def)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   309
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
   310
apply (fastforce)
14074
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   311
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   312
apply clarsimp
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   313
apply(rule conjI)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   314
apply(clarsimp)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   315
apply(rule conjI)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   316
apply(clarsimp simp:neq_commute)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   317
apply(clarsimp simp:neq_commute)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   318
apply(clarsimp simp:neq_commute)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   319
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   320
apply(clarsimp simp add:List_app)
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   321
done
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   322
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   323
text\<open>The proof is a LOT simpler because it does not need
14074
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   324
instantiations anymore, but it is still not quite automatic, probably
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   325
because of this wrong orientation business.\<close>
14074
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   326
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   327
text\<open>More of the previous proof without ghost variables can be
14074
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   328
automated, but the runtime goes up drastically. In general it is
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   329
usually more efficient to give the witness directly than to have it
93dfce3b6f86 *** empty log message ***
nipkow
parents: 14062
diff changeset
   330
found by proof.
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   331
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   332
Now we try a functional version of the abstraction relation \<^term>\<open>Path\<close>. Since the result is not that convincing, we do not prove any of
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   333
the lemmas.\<close>
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   334
38353
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   335
axiomatization
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   336
  ispath :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a ref \<Rightarrow> bool" and
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   337
  path :: "('a \<Rightarrow> 'a ref) \<Rightarrow> 'a ref \<Rightarrow> 'a ref \<Rightarrow> 'a list"
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   338
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   339
text"First some basic lemmas:"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   340
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   341
lemma [simp]: "ispath f p p"
24499
5a3ee202e0b0 do not touch quick_and_dirty;
wenzelm
parents: 19399
diff changeset
   342
by (rule unproven)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   343
lemma [simp]: "path f p p = []"
24499
5a3ee202e0b0 do not touch quick_and_dirty;
wenzelm
parents: 19399
diff changeset
   344
by (rule unproven)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   345
lemma [simp]: "ispath f p q \<Longrightarrow> a \<notin> set(path f p q) \<Longrightarrow> ispath (f(a := r)) p q"
24499
5a3ee202e0b0 do not touch quick_and_dirty;
wenzelm
parents: 19399
diff changeset
   346
by (rule unproven)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   347
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
   348
 path (f(a := r)) p q = path f p q"
24499
5a3ee202e0b0 do not touch quick_and_dirty;
wenzelm
parents: 19399
diff changeset
   349
by (rule unproven)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   350
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   351
text"Some more specific lemmas needed by the example:"
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   352
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   353
lemma [simp]: "ispath (f(a := q)) p (Ref a) \<Longrightarrow> ispath (f(a := q)) p q"
24499
5a3ee202e0b0 do not touch quick_and_dirty;
wenzelm
parents: 19399
diff changeset
   354
by (rule unproven)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   355
lemma [simp]: "ispath (f(a := q)) p (Ref a) \<Longrightarrow>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   356
 path (f(a := q)) p q = path (f(a := q)) p (Ref a) @ [a]"
24499
5a3ee202e0b0 do not touch quick_and_dirty;
wenzelm
parents: 19399
diff changeset
   357
by (rule unproven)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   358
lemma [simp]: "ispath f p (Ref a) \<Longrightarrow> f a = Ref b \<Longrightarrow>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   359
 b \<notin> set (path f p (Ref a))"
24499
5a3ee202e0b0 do not touch quick_and_dirty;
wenzelm
parents: 19399
diff changeset
   360
by (rule unproven)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   361
lemma [simp]: "ispath f p (Ref a) \<Longrightarrow> f a = Null \<Longrightarrow> islist f p"
24499
5a3ee202e0b0 do not touch quick_and_dirty;
wenzelm
parents: 19399
diff changeset
   362
by (rule unproven)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   363
lemma [simp]: "ispath f p (Ref a) \<Longrightarrow> f a = Null \<Longrightarrow> list f p = path f p (Ref a) @ [a]"
24499
5a3ee202e0b0 do not touch quick_and_dirty;
wenzelm
parents: 19399
diff changeset
   364
by (rule unproven)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   365
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   366
lemma [simp]: "islist f p \<Longrightarrow> distinct (list f p)"
24499
5a3ee202e0b0 do not touch quick_and_dirty;
wenzelm
parents: 19399
diff changeset
   367
by (rule unproven)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   368
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   369
lemma "VARS hd tl p q r s
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67444
diff changeset
   370
 {islist tl p \<and> Ps = list tl p \<and> islist tl q \<and> Qs = list tl q \<and>
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   371
  set Ps \<inter> set Qs = {} \<and>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   372
  (p \<noteq> Null \<or> q \<noteq> Null)}
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   373
 IF cor (q = Null) (cand (p \<noteq> Null) (p^.hd \<le> q^.hd))
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   374
 THEN r := p; p := p^.tl ELSE r := q; q := q^.tl FI;
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   375
 s := r;
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   376
 WHILE p \<noteq> Null \<or> q \<noteq> Null
67613
ce654b0e6d69 more symbols;
wenzelm
parents: 67444
diff changeset
   377
 INV {\<exists>rs ps qs a. ispath tl r s \<and> rs = path tl r s \<and>
ce654b0e6d69 more symbols;
wenzelm
parents: 67444
diff changeset
   378
      islist tl p \<and> ps = list tl p \<and> islist tl q \<and> qs = list tl q \<and>
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   379
      distinct(a # ps @ qs @ rs) \<and> s = Ref a \<and>
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   380
      merge(Ps,Qs,\<lambda>x y. hd x \<le> hd y) =
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   381
      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
   382
      (tl a = p \<or> tl a = q)}
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   383
 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
   384
    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
   385
    s := s^.tl
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   386
 OD
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   387
 {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
   388
apply vcg_simp
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   389
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   390
apply (simp_all add: cand_def cor_def)
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
   391
  apply (fastforce)
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
   392
 apply (fastforce simp: eq_sym_conv)
13773
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   393
apply(clarsimp)
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   394
done
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   395
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   396
text"The proof is automatic, but requires a numbet of special lemmas."
58dc4ab362d0 new versions of merge-example
nipkow
parents: 13772
diff changeset
   397
19399
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   398
72990
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
   399
subsubsection "Cyclic list reversal"
19399
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   400
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   401
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   402
text\<open>We consider two algorithms for the reversal of circular lists.
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   403
\<close>
19399
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   404
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   405
lemma circular_list_rev_I:
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   406
  "VARS next root p q tmp
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   407
  {root = Ref r \<and> distPath next root (r#Ps) root}
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   408
  p := root; q := root^.next;
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   409
  WHILE q \<noteq> root
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   410
  INV {\<exists> ps qs. distPath next p ps root \<and> distPath next q qs root \<and> 
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   411
             root = Ref r \<and> r \<notin> set Ps  \<and> set ps \<inter> set qs = {} \<and> 
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   412
             Ps = (rev ps) @ qs  }
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   413
  DO tmp := q; q := q^.next; tmp^.next := p; p:=tmp OD;
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   414
  root^.next := p
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   415
  { root = Ref r \<and> distPath next root (r#rev Ps) root}"
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   416
apply (simp only:distPath_def)
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   417
apply vcg_simp
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   418
  apply (rule_tac x="[]" in exI)
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   419
  apply auto
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   420
 apply (drule (2) neq_dP)
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   421
 apply clarsimp
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   422
 apply(rule_tac x="a # ps" in exI)
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   423
apply clarsimp
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   424
done
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   425
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   426
text\<open>In the beginning, we are able to assert \<^term>\<open>distPath next
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   427
root as root\<close>, with \<^term>\<open>as\<close> set to \<^term>\<open>[]\<close> or
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   428
\<^term>\<open>[r,a,b,c]\<close>. Note that \<^term>\<open>Path next root as root\<close> would
19399
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   429
additionally give us an infinite number of lists with the recurring
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   430
sequence \<^term>\<open>[r,a,b,c]\<close>.
19399
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   431
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   432
The precondition states that there exists a non-empty non-repeating
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   433
path \mbox{\<^term>\<open>r # Ps\<close>} from pointer \<^term>\<open>root\<close> to itself, given that
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   434
\<^term>\<open>root\<close> points to location \<^term>\<open>r\<close>. Pointers \<^term>\<open>p\<close> and
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   435
\<^term>\<open>q\<close> are then set to \<^term>\<open>root\<close> and the successor of \<^term>\<open>root\<close> respectively. If \<^term>\<open>q = root\<close>, we have circled the loop,
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   436
otherwise we set the \<^term>\<open>next\<close> pointer field of \<^term>\<open>q\<close> to point
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   437
to \<^term>\<open>p\<close>, and shift \<^term>\<open>p\<close> and \<^term>\<open>q\<close> one step forward. The
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   438
invariant thus states that \<^term>\<open>p\<close> and \<^term>\<open>q\<close> point to two
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   439
disjoint lists \<^term>\<open>ps\<close> and \<^term>\<open>qs\<close>, such that \<^term>\<open>Ps = rev ps
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   440
@ qs\<close>. After the loop terminates, one
19399
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   441
extra step is needed to close the loop. As expected, the postcondition
69597
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   442
states that the \<^term>\<open>distPath\<close> from \<^term>\<open>root\<close> to itself is now
ff784d5a5bfb isabelle update -u control_cartouches;
wenzelm
parents: 67613
diff changeset
   443
\<^term>\<open>r # (rev Ps)\<close>.
19399
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   444
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   445
It may come as a surprise to the reader that the simple algorithm for
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   446
acyclic list reversal, with modified annotations, works for cyclic
62042
6c6ccf573479 isabelle update_cartouches -c -t;
wenzelm
parents: 44890
diff changeset
   447
lists as well:\<close>
19399
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   448
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   449
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   450
lemma circular_list_rev_II:
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   451
"VARS next p q tmp
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   452
{p = Ref r \<and> distPath next p (r#Ps) p}
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   453
q:=Null;
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   454
WHILE p \<noteq> Null
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   455
INV
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   456
{ ((q = Null) \<longrightarrow> (\<exists>ps. distPath next p (ps) (Ref r) \<and> ps = r#Ps)) \<and>
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   457
  ((q \<noteq> Null) \<longrightarrow> (\<exists>ps qs. distPath next q (qs) (Ref r) \<and> List next p ps  \<and>
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   458
                   set ps \<inter> set qs = {} \<and> rev qs @ ps = Ps@[r])) \<and>
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   459
  \<not> (p = Null \<and> q = Null) }
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   460
DO tmp := p; p := p^.next; tmp^.next := q; q:=tmp OD
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   461
{q = Ref r \<and> distPath next q (r # rev Ps) q}"
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   462
apply (simp only:distPath_def)
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   463
apply vcg_simp
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   464
  apply clarsimp
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   465
 apply (case_tac "(q = Null)")
44890
22f665a2e91c new fastforce replacing fastsimp - less confusing name
nipkow
parents: 41959
diff changeset
   466
  apply (fastforce intro: Path_is_List)
19399
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   467
 apply clarsimp
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   468
 apply (rule_tac x= "bs" in exI)
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   469
 apply (rule_tac x= "y # qs" in exI)
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   470
 apply clarsimp
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   471
apply (auto simp:fun_upd_apply)
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   472
done
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   473
fd2ba98056a2 Included cyclic list examples
nipkow
parents: 19397
diff changeset
   474
72990
db8f94656024 tuned document, notably authors and sections;
wenzelm
parents: 71989
diff changeset
   475
subsubsection "Storage allocation"
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   476
38353
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   477
definition new :: "'a set \<Rightarrow> 'a"
d98baa2cf589 modernized specifications;
wenzelm
parents: 35419
diff changeset
   478
  where "new A = (SOME a. a \<notin> A)"
13772
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   479
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   480
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   481
lemma new_notin:
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   482
 "\<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
   483
apply(unfold new_def)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   484
apply(rule someI2_ex)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   485
 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
   486
apply (fast)
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   487
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   488
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   489
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   490
lemma "~finite(UNIV::'a set) \<Longrightarrow>
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   491
  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
   492
  {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
   493
  WHILE xs \<noteq> []
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   494
  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
   495
       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
   496
  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
   497
     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
   498
  OD
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   499
  {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
   500
apply vcg_simp
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   501
 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
   502
done
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   503
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   504
73d041cc6a66 Split Pointers.thy and automated one proof, which caused the runtime to explode
nipkow
parents:
diff changeset
   505
end