(* * Copyright 2020, Data61, CSIRO (ABN 41 687 119 230) * * SPDX-License-Identifier: BSD-2-Clause *) (* Title: Adaptation of example from HOL/Hoare/Separation Author: Gerwin Klein, 2012 Maintainers: Gerwin Klein Rafal Kolanski *) chapter "Example from HOL/Hoare/Separation" theory Simple_Separation_Example imports "HOL-Hoare.Hoare_Logic_Abort" Sep_Heap_Instance Sep_Tactics begin declare [[syntax_ambiguity_warning = false]] type_synonym heap = "(nat \ nat option)" (* This syntax isn't ideal, but this is what is used in the original *) definition maps_to:: "nat \ nat \ heap \ bool" ("_ \ _" [56,51] 56) where "x \ y \ \h. h = [x \ y]" (* If you don't mind syntax ambiguity: *) notation pred_ex (binder "\" 10) (* could be generated automatically *) definition maps_to_ex :: "nat \ heap \ bool" ("_ \ -" [56] 56) where "x \ - \ \y. x \ y" (* could be generated automatically *) lemma maps_to_maps_to_ex [elim!]: "(p \ v) s \ (p \ -) s" by (auto simp: maps_to_ex_def) (* The basic properties of maps_to: *) lemma maps_to_write: "(p \ - ** P) H \ (p \ v ** P) (H (p \ v))" apply (clarsimp simp: sep_conj_def maps_to_def maps_to_ex_def split: option.splits) apply (rule_tac x=y in exI) apply (auto simp: sep_disj_fun_def map_convs map_add_def split: option.splits) done lemma points_to: "(p \ v ** P) H \ the (H p) = v" by (auto elim!: sep_conjE simp: sep_disj_fun_def maps_to_def map_convs map_add_def split: option.splits) (* This differs from the original and uses separation logic for the definition. *) primrec list :: "nat \ nat list \ heap \ bool" where "list i [] = (\i=0\ and \)" | "list i (x#xs) = (\i=x \ i\0\ and (EXS j. i \ j ** list j xs))" lemma list_empty [simp]: shows "list 0 xs = (\s. xs = [] \ \ s)" by (cases xs) auto (* The examples from Hoare/Separation.thy *) lemma "VARS x y z w h {(x \ y ** z \ w) h} SKIP {x \ z}" apply vcg apply(auto elim!: sep_conjE simp: maps_to_def sep_disj_fun_def domain_conv) done lemma "VARS H x y z w {(P ** Q) H} SKIP {(Q ** P) H}" apply vcg apply(simp add: sep_conj_commute) done lemma "VARS H {p\0 \ (p \ - ** list q qs) H} H := H(p \ q) {list p (p#qs) H}" apply vcg apply (auto intro: maps_to_write) done text \First a proof without using tactics\ lemma "VARS H p q r {(list p Ps ** list q Qs) H} WHILE p \ 0 INV {\ps qs. (list p ps ** list q qs) H \ rev ps @ qs = rev Ps @ Qs} DO r := p; p := the(H p); H := H(r \ q); q := r OD {list q (rev Ps @ Qs) H}" apply vcg apply fastforce apply clarsimp apply (case_tac ps, simp) apply (rename_tac p ps') apply (rule_tac x = "ps'" in exI) apply (rule_tac x = "p # qs" in exI) apply clarsimp apply (clarsimp simp: sep_conj_exists sep_conj_ac) apply (rule_tac x = "q" in exI) thm points_to apply (subst (asm) sep_conj_ac) apply (subst (asm) sep_conj_commute) apply (simp add: sep_conj_assoc) apply (frule points_to) apply simp thm maps_to_write apply (subst sep_conj_assoc[symmetric]) apply (subst sep_conj_commute) apply (simp add: sep_conj_assoc) apply (rule maps_to_write) thm maps_to_maps_to_ex apply (drule sep_conj_impl1[rotated]) apply (erule maps_to_maps_to_ex) apply (rule sep_conj_impl1[rotated]) apply (subst sep_conj_ac, assumption) apply assumption apply clarsimp done text \Now with tactics\ lemma "VARS H p q r {(list p Ps ** list q Qs) H} WHILE p \ 0 INV {\ps qs. (list p ps ** list q qs) H \ rev ps @ qs = rev Ps @ Qs} DO r := p; p := the(H p); H := H(r \ q); q := r OD {list q (rev Ps @ Qs) H}" apply vcg apply fastforce apply clarsimp apply (case_tac ps, simp) apply (rename_tac p ps') apply (clarsimp simp: sep_conj_exists sep_conj_ac) apply (subst points_to, sep_solve) apply (rule_tac x = "ps'" in exI) apply (rule_tac x = "p # qs" in exI) apply (simp add: sep_conj_exists sep_conj_ac) apply (rule exI) apply (sep_rule (direct) maps_to_write) \ \note: demonstrates computation\ apply (sep_solve add: maps_to_maps_to_ex) apply clarsimp done end