| 
25
 | 
     1  | 
(* Authors: Gerwin Klein and Rafal Kolanski, 2012
  | 
| 
 | 
     2  | 
   Maintainers: Gerwin Klein <kleing at cse.unsw.edu.au>
  | 
| 
 | 
     3  | 
                Rafal Kolanski <rafal.kolanski at nicta.com.au>
  | 
| 
 | 
     4  | 
*)
  | 
| 
 | 
     5  | 
  | 
| 
 | 
     6  | 
header "Abstract Separation Algebra"
  | 
| 
 | 
     7  | 
  | 
| 
 | 
     8  | 
theory Separation_Algebra
  | 
| 
 | 
     9  | 
imports Main
  | 
| 
 | 
    10  | 
begin
  | 
| 
 | 
    11  | 
  | 
| 
 | 
    12  | 
  | 
| 
 | 
    13  | 
text {* This theory is the main abstract separation algebra development *}
 | 
| 
 | 
    14  | 
  | 
| 
 | 
    15  | 
  | 
| 
 | 
    16  | 
section {* Input syntax for lifting boolean predicates to separation predicates *}
 | 
| 
 | 
    17  | 
  | 
| 
 | 
    18  | 
abbreviation (input)
  | 
| 
 | 
    19  | 
  pred_and :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" (infixr "and" 35) where
 | 
| 
 | 
    20  | 
  "a and b \<equiv> \<lambda>s. a s \<and> b s"
  | 
| 
 | 
    21  | 
  | 
| 
 | 
    22  | 
abbreviation (input)
  | 
| 
 | 
    23  | 
  pred_or :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" (infixr "or" 30) where
 | 
| 
 | 
    24  | 
  "a or b \<equiv> \<lambda>s. a s \<or> b s"
  | 
| 
 | 
    25  | 
  | 
| 
 | 
    26  | 
abbreviation (input)
  | 
| 
 | 
    27  | 
  pred_not :: "('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" ("not _" [40] 40) where
 | 
| 
 | 
    28  | 
  "not a \<equiv> \<lambda>s. \<not>a s"
  | 
| 
 | 
    29  | 
  | 
| 
 | 
    30  | 
abbreviation (input)
  | 
| 
 | 
    31  | 
  pred_imp :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" (infixr "imp" 25) where
 | 
| 
 | 
    32  | 
  "a imp b \<equiv> \<lambda>s. a s \<longrightarrow> b s"
  | 
| 
 | 
    33  | 
  | 
| 
 | 
    34  | 
abbreviation (input)
  | 
| 
 | 
    35  | 
  pred_K :: "'b \<Rightarrow> 'a \<Rightarrow> 'b" ("\<langle>_\<rangle>") where
 | 
| 
 | 
    36  | 
  "\<langle>f\<rangle> \<equiv> \<lambda>s. f"
  | 
| 
 | 
    37  | 
  | 
| 
 | 
    38  | 
(* abbreviation *)
  | 
| 
 | 
    39  | 
  definition pred_ex :: "('b \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" (binder "EXS " 10) where
 | 
| 
 | 
    40  | 
  "EXS x. P x \<equiv> \<lambda>s. \<exists>x. P x s"
  | 
| 
 | 
    41  | 
  | 
| 
 | 
    42  | 
abbreviation (input)
  | 
| 
 | 
    43  | 
  pred_all :: "('b \<Rightarrow> 'a \<Rightarrow> bool) \<Rightarrow> 'a \<Rightarrow> bool" (binder "ALLS " 10) where
 | 
| 
 | 
    44  | 
  "ALLS x. P x \<equiv> \<lambda>s. \<forall>x. P x s"
  | 
| 
 | 
    45  | 
  | 
| 
 | 
    46  | 
  | 
| 
 | 
    47  | 
section {* Associative/Commutative Monoid Basis of Separation Algebras *}
 | 
| 
 | 
    48  | 
  | 
| 
 | 
    49  | 
class pre_sep_algebra = zero + plus +
  | 
| 
 | 
    50  | 
  fixes sep_disj :: "'a => 'a => bool" (infix "##" 60)
  | 
| 
 | 
    51  | 
  | 
| 
 | 
    52  | 
  assumes sep_disj_zero [simp]: "x ## 0"
  | 
| 
 | 
    53  | 
  assumes sep_disj_commuteI: "x ## y \<Longrightarrow> y ## x"
  | 
| 
 | 
    54  | 
  | 
| 
 | 
    55  | 
  assumes sep_add_zero [simp]: "x + 0 = x"
  | 
| 
 | 
    56  | 
  assumes sep_add_commute: "x ## y \<Longrightarrow> x + y = y + x"
  | 
| 
 | 
    57  | 
  | 
| 
 | 
    58  | 
  assumes sep_add_assoc:
  | 
| 
 | 
    59  | 
    "\<lbrakk> x ## y; y ## z; x ## z \<rbrakk> \<Longrightarrow> (x + y) + z = x + (y + z)"
  | 
| 
 | 
    60  | 
begin
  | 
| 
 | 
    61  | 
  | 
| 
 | 
    62  | 
lemma sep_disj_commute: "x ## y = y ## x"
  | 
| 
 | 
    63  | 
  by (blast intro: sep_disj_commuteI)
  | 
| 
 | 
    64  | 
  | 
| 
 | 
    65  | 
lemma sep_add_left_commute:
  | 
| 
 | 
    66  | 
  assumes a: "a ## b" "b ## c" "a ## c"
  | 
| 
 | 
    67  | 
  shows "b + (a + c) = a + (b + c)" (is "?lhs = ?rhs")
  | 
| 
 | 
    68  | 
proof -
  | 
| 
 | 
    69  | 
  have "?lhs = b + a + c" using a
  | 
| 
 | 
    70  | 
    by (simp add: sep_add_assoc[symmetric] sep_disj_commute)
  | 
| 
 | 
    71  | 
  also have "... = a + b + c" using a
  | 
| 
 | 
    72  | 
    by (simp add: sep_add_commute sep_disj_commute)
  | 
| 
 | 
    73  | 
  also have "... = ?rhs" using a
  | 
| 
 | 
    74  | 
    by (simp add: sep_add_assoc sep_disj_commute)
  | 
| 
 | 
    75  | 
  finally show ?thesis .
  | 
| 
 | 
    76  | 
qed
  | 
| 
 | 
    77  | 
  | 
| 
 | 
    78  | 
lemmas sep_add_ac = sep_add_assoc sep_add_commute sep_add_left_commute
  | 
| 
 | 
    79  | 
                    sep_disj_commute (* nearly always necessary *)
  | 
| 
 | 
    80  | 
  | 
| 
 | 
    81  | 
end
  | 
| 
 | 
    82  | 
  | 
| 
 | 
    83  | 
  | 
| 
 | 
    84  | 
section {* Separation Algebra as Defined by Calcagno et al. *}
 | 
| 
 | 
    85  | 
  | 
| 
 | 
    86  | 
class sep_algebra = pre_sep_algebra +
  | 
| 
 | 
    87  | 
  assumes sep_disj_addD1: "\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x ## y"
  | 
| 
 | 
    88  | 
  assumes sep_disj_addI1: "\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x + y ##  z"
  | 
| 
 | 
    89  | 
begin
  | 
| 
 | 
    90  | 
  | 
| 
 | 
    91  | 
subsection {* Basic Construct Definitions and Abbreviations *}
 | 
| 
 | 
    92  | 
  | 
| 
 | 
    93  | 
definition
  | 
| 
 | 
    94  | 
  sep_conj :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool)" (infixr "**" 35)
 | 
| 
 | 
    95  | 
  where
  | 
| 
 | 
    96  | 
  "P ** Q \<equiv> \<lambda>h. \<exists>x y. x ## y \<and> h = x + y \<and> P x \<and> Q y"
  | 
| 
 | 
    97  | 
  | 
| 
 | 
    98  | 
notation
  | 
| 
 | 
    99  | 
  sep_conj (infixr "\<and>*" 35)
  | 
| 
 | 
   100  | 
  | 
| 
 | 
   101  | 
definition
  | 
| 
 | 
   102  | 
  sep_empty :: "'a \<Rightarrow> bool" ("\<box>") where
 | 
| 
 | 
   103  | 
  "\<box> \<equiv> \<lambda>h. h = 0"
  | 
| 
 | 
   104  | 
  | 
| 
 | 
   105  | 
definition
  | 
| 
 | 
   106  | 
  sep_impl :: "('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool) \<Rightarrow> ('a \<Rightarrow> bool)" (infixr "\<longrightarrow>*" 25)
 | 
| 
 | 
   107  | 
  where
  | 
| 
 | 
   108  | 
  "P \<longrightarrow>* Q \<equiv> \<lambda>h. \<forall>h'. h ## h' \<and> P h' \<longrightarrow> Q (h + h')"
  | 
| 
 | 
   109  | 
  | 
| 
 | 
   110  | 
definition
  | 
| 
 | 
   111  | 
  sep_substate :: "'a => 'a => bool" (infix "\<preceq>" 60) where
  | 
| 
 | 
   112  | 
  "x \<preceq> y \<equiv> \<exists>z. x ## z \<and> x + z = y"
  | 
| 
 | 
   113  | 
  | 
| 
 | 
   114  | 
(* We want these to be abbreviations not definitions, because basic True and
  | 
| 
 | 
   115  | 
   False will occur by simplification in sep_conj terms *)
  | 
| 
 | 
   116  | 
abbreviation
  | 
| 
 | 
   117  | 
  "sep_true \<equiv> \<langle>True\<rangle>"
  | 
| 
 | 
   118  | 
  | 
| 
 | 
   119  | 
abbreviation
  | 
| 
 | 
   120  | 
  "sep_false \<equiv> \<langle>False\<rangle>"
  | 
| 
 | 
   121  | 
  | 
| 
 | 
   122  | 
definition
  | 
| 
 | 
   123  | 
  sep_list_conj :: "('a \<Rightarrow> bool) list \<Rightarrow> ('a \<Rightarrow> bool)"  ("\<And>* _" [60] 90) where
 | 
| 
 | 
   124  | 
  "sep_list_conj Ps \<equiv> foldl (op **) \<box> Ps"
  | 
| 
 | 
   125  | 
  | 
| 
 | 
   126  | 
  | 
| 
 | 
   127  | 
subsection {* Disjunction/Addition Properties *}
 | 
| 
 | 
   128  | 
  | 
| 
 | 
   129  | 
lemma disjoint_zero_sym [simp]: "0 ## x"
  | 
| 
 | 
   130  | 
  by (simp add: sep_disj_commute)
  | 
| 
 | 
   131  | 
  | 
| 
 | 
   132  | 
lemma sep_add_zero_sym [simp]: "0 + x = x"
  | 
| 
 | 
   133  | 
  by (simp add: sep_add_commute)
  | 
| 
 | 
   134  | 
  | 
| 
 | 
   135  | 
lemma sep_disj_addD2: "\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x ## z"
  | 
| 
 | 
   136  | 
  by (metis sep_disj_addD1 sep_add_ac)
  | 
| 
 | 
   137  | 
  | 
| 
 | 
   138  | 
lemma sep_disj_addD: "\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x ## y \<and> x ## z"
  | 
| 
 | 
   139  | 
  by (metis sep_disj_addD1 sep_disj_addD2)
  | 
| 
 | 
   140  | 
  | 
| 
 | 
   141  | 
lemma sep_add_disjD: "\<lbrakk> x + y ## z; x ## y \<rbrakk> \<Longrightarrow> x ## z \<and> y ## z"
  | 
| 
 | 
   142  | 
  by (metis sep_disj_addD sep_disj_commuteI)
  | 
| 
 | 
   143  | 
  | 
| 
 | 
   144  | 
lemma sep_disj_addI2:
  | 
| 
 | 
   145  | 
  "\<lbrakk> x ## y + z; y ## z \<rbrakk> \<Longrightarrow> x + z ## y"
  | 
| 
 | 
   146  | 
  by (metis sep_add_ac sep_disj_addI1)
  | 
| 
 | 
   147  | 
  | 
| 
 | 
   148  | 
lemma sep_add_disjI1:
  | 
| 
 | 
   149  | 
  "\<lbrakk> x + y ## z; x ## y \<rbrakk> \<Longrightarrow> x + z ## y"
  | 
| 
 | 
   150  | 
  by (metis sep_add_ac sep_add_disjD sep_disj_addI2)
  | 
| 
 | 
   151  | 
  | 
| 
 | 
   152  | 
lemma sep_add_disjI2:
  | 
| 
 | 
   153  | 
  "\<lbrakk> x + y ## z; x ## y \<rbrakk> \<Longrightarrow> z + y ## x"
  | 
| 
 | 
   154  | 
  by (metis sep_add_ac sep_add_disjD sep_disj_addI2)
  | 
| 
 | 
   155  | 
  | 
| 
 | 
   156  | 
lemma sep_disj_addI3:
  | 
| 
 | 
   157  | 
   "x + y ## z \<Longrightarrow> x ## y \<Longrightarrow> x ## y + z"
  | 
| 
 | 
   158  | 
   by (metis sep_add_ac sep_add_disjD sep_add_disjI2)
  | 
| 
 | 
   159  | 
  | 
| 
 | 
   160  | 
lemma sep_disj_add:
  | 
| 
 | 
   161  | 
  "\<lbrakk> y ## z; x ## y \<rbrakk> \<Longrightarrow> x ## y + z = x + y ## z"
  | 
| 
 | 
   162  | 
  by (metis sep_disj_addI1 sep_disj_addI3)
  | 
| 
 | 
   163  | 
  | 
| 
 | 
   164  | 
  | 
| 
 | 
   165  | 
subsection {* Substate Properties *}
 | 
| 
 | 
   166  | 
  | 
| 
 | 
   167  | 
lemma sep_substate_disj_add:
  | 
| 
 | 
   168  | 
  "x ## y \<Longrightarrow> x \<preceq> x + y"
  | 
| 
 | 
   169  | 
  unfolding sep_substate_def by blast
  | 
| 
 | 
   170  | 
  | 
| 
 | 
   171  | 
lemma sep_substate_disj_add':
  | 
| 
 | 
   172  | 
  "x ## y \<Longrightarrow> x \<preceq> y + x"
  | 
| 
 | 
   173  | 
  by (simp add: sep_add_ac sep_substate_disj_add)
  | 
| 
 | 
   174  | 
  | 
| 
 | 
   175  | 
  | 
| 
 | 
   176  | 
subsection {* Separating Conjunction Properties *}
 | 
| 
 | 
   177  | 
  | 
| 
 | 
   178  | 
lemma sep_conjD:
  | 
| 
 | 
   179  | 
  "(P \<and>* Q) h \<Longrightarrow> \<exists>x y. x ## y \<and> h = x + y \<and> P x \<and> Q y"
  | 
| 
 | 
   180  | 
  by (simp add: sep_conj_def)
  | 
| 
 | 
   181  | 
  | 
| 
 | 
   182  | 
lemma sep_conjE:
  | 
| 
 | 
   183  | 
  "\<lbrakk> (P ** Q) h; \<And>x y. \<lbrakk> P x; Q y; x ## y; h = x + y \<rbrakk> \<Longrightarrow> X \<rbrakk> \<Longrightarrow> X"
  | 
| 
 | 
   184  | 
  by (auto simp: sep_conj_def)
  | 
| 
 | 
   185  | 
  | 
| 
 | 
   186  | 
lemma sep_conjI:
  | 
| 
 | 
   187  | 
  "\<lbrakk> P x; Q y; x ## y; h = x + y \<rbrakk> \<Longrightarrow> (P ** Q) h"
  | 
| 
 | 
   188  | 
  by (auto simp: sep_conj_def)
  | 
| 
 | 
   189  | 
  | 
| 
 | 
   190  | 
lemma sep_conj_commuteI:
  | 
| 
 | 
   191  | 
  "(P ** Q) h \<Longrightarrow> (Q ** P) h"
  | 
| 
 | 
   192  | 
  by (auto intro!: sep_conjI elim!: sep_conjE simp: sep_add_ac)
  | 
| 
 | 
   193  | 
  | 
| 
 | 
   194  | 
lemma sep_conj_commute:
  | 
| 
 | 
   195  | 
  "(P ** Q) = (Q ** P)"
  | 
| 
 | 
   196  | 
  by (rule ext) (auto intro: sep_conj_commuteI)
  | 
| 
 | 
   197  | 
  | 
| 
 | 
   198  | 
lemma sep_conj_assoc:
  | 
| 
 | 
   199  | 
  "((P ** Q) ** R) = (P ** Q ** R)" (is "?lhs = ?rhs")
  | 
| 
 | 
   200  | 
proof (rule ext, rule iffI)
  | 
| 
 | 
   201  | 
  fix h
  | 
| 
 | 
   202  | 
  assume a: "?lhs h"
  | 
| 
 | 
   203  | 
  then obtain x y z where "P x" and "Q y" and "R z"
  | 
| 
 | 
   204  | 
                      and "x ## y" and "x ## z" and "y ## z" and "x + y ## z"
  | 
| 
 | 
   205  | 
                      and "h = x + y + z"
  | 
| 
 | 
   206  | 
    by (auto dest!: sep_conjD dest: sep_add_disjD)
  | 
| 
 | 
   207  | 
  moreover
  | 
| 
 | 
   208  | 
  then have "x ## y + z"
  | 
| 
 | 
   209  | 
    by (simp add: sep_disj_add)
  | 
| 
 | 
   210  | 
  ultimately
  | 
| 
 | 
   211  | 
  show "?rhs h"
  | 
| 
 | 
   212  | 
    by (auto simp: sep_add_ac intro!: sep_conjI)
  | 
| 
 | 
   213  | 
next
  | 
| 
 | 
   214  | 
  fix h
  | 
| 
 | 
   215  | 
  assume a: "?rhs h"
  | 
| 
 | 
   216  | 
  then obtain x y z where "P x" and "Q y" and "R z"
  | 
| 
 | 
   217  | 
                      and "x ## y" and "x ## z" and "y ## z" and "x ## y + z"
  | 
| 
 | 
   218  | 
                      and "h = x + y + z"
  | 
| 
 | 
   219  | 
    by (fastforce elim!: sep_conjE simp: sep_add_ac dest: sep_disj_addD)
  | 
| 
 | 
   220  | 
  thus "?lhs h"
  | 
| 
 | 
   221  | 
    by (metis sep_conj_def sep_disj_addI1)
  | 
| 
 | 
   222  | 
qed
  | 
| 
 | 
   223  | 
  | 
| 
 | 
   224  | 
lemma sep_conj_impl:
  | 
| 
 | 
   225  | 
  "\<lbrakk> (P ** Q) h; \<And>h. P h \<Longrightarrow> P' h; \<And>h. Q h \<Longrightarrow> Q' h \<rbrakk> \<Longrightarrow> (P' ** Q') h"
  | 
| 
 | 
   226  | 
  by (erule sep_conjE, auto intro!: sep_conjI)
  | 
| 
 | 
   227  | 
  | 
| 
 | 
   228  | 
lemma sep_conj_impl1:
  | 
| 
 | 
   229  | 
  assumes P: "\<And>h. P h \<Longrightarrow> I h"
  | 
| 
 | 
   230  | 
  shows "(P ** R) h \<Longrightarrow> (I ** R) h"
  | 
| 
 | 
   231  | 
  by (auto intro: sep_conj_impl P)
  | 
| 
 | 
   232  | 
  | 
| 
 | 
   233  | 
lemma sep_globalise:
  | 
| 
 | 
   234  | 
  "\<lbrakk> (P ** R) h; (\<And>h. P h \<Longrightarrow> Q h) \<rbrakk> \<Longrightarrow> (Q ** R) h"
  | 
| 
 | 
   235  | 
  by (fast elim: sep_conj_impl)
  | 
| 
 | 
   236  | 
  | 
| 
 | 
   237  | 
lemma sep_conj_trivial_strip2:
  | 
| 
 | 
   238  | 
  "Q = R \<Longrightarrow> (Q ** P) = (R ** P)" by simp
  | 
| 
 | 
   239  | 
  | 
| 
 | 
   240  | 
lemma disjoint_subheaps_exist:
  | 
| 
 | 
   241  | 
  "\<exists>x y. x ## y \<and> h = x + y"
  | 
| 
 | 
   242  | 
  by (rule_tac x=0 in exI, auto)
  | 
| 
 | 
   243  | 
  | 
| 
 | 
   244  | 
lemma sep_conj_left_commute: (* for permutative rewriting *)
  | 
| 
 | 
   245  | 
  "(P ** (Q ** R)) = (Q ** (P ** R))" (is "?x = ?y")
  | 
| 
 | 
   246  | 
proof -
  | 
| 
 | 
   247  | 
  have "?x = ((Q ** R) ** P)" by (simp add: sep_conj_commute)
  | 
| 
 | 
   248  | 
  also have "\<dots> = (Q ** (R ** P))" by (subst sep_conj_assoc, simp)
  | 
| 
 | 
   249  | 
  finally show ?thesis by (simp add: sep_conj_commute)
  | 
| 
 | 
   250  | 
qed
  | 
| 
 | 
   251  | 
  | 
| 
 | 
   252  | 
lemmas sep_conj_ac = sep_conj_commute sep_conj_assoc sep_conj_left_commute
  | 
| 
 | 
   253  | 
  | 
| 
 | 
   254  | 
lemma ab_semigroup_mult_sep_conj: "class.ab_semigroup_mult op **"
  | 
| 
 | 
   255  | 
  by (unfold_locales)
  | 
| 
 | 
   256  | 
     (auto simp: sep_conj_ac)
  | 
| 
 | 
   257  | 
  | 
| 
 | 
   258  | 
lemma sep_empty_zero [simp,intro!]: "\<box> 0"
  | 
| 
 | 
   259  | 
  by (simp add: sep_empty_def)
  | 
| 
 | 
   260  | 
  | 
| 
 | 
   261  | 
  | 
| 
 | 
   262  | 
subsection {* Properties of @{text sep_true} and @{text sep_false} *}
 | 
| 
 | 
   263  | 
  | 
| 
 | 
   264  | 
lemma sep_conj_sep_true:
  | 
| 
 | 
   265  | 
  "P h \<Longrightarrow> (P ** sep_true) h"
  | 
| 
 | 
   266  | 
  by (simp add: sep_conjI[where y=0])
  | 
| 
 | 
   267  | 
  | 
| 
 | 
   268  | 
lemma sep_conj_sep_true':
  | 
| 
 | 
   269  | 
  "P h \<Longrightarrow> (sep_true ** P) h"
  | 
| 
 | 
   270  | 
  by (simp add: sep_conjI[where x=0])
  | 
| 
 | 
   271  | 
  | 
| 
 | 
   272  | 
lemma sep_conj_true [simp]:
  | 
| 
 | 
   273  | 
  "(sep_true ** sep_true) = sep_true"
  | 
| 
 | 
   274  | 
  unfolding sep_conj_def
  | 
| 
 | 
   275  | 
  by (auto intro!: ext intro: disjoint_subheaps_exist)
  | 
| 
 | 
   276  | 
  | 
| 
 | 
   277  | 
lemma sep_conj_false_right [simp]:
  | 
| 
 | 
   278  | 
  "(P ** sep_false) = sep_false"
  | 
| 
 | 
   279  | 
  by (force elim: sep_conjE intro!: ext)
  | 
| 
 | 
   280  | 
  | 
| 
 | 
   281  | 
lemma sep_conj_false_left [simp]:
  | 
| 
 | 
   282  | 
  "(sep_false ** P) = sep_false"
  | 
| 
 | 
   283  | 
  by (subst sep_conj_commute) (rule sep_conj_false_right)
  | 
| 
 | 
   284  | 
  | 
| 
 | 
   285  | 
  | 
| 
 | 
   286  | 
  | 
| 
 | 
   287  | 
subsection {* Properties of zero (@{const sep_empty}) *}
 | 
| 
 | 
   288  | 
  | 
| 
 | 
   289  | 
lemma sep_conj_empty [simp]:
  | 
| 
 | 
   290  | 
  "(P ** \<box>) = P"
  | 
| 
 | 
   291  | 
  by (simp add: sep_conj_def sep_empty_def)
  | 
| 
 | 
   292  | 
  | 
| 
 | 
   293  | 
lemma sep_conj_empty'[simp]:
  | 
| 
 | 
   294  | 
  "(\<box> ** P) = P"
  | 
| 
 | 
   295  | 
  by (subst sep_conj_commute, rule sep_conj_empty)
  | 
| 
 | 
   296  | 
  | 
| 
 | 
   297  | 
lemma sep_conj_sep_emptyI:
  | 
| 
 | 
   298  | 
  "P h \<Longrightarrow> (P ** \<box>) h"
  | 
| 
 | 
   299  | 
  by simp
  | 
| 
 | 
   300  | 
  | 
| 
 | 
   301  | 
lemma sep_conj_sep_emptyE:
  | 
| 
 | 
   302  | 
  "\<lbrakk> P s; (P ** \<box>) s \<Longrightarrow> (Q ** R) s \<rbrakk> \<Longrightarrow> (Q ** R) s"
  | 
| 
 | 
   303  | 
  by simp
  | 
| 
 | 
   304  | 
  | 
| 
 | 
   305  | 
lemma monoid_add: "class.monoid_add (op **) \<box>"
  | 
| 
 | 
   306  | 
  by (unfold_locales) (auto simp: sep_conj_ac)
  | 
| 
 | 
   307  | 
  | 
| 
 | 
   308  | 
lemma comm_monoid_add: "class.comm_monoid_add op ** \<box>"
  | 
| 
 | 
   309  | 
  by (unfold_locales) (auto simp: sep_conj_ac)
  | 
| 
 | 
   310  | 
  | 
| 
 | 
   311  | 
  | 
| 
 | 
   312  | 
subsection {* Properties of top (@{text sep_true}) *}
 | 
| 
 | 
   313  | 
  | 
| 
 | 
   314  | 
lemma sep_conj_true_P [simp]:
  | 
| 
 | 
   315  | 
  "(sep_true ** (sep_true ** P)) = (sep_true ** P)"
  | 
| 
 | 
   316  | 
  by (simp add: sep_conj_assoc[symmetric])
  | 
| 
 | 
   317  | 
  | 
| 
 | 
   318  | 
lemma sep_conj_disj:
  | 
| 
 | 
   319  | 
  "((P or Q) ** R) = ((P ** R) or (Q ** R))"
  | 
| 
 | 
   320  | 
  by (auto simp: sep_conj_def intro!: ext)
  | 
| 
 | 
   321  | 
  | 
| 
 | 
   322  | 
lemma sep_conj_sep_true_left:
  | 
| 
 | 
   323  | 
  "(P ** Q) h \<Longrightarrow> (sep_true ** Q) h"
  | 
| 
 | 
   324  | 
  by (erule sep_conj_impl, simp+)
  | 
| 
 | 
   325  | 
  | 
| 
 | 
   326  | 
lemma sep_conj_sep_true_right:
  | 
| 
 | 
   327  | 
  "(P ** Q) h \<Longrightarrow> (P ** sep_true) h"
  | 
| 
 | 
   328  | 
  by (subst (asm) sep_conj_commute, drule sep_conj_sep_true_left,
  | 
| 
 | 
   329  | 
      simp add: sep_conj_ac)
  | 
| 
 | 
   330  | 
  | 
| 
 | 
   331  | 
  | 
| 
 | 
   332  | 
subsection {* Separating Conjunction with Quantifiers *}
 | 
| 
 | 
   333  | 
  | 
| 
 | 
   334  | 
lemma sep_conj_conj:
  | 
| 
 | 
   335  | 
  "((P and Q) ** R) h \<Longrightarrow> ((P ** R) and (Q ** R)) h"
  | 
| 
 | 
   336  | 
  by (force intro: sep_conjI elim!: sep_conjE)
  | 
| 
 | 
   337  | 
  | 
| 
 | 
   338  | 
lemma sep_conj_exists1:
  | 
| 
 | 
   339  | 
  "((EXS x. P x) ** Q) = (EXS x. (P x ** Q))"
  | 
| 
 | 
   340  | 
  by (unfold pred_ex_def, force intro!: ext intro: sep_conjI elim: sep_conjE)
  | 
| 
 | 
   341  | 
  | 
| 
 | 
   342  | 
lemma sep_conj_exists2:
  | 
| 
 | 
   343  | 
  "(P ** (EXS x. Q x)) = (EXS x. P ** Q x)"
  | 
| 
 | 
   344  | 
  by (unfold pred_ex_def, force intro!: sep_conjI ext elim!: sep_conjE)
  | 
| 
 | 
   345  | 
  | 
| 
 | 
   346  | 
lemmas sep_conj_exists = sep_conj_exists1 sep_conj_exists2
  | 
| 
 | 
   347  | 
  | 
| 
 | 
   348  | 
lemma sep_conj_spec:
  | 
| 
 | 
   349  | 
  "((ALLS x. P x) ** Q) h \<Longrightarrow> (P x ** Q) h"
  | 
| 
 | 
   350  | 
  by (force intro: sep_conjI elim: sep_conjE)
  | 
| 
 | 
   351  | 
  | 
| 
 | 
   352  | 
  | 
| 
 | 
   353  | 
subsection {* Properties of Separating Implication *}
 | 
| 
 | 
   354  | 
  | 
| 
 | 
   355  | 
lemma sep_implI:
  | 
| 
 | 
   356  | 
  assumes a: "\<And>h'. \<lbrakk> h ## h'; P h' \<rbrakk> \<Longrightarrow> Q (h + h')"
  | 
| 
 | 
   357  | 
  shows "(P \<longrightarrow>* Q) h"
  | 
| 
 | 
   358  | 
  unfolding sep_impl_def by (auto elim: a)
  | 
| 
 | 
   359  | 
  | 
| 
 | 
   360  | 
lemma sep_implD:
  | 
| 
 | 
   361  | 
  "(x \<longrightarrow>* y) h \<Longrightarrow> \<forall>h'. h ## h' \<and> x h' \<longrightarrow> y (h + h')"
  | 
| 
 | 
   362  | 
  by (force simp: sep_impl_def)
  | 
| 
 | 
   363  | 
  | 
| 
 | 
   364  | 
lemma sep_implE:
  | 
| 
 | 
   365  | 
  "(x \<longrightarrow>* y) h \<Longrightarrow> (\<forall>h'. h ## h' \<and> x h' \<longrightarrow> y (h + h') \<Longrightarrow> Q) \<Longrightarrow> Q"
  | 
| 
 | 
   366  | 
  by (auto dest: sep_implD)
  | 
| 
 | 
   367  | 
  | 
| 
 | 
   368  | 
lemma sep_impl_sep_true [simp]:
  | 
| 
 | 
   369  | 
  "(P \<longrightarrow>* sep_true) = sep_true"
  | 
| 
 | 
   370  | 
  by (force intro!: sep_implI ext)
  | 
| 
 | 
   371  | 
  | 
| 
 | 
   372  | 
lemma sep_impl_sep_false [simp]:
  | 
| 
 | 
   373  | 
  "(sep_false \<longrightarrow>* P) = sep_true"
  | 
| 
 | 
   374  | 
  by (force intro!: sep_implI ext)
  | 
| 
 | 
   375  | 
  | 
| 
 | 
   376  | 
lemma sep_impl_sep_true_P:
  | 
| 
 | 
   377  | 
  "(sep_true \<longrightarrow>* P) h \<Longrightarrow> P h"
  | 
| 
 | 
   378  | 
  by (clarsimp dest!: sep_implD elim!: allE[where x=0])
  | 
| 
 | 
   379  | 
  | 
| 
 | 
   380  | 
lemma sep_impl_sep_true_false [simp]:
  | 
| 
 | 
   381  | 
  "(sep_true \<longrightarrow>* sep_false) = sep_false"
  | 
| 
 | 
   382  | 
  by (force intro!: ext dest: sep_impl_sep_true_P)
  | 
| 
 | 
   383  | 
  | 
| 
 | 
   384  | 
lemma sep_conj_sep_impl:
  | 
| 
 | 
   385  | 
  "\<lbrakk> P h; \<And>h. (P ** Q) h \<Longrightarrow> R h \<rbrakk> \<Longrightarrow> (Q \<longrightarrow>* R) h"
  | 
| 
 | 
   386  | 
proof (rule sep_implI)
  | 
| 
 | 
   387  | 
  fix h' h
  | 
| 
 | 
   388  | 
  assume "P h" and "h ## h'" and "Q h'"
  | 
| 
 | 
   389  | 
  hence "(P ** Q) (h + h')" by (force intro: sep_conjI)
  | 
| 
 | 
   390  | 
  moreover assume "\<And>h. (P ** Q) h \<Longrightarrow> R h"
  | 
| 
 | 
   391  | 
  ultimately show "R (h + h')" by simp
  | 
| 
 | 
   392  | 
qed
  | 
| 
 | 
   393  | 
  | 
| 
 | 
   394  | 
lemma sep_conj_sep_impl2:
  | 
| 
 | 
   395  | 
  "\<lbrakk> (P ** Q) h; \<And>h. P h \<Longrightarrow> (Q \<longrightarrow>* R) h \<rbrakk> \<Longrightarrow> R h"
  | 
| 
 | 
   396  | 
  by (force dest: sep_implD elim: sep_conjE)
  | 
| 
 | 
   397  | 
  | 
| 
 | 
   398  | 
lemma sep_conj_sep_impl_sep_conj2:
  | 
| 
 | 
   399  | 
  "(P ** R) h \<Longrightarrow> (P ** (Q \<longrightarrow>* (Q ** R))) h"
  | 
| 
 | 
   400  | 
  by (erule (1) sep_conj_impl, erule sep_conj_sep_impl, simp add: sep_conj_ac)
  | 
| 
 | 
   401  | 
  | 
| 
 | 
   402  | 
  | 
| 
 | 
   403  | 
subsection {* Pure assertions *}
 | 
| 
 | 
   404  | 
  | 
| 
 | 
   405  | 
definition
  | 
| 
 | 
   406  | 
  pure :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
| 
 | 
   407  | 
  "pure P \<equiv> \<forall>h h'. P h = P h'"
  | 
| 
 | 
   408  | 
  | 
| 
 | 
   409  | 
lemma pure_sep_true:
  | 
| 
 | 
   410  | 
  "pure sep_true"
  | 
| 
 | 
   411  | 
  by (simp add: pure_def)
  | 
| 
 | 
   412  | 
  | 
| 
 | 
   413  | 
lemma pure_sep_false:
  | 
| 
 | 
   414  | 
  "pure sep_true"
  | 
| 
 | 
   415  | 
  by (simp add: pure_def)
  | 
| 
 | 
   416  | 
  | 
| 
 | 
   417  | 
lemma pure_split:
  | 
| 
 | 
   418  | 
  "pure P = (P = sep_true \<or> P = sep_false)"
  | 
| 
 | 
   419  | 
  by (force simp: pure_def intro!: ext)
  | 
| 
 | 
   420  | 
  | 
| 
 | 
   421  | 
lemma pure_sep_conj:
  | 
| 
 | 
   422  | 
  "\<lbrakk> pure P; pure Q \<rbrakk> \<Longrightarrow> pure (P \<and>* Q)"
  | 
| 
 | 
   423  | 
  by (force simp: pure_split)
  | 
| 
 | 
   424  | 
  | 
| 
 | 
   425  | 
lemma pure_sep_impl:
  | 
| 
 | 
   426  | 
  "\<lbrakk> pure P; pure Q \<rbrakk> \<Longrightarrow> pure (P \<longrightarrow>* Q)"
  | 
| 
 | 
   427  | 
  by (force simp: pure_split)
  | 
| 
 | 
   428  | 
  | 
| 
 | 
   429  | 
lemma pure_conj_sep_conj:
  | 
| 
 | 
   430  | 
  "\<lbrakk> (P and Q) h; pure P \<or> pure Q \<rbrakk> \<Longrightarrow> (P \<and>* Q) h"
  | 
| 
 | 
   431  | 
  by (metis pure_def sep_add_zero sep_conjI sep_conj_commute sep_disj_zero)
  | 
| 
 | 
   432  | 
  | 
| 
 | 
   433  | 
lemma pure_sep_conj_conj:
  | 
| 
 | 
   434  | 
  "\<lbrakk> (P \<and>* Q) h; pure P; pure Q \<rbrakk> \<Longrightarrow> (P and Q) h"
  | 
| 
 | 
   435  | 
  by (force simp: pure_split)
  | 
| 
 | 
   436  | 
  | 
| 
 | 
   437  | 
lemma pure_conj_sep_conj_assoc:
  | 
| 
 | 
   438  | 
  "pure P \<Longrightarrow> ((P and Q) \<and>* R) = (P and (Q \<and>* R))"
  | 
| 
 | 
   439  | 
  by (auto simp: pure_split)
  | 
| 
 | 
   440  | 
  | 
| 
 | 
   441  | 
lemma pure_sep_impl_impl:
  | 
| 
 | 
   442  | 
  "\<lbrakk> (P \<longrightarrow>* Q) h; pure P \<rbrakk> \<Longrightarrow> P h \<longrightarrow> Q h"
  | 
| 
 | 
   443  | 
  by (force simp: pure_split dest: sep_impl_sep_true_P)
  | 
| 
 | 
   444  | 
  | 
| 
 | 
   445  | 
lemma pure_impl_sep_impl:
  | 
| 
 | 
   446  | 
  "\<lbrakk> P h \<longrightarrow> Q h; pure P; pure Q \<rbrakk> \<Longrightarrow> (P \<longrightarrow>* Q) h"
  | 
| 
 | 
   447  | 
  by (force simp: pure_split)
  | 
| 
 | 
   448  | 
  | 
| 
 | 
   449  | 
lemma pure_conj_right: "(Q \<and>* (\<langle>P'\<rangle> and Q')) = (\<langle>P'\<rangle> and (Q \<and>* Q'))"
  | 
| 
 | 
   450  | 
  by (rule ext, rule, rule, clarsimp elim!: sep_conjE)
  | 
| 
 | 
   451  | 
     (erule sep_conj_impl, auto)
  | 
| 
 | 
   452  | 
  | 
| 
 | 
   453  | 
lemma pure_conj_right': "(Q \<and>* (P' and \<langle>Q'\<rangle>)) = (\<langle>Q'\<rangle> and (Q \<and>* P'))"
  | 
| 
 | 
   454  | 
  by (simp add: conj_comms pure_conj_right)
  | 
| 
 | 
   455  | 
  | 
| 
 | 
   456  | 
lemma pure_conj_left: "((\<langle>P'\<rangle> and Q') \<and>* Q) = (\<langle>P'\<rangle> and (Q' \<and>* Q))"
  | 
| 
 | 
   457  | 
  by (simp add: pure_conj_right sep_conj_ac)
  | 
| 
 | 
   458  | 
  | 
| 
 | 
   459  | 
lemma pure_conj_left': "((P' and \<langle>Q'\<rangle>) \<and>* Q) = (\<langle>Q'\<rangle> and (P' \<and>* Q))"
  | 
| 
 | 
   460  | 
  by (subst conj_comms, subst pure_conj_left, simp)
  | 
| 
 | 
   461  | 
  | 
| 
 | 
   462  | 
lemmas pure_conj = pure_conj_right pure_conj_right' pure_conj_left
  | 
| 
 | 
   463  | 
    pure_conj_left'
  | 
| 
 | 
   464  | 
  | 
| 
 | 
   465  | 
declare pure_conj[simp add]
  | 
| 
 | 
   466  | 
  | 
| 
 | 
   467  | 
  | 
| 
 | 
   468  | 
subsection {* Intuitionistic assertions *}
 | 
| 
 | 
   469  | 
  | 
| 
 | 
   470  | 
definition intuitionistic :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
| 
 | 
   471  | 
  "intuitionistic P \<equiv> \<forall>h h'. P h \<and> h \<preceq> h' \<longrightarrow> P h'"
  | 
| 
 | 
   472  | 
  | 
| 
 | 
   473  | 
lemma intuitionisticI:
  | 
| 
 | 
   474  | 
  "(\<And>h h'. \<lbrakk> P h; h \<preceq> h' \<rbrakk> \<Longrightarrow> P h') \<Longrightarrow> intuitionistic P"
  | 
| 
 | 
   475  | 
  by (unfold intuitionistic_def, fast)
  | 
| 
 | 
   476  | 
  | 
| 
 | 
   477  | 
lemma intuitionisticD:
  | 
| 
 | 
   478  | 
  "\<lbrakk> intuitionistic P; P h; h \<preceq> h' \<rbrakk> \<Longrightarrow> P h'"
  | 
| 
 | 
   479  | 
  by (unfold intuitionistic_def, fast)
  | 
| 
 | 
   480  | 
  | 
| 
 | 
   481  | 
lemma pure_intuitionistic:
  | 
| 
 | 
   482  | 
  "pure P \<Longrightarrow> intuitionistic P"
  | 
| 
 | 
   483  | 
  by (clarsimp simp: intuitionistic_def pure_def, fast)
  | 
| 
 | 
   484  | 
  | 
| 
 | 
   485  | 
lemma intuitionistic_conj:
  | 
| 
 | 
   486  | 
  "\<lbrakk> intuitionistic P; intuitionistic Q \<rbrakk> \<Longrightarrow> intuitionistic (P and Q)"
  | 
| 
 | 
   487  | 
  by (force intro: intuitionisticI dest: intuitionisticD)
  | 
| 
 | 
   488  | 
  | 
| 
 | 
   489  | 
lemma intuitionistic_disj:
  | 
| 
 | 
   490  | 
  "\<lbrakk> intuitionistic P; intuitionistic Q \<rbrakk> \<Longrightarrow> intuitionistic (P or Q)"
  | 
| 
 | 
   491  | 
  by (force intro: intuitionisticI dest: intuitionisticD)
  | 
| 
 | 
   492  | 
  | 
| 
 | 
   493  | 
lemma intuitionistic_forall:
  | 
| 
 | 
   494  | 
  "(\<And>x. intuitionistic (P x)) \<Longrightarrow> intuitionistic (ALLS x. P x)"
  | 
| 
 | 
   495  | 
  by (force intro: intuitionisticI dest: intuitionisticD)
  | 
| 
 | 
   496  | 
  | 
| 
 | 
   497  | 
lemma intuitionistic_exists:
  | 
| 
 | 
   498  | 
  "(\<And>x. intuitionistic (P x)) \<Longrightarrow> intuitionistic (EXS x. P x)"
  | 
| 
 | 
   499  | 
  by (unfold pred_ex_def, force intro: intuitionisticI dest: intuitionisticD)
  | 
| 
 | 
   500  | 
  | 
| 
 | 
   501  | 
lemma intuitionistic_sep_conj_sep_true:
  | 
| 
 | 
   502  | 
  "intuitionistic (sep_true \<and>* P)"
  | 
| 
 | 
   503  | 
proof (rule intuitionisticI)
  | 
| 
 | 
   504  | 
  fix h h' r
  | 
| 
 | 
   505  | 
  assume a: "(sep_true \<and>* P) h"
  | 
| 
 | 
   506  | 
  then obtain x y where P: "P y" and h: "h = x + y" and xyd: "x ## y"
  | 
| 
 | 
   507  | 
    by - (drule sep_conjD, clarsimp)
  | 
| 
 | 
   508  | 
  moreover assume a2: "h \<preceq> h'"
  | 
| 
 | 
   509  | 
  then obtain z where h': "h' = h + z" and hzd: "h ## z"
  | 
| 
 | 
   510  | 
    by (clarsimp simp: sep_substate_def)
  | 
| 
 | 
   511  | 
  | 
| 
 | 
   512  | 
  moreover have "(P \<and>* sep_true) (y + (x + z))"
  | 
| 
 | 
   513  | 
    using P h hzd xyd
  | 
| 
 | 
   514  | 
    by (metis sep_add_disjI1 sep_disj_commute sep_conjI)
  | 
| 
 | 
   515  | 
  ultimately show "(sep_true \<and>* P) h'" using hzd
  | 
| 
 | 
   516  | 
    by (auto simp: sep_conj_commute sep_add_ac dest!: sep_disj_addD)
  | 
| 
 | 
   517  | 
qed
  | 
| 
 | 
   518  | 
  | 
| 
 | 
   519  | 
lemma intuitionistic_sep_impl_sep_true:
  | 
| 
 | 
   520  | 
  "intuitionistic (sep_true \<longrightarrow>* P)"
  | 
| 
 | 
   521  | 
proof (rule intuitionisticI)
  | 
| 
 | 
   522  | 
  fix h h'
  | 
| 
 | 
   523  | 
  assume imp: "(sep_true \<longrightarrow>* P) h" and hh': "h \<preceq> h'"
  | 
| 
 | 
   524  | 
  | 
| 
 | 
   525  | 
  from hh' obtain z where h': "h' = h + z" and hzd: "h ## z"
  | 
| 
 | 
   526  | 
    by (clarsimp simp: sep_substate_def)
  | 
| 
 | 
   527  | 
  show "(sep_true \<longrightarrow>* P) h'" using imp h' hzd
  | 
| 
 | 
   528  | 
    apply (clarsimp dest!: sep_implD)
  | 
| 
 | 
   529  | 
    apply (metis sep_add_assoc sep_add_disjD sep_disj_addI3 sep_implI)
  | 
| 
 | 
   530  | 
    done
  | 
| 
 | 
   531  | 
qed
  | 
| 
 | 
   532  | 
  | 
| 
 | 
   533  | 
lemma intuitionistic_sep_conj:
  | 
| 
 | 
   534  | 
  assumes ip: "intuitionistic (P::('a \<Rightarrow> bool))"
 | 
| 
 | 
   535  | 
  shows "intuitionistic (P \<and>* Q)"
  | 
| 
 | 
   536  | 
proof (rule intuitionisticI)
  | 
| 
 | 
   537  | 
  fix h h'
  | 
| 
 | 
   538  | 
  assume sc: "(P \<and>* Q) h" and hh': "h \<preceq> h'"
  | 
| 
 | 
   539  | 
  | 
| 
 | 
   540  | 
  from hh' obtain z where h': "h' = h + z" and hzd: "h ## z"
  | 
| 
 | 
   541  | 
    by (clarsimp simp: sep_substate_def)
  | 
| 
 | 
   542  | 
  | 
| 
 | 
   543  | 
  from sc obtain x y where px: "P x" and qy: "Q y"
  | 
| 
 | 
   544  | 
                       and h: "h = x + y" and xyd: "x ## y"
  | 
| 
 | 
   545  | 
    by (clarsimp simp: sep_conj_def)
  | 
| 
 | 
   546  | 
  | 
| 
 | 
   547  | 
  have "x ## z" using hzd h xyd
  | 
| 
 | 
   548  | 
    by (metis sep_add_disjD)
  | 
| 
 | 
   549  | 
  | 
| 
 | 
   550  | 
  with ip px have "P (x + z)"
  | 
| 
 | 
   551  | 
    by (fastforce elim: intuitionisticD sep_substate_disj_add)
  | 
| 
 | 
   552  | 
  | 
| 
 | 
   553  | 
  thus "(P \<and>* Q) h'" using h' h hzd qy xyd
  | 
| 
 | 
   554  | 
    by (metis (full_types) sep_add_commute sep_add_disjD sep_add_disjI2
  | 
| 
 | 
   555  | 
              sep_add_left_commute sep_conjI)
  | 
| 
 | 
   556  | 
qed
  | 
| 
 | 
   557  | 
  | 
| 
 | 
   558  | 
lemma intuitionistic_sep_impl:
  | 
| 
 | 
   559  | 
  assumes iq: "intuitionistic Q"
  | 
| 
 | 
   560  | 
  shows "intuitionistic (P \<longrightarrow>* Q)"
  | 
| 
 | 
   561  | 
proof (rule intuitionisticI)
  | 
| 
 | 
   562  | 
  fix h h'
  | 
| 
 | 
   563  | 
  assume imp: "(P \<longrightarrow>* Q) h" and hh': "h \<preceq> h'"
  | 
| 
 | 
   564  | 
  | 
| 
 | 
   565  | 
  from hh' obtain z where h': "h' = h + z" and hzd: "h ## z"
  | 
| 
 | 
   566  | 
    by (clarsimp simp: sep_substate_def)
  | 
| 
 | 
   567  | 
  | 
| 
 | 
   568  | 
  {
 | 
| 
 | 
   569  | 
    fix x
  | 
| 
 | 
   570  | 
    assume px: "P x" and hzx: "h + z ## x"
  | 
| 
 | 
   571  | 
  | 
| 
 | 
   572  | 
    have "h + x \<preceq> h + x + z" using hzx hzd
  | 
| 
 | 
   573  | 
    by (metis sep_add_disjI1 sep_substate_def)
  | 
| 
 | 
   574  | 
  | 
| 
 | 
   575  | 
    with imp hzd iq px hzx
  | 
| 
 | 
   576  | 
    have "Q (h + z + x)"
  | 
| 
 | 
   577  | 
    by (metis intuitionisticD sep_add_assoc sep_add_ac sep_add_disjD sep_implE)
  | 
| 
 | 
   578  | 
  }
  | 
| 
 | 
   579  | 
  | 
| 
 | 
   580  | 
  with imp h' hzd iq show "(P \<longrightarrow>* Q) h'"
  | 
| 
 | 
   581  | 
    by (fastforce intro: sep_implI)
  | 
| 
 | 
   582  | 
qed
  | 
| 
 | 
   583  | 
  | 
| 
 | 
   584  | 
lemma strongest_intuitionistic:
  | 
| 
 | 
   585  | 
  "\<not> (\<exists>Q. (\<forall>h. (Q h \<longrightarrow> (P \<and>* sep_true) h)) \<and> intuitionistic Q \<and>
  | 
| 
 | 
   586  | 
      Q \<noteq> (P \<and>* sep_true) \<and> (\<forall>h. P h \<longrightarrow> Q h))"
  | 
| 
 | 
   587  | 
  by (fastforce intro!: ext sep_substate_disj_add
  | 
| 
 | 
   588  | 
                dest!: sep_conjD intuitionisticD)
  | 
| 
 | 
   589  | 
  | 
| 
 | 
   590  | 
lemma weakest_intuitionistic:
  | 
| 
 | 
   591  | 
  "\<not> (\<exists>Q. (\<forall>h. ((sep_true \<longrightarrow>* P) h \<longrightarrow> Q h)) \<and> intuitionistic Q \<and>
  | 
| 
 | 
   592  | 
      Q \<noteq> (sep_true \<longrightarrow>* P) \<and> (\<forall>h. Q h \<longrightarrow> P h))"
  | 
| 
 | 
   593  | 
  apply (clarsimp intro!: ext)
  | 
| 
 | 
   594  | 
  apply (rule iffI)
  | 
| 
 | 
   595  | 
   apply (rule sep_implI)
  | 
| 
 | 
   596  | 
   apply (drule_tac h="x" and h'="x + h'" in intuitionisticD)
  | 
| 
 | 
   597  | 
     apply (clarsimp simp: sep_add_ac sep_substate_disj_add)+
  | 
| 
 | 
   598  | 
  done
  | 
| 
 | 
   599  | 
  | 
| 
 | 
   600  | 
lemma intuitionistic_sep_conj_sep_true_P:
  | 
| 
 | 
   601  | 
  "\<lbrakk> (P \<and>* sep_true) s; intuitionistic P \<rbrakk> \<Longrightarrow> P s"
  | 
| 
 | 
   602  | 
  by (force dest: intuitionisticD elim: sep_conjE sep_substate_disj_add)
  | 
| 
 | 
   603  | 
  | 
| 
 | 
   604  | 
lemma intuitionistic_sep_conj_sep_true_simp:
  | 
| 
 | 
   605  | 
  "intuitionistic P \<Longrightarrow> (P \<and>* sep_true) = P"
  | 
| 
 | 
   606  | 
  by (fast intro!: sep_conj_sep_true ext
  | 
| 
 | 
   607  | 
           elim: intuitionistic_sep_conj_sep_true_P)
  | 
| 
 | 
   608  | 
  | 
| 
 | 
   609  | 
lemma intuitionistic_sep_impl_sep_true_P:
  | 
| 
 | 
   610  | 
  "\<lbrakk> P h; intuitionistic P \<rbrakk> \<Longrightarrow> (sep_true \<longrightarrow>* P) h"
  | 
| 
 | 
   611  | 
  by (force intro!: sep_implI dest: intuitionisticD
  | 
| 
 | 
   612  | 
            intro: sep_substate_disj_add)
  | 
| 
 | 
   613  | 
  | 
| 
 | 
   614  | 
lemma intuitionistic_sep_impl_sep_true_simp:
  | 
| 
 | 
   615  | 
  "intuitionistic P \<Longrightarrow> (sep_true \<longrightarrow>* P) = P"
  | 
| 
 | 
   616  | 
  by (fast intro!: ext
  | 
| 
 | 
   617  | 
           elim: sep_impl_sep_true_P intuitionistic_sep_impl_sep_true_P)
  | 
| 
 | 
   618  | 
  | 
| 
 | 
   619  | 
  | 
| 
 | 
   620  | 
subsection {* Strictly exact assertions *}
 | 
| 
 | 
   621  | 
  | 
| 
 | 
   622  | 
definition strictly_exact :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
| 
 | 
   623  | 
  "strictly_exact P \<equiv> \<forall>h h'. P h \<and> P h' \<longrightarrow> h = h'"
  | 
| 
 | 
   624  | 
  | 
| 
 | 
   625  | 
lemma strictly_exactD:
  | 
| 
 | 
   626  | 
  "\<lbrakk> strictly_exact P; P h; P h' \<rbrakk> \<Longrightarrow> h = h'"
  | 
| 
 | 
   627  | 
  by (unfold strictly_exact_def, fast)
  | 
| 
 | 
   628  | 
  | 
| 
 | 
   629  | 
lemma strictly_exactI:
  | 
| 
 | 
   630  | 
  "(\<And>h h'. \<lbrakk> P h; P h' \<rbrakk> \<Longrightarrow> h = h') \<Longrightarrow> strictly_exact P"
  | 
| 
 | 
   631  | 
  by (unfold strictly_exact_def, fast)
  | 
| 
 | 
   632  | 
  | 
| 
 | 
   633  | 
lemma strictly_exact_sep_conj:
  | 
| 
 | 
   634  | 
  "\<lbrakk> strictly_exact P; strictly_exact Q \<rbrakk> \<Longrightarrow> strictly_exact (P \<and>* Q)"
  | 
| 
 | 
   635  | 
  apply (rule strictly_exactI)
  | 
| 
 | 
   636  | 
  apply (erule sep_conjE)+
  | 
| 
 | 
   637  | 
  apply (drule_tac h="x" and h'="xa" in strictly_exactD, assumption+)
  | 
| 
 | 
   638  | 
  apply (drule_tac h="y" and h'="ya" in strictly_exactD, assumption+)
  | 
| 
 | 
   639  | 
  apply clarsimp
  | 
| 
 | 
   640  | 
  done
  | 
| 
 | 
   641  | 
  | 
| 
 | 
   642  | 
lemma strictly_exact_conj_impl:
  | 
| 
 | 
   643  | 
  "\<lbrakk> (Q \<and>* sep_true) h; P h; strictly_exact Q \<rbrakk> \<Longrightarrow> (Q \<and>* (Q \<longrightarrow>* P)) h"
  | 
| 
 | 
   644  | 
  by (force intro: sep_conjI sep_implI dest: strictly_exactD elim!: sep_conjE
  | 
| 
 | 
   645  | 
            simp: sep_add_commute sep_add_assoc)
  | 
| 
 | 
   646  | 
  | 
| 
 | 
   647  | 
end
  | 
| 
 | 
   648  | 
  | 
| 
 | 
   649  | 
interpretation sep: ab_semigroup_mult "op **"
  | 
| 
 | 
   650  | 
  by (rule ab_semigroup_mult_sep_conj)
  | 
| 
 | 
   651  | 
  | 
| 
 | 
   652  | 
interpretation sep: comm_monoid_add "op **" \<box>
  | 
| 
 | 
   653  | 
  by (rule comm_monoid_add)
  | 
| 
 | 
   654  | 
  | 
| 
 | 
   655  | 
  | 
| 
 | 
   656  | 
section {* Separation Algebra with Stronger, but More Intuitive Disjunction Axiom *}
 | 
| 
 | 
   657  | 
  | 
| 
 | 
   658  | 
class stronger_sep_algebra = pre_sep_algebra +
  | 
| 
 | 
   659  | 
  assumes sep_add_disj_eq [simp]: "y ## z \<Longrightarrow> x ## y + z = (x ## y \<and> x ## z)"
  | 
| 
 | 
   660  | 
begin
  | 
| 
 | 
   661  | 
  | 
| 
 | 
   662  | 
lemma sep_disj_add_eq [simp]: "x ## y \<Longrightarrow> x + y ## z = (x ## z \<and> y ## z)"
  | 
| 
 | 
   663  | 
  by (metis sep_add_disj_eq sep_disj_commute)
  | 
| 
 | 
   664  | 
  | 
| 
 | 
   665  | 
subclass sep_algebra by default auto
  | 
| 
 | 
   666  | 
  | 
| 
 | 
   667  | 
end
  | 
| 
 | 
   668  | 
  | 
| 
 | 
   669  | 
  | 
| 
 | 
   670  | 
section {* Folding separating conjunction over lists of predicates *}
 | 
| 
 | 
   671  | 
  | 
| 
 | 
   672  | 
lemma sep_list_conj_Nil [simp]: "\<And>* [] = \<box>"
  | 
| 
 | 
   673  | 
  by (simp add: sep_list_conj_def)
  | 
| 
 | 
   674  | 
  | 
| 
 | 
   675  | 
(* apparently these two are rarely used and had to be removed from List.thy *)
  | 
| 
 | 
   676  | 
lemma (in semigroup_add) foldl_assoc:
  | 
| 
 | 
   677  | 
shows "foldl op+ (x+y) zs = x + (foldl op+ y zs)"
  | 
| 
 | 
   678  | 
by (induct zs arbitrary: y) (simp_all add:add_assoc)
  | 
| 
 | 
   679  | 
  | 
| 
 | 
   680  | 
lemma (in monoid_add) foldl_absorb0:
  | 
| 
 | 
   681  | 
shows "x + (foldl op+ 0 zs) = foldl op+ x zs"
  | 
| 
 | 
   682  | 
by (induct zs) (simp_all add:foldl_assoc)
  | 
| 
 | 
   683  | 
  | 
| 
 | 
   684  | 
lemma sep_list_conj_Cons [simp]: "\<And>* (x#xs) = (x ** \<And>* xs)"
  | 
| 
 | 
   685  | 
  by (simp add: sep_list_conj_def sep.foldl_absorb0)
  | 
| 
 | 
   686  | 
  | 
| 
 | 
   687  | 
lemma sep_list_conj_append [simp]: "\<And>* (xs @ ys) = (\<And>* xs ** \<And>* ys)"
  | 
| 
 | 
   688  | 
  by (simp add: sep_list_conj_def sep.foldl_absorb0)
  | 
| 
 | 
   689  | 
  | 
| 
 | 
   690  | 
lemma (in comm_monoid_add) foldl_map_filter:
  | 
| 
 | 
   691  | 
  "foldl op + 0 (map f (filter P xs)) +
  | 
| 
 | 
   692  | 
     foldl op + 0 (map f (filter (not P) xs))
  | 
| 
 | 
   693  | 
   = foldl op + 0 (map f xs)"
  | 
| 
 | 
   694  | 
proof (induct xs)
  | 
| 
 | 
   695  | 
  case Nil thus ?case by clarsimp
  | 
| 
 | 
   696  | 
next
  | 
| 
 | 
   697  | 
  case (Cons x xs)
  | 
| 
 | 
   698  | 
  hence IH: "foldl op + 0 (map f xs) =
  | 
| 
 | 
   699  | 
               foldl op + 0 (map f (filter P xs)) +
  | 
| 
 | 
   700  | 
               foldl op + 0 (map f [x\<leftarrow>xs . \<not> P x])"
  | 
| 
 | 
   701  | 
               by (simp only: eq_commute)
  | 
| 
 | 
   702  | 
  | 
| 
 | 
   703  | 
  have foldl_Cons':
  | 
| 
 | 
   704  | 
    "\<And>x xs. foldl op + 0 (x # xs) = x + (foldl op + 0 xs)"
  | 
| 
 | 
   705  | 
    by (simp, subst foldl_absorb0[symmetric], rule refl)
  | 
| 
 | 
   706  | 
  | 
| 
 | 
   707  | 
  { assume "P x"
 | 
| 
 | 
   708  | 
    hence ?case by (auto simp del: foldl_Cons simp add: foldl_Cons' IH add_ac)
  | 
| 
 | 
   709  | 
  } moreover {
 | 
| 
 | 
   710  | 
    assume "\<not> P x"
  | 
| 
 | 
   711  | 
    hence ?case by (auto simp del: foldl_Cons simp add: foldl_Cons' IH add_ac)
  | 
| 
 | 
   712  | 
  }
  | 
| 
 | 
   713  | 
  ultimately show ?case by blast
  | 
| 
 | 
   714  | 
qed
  | 
| 
 | 
   715  | 
  | 
| 
 | 
   716  | 
  | 
| 
 | 
   717  | 
section {* Separation Algebra with a Cancellative Monoid (for completeness) *}
 | 
| 
 | 
   718  | 
  | 
| 
 | 
   719  | 
text {*
 | 
| 
 | 
   720  | 
  Separation algebra with a cancellative monoid. The results of being a precise
  | 
| 
 | 
   721  | 
  assertion (distributivity over separating conjunction) require this.
  | 
| 
 | 
   722  | 
  although we never actually use this property in our developments, we keep
  | 
| 
 | 
   723  | 
  it here for completeness.
  | 
| 
 | 
   724  | 
  *}
  | 
| 
 | 
   725  | 
class cancellative_sep_algebra = sep_algebra +
  | 
| 
 | 
   726  | 
  assumes sep_add_cancelD: "\<lbrakk> x + z = y + z ; x ## z ; y ## z \<rbrakk> \<Longrightarrow> x = y"
  | 
| 
 | 
   727  | 
begin
  | 
| 
 | 
   728  | 
  | 
| 
 | 
   729  | 
definition
  | 
| 
 | 
   730  | 
  (* In any heap, there exists at most one subheap for which P holds *)
  | 
| 
 | 
   731  | 
  precise :: "('a \<Rightarrow> bool) \<Rightarrow> bool" where
 | 
| 
 | 
   732  | 
  "precise P = (\<forall>h hp hp'. hp \<preceq> h \<and> P hp \<and> hp' \<preceq> h \<and> P hp' \<longrightarrow> hp = hp')"
  | 
| 
 | 
   733  | 
  | 
| 
 | 
   734  | 
lemma "precise (op = s)"
  | 
| 
 | 
   735  | 
  by (metis (full_types) precise_def)
  | 
| 
 | 
   736  | 
  | 
| 
 | 
   737  | 
lemma sep_add_cancel:
  | 
| 
 | 
   738  | 
  "x ## z \<Longrightarrow> y ## z \<Longrightarrow> (x + z = y + z) = (x = y)"
  | 
| 
 | 
   739  | 
  by (metis sep_add_cancelD)
  | 
| 
 | 
   740  | 
  | 
| 
 | 
   741  | 
lemma precise_distribute:
  | 
| 
 | 
   742  | 
  "precise P = (\<forall>Q R. ((Q and R) \<and>* P) = ((Q \<and>* P) and (R \<and>* P)))"
  | 
| 
 | 
   743  | 
proof (rule iffI)
  | 
| 
 | 
   744  | 
  assume pp: "precise P"
  | 
| 
 | 
   745  | 
  {
 | 
| 
 | 
   746  | 
    fix Q R
  | 
| 
 | 
   747  | 
    fix h hp hp' s
  | 
| 
 | 
   748  | 
  | 
| 
 | 
   749  | 
    { assume a: "((Q and R) \<and>* P) s"
 | 
| 
 | 
   750  | 
      hence "((Q \<and>* P) and (R \<and>* P)) s"
  | 
| 
 | 
   751  | 
        by (fastforce dest!: sep_conjD elim: sep_conjI)
  | 
| 
 | 
   752  | 
    }
  | 
| 
 | 
   753  | 
    moreover
  | 
| 
 | 
   754  | 
    { assume qs: "(Q \<and>* P) s" and qr: "(R \<and>* P) s"
 | 
| 
 | 
   755  | 
  | 
| 
 | 
   756  | 
      from qs obtain x y where sxy: "s = x + y" and xy: "x ## y"
  | 
| 
 | 
   757  | 
                           and x: "Q x" and y: "P y"
  | 
| 
 | 
   758  | 
        by (fastforce dest!: sep_conjD)
  | 
| 
 | 
   759  | 
      from qr obtain x' y' where sxy': "s = x' + y'" and xy': "x' ## y'"
  | 
| 
 | 
   760  | 
                           and x': "R x'" and y': "P y'"
  | 
| 
 | 
   761  | 
        by (fastforce dest!: sep_conjD)
  | 
| 
 | 
   762  | 
  | 
| 
 | 
   763  | 
      from sxy have ys: "y \<preceq> x + y" using xy
  | 
| 
 | 
   764  | 
        by (fastforce simp: sep_substate_disj_add' sep_disj_commute)
  | 
| 
 | 
   765  | 
      from sxy' have ys': "y' \<preceq> x' + y'" using xy'
  | 
| 
 | 
   766  | 
        by (fastforce simp: sep_substate_disj_add' sep_disj_commute)
  | 
| 
 | 
   767  | 
  | 
| 
 | 
   768  | 
      from pp have yy: "y = y'" using sxy sxy' xy xy' y y' ys ys'
  | 
| 
 | 
   769  | 
        by (fastforce simp: precise_def)
  | 
| 
 | 
   770  | 
  | 
| 
 | 
   771  | 
      hence "x = x'" using sxy sxy' xy xy'
  | 
| 
 | 
   772  | 
        by (fastforce dest!: sep_add_cancelD)
  | 
| 
 | 
   773  | 
  | 
| 
 | 
   774  | 
      hence "((Q and R) \<and>* P) s" using sxy x x' yy y' xy'
  | 
| 
 | 
   775  | 
        by (fastforce intro: sep_conjI)
  | 
| 
 | 
   776  | 
    }
  | 
| 
 | 
   777  | 
    ultimately
  | 
| 
 | 
   778  | 
    have "((Q and R) \<and>* P) s = ((Q \<and>* P) and (R \<and>* P)) s" using pp by blast
  | 
| 
 | 
   779  | 
  }
  | 
| 
 | 
   780  | 
  thus "\<forall>Q R. ((Q and R) \<and>* P) = ((Q \<and>* P) and (R \<and>* P))" by (blast intro!: ext)
  | 
| 
 | 
   781  | 
  | 
| 
 | 
   782  | 
next
  | 
| 
 | 
   783  | 
  assume a: "\<forall>Q R. ((Q and R) \<and>* P) = ((Q \<and>* P) and (R \<and>* P))"
  | 
| 
 | 
   784  | 
  thus "precise P"
  | 
| 
 | 
   785  | 
  proof (clarsimp simp: precise_def)
  | 
| 
 | 
   786  | 
    fix h hp hp' Q R
  | 
| 
 | 
   787  | 
    assume hp: "hp \<preceq> h" and hp': "hp' \<preceq> h" and php: "P hp" and php': "P hp'"
  | 
| 
 | 
   788  | 
  | 
| 
 | 
   789  | 
    obtain z where hhp: "h = hp + z" and hpz: "hp ## z" using hp
  | 
| 
 | 
   790  | 
      by (clarsimp simp: sep_substate_def)
  | 
| 
 | 
   791  | 
    obtain z' where hhp': "h = hp' + z'" and hpz': "hp' ## z'" using hp'
  | 
| 
 | 
   792  | 
      by (clarsimp simp: sep_substate_def)
  | 
| 
 | 
   793  | 
  | 
| 
 | 
   794  | 
    have h_eq: "z' + hp' = z + hp" using hhp hhp' hpz hpz'
  | 
| 
 | 
   795  | 
      by (fastforce simp: sep_add_ac)
  | 
| 
 | 
   796  | 
  | 
| 
 | 
   797  | 
    from hhp hhp' a hpz hpz' h_eq
  | 
| 
 | 
   798  | 
    have "\<forall>Q R. ((Q and R) \<and>* P) (z + hp) = ((Q \<and>* P) and (R \<and>* P)) (z' + hp')"
  | 
| 
 | 
   799  | 
      by (fastforce simp: h_eq sep_add_ac sep_conj_commute)
  | 
| 
 | 
   800  | 
  | 
| 
 | 
   801  | 
    hence "((op = z and op = z') \<and>* P) (z + hp) =
  | 
| 
 | 
   802  | 
           ((op = z \<and>* P) and (op = z' \<and>* P)) (z' + hp')" by blast
  | 
| 
 | 
   803  | 
  | 
| 
 | 
   804  | 
    thus  "hp = hp'" using php php' hpz hpz' h_eq
  | 
| 
 | 
   805  | 
      by (fastforce dest!: iffD2 cong: conj_cong
  | 
| 
 | 
   806  | 
                    simp: sep_add_ac sep_add_cancel sep_conj_def)
  | 
| 
 | 
   807  | 
  qed
  | 
| 
 | 
   808  | 
qed
  | 
| 
 | 
   809  | 
  | 
| 
 | 
   810  | 
lemma strictly_precise: "strictly_exact P \<Longrightarrow> precise P"
  | 
| 
 | 
   811  | 
  by (metis precise_def strictly_exactD)
  | 
| 
 | 
   812  | 
  | 
| 
 | 
   813  | 
end
  | 
| 
 | 
   814  | 
  | 
| 
 | 
   815  | 
end
  |