Theory ShareRepProof

(*  Title:       BDD

    Author:      Veronika Ortner and Norbert Schirmer, 2004
    Maintainer:  Norbert Schirmer,  norbert.schirmer at web de
    License:     LGPL
*)

(*  
ShareRepProof.thy

Copyright (C) 2004-2008 Veronika Ortner and Norbert Schirmer 
Some rights reserved, TU Muenchen

This library is free software; you can redistribute it and/or modify
it under the terms of the GNU Lesser General Public License as
published by the Free Software Foundation; either version 2.1 of the
License, or (at your option) any later version.

This library is distributed in the hope that it will be useful, but
WITHOUT ANY WARRANTY; without even the implied warranty of
MERCHANTABILITY or FITNESS FOR A PARTICULAR PURPOSE. See the GNU
Lesser General Public License for more details.

You should have received a copy of the GNU Lesser General Public
License along with this library; if not, write to the Free Software
Foundation, Inc., 59 Temple Place, Suite 330, Boston, MA 02111-1307
USA
*)

section ‹Proof of Procedure ShareRep›
theory ShareRepProof imports ProcedureSpecs Simpl.HeapList begin

lemma (in ShareRep_impl) ShareRep_modifies:
  shows "σ. Γ{σ}  PROC ShareRep (´nodeslist, ´p) 
             {t. t may_only_modify_globals σ in [rep]}"
  apply (hoare_rule HoarePartial.ProcRec1)
  apply (vcg spec=modifies)
  done


lemma hd_filter_cons: 
" i.  P (xs ! i) p; i < length xs;  no  set (take i xs). ¬ P no p;  a b.  P a b = P b a
   xs ! i = hd (filter (P p) xs)"
apply (induct xs)
apply simp
apply (case_tac "P a p")
apply simp
apply (case_tac i)
apply simp
apply simp
apply (case_tac i)
apply simp
apply auto
done

lemma (in ShareRep_impl) ShareRep_spec_total:
shows 
  "σ ns. Γ,Θt 
  σ. List ´nodeslist ´next ns 
     (no  set ns. no  Null   
       ((no´low = Null) = (no´high = Null)) 
       (isLeaf_pt ´p ´low ´high  isLeaf_pt no ´low ´high) 
       no´var = ´p´var)  
       ´p  set ns 
  PROC ShareRep (´nodeslist, ´p)
   (σp  ´rep = hd (filter (λ sn. repNodes_eq sn σp σlow σhigh σrep) ns)) 
    (pt.  pt  σp  ptσrep = pt´rep)  
    (σp´repσvar = σp  σvar)"
apply (hoare_rule HoareTotal.ProcNoRec1)
apply (hoare_rule anno=  
  "IF (isLeaf_pt ´p ´low ´high) 
   THEN ´p  ´rep :== ´nodeslist
   ELSE
     WHILE (´nodeslist  Null)  
     INV prx sfx. List ´nodeslist ´next sfx  ns=prx@sfx  
           ¬ isLeaf_pt ´p ´low ´high 
           (no  set ns. no  Null  
             ((noσlow = Null) = (noσhigh = Null)) 
             (isLeaf_pt σp σlow σhigh  isLeaf_pt no σlow σhigh) 
             noσvar = σpσvar)  
        σp  set ns  
        ((pt  set prx.  repNodes_eq pt σp σlow σhigh σrep) 
          ´rep  σp =  hd (filter (λ sn. repNodes_eq sn σp σlow σhigh σrep) prx) 
             (pt. pt  σp  ptσrep = pt´rep)) 
        ((pt  set prx.  ¬ repNodes_eq pt σp σlow σhigh σrep)   σrep = ´rep) 
        (´nodeslist  Null  
           (pt  set prx. ¬ repNodes_eq pt σp σlow σhigh σrep))   
        (´p = σp  ´high = σhigh  ´low = σlow)
     VAR MEASURE (length (list ´nodeslist ´next))  
     DO
       IF (repNodes_eq ´nodeslist ´p ´low ´high ´rep)
       THEN ´p´rep :== ´nodeslist;; ´nodeslist :== Null
       ELSE ´nodeslist :== ´nodeslist´next
       FI
     OD
  FI" in HoareTotal.annotateI)
apply vcg
using  [[simp_depth_limit = 2]]
apply   (rule conjI)
apply    clarify
apply    (simp (no_asm_use))
prefer 2
apply    clarify
apply    (rule_tac x="[]" in exI)
apply    (rule_tac x=ns in exI)
apply    (simp (no_asm_use))
prefer 2
apply   clarify
apply   (rule conjI)
apply    clarify
apply    (rule conjI)
apply     (clarsimp simp add: List_list) (* solving termination contraint *)
apply    (simp (no_asm_use))
apply    (rule conjI)
apply    assumption
prefer 2
apply    clarify
apply    (simp (no_asm_use))
apply    (rule conjI)
apply    (clarsimp simp add: List_list) (* solving termination constraint *)
apply    (simp only: List_not_Null simp_thms triv_forall_equality)
apply    clarify
apply    (simp only: triv_forall_equality)
apply    (rename_tac sfx)
apply    (rule_tac x="prx@[nodeslist]" in exI)
apply    (rule_tac x="sfx" in exI)
apply    (rule conjI)
apply     assumption
apply    (rule conjI)
apply     simp
prefer 4
apply   (elim exE conjE)
apply   (simp (no_asm_use))
apply   hypsubst
using  [[simp_depth_limit = 100]]
proof -
  (* IF-THEN to postcondition *)
  fix ns var low high rep "next" p nodeslist
  assume ns: "List nodeslist next ns"
  assume no_prop:  "noset ns.
           no  Null 
           (low no = Null) = (high no = Null) 
           (isLeaf_pt p low high  isLeaf_pt no low high)  var no = var p"
  assume p_in_ns: "p  set ns" 
  assume p_Leaf: "isLeaf_pt p low high"
  show "nodeslist = hd [snns . repNodes_eq sn p low high rep] 
        var nodeslist = var p"
  proof -
    from p_in_ns no_prop have p_not_Null: "pNull"
      using [[simp_depth_limit=2]]
      by auto
    from p_in_ns have "ns  []"
      by (cases ns) auto
    with ns obtain ns' where ns': "ns = nodeslist#ns'" 
      by (cases "nodeslist=Null") auto
    with no_prop p_Leaf obtain 
      "isLeaf_pt nodeslist low high" and
      var_eq: "var nodeslist = var p" and
      "nodeslistNull"
      using [[simp_depth_limit=2]]
      by auto
    with p_not_Null p_Leaf have "repNodes_eq nodeslist p low high rep"
      by (simp add: repNodes_eq_def isLeaf_pt_def null_comp_def)
    with ns' var_eq
    show ?thesis
      by simp
  qed
next
  (* From invariant to postcondition *)
  fix var::"refnat" and low high rep repa p prx sfx "next"
  assume sfx: "List Null next sfx"
  assume p_in_ns: "p  set (prx @ sfx)"
  assume no_props: "noset (prx @ sfx).
           no  Null 
           (low no = Null) = (high no = Null) 
           (isLeaf_pt p low high  isLeaf_pt no low high)  var no = var p"
  assume match_prx: "(ptset prx. repNodes_eq pt p low high rep) 
                       repa p = hd [snprx . repNodes_eq sn p low high rep] 
                      (pt. pt  p  rep pt = repa pt)"
  show "repa p = hd [snprx @ sfx . repNodes_eq sn p low high rep] 
          (pt. pt  p  rep pt = repa pt)  var (repa p) = var p"
  proof -
    from sfx
    have sfx_Nil: "sfx=[]"
      by simp
    with p_in_ns have ex_match: "(ptset prx. repNodes_eq pt p low high rep)"
      apply -
      apply (rule_tac x=p in bexI)
      apply  (simp add: repNodes_eq_def)
      apply simp
      done
    hence not_empty: "[snprx . repNodes_eq sn p low high rep]  []"
      apply -
      apply (erule bexE)
      apply (rule filter_not_empty)
      apply auto
      done
    from ex_match match_prx obtain
      found: "repa p = hd [snprx . repNodes_eq sn p low high rep]" and
      unmodif: "pt. pt  p  rep pt = repa pt"
      by blast
    from hd_filter_in_list [OF not_empty] found
    have "repa p  set prx"
      by simp
    with no_props
    have "var (repa p) = var p"
      using [[simp_depth_limit=2]]
      by simp
    with found unmodif sfx_Nil
    show ?thesis
      by simp
  qed
next
  (* Invariant to invariant; ELSE part *)
  fix var low high p repa "next" nodeslist prx sfx
  assume nodeslist_not_Null: "nodeslist  Null" 
  assume p_no_Leaf: "¬ isLeaf_pt p low high"
  assume no_props: "noset prx  set (nodeslist # sfx).
           no  Null  (low no = Null) = (high no = Null)  var no = var p"
  assume p_in_ns: "p  set prx  p  set (nodeslist # sfx)"
  assume match_prx: "(ptset prx. repNodes_eq pt p low high repa) 
            repa p = hd [snprx . repNodes_eq sn p low high repa]"
  assume nomatch_prx: "ptset prx. ¬ repNodes_eq pt p low high repa"
  assume nomatch_nodeslist: "¬ repNodes_eq nodeslist p low high repa"
  assume sfx: "List (next nodeslist) next sfx"
  show "(noset prx  set (nodeslist # sfx).
              no  Null  (low no = Null) = (high no = Null)  var no = var p) 
        ((ptset (prx @ [nodeslist]). repNodes_eq pt p low high repa) 
           repa p = hd [snprx @ [nodeslist] . repNodes_eq sn p low high repa]) 
        (next nodeslist  Null 
            (ptset (prx @ [nodeslist]). ¬ repNodes_eq pt p low high repa))"
  proof -
    from nomatch_prx nomatch_nodeslist
    have "((ptset (prx @ [nodeslist]). repNodes_eq pt p low high repa) 
           repa p = hd [snprx @ [nodeslist] . repNodes_eq sn p low high repa])" 
      by auto
    moreover
    from nomatch_prx nomatch_nodeslist
    have "(next nodeslist  Null 
            (ptset (prx @ [nodeslist]). ¬ repNodes_eq pt p low high repa))"
      by auto
    ultimately show ?thesis
      using no_props
      by (intro conjI)
  qed
next
  (* Invariant to invariant: THEN part *)
  fix var low high p repa "next" nodeslist prx sfx
  assume nodeslist_not_Null: "nodeslist  Null" 
  assume sfx: "List nodeslist next sfx" 
  assume p_not_Leaf: "¬ isLeaf_pt p low high"
  assume no_props: "noset prx  set sfx.
           no  Null 
           (low no = Null) = (high no = Null) 
           (isLeaf_pt p low high  isLeaf_pt no low high)  var no = var p"
  assume p_in_ns: "p  set prx  p  set sfx"
  assume match_prx: "(ptset prx. repNodes_eq pt p low high repa) 
        repa p = hd [snprx . repNodes_eq sn p low high repa]"
  assume nomatch_prx: "ptset prx. ¬ repNodes_eq pt p low high repa"
  assume match: "repNodes_eq nodeslist p low high repa"
  show "(noset prx  set sfx.
              no  Null 
              (low no = Null) = (high no = Null) 
              (isLeaf_pt p low high  isLeaf_pt no low high)  var no = var p) 
        (p  set prx  p  set sfx) 
        ((ptset prx  set sfx. repNodes_eq pt p low high repa) 
           nodeslist =
           hd ([snprx . repNodes_eq sn p low high repa] @
               [snsfx . repNodes_eq sn p low high repa])) 
        ((ptset prx  set sfx. ¬ repNodes_eq pt p low high repa) 
           repa = repa(p := nodeslist))"
  proof -
    from nodeslist_not_Null sfx
    obtain sfx' where sfx': "sfx=nodeslist#sfx'"
      by (cases "nodeslist=Null") auto
    from nomatch_prx match sfx'
    have hd: "hd ([snprx . repNodes_eq sn p low high repa] @
               [snsfx . repNodes_eq sn p low high repa]) = nodeslist"
      by simp
    from match sfx'
    have triv: "((ptset prx  set sfx. ¬ repNodes_eq pt p low high repa) 
           repa = repa(p := nodeslist))" 
      by simp
    show ?thesis
      apply (rule conjI)
      apply (rule no_props)
      apply (intro conjI)
      apply   (rule p_in_ns)
      apply  (simp add: hd)
      apply (rule triv)
      done
  qed
qed

end