# HG changeset patch # User Thomas Sewell # Date 1253591539 -36000 # Node ID 5b65449d76692d31de744a1b3269468e0e0d8703 # Parent c876bcb601fc54abf00c58ff1090a632381e9ccd# Parent a6909ef949aa853d8df9d7ebc06a50722fba29eb Branch merge for isabelle mainline updates. diff -r c876bcb601fc -r 5b65449d7669 Admin/E/eproof --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/Admin/E/eproof Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,92 @@ +#!/usr/bin/perl -w +# +# eproof - run E and translate its output into TSTP format +# +# Author: Sascha Boehme, TU Muenchen +# +# This script is a port of a Bash script with the same name coming with +# E 1.0-004 (written by Stephan Schulz). +# + + +use File::Basename qw/ dirname /; +use File::Temp qw/ tempfile /; + + +# E executables + +my $edir = dirname($0); +my $eprover = "$edir/eprover"; +my $epclextract = "$edir/epclextract"; + + +# build E command from given commandline arguments + +my $format = ""; +my $timelimit = 2000000000; # effectively unlimited + +my $eprover_cmd = "'$eprover'"; +foreach (@ARGV) { + if (m/--cpu-limit=([0-9]+)/) { + $timelimit = $1; + } + + if (m/--tstp-out/) { + $format = $_; + } + else { + $eprover_cmd = "$eprover_cmd '$_'"; + } +} +$eprover_cmd = "$eprover_cmd -l4 -R -o- --pcl-terms-compressed --pcl-compact"; + + +# run E, redirecting output into a temporary file + +my ($fh, $filename) = tempfile(UNLINK => 1); +my $r = system "$eprover_cmd > $filename"; +exit ($r >> 8) if $r != 0; + + +# analyze E output + +my @lines = <$fh>; +my $content = join "", @lines[-60 .. -1]; + # Note: Like the original eproof, we only look at the last 60 lines. + +if ($content =~ m/Total time\s*:\s*([0-9]+\.[0-9]+)/) { + $timelimit = $timelimit - $1 - 1; + + if ($content =~ m/No proof found!/) { + print "# Problem is satisfiable (or invalid), " . + "generating saturation derivation\n"; + } + elsif ($content =~ m/Proof found!/) { + print "# Problem is unsatisfiable (or provable), " . + "constructing proof object\n"; + } + elsif ($content =~ m/Watchlist is empty!/) { + print "# All watchlist clauses generated, constructing derivation\n"; + } + else { + print "# Cannot determine problem status\n"; + exit $r; + } +} +else { + print "# Cannot determine problem status within resource limit\n"; + exit $r; +} + + +# translate E output + +foreach (@lines) { + print if (m/# SZS status/ or m/"# Failure"/); +} +$r = system ("exec bash -c \"ulimit -S -t $timelimit; " . + "'$epclextract' $format -f --competition-framing $filename\""); + # Note: Setting the user time is not supported on Cygwin, i.e., ulimit fails + # and prints and error message. How could we then limit the execution time? +exit ($r >> 8); + diff -r c876bcb601fc -r 5b65449d7669 Admin/isatest/isatest-stats --- a/Admin/isatest/isatest-stats Thu Sep 17 14:17:37 2009 +1000 +++ b/Admin/isatest/isatest-stats Tue Sep 22 13:52:19 2009 +1000 @@ -24,9 +24,9 @@ HOL-MetisExamples \ HOL-MicroJava \ HOL-NSA \ - HOL-NewNumberTheory \ HOL-Nominal-Examples \ - HOL-NumberTheory \ + HOL-Number_Theory \ + HOL-Old_Number_Theory \ HOL-SET-Protocol \ HOL-UNITY \ HOL-Word \ diff -r c876bcb601fc -r 5b65449d7669 CONTRIBUTORS --- a/CONTRIBUTORS Thu Sep 17 14:17:37 2009 +1000 +++ b/CONTRIBUTORS Tue Sep 22 13:52:19 2009 +1000 @@ -7,6 +7,15 @@ Contributions to this Isabelle version -------------------------------------- +* September 2009: Sascha Boehme, TUM + SMT method using external SMT solvers + +* September 2009: Florian Haftmann, TUM + Refinement of Sets and Lattices + +* July 2009: Jeremy Avigad and Amine Chaieb + New number theory + * July 2009: Philipp Meyer, TUM HOL/Library/Sum_of_Squares: functionality to call a remote csdp prover diff -r c876bcb601fc -r 5b65449d7669 NEWS --- a/NEWS Thu Sep 17 14:17:37 2009 +1000 +++ b/NEWS Tue Sep 22 13:52:19 2009 +1000 @@ -18,15 +18,36 @@ *** HOL *** +* New proof method "smt" for a combination of first-order logic +with equality, linear and nonlinear (natural/integer/real) +arithmetic, and fixed-size bitvectors; there is also basic +support for higher-order features (esp. lambda abstractions). +It is an incomplete decision procedure based on external SMT +solvers using the oracle mechanism. + * Reorganization of number theory: - * former session NumberTheory now named Old_Number_Theory; former session NewNumberTheory - named NumberTheory; - * split off prime number ingredients from theory GCD to theory Number_Theory/Primes; + * former session NumberTheory now named Old_Number_Theory + * new session Number_Theory by Jeremy Avigad; if possible, prefer this. * moved legacy theories Legacy_GCD and Primes from Library/ to Old_Number_Theory/; * moved theory Pocklington from Library/ to Old_Number_Theory/; * removed various references to Old_Number_Theory from HOL distribution. INCOMPATIBILITY. +* Theory GCD now has functions Gcd/GCD and Lcm/LCM for the gcd and lcm +of finite and infinite sets. It is shown that they form a complete +lattice. + +* Split off prime number ingredients from theory GCD +to theory Number_Theory/Primes; + +* Class semiring_div requires superclass no_zero_divisors and proof of +div_mult_mult1; theorems div_mult_mult1, div_mult_mult2, +div_mult_mult1_if, div_mult_mult1 and div_mult_mult2 have been +generalized to class semiring_div, subsuming former theorems +zdiv_zmult_zmult1, zdiv_zmult_zmult1_if, zdiv_zmult_zmult1 and +zdiv_zmult_zmult2. div_mult_mult1 is now [simp] by default. +INCOMPATIBILITY. + * New testing tool "Mirabelle" for automated (proof) tools. Applies several tools and tactics like sledgehammer, metis, or quickcheck, to every proof step in a theory. To be used in batch mode via the @@ -47,16 +68,15 @@ etc. INCOMPATIBILITY. -* New class "boolean_algebra". - * Refinements to lattice classes and sets: - less default intro/elim rules in locale variant, more default intro/elim rules in class variant: more uniformity - lemma ge_sup_conv renamed to le_sup_iff, in accordance with le_inf_iff - dropped lemma alias inf_ACI for inf_aci (same for sup_ACI and sup_aci) - renamed ACI to inf_sup_aci + - new class "boolean_algebra" - class "complete_lattice" moved to separate theory "complete_lattice"; - corresponding constants (and abbreviations) renamed: + corresponding constants (and abbreviations) renamed and with authentic syntax: Set.Inf ~> Complete_Lattice.Inf Set.Sup ~> Complete_Lattice.Sup Set.INFI ~> Complete_Lattice.INFI @@ -66,24 +86,22 @@ Set.INTER ~> Complete_Lattice.INTER Set.UNION ~> Complete_Lattice.UNION - more convenient names for set intersection and union: - Set.Int ~> Set.inter - Set.Un ~> Set.union + Set.Int ~> Set.inter + Set.Un ~> Set.union + - authentic syntax for + Set.Pow + Set.image - mere abbreviations: Set.empty (for bot) Set.UNIV (for top) Complete_Lattice.Inter (for Inf) Complete_Lattice.Union (for Sup) + Complete_Lattice.INTER (for INFI) + Complete_Lattice.UNION (for SUPR) + - object-logic definitions as far as appropriate INCOMPATIBILITY. -* Class semiring_div requires superclass no_zero_divisors and proof of -div_mult_mult1; theorems div_mult_mult1, div_mult_mult2, -div_mult_mult1_if, div_mult_mult1 and div_mult_mult2 have been -generalized to class semiring_div, subsuming former theorems -zdiv_zmult_zmult1, zdiv_zmult_zmult1_if, zdiv_zmult_zmult1 and -zdiv_zmult_zmult2. div_mult_mult1 is now [simp] by default. -INCOMPATIBILITY. - * Power operations on relations and functions are now one dedicate constant "compow" with infix syntax "^^". Power operations on multiplicative monoids retains syntax "^" and is now defined generic @@ -96,10 +114,6 @@ this. Fix using O_assoc[symmetric]. The same applies to the curried version "R OO S". -* Theory GCD now has functions Gcd/GCD and Lcm/LCM for the gcd and lcm -of finite and infinite sets. It is shown that they form a complete -lattice. - * ML antiquotation @{code_datatype} inserts definition of a datatype generated by the code generator; see Predicate.thy for an example. @@ -110,10 +124,6 @@ default generators are provided for all suitable HOL types, records and datatypes. -* Constants Set.Pow and Set.image now with authentic syntax; -object-logic definitions Set.Pow_def and Set.image_def. -INCOMPATIBILITY. - * Renamed theorems: Suc_eq_add_numeral_1 -> Suc_eq_plus1 Suc_eq_add_numeral_1_left -> Suc_eq_plus1_left @@ -127,6 +137,9 @@ ATPs down considerably but eliminates a source of unsound "proofs" that fail later. +* New method metisFT: A version of metis that uses full type information +in order to avoid failures of proof reconstruction. + * Discontinued ancient tradition to suffix certain ML module names with "_package", e.g.: @@ -135,9 +148,6 @@ INCOMPATIBILITY. -* NewNumberTheory: Jeremy Avigad's new version of part of -NumberTheory. If possible, use NewNumberTheory, not NumberTheory. - * Discontinued abbreviation "arbitrary" of constant "undefined". INCOMPATIBILITY, use "undefined" directly. diff -r c876bcb601fc -r 5b65449d7669 etc/components --- a/etc/components Thu Sep 17 14:17:37 2009 +1000 +++ b/etc/components Tue Sep 22 13:52:19 2009 +1000 @@ -15,3 +15,4 @@ src/HOL/Tools/ATP_Manager src/HOL/Mirabelle src/HOL/Library/Sum_Of_Squares +src/HOL/SMT diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Auth/All_Symmetric.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Auth/All_Symmetric.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,12 @@ +theory All_Symmetric +imports Message +begin + +text {* All keys are symmetric *} + +defs all_symmetric_def: "all_symmetric \ True" + +lemma isSym_keys: "K \ symKeys" + by (simp add: symKeys_def all_symmetric_def invKey_symmetric) + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Auth/Auth_Public.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Auth/Auth_Public.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,15 @@ +(* Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1996 University of Cambridge +*) + +header {* Conventional protocols: rely on conventional Message, Event and Public -- Public-key protocols *} + +theory Auth_Public +imports + "NS_Public_Bad" + "NS_Public" + "TLS" + "CertifiedEmail" +begin + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Auth/Auth_Shared.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Auth/Auth_Shared.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,27 @@ +(* Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1996 University of Cambridge +*) + +header {* Conventional protocols: rely on conventional Message, Event and Public -- Shared-key protocols *} + +theory Auth_Shared +imports + "NS_Shared" + "Kerberos_BAN" + "Kerberos_BAN_Gets" + "KerberosIV" + "KerberosIV_Gets" + "KerberosV" + "OtwayRees" + "OtwayRees_AN" + "OtwayRees_Bad" + "OtwayReesBella" + "WooLam" + "Recur" + "Yahalom" + "Yahalom2" + "Yahalom_Bad" + "ZhouGollmann" +begin + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Auth/Guard/Auth_Guard_Public.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Auth/Guard/Auth_Guard_Public.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,15 @@ +(* Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1996 University of Cambridge +*) + +header {* Blanqui's "guard" concept: protocol-independent secrecy *} + +theory Auth_Guard_Public +imports + "P1" + "P2" + "Guard_NS_Public" + "Proto" +begin + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Auth/Guard/Auth_Guard_Shared.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Auth/Guard/Auth_Guard_Shared.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,13 @@ +(* Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1996 University of Cambridge +*) + +header {* Blanqui's "guard" concept: protocol-independent secrecy *} + +theory Auth_Guard_Shared +imports + "Guard_OtwayRees" + "Guard_Yahalom" +begin + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Auth/Public.thy --- a/src/HOL/Auth/Public.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Auth/Public.thy Tue Sep 22 13:52:19 2009 +1000 @@ -1,5 +1,4 @@ (* Title: HOL/Auth/Public - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1996 University of Cambridge @@ -8,7 +7,9 @@ Private and public keys; initial states of agents *) -theory Public imports Event begin +theory Public +imports Event +begin lemma invKey_K: "K \ symKeys ==> invKey K = K" by (simp add: symKeys_def) diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Auth/ROOT.ML --- a/src/HOL/Auth/ROOT.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Auth/ROOT.ML Tue Sep 22 13:52:19 2009 +1000 @@ -1,51 +1,2 @@ -(* Title: HOL/Auth/ROOT.ML - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory - Copyright 1996 University of Cambridge -Root file for protocol proofs. -*) - -use_thys [ - -(* Conventional protocols: rely on - conventional Message, Event and Public *) - -(*Shared-key protocols*) - "NS_Shared", - "Kerberos_BAN", - "Kerberos_BAN_Gets", - "KerberosIV", - "KerberosIV_Gets", - "KerberosV", - "OtwayRees", - "OtwayRees_AN", - "OtwayRees_Bad", - "OtwayReesBella", - "WooLam", - "Recur", - "Yahalom", - "Yahalom2", - "Yahalom_Bad", - "ZhouGollmann", - -(*Public-key protocols*) - "NS_Public_Bad", - "NS_Public", - "TLS", - "CertifiedEmail", - -(*Smartcard protocols: rely on conventional Message and on new - EventSC and Smartcard *) - - "Smartcard/ShoupRubin", - "Smartcard/ShoupRubinBella", - -(*Blanqui's "guard" concept: protocol-independent secrecy*) - "Guard/P1", - "Guard/P2", - "Guard/Guard_NS_Public", - "Guard/Guard_OtwayRees", - "Guard/Guard_Yahalom", - "Guard/Proto" -]; +use_thys ["Auth_Shared", "Auth_Public", "Smartcard/Auth_Smartcard", "Guard/Auth_Guard_Shared", "Guard/Auth_Guard_Public"]; diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Auth/Shared.thy --- a/src/HOL/Auth/Shared.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Auth/Shared.thy Tue Sep 22 13:52:19 2009 +1000 @@ -1,5 +1,4 @@ (* Title: HOL/Auth/Shared - ID: $Id$ Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1996 University of Cambridge @@ -8,7 +7,9 @@ Shared, long-term keys; initial states of agents *) -theory Shared imports Event begin +theory Shared +imports Event All_Symmetric +begin consts shrK :: "agent => key" (*symmetric keys*); @@ -20,13 +21,6 @@ apply (simp add: inj_on_def split: agent.split) done -text{*All keys are symmetric*} - -defs all_symmetric_def: "all_symmetric == True" - -lemma isSym_keys: "K \ symKeys" -by (simp add: symKeys_def all_symmetric_def invKey_symmetric) - text{*Server knows all long-term keys; other agents know only their own*} primrec initState_Server: "initState Server = Key ` range shrK" diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Auth/Smartcard/Auth_Smartcard.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Auth/Smartcard/Auth_Smartcard.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,13 @@ +(* Author: Lawrence C Paulson, Cambridge University Computer Laboratory + Copyright 1996 University of Cambridge +*) + +header {* Smartcard protocols: rely on conventional Message and on new EventSC and Smartcard *} + +theory Auth_Smartcard +imports + "ShoupRubin" + "ShoupRubinBella" +begin + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Auth/Smartcard/Smartcard.thy --- a/src/HOL/Auth/Smartcard/Smartcard.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Auth/Smartcard/Smartcard.thy Tue Sep 22 13:52:19 2009 +1000 @@ -1,10 +1,11 @@ -(* ID: $Id$ - Author: Giampaolo Bella, Catania University +(* Author: Giampaolo Bella, Catania University *) header{*Theory of smartcards*} -theory Smartcard imports EventSC begin +theory Smartcard +imports EventSC All_Symmetric +begin text{* As smartcards handle long-term (symmetric) keys, this theoy extends and @@ -42,14 +43,6 @@ shrK_disj_pin [iff]: "shrK P \ pin Q" crdK_disj_pin [iff]: "crdK C \ pin P" - -text{*All keys are symmetric*} -defs all_symmetric_def: "all_symmetric == True" - -lemma isSym_keys: "K \ symKeys" -by (simp add: symKeys_def all_symmetric_def invKey_symmetric) - - constdefs legalUse :: "card => bool" ("legalUse (_)") "legalUse C == C \ stolen" diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Bali/Bali.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Bali/Bali.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,11 @@ +(* Author: David von Oheimb + Copyright 1999 Technische Universitaet Muenchen +*) + +header {* The Hoare logic for Bali. *} + +theory Bali +imports AxExample AxSound AxCompl Trans +begin + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Bali/ROOT.ML --- a/src/HOL/Bali/ROOT.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Bali/ROOT.ML Tue Sep 22 13:52:19 2009 +1000 @@ -1,9 +1,2 @@ -(* Title: HOL/Bali/ROOT.ML - ID: $Id$ - Author: David von Oheimb - Copyright 1999 Technische Universitaet Muenchen -The Hoare logic for Bali. -*) - -use_thys ["AxExample", "AxSound", "AxCompl", "Trans"]; +use_thy "Bali" diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Complete_Lattice.thy --- a/src/HOL/Complete_Lattice.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Complete_Lattice.thy Tue Sep 22 13:52:19 2009 +1000 @@ -278,8 +278,8 @@ subsection {* Unions of families *} -definition UNION :: "'a set \ ('a \ 'b set) \ 'b set" where - SUPR_set_eq [symmetric]: "UNION S f = (SUP x:S. f x)" +abbreviation UNION :: "'a set \ ('a \ 'b set) \ 'b set" where + "UNION \ SUPR" syntax "@UNION1" :: "pttrns => 'b set => 'b set" ("(3UN _./ _)" [0, 10] 10) @@ -314,7 +314,7 @@ lemma UNION_eq_Union_image: "(\x\A. B x) = \(B`A)" - by (simp add: SUPR_def SUPR_set_eq [symmetric]) + by (fact SUPR_def) lemma Union_def: "\S = (\x\S. x)" @@ -351,7 +351,7 @@ by blast lemma UN_upper: "a \ A ==> B a \ (\x\A. B x)" - by blast + by (fact le_SUPI) lemma UN_least: "(!!x. x \ A ==> B x \ C) ==> (\x\A. B x) \ C" by (iprover intro: subsetI elim: UN_E dest: subsetD) @@ -514,8 +514,8 @@ subsection {* Intersections of families *} -definition INTER :: "'a set \ ('a \ 'b set) \ 'b set" where - INFI_set_eq [symmetric]: "INTER S f = (INF x:S. f x)" +abbreviation INTER :: "'a set \ ('a \ 'b set) \ 'b set" where + "INTER \ INFI" syntax "@INTER1" :: "pttrns => 'b set => 'b set" ("(3INT _./ _)" [0, 10] 10) @@ -541,7 +541,7 @@ lemma INTER_eq_Inter_image: "(\x\A. B x) = \(B`A)" - by (simp add: INFI_def INFI_set_eq [symmetric]) + by (fact INFI_def) lemma Inter_def: "\S = (\x\S. x)" @@ -579,10 +579,10 @@ by blast lemma INT_lower: "a \ A ==> (\x\A. B x) \ B a" - by blast + by (fact INF_leI) lemma INT_greatest: "(!!x. x \ A ==> C \ B x) ==> C \ (\x\A. B x)" - by (iprover intro: INT_I subsetI dest: subsetD) + by (fact le_INFI) lemma INT_empty [simp]: "(\x\{}. B x) = UNIV" by blast diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/Gar_Coll.thy --- a/src/HOL/HoareParallel/Gar_Coll.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,846 +0,0 @@ - -header {* \section{The Single Mutator Case} *} - -theory Gar_Coll imports Graph OG_Syntax begin - -declare psubsetE [rule del] - -text {* Declaration of variables: *} - -record gar_coll_state = - M :: nodes - E :: edges - bc :: "nat set" - obc :: "nat set" - Ma :: nodes - ind :: nat - k :: nat - z :: bool - -subsection {* The Mutator *} - -text {* The mutator first redirects an arbitrary edge @{text "R"} from -an arbitrary accessible node towards an arbitrary accessible node -@{text "T"}. It then colors the new target @{text "T"} black. - -We declare the arbitrarily selected node and edge as constants:*} - -consts R :: nat T :: nat - -text {* \noindent The following predicate states, given a list of -nodes @{text "m"} and a list of edges @{text "e"}, the conditions -under which the selected edge @{text "R"} and node @{text "T"} are -valid: *} - -constdefs - Mut_init :: "gar_coll_state \ bool" - "Mut_init \ \ T \ Reach \E \ R < length \E \ T < length \M \" - -text {* \noindent For the mutator we -consider two modules, one for each action. An auxiliary variable -@{text "\z"} is set to false if the mutator has already redirected an -edge but has not yet colored the new target. *} - -constdefs - Redirect_Edge :: "gar_coll_state ann_com" - "Redirect_Edge \ .{\Mut_init \ \z}. \\E:=\E[R:=(fst(\E!R), T)],, \z:= (\\z)\" - - Color_Target :: "gar_coll_state ann_com" - "Color_Target \ .{\Mut_init \ \\z}. \\M:=\M[T:=Black],, \z:= (\\z)\" - - Mutator :: "gar_coll_state ann_com" - "Mutator \ - .{\Mut_init \ \z}. - WHILE True INV .{\Mut_init \ \z}. - DO Redirect_Edge ;; Color_Target OD" - -subsubsection {* Correctness of the mutator *} - -lemmas mutator_defs = Mut_init_def Redirect_Edge_def Color_Target_def - -lemma Redirect_Edge: - "\ Redirect_Edge pre(Color_Target)" -apply (unfold mutator_defs) -apply annhoare -apply(simp_all) -apply(force elim:Graph2) -done - -lemma Color_Target: - "\ Color_Target .{\Mut_init \ \z}." -apply (unfold mutator_defs) -apply annhoare -apply(simp_all) -done - -lemma Mutator: - "\ Mutator .{False}." -apply(unfold Mutator_def) -apply annhoare -apply(simp_all add:Redirect_Edge Color_Target) -apply(simp add:mutator_defs Redirect_Edge_def) -done - -subsection {* The Collector *} - -text {* \noindent A constant @{text "M_init"} is used to give @{text "\Ma"} a -suitable first value, defined as a list of nodes where only the @{text -"Roots"} are black. *} - -consts M_init :: nodes - -constdefs - Proper_M_init :: "gar_coll_state \ bool" - "Proper_M_init \ \ Blacks M_init=Roots \ length M_init=length \M \" - - Proper :: "gar_coll_state \ bool" - "Proper \ \ Proper_Roots \M \ Proper_Edges(\M, \E) \ \Proper_M_init \" - - Safe :: "gar_coll_state \ bool" - "Safe \ \ Reach \E \ Blacks \M \" - -lemmas collector_defs = Proper_M_init_def Proper_def Safe_def - -subsubsection {* Blackening the roots *} - -constdefs - Blacken_Roots :: " gar_coll_state ann_com" - "Blacken_Roots \ - .{\Proper}. - \ind:=0;; - .{\Proper \ \ind=0}. - WHILE \indM - INV .{\Proper \ (\i<\ind. i \ Roots \ \M!i=Black) \ \ind\length \M}. - DO .{\Proper \ (\i<\ind. i \ Roots \ \M!i=Black) \ \indM}. - IF \ind\Roots THEN - .{\Proper \ (\i<\ind. i \ Roots \ \M!i=Black) \ \indM \ \ind\Roots}. - \M:=\M[\ind:=Black] FI;; - .{\Proper \ (\i<\ind+1. i \ Roots \ \M!i=Black) \ \indM}. - \ind:=\ind+1 - OD" - -lemma Blacken_Roots: - "\ Blacken_Roots .{\Proper \ Roots\Blacks \M}." -apply (unfold Blacken_Roots_def) -apply annhoare -apply(simp_all add:collector_defs Graph_defs) -apply safe -apply(simp_all add:nth_list_update) - apply (erule less_SucE) - apply simp+ - apply force -apply force -done - -subsubsection {* Propagating black *} - -constdefs - PBInv :: "gar_coll_state \ nat \ bool" - "PBInv \ \ \ind. \obc < Blacks \M \ (\i BtoW (\E!i, \M) \ - (\\z \ i=R \ (snd(\E!R)) = T \ (\r. ind \ r \ r < length \E \ BtoW(\E!r,\M))))\" - -constdefs - Propagate_Black_aux :: "gar_coll_state ann_com" - "Propagate_Black_aux \ - .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M}. - \ind:=0;; - .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M \ \ind=0}. - WHILE \indE - INV .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ \PBInv \ind \ \ind\length \E}. - DO .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ \PBInv \ind \ \indE}. - IF \M!(fst (\E!\ind)) = Black THEN - .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ \PBInv \ind \ \indE \ \M!fst(\E!\ind)=Black}. - \M:=\M[snd(\E!\ind):=Black];; - .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ \PBInv (\ind + 1) \ \indE}. - \ind:=\ind+1 - FI - OD" - -lemma Propagate_Black_aux: - "\ Propagate_Black_aux - .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ ( \obc < Blacks \M \ \Safe)}." -apply (unfold Propagate_Black_aux_def PBInv_def collector_defs) -apply annhoare -apply(simp_all add:Graph6 Graph7 Graph8 Graph12) - apply force - apply force - apply force ---{* 4 subgoals left *} -apply clarify -apply(simp add:Proper_Edges_def Proper_Roots_def Graph6 Graph7 Graph8 Graph12) -apply (erule disjE) - apply(rule disjI1) - apply(erule Graph13) - apply force -apply (case_tac "M x ! snd (E x ! ind x)=Black") - apply (simp add: Graph10 BtoW_def) - apply (rule disjI2) - apply clarify - apply (erule less_SucE) - apply (erule_tac x=i in allE , erule (1) notE impE) - apply simp - apply clarify - apply (drule_tac y = r in le_imp_less_or_eq) - apply (erule disjE) - apply (subgoal_tac "Suc (ind x)\r") - apply fast - apply arith - apply fast - apply fast -apply(rule disjI1) -apply(erule subset_psubset_trans) -apply(erule Graph11) -apply fast ---{* 3 subgoals left *} -apply force -apply force ---{* last *} -apply clarify -apply simp -apply(subgoal_tac "ind x = length (E x)") - apply (rotate_tac -1) - apply (simp (asm_lr)) - apply(drule Graph1) - apply simp - apply clarify - apply(erule allE, erule impE, assumption) - apply force - apply force -apply arith -done - -subsubsection {* Refining propagating black *} - -constdefs - Auxk :: "gar_coll_state \ bool" - "Auxk \ \\kM \ (\M!\k\Black \ \BtoW(\E!\ind, \M) \ - \obcM \ (\\z \ \ind=R \ snd(\E!R)=T - \ (\r. \ind rE \ BtoW(\E!r, \M))))\" - -constdefs - Propagate_Black :: " gar_coll_state ann_com" - "Propagate_Black \ - .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M}. - \ind:=0;; - .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M \ \ind=0}. - WHILE \indE - INV .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ \PBInv \ind \ \ind\length \E}. - DO .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ \PBInv \ind \ \indE}. - IF (\M!(fst (\E!\ind)))=Black THEN - .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ \PBInv \ind \ \indE \ (\M!fst(\E!\ind))=Black}. - \k:=(snd(\E!\ind));; - .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ \PBInv \ind \ \indE \ (\M!fst(\E!\ind))=Black - \ \Auxk}. - \\M:=\M[\k:=Black],, \ind:=\ind+1\ - ELSE .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ \PBInv \ind \ \indE}. - \IF (\M!(fst (\E!\ind)))\Black THEN \ind:=\ind+1 FI\ - FI - OD" - -lemma Propagate_Black: - "\ Propagate_Black - .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ ( \obc < Blacks \M \ \Safe)}." -apply (unfold Propagate_Black_def PBInv_def Auxk_def collector_defs) -apply annhoare -apply(simp_all add:Graph6 Graph7 Graph8 Graph12) - apply force - apply force - apply force ---{* 5 subgoals left *} -apply clarify -apply(simp add:BtoW_def Proper_Edges_def) ---{* 4 subgoals left *} -apply clarify -apply(simp add:Proper_Edges_def Graph6 Graph7 Graph8 Graph12) -apply (erule disjE) - apply (rule disjI1) - apply (erule psubset_subset_trans) - apply (erule Graph9) -apply (case_tac "M x!k x=Black") - apply (case_tac "M x ! snd (E x ! ind x)=Black") - apply (simp add: Graph10 BtoW_def) - apply (rule disjI2) - apply clarify - apply (erule less_SucE) - apply (erule_tac x=i in allE , erule (1) notE impE) - apply simp - apply clarify - apply (drule_tac y = r in le_imp_less_or_eq) - apply (erule disjE) - apply (subgoal_tac "Suc (ind x)\r") - apply fast - apply arith - apply fast - apply fast - apply (simp add: Graph10 BtoW_def) - apply (erule disjE) - apply (erule disjI1) - apply clarify - apply (erule less_SucE) - apply force - apply simp - apply (subgoal_tac "Suc R\r") - apply fast - apply arith -apply(rule disjI1) -apply(erule subset_psubset_trans) -apply(erule Graph11) -apply fast ---{* 3 subgoals left *} -apply force ---{* 2 subgoals left *} -apply clarify -apply(simp add:Proper_Edges_def Graph6 Graph7 Graph8 Graph12) -apply (erule disjE) - apply fast -apply clarify -apply (erule less_SucE) - apply (erule_tac x=i in allE , erule (1) notE impE) - apply simp - apply clarify - apply (drule_tac y = r in le_imp_less_or_eq) - apply (erule disjE) - apply (subgoal_tac "Suc (ind x)\r") - apply fast - apply arith - apply (simp add: BtoW_def) -apply (simp add: BtoW_def) ---{* last *} -apply clarify -apply simp -apply(subgoal_tac "ind x = length (E x)") - apply (rotate_tac -1) - apply (simp (asm_lr)) - apply(drule Graph1) - apply simp - apply clarify - apply(erule allE, erule impE, assumption) - apply force - apply force -apply arith -done - -subsubsection {* Counting black nodes *} - -constdefs - CountInv :: "gar_coll_state \ nat \ bool" - "CountInv \ \ \ind. {i. i \Ma!i=Black}\\bc \" - -constdefs - Count :: " gar_coll_state ann_com" - "Count \ - .{\Proper \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ (\obc < Blacks \Ma \ \Safe) \ \bc={}}. - \ind:=0;; - .{\Proper \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ (\obc < Blacks \Ma \ \Safe) \ \bc={} - \ \ind=0}. - WHILE \indM - INV .{\Proper \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ \CountInv \ind - \ ( \obc < Blacks \Ma \ \Safe) \ \ind\length \M}. - DO .{\Proper \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ \CountInv \ind - \ ( \obc < Blacks \Ma \ \Safe) \ \indM}. - IF \M!\ind=Black - THEN .{\Proper \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ \CountInv \ind - \ ( \obc < Blacks \Ma \ \Safe) \ \indM \ \M!\ind=Black}. - \bc:=insert \ind \bc - FI;; - .{\Proper \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ \CountInv (\ind+1) - \ ( \obc < Blacks \Ma \ \Safe) \ \indM}. - \ind:=\ind+1 - OD" - -lemma Count: - "\ Count - .{\Proper \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\\bc \ \bc\Blacks \M \ length \Ma=length \M - \ (\obc < Blacks \Ma \ \Safe)}." -apply(unfold Count_def) -apply annhoare -apply(simp_all add:CountInv_def Graph6 Graph7 Graph8 Graph12 Blacks_def collector_defs) - apply force - apply force - apply force - apply clarify - apply simp - apply(fast elim:less_SucE) - apply clarify - apply simp - apply(fast elim:less_SucE) - apply force -apply force -done - -subsubsection {* Appending garbage nodes to the free list *} - -consts Append_to_free :: "nat \ edges \ edges" - -axioms - Append_to_free0: "length (Append_to_free (i, e)) = length e" - Append_to_free1: "Proper_Edges (m, e) - \ Proper_Edges (m, Append_to_free(i, e))" - Append_to_free2: "i \ Reach e - \ n \ Reach (Append_to_free(i, e)) = ( n = i \ n \ Reach e)" - -constdefs - AppendInv :: "gar_coll_state \ nat \ bool" - "AppendInv \ \\ind. \iM. ind\i \ i\Reach \E \ \M!i=Black\" - -constdefs - Append :: " gar_coll_state ann_com" - "Append \ - .{\Proper \ Roots\Blacks \M \ \Safe}. - \ind:=0;; - .{\Proper \ Roots\Blacks \M \ \Safe \ \ind=0}. - WHILE \indM - INV .{\Proper \ \AppendInv \ind \ \ind\length \M}. - DO .{\Proper \ \AppendInv \ind \ \indM}. - IF \M!\ind=Black THEN - .{\Proper \ \AppendInv \ind \ \indM \ \M!\ind=Black}. - \M:=\M[\ind:=White] - ELSE .{\Proper \ \AppendInv \ind \ \indM \ \ind\Reach \E}. - \E:=Append_to_free(\ind,\E) - FI;; - .{\Proper \ \AppendInv (\ind+1) \ \indM}. - \ind:=\ind+1 - OD" - -lemma Append: - "\ Append .{\Proper}." -apply(unfold Append_def AppendInv_def) -apply annhoare -apply(simp_all add:collector_defs Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) - apply(force simp:Blacks_def nth_list_update) - apply force - apply force - apply(force simp add:Graph_defs) - apply force - apply clarify - apply simp - apply(rule conjI) - apply (erule Append_to_free1) - apply clarify - apply (drule_tac n = "i" in Append_to_free2) - apply force - apply force -apply force -done - -subsubsection {* Correctness of the Collector *} - -constdefs - Collector :: " gar_coll_state ann_com" - "Collector \ -.{\Proper}. - WHILE True INV .{\Proper}. - DO - Blacken_Roots;; - .{\Proper \ Roots\Blacks \M}. - \obc:={};; - .{\Proper \ Roots\Blacks \M \ \obc={}}. - \bc:=Roots;; - .{\Proper \ Roots\Blacks \M \ \obc={} \ \bc=Roots}. - \Ma:=M_init;; - .{\Proper \ Roots\Blacks \M \ \obc={} \ \bc=Roots \ \Ma=M_init}. - WHILE \obc\\bc - INV .{\Proper \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\\bc \ \bc\Blacks \M - \ length \Ma=length \M \ (\obc < Blacks \Ma \ \Safe)}. - DO .{\Proper \ Roots\Blacks \M \ \bc\Blacks \M}. - \obc:=\bc;; - Propagate_Black;; - .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ (\obc < Blacks \M \ \Safe)}. - \Ma:=\M;; - .{\Proper \ Roots\Blacks \M \ \obc\Blacks \Ma - \ Blacks \Ma\Blacks \M \ \bc\Blacks \M \ length \Ma=length \M - \ ( \obc < Blacks \Ma \ \Safe)}. - \bc:={};; - Count - OD;; - Append - OD" - -lemma Collector: - "\ Collector .{False}." -apply(unfold Collector_def) -apply annhoare -apply(simp_all add: Blacken_Roots Propagate_Black Count Append) -apply(simp_all add:Blacken_Roots_def Propagate_Black_def Count_def Append_def collector_defs) - apply (force simp add: Proper_Roots_def) - apply force - apply force -apply clarify -apply (erule disjE) -apply(simp add:psubsetI) - apply(force dest:subset_antisym) -done - -subsection {* Interference Freedom *} - -lemmas modules = Redirect_Edge_def Color_Target_def Blacken_Roots_def - Propagate_Black_def Count_def Append_def -lemmas Invariants = PBInv_def Auxk_def CountInv_def AppendInv_def -lemmas abbrev = collector_defs mutator_defs Invariants - -lemma interfree_Blacken_Roots_Redirect_Edge: - "interfree_aux (Some Blacken_Roots, {}, Some Redirect_Edge)" -apply (unfold modules) -apply interfree_aux -apply safe -apply (simp_all add:Graph6 Graph12 abbrev) -done - -lemma interfree_Redirect_Edge_Blacken_Roots: - "interfree_aux (Some Redirect_Edge, {}, Some Blacken_Roots)" -apply (unfold modules) -apply interfree_aux -apply safe -apply(simp add:abbrev)+ -done - -lemma interfree_Blacken_Roots_Color_Target: - "interfree_aux (Some Blacken_Roots, {}, Some Color_Target)" -apply (unfold modules) -apply interfree_aux -apply safe -apply(simp_all add:Graph7 Graph8 nth_list_update abbrev) -done - -lemma interfree_Color_Target_Blacken_Roots: - "interfree_aux (Some Color_Target, {}, Some Blacken_Roots)" -apply (unfold modules ) -apply interfree_aux -apply safe -apply(simp add:abbrev)+ -done - -lemma interfree_Propagate_Black_Redirect_Edge: - "interfree_aux (Some Propagate_Black, {}, Some Redirect_Edge)" -apply (unfold modules ) -apply interfree_aux ---{* 11 subgoals left *} -apply(clarify, simp add:abbrev Graph6 Graph12) -apply(clarify, simp add:abbrev Graph6 Graph12) -apply(clarify, simp add:abbrev Graph6 Graph12) -apply(clarify, simp add:abbrev Graph6 Graph12) -apply(erule conjE)+ -apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) - apply(erule Graph4) - apply(simp)+ - apply (simp add:BtoW_def) - apply (simp add:BtoW_def) -apply(rule conjI) - apply (force simp add:BtoW_def) -apply(erule Graph4) - apply simp+ ---{* 7 subgoals left *} -apply(clarify, simp add:abbrev Graph6 Graph12) -apply(erule conjE)+ -apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) - apply(erule Graph4) - apply(simp)+ - apply (simp add:BtoW_def) - apply (simp add:BtoW_def) -apply(rule conjI) - apply (force simp add:BtoW_def) -apply(erule Graph4) - apply simp+ ---{* 6 subgoals left *} -apply(clarify, simp add:abbrev Graph6 Graph12) -apply(erule conjE)+ -apply(rule conjI) - apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) - apply(erule Graph4) - apply(simp)+ - apply (simp add:BtoW_def) - apply (simp add:BtoW_def) - apply(rule conjI) - apply (force simp add:BtoW_def) - apply(erule Graph4) - apply simp+ -apply(simp add:BtoW_def nth_list_update) -apply force ---{* 5 subgoals left *} -apply(clarify, simp add:abbrev Graph6 Graph12) ---{* 4 subgoals left *} -apply(clarify, simp add:abbrev Graph6 Graph12) -apply(rule conjI) - apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) - apply(erule Graph4) - apply(simp)+ - apply (simp add:BtoW_def) - apply (simp add:BtoW_def) - apply(rule conjI) - apply (force simp add:BtoW_def) - apply(erule Graph4) - apply simp+ -apply(rule conjI) - apply(simp add:nth_list_update) - apply force -apply(rule impI, rule impI, erule disjE, erule disjI1, case_tac "R = (ind x)" ,case_tac "M x ! T = Black") - apply(force simp add:BtoW_def) - apply(case_tac "M x !snd (E x ! ind x)=Black") - apply(rule disjI2) - apply simp - apply (erule Graph5) - apply simp+ - apply(force simp add:BtoW_def) -apply(force simp add:BtoW_def) ---{* 3 subgoals left *} -apply(clarify, simp add:abbrev Graph6 Graph12) ---{* 2 subgoals left *} -apply(clarify, simp add:abbrev Graph6 Graph12) -apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) - apply clarify - apply(erule Graph4) - apply(simp)+ - apply (simp add:BtoW_def) - apply (simp add:BtoW_def) -apply(rule conjI) - apply (force simp add:BtoW_def) -apply(erule Graph4) - apply simp+ -done - -lemma interfree_Redirect_Edge_Propagate_Black: - "interfree_aux (Some Redirect_Edge, {}, Some Propagate_Black)" -apply (unfold modules ) -apply interfree_aux -apply(clarify, simp add:abbrev)+ -done - -lemma interfree_Propagate_Black_Color_Target: - "interfree_aux (Some Propagate_Black, {}, Some Color_Target)" -apply (unfold modules ) -apply interfree_aux ---{* 11 subgoals left *} -apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)+ -apply(erule conjE)+ -apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, - case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, - erule allE, erule impE, assumption, erule impE, assumption, - simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) ---{* 7 subgoals left *} -apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) -apply(erule conjE)+ -apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, - case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, - erule allE, erule impE, assumption, erule impE, assumption, - simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) ---{* 6 subgoals left *} -apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) -apply clarify -apply (rule conjI) - apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, - case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, - erule allE, erule impE, assumption, erule impE, assumption, - simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) -apply(simp add:nth_list_update) ---{* 5 subgoals left *} -apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) ---{* 4 subgoals left *} -apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) -apply (rule conjI) - apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, - case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, - erule allE, erule impE, assumption, erule impE, assumption, - simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) -apply(rule conjI) -apply(simp add:nth_list_update) -apply(rule impI,rule impI, case_tac "M x!T=Black",rotate_tac -1, force simp add: BtoW_def Graph10, - erule subset_psubset_trans, erule Graph11, force) ---{* 3 subgoals left *} -apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) ---{* 2 subgoals left *} -apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) -apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, - case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, - erule allE, erule impE, assumption, erule impE, assumption, - simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) ---{* 3 subgoals left *} -apply(simp add:abbrev) -done - -lemma interfree_Color_Target_Propagate_Black: - "interfree_aux (Some Color_Target, {}, Some Propagate_Black)" -apply (unfold modules ) -apply interfree_aux -apply(clarify, simp add:abbrev)+ -done - -lemma interfree_Count_Redirect_Edge: - "interfree_aux (Some Count, {}, Some Redirect_Edge)" -apply (unfold modules) -apply interfree_aux ---{* 9 subgoals left *} -apply(simp_all add:abbrev Graph6 Graph12) ---{* 6 subgoals left *} -apply(clarify, simp add:abbrev Graph6 Graph12, - erule disjE,erule disjI1,rule disjI2,rule subset_trans, erule Graph3,force,force)+ -done - -lemma interfree_Redirect_Edge_Count: - "interfree_aux (Some Redirect_Edge, {}, Some Count)" -apply (unfold modules ) -apply interfree_aux -apply(clarify,simp add:abbrev)+ -apply(simp add:abbrev) -done - -lemma interfree_Count_Color_Target: - "interfree_aux (Some Count, {}, Some Color_Target)" -apply (unfold modules ) -apply interfree_aux ---{* 9 subgoals left *} -apply(simp_all add:abbrev Graph7 Graph8 Graph12) ---{* 6 subgoals left *} -apply(clarify,simp add:abbrev Graph7 Graph8 Graph12, - erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9)+ ---{* 2 subgoals left *} -apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) -apply(rule conjI) - apply(erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9) -apply(simp add:nth_list_update) ---{* 1 subgoal left *} -apply(clarify, simp add:abbrev Graph7 Graph8 Graph12, - erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9) -done - -lemma interfree_Color_Target_Count: - "interfree_aux (Some Color_Target, {}, Some Count)" -apply (unfold modules ) -apply interfree_aux -apply(clarify, simp add:abbrev)+ -apply(simp add:abbrev) -done - -lemma interfree_Append_Redirect_Edge: - "interfree_aux (Some Append, {}, Some Redirect_Edge)" -apply (unfold modules ) -apply interfree_aux -apply( simp_all add:abbrev Graph6 Append_to_free0 Append_to_free1 Graph12) -apply(clarify, simp add:abbrev Graph6 Append_to_free0 Append_to_free1 Graph12, force dest:Graph3)+ -done - -lemma interfree_Redirect_Edge_Append: - "interfree_aux (Some Redirect_Edge, {}, Some Append)" -apply (unfold modules ) -apply interfree_aux -apply(clarify, simp add:abbrev Append_to_free0)+ -apply (force simp add: Append_to_free2) -apply(clarify, simp add:abbrev Append_to_free0)+ -done - -lemma interfree_Append_Color_Target: - "interfree_aux (Some Append, {}, Some Color_Target)" -apply (unfold modules ) -apply interfree_aux -apply(clarify, simp add:abbrev Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12 nth_list_update)+ -apply(simp add:abbrev Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12 nth_list_update) -done - -lemma interfree_Color_Target_Append: - "interfree_aux (Some Color_Target, {}, Some Append)" -apply (unfold modules ) -apply interfree_aux -apply(clarify, simp add:abbrev Append_to_free0)+ -apply (force simp add: Append_to_free2) -apply(clarify,simp add:abbrev Append_to_free0)+ -done - -lemmas collector_mutator_interfree = - interfree_Blacken_Roots_Redirect_Edge interfree_Blacken_Roots_Color_Target - interfree_Propagate_Black_Redirect_Edge interfree_Propagate_Black_Color_Target - interfree_Count_Redirect_Edge interfree_Count_Color_Target - interfree_Append_Redirect_Edge interfree_Append_Color_Target - interfree_Redirect_Edge_Blacken_Roots interfree_Color_Target_Blacken_Roots - interfree_Redirect_Edge_Propagate_Black interfree_Color_Target_Propagate_Black - interfree_Redirect_Edge_Count interfree_Color_Target_Count - interfree_Redirect_Edge_Append interfree_Color_Target_Append - -subsubsection {* Interference freedom Collector-Mutator *} - -lemma interfree_Collector_Mutator: - "interfree_aux (Some Collector, {}, Some Mutator)" -apply(unfold Collector_def Mutator_def) -apply interfree_aux -apply(simp_all add:collector_mutator_interfree) -apply(unfold modules collector_defs mutator_defs) -apply(tactic {* TRYALL (interfree_aux_tac) *}) ---{* 32 subgoals left *} -apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) ---{* 20 subgoals left *} -apply(tactic{* TRYALL (clarify_tac @{claset}) *}) -apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) -apply(tactic {* TRYALL (etac disjE) *}) -apply simp_all -apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac subset_trans,etac @{thm Graph3},force_tac @{clasimpset}, assume_tac]) *}) -apply(tactic {* TRYALL(EVERY'[rtac disjI2,etac subset_trans,rtac @{thm Graph9},force_tac @{clasimpset}]) *}) -apply(tactic {* TRYALL(EVERY'[rtac disjI1,etac @{thm psubset_subset_trans},rtac @{thm Graph9},force_tac @{clasimpset}]) *}) -done - -subsubsection {* Interference freedom Mutator-Collector *} - -lemma interfree_Mutator_Collector: - "interfree_aux (Some Mutator, {}, Some Collector)" -apply(unfold Collector_def Mutator_def) -apply interfree_aux -apply(simp_all add:collector_mutator_interfree) -apply(unfold modules collector_defs mutator_defs) -apply(tactic {* TRYALL (interfree_aux_tac) *}) ---{* 64 subgoals left *} -apply(simp_all add:nth_list_update Invariants Append_to_free0)+ -apply(tactic{* TRYALL (clarify_tac @{claset}) *}) ---{* 4 subgoals left *} -apply force -apply(simp add:Append_to_free2) -apply force -apply(simp add:Append_to_free2) -done - -subsubsection {* The Garbage Collection algorithm *} - -text {* In total there are 289 verification conditions. *} - -lemma Gar_Coll: - "\- .{\Proper \ \Mut_init \ \z}. - COBEGIN - Collector - .{False}. - \ - Mutator - .{False}. - COEND - .{False}." -apply oghoare -apply(force simp add: Mutator_def Collector_def modules) -apply(rule Collector) -apply(rule Mutator) -apply(simp add:interfree_Collector_Mutator) -apply(simp add:interfree_Mutator_Collector) -apply force -done - -end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/Graph.thy --- a/src/HOL/HoareParallel/Graph.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,410 +0,0 @@ -header {* \chapter{Case Study: Single and Multi-Mutator Garbage Collection Algorithms} - -\section {Formalization of the Memory} *} - -theory Graph imports Main begin - -datatype node = Black | White - -types - nodes = "node list" - edge = "nat \ nat" - edges = "edge list" - -consts Roots :: "nat set" - -constdefs - Proper_Roots :: "nodes \ bool" - "Proper_Roots M \ Roots\{} \ Roots \ {i. i edges) \ bool" - "Proper_Edges \ (\(M,E). \i snd(E!i) nodes) \ bool" - "BtoW \ (\(e,M). (M!fst e)=Black \ (M!snd e)\Black)" - - Blacks :: "nodes \ nat set" - "Blacks M \ {i. i M!i=Black}" - - Reach :: "edges \ nat set" - "Reach E \ {x. (\path. 1 path!(length path - 1)\Roots \ x=path!0 - \ (\ij x\Roots}" - -text{* Reach: the set of reachable nodes is the set of Roots together with the -nodes reachable from some Root by a path represented by a list of - nodes (at least two since we traverse at least one edge), where two -consecutive nodes correspond to an edge in E. *} - -subsection {* Proofs about Graphs *} - -lemmas Graph_defs= Blacks_def Proper_Roots_def Proper_Edges_def BtoW_def -declare Graph_defs [simp] - -subsubsection{* Graph 1 *} - -lemma Graph1_aux [rule_format]: - "\ Roots\Blacks M; \iBtoW(E!i,M)\ - \ 1< length path \ (path!(length path - 1))\Roots \ - (\ij. j < length E \ E!j=(path!(Suc i), path!i))) - \ M!(path!0) = Black" -apply(induct_tac "path") - apply force -apply clarify -apply simp -apply(case_tac "list") - apply force -apply simp -apply(rotate_tac -2) -apply(erule_tac x = "0" in all_dupE) -apply simp -apply clarify -apply(erule allE , erule (1) notE impE) -apply simp -apply(erule mp) -apply(case_tac "lista") - apply force -apply simp -apply(erule mp) -apply clarify -apply(erule_tac x = "Suc i" in allE) -apply force -done - -lemma Graph1: - "\Roots\Blacks M; Proper_Edges(M, E); \iBtoW(E!i,M) \ - \ Reach E\Blacks M" -apply (unfold Reach_def) -apply simp -apply clarify -apply(erule disjE) - apply clarify - apply(rule conjI) - apply(subgoal_tac "0< length path - Suc 0") - apply(erule allE , erule (1) notE impE) - apply force - apply simp - apply(rule Graph1_aux) -apply auto -done - -subsubsection{* Graph 2 *} - -lemma Ex_first_occurrence [rule_format]: - "P (n::nat) \ (\m. P m \ (\i. i \ P i))"; -apply(rule nat_less_induct) -apply clarify -apply(case_tac "\m. m \ P m") -apply auto -done - -lemma Compl_lemma: "(n::nat)\l \ (\m. m\l \ n=l - m)" -apply(rule_tac x = "l - n" in exI) -apply arith -done - -lemma Ex_last_occurrence: - "\P (n::nat); n\l\ \ (\m. P (l - m) \ (\i. i \P (l - i)))" -apply(drule Compl_lemma) -apply clarify -apply(erule Ex_first_occurrence) -done - -lemma Graph2: - "\T \ Reach E; R \ T \ Reach (E[R:=(fst(E!R), T)])" -apply (unfold Reach_def) -apply clarify -apply simp -apply(case_tac "\zpath!z") - apply(rule_tac x = "path" in exI) - apply simp - apply clarify - apply(erule allE , erule (1) notE impE) - apply clarify - apply(rule_tac x = "j" in exI) - apply(case_tac "j=R") - apply(erule_tac x = "Suc i" in allE) - apply simp - apply (force simp add:nth_list_update) -apply simp -apply(erule exE) -apply(subgoal_tac "z \ length path - Suc 0") - prefer 2 apply arith -apply(drule_tac P = "\m. m fst(E!R)=path!m" in Ex_last_occurrence) - apply assumption -apply clarify -apply simp -apply(rule_tac x = "(path!0)#(drop (length path - Suc m) path)" in exI) -apply simp -apply(case_tac "length path - (length path - Suc m)") - apply arith -apply simp -apply(subgoal_tac "(length path - Suc m) + nat \ length path") - prefer 2 apply arith -apply(drule nth_drop) -apply simp -apply(subgoal_tac "length path - Suc m + nat = length path - Suc 0") - prefer 2 apply arith -apply simp -apply clarify -apply(case_tac "i") - apply(force simp add: nth_list_update) -apply simp -apply(subgoal_tac "(length path - Suc m) + nata \ length path") - prefer 2 apply arith -apply(subgoal_tac "(length path - Suc m) + (Suc nata) \ length path") - prefer 2 apply arith -apply simp -apply(erule_tac x = "length path - Suc m + nata" in allE) -apply simp -apply clarify -apply(rule_tac x = "j" in exI) -apply(case_tac "R=j") - prefer 2 apply force -apply simp -apply(drule_tac t = "path ! (length path - Suc m)" in sym) -apply simp -apply(case_tac " length path - Suc 0 < m") - apply(subgoal_tac "(length path - Suc m)=0") - prefer 2 apply arith - apply(simp del: diff_is_0_eq) - apply(subgoal_tac "Suc nata\nat") - prefer 2 apply arith - apply(drule_tac n = "Suc nata" in Compl_lemma) - apply clarify - using [[linarith_split_limit = 0]] - apply force - using [[linarith_split_limit = 9]] -apply(drule leI) -apply(subgoal_tac "Suc (length path - Suc m + nata)=(length path - Suc 0) - (m - Suc nata)") - apply(erule_tac x = "m - (Suc nata)" in allE) - apply(case_tac "m") - apply simp - apply simp -apply simp -done - - -subsubsection{* Graph 3 *} - -lemma Graph3: - "\ T\Reach E; R \ Reach(E[R:=(fst(E!R),T)]) \ Reach E" -apply (unfold Reach_def) -apply clarify -apply simp -apply(case_tac "\i(length path\m)") - prefer 2 apply arith - apply(simp) - apply(rule conjI) - apply(subgoal_tac "\(m + length patha - 1 < m)") - prefer 2 apply arith - apply(simp add: nth_append) - apply(rule conjI) - apply(case_tac "m") - apply force - apply(case_tac "path") - apply force - apply force - apply clarify - apply(case_tac "Suc i\m") - apply(erule_tac x = "i" in allE) - apply simp - apply clarify - apply(rule_tac x = "j" in exI) - apply(case_tac "Suc i(length path\Suc m)" ) - prefer 2 apply arith - apply clarsimp - apply(erule_tac x = "i" in allE) - apply simp - apply clarify - apply(case_tac "R=j") - apply(force simp add: nth_list_update) - apply(force simp add: nth_list_update) ---{* the changed edge is not part of the path *} -apply(rule_tac x = "path" in exI) -apply simp -apply clarify -apply(erule_tac x = "i" in allE) -apply clarify -apply(case_tac "R=j") - apply(erule_tac x = "i" in allE) - apply simp -apply(force simp add: nth_list_update) -done - -subsubsection{* Graph 4 *} - -lemma Graph4: - "\T \ Reach E; Roots\Blacks M; I\length E; TiBtoW(E!i,M); RBlack\ \ - (\r. I\r \ r BtoW(E[R:=(fst(E!R),T)]!r,M))" -apply (unfold Reach_def) -apply simp -apply(erule disjE) - prefer 2 apply force -apply clarify ---{* there exist a black node in the path to T *} -apply(case_tac "\m T \ Reach E ; Roots \ Blacks M; \iBtoW(E!i,M); T Black\ - \ (\r. R r BtoW(E[R:=(fst(E!R),T)]!r,M))" -apply (unfold Reach_def) -apply simp -apply(erule disjE) - prefer 2 apply force -apply clarify ---{* there exist a black node in the path to T*} -apply(case_tac "\mR") - apply(drule le_imp_less_or_eq [of _ R]) - apply(erule disjE) - apply(erule allE , erule (1) notE impE) - apply force - apply force - apply(rule_tac x = "j" in exI) - apply(force simp add: nth_list_update) -apply simp -apply(rotate_tac -1) -apply(erule_tac x = "length path - 1" in allE) -apply(case_tac "length path") - apply force -apply force -done - -subsubsection {* Other lemmas about graphs *} - -lemma Graph6: - "\Proper_Edges(M,E); R \ Proper_Edges(M,E[R:=(fst(E!R),T)])" -apply (unfold Proper_Edges_def) - apply(force simp add: nth_list_update) -done - -lemma Graph7: - "\Proper_Edges(M,E)\ \ Proper_Edges(M[T:=a],E)" -apply (unfold Proper_Edges_def) -apply force -done - -lemma Graph8: - "\Proper_Roots(M)\ \ Proper_Roots(M[T:=a])" -apply (unfold Proper_Roots_def) -apply force -done - -text{* Some specific lemmata for the verification of garbage collection algorithms. *} - -lemma Graph9: "j Blacks M\Blacks (M[j := Black])" -apply (unfold Blacks_def) - apply(force simp add: nth_list_update) -done - -lemma Graph10 [rule_format (no_asm)]: "\i. M!i=a \M[i:=a]=M" -apply(induct_tac "M") -apply auto -apply(case_tac "i") -apply auto -done - -lemma Graph11 [rule_format (no_asm)]: - "\ M!j\Black;j \ Blacks M \ Blacks (M[j := Black])" -apply (unfold Blacks_def) -apply(rule psubsetI) - apply(force simp add: nth_list_update) -apply safe -apply(erule_tac c = "j" in equalityCE) -apply auto -done - -lemma Graph12: "\a\Blacks M;j \ a\Blacks (M[j := Black])" -apply (unfold Blacks_def) -apply(force simp add: nth_list_update) -done - -lemma Graph13: "\a\ Blacks M;j \ a \ Blacks (M[j := Black])" -apply (unfold Blacks_def) -apply(erule psubset_subset_trans) -apply(force simp add: nth_list_update) -done - -declare Graph_defs [simp del] - -end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/Mul_Gar_Coll.thy --- a/src/HOL/HoareParallel/Mul_Gar_Coll.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1283 +0,0 @@ - -header {* \section{The Multi-Mutator Case} *} - -theory Mul_Gar_Coll imports Graph OG_Syntax begin - -text {* The full theory takes aprox. 18 minutes. *} - -record mut = - Z :: bool - R :: nat - T :: nat - -text {* Declaration of variables: *} - -record mul_gar_coll_state = - M :: nodes - E :: edges - bc :: "nat set" - obc :: "nat set" - Ma :: nodes - ind :: nat - k :: nat - q :: nat - l :: nat - Muts :: "mut list" - -subsection {* The Mutators *} - -constdefs - Mul_mut_init :: "mul_gar_coll_state \ nat \ bool" - "Mul_mut_init \ \ \n. n=length \Muts \ (\iMuts!i)E - \ T (\Muts!i)M) \" - - Mul_Redirect_Edge :: "nat \ nat \ mul_gar_coll_state ann_com" - "Mul_Redirect_Edge j n \ - .{\Mul_mut_init n \ Z (\Muts!j)}. - \IF T(\Muts!j) \ Reach \E THEN - \E:= \E[R (\Muts!j):= (fst (\E!R(\Muts!j)), T (\Muts!j))] FI,, - \Muts:= \Muts[j:= (\Muts!j) \Z:=False\]\" - - Mul_Color_Target :: "nat \ nat \ mul_gar_coll_state ann_com" - "Mul_Color_Target j n \ - .{\Mul_mut_init n \ \ Z (\Muts!j)}. - \\M:=\M[T (\Muts!j):=Black],, \Muts:=\Muts[j:= (\Muts!j) \Z:=True\]\" - - Mul_Mutator :: "nat \ nat \ mul_gar_coll_state ann_com" - "Mul_Mutator j n \ - .{\Mul_mut_init n \ Z (\Muts!j)}. - WHILE True - INV .{\Mul_mut_init n \ Z (\Muts!j)}. - DO Mul_Redirect_Edge j n ;; - Mul_Color_Target j n - OD" - -lemmas mul_mutator_defs = Mul_mut_init_def Mul_Redirect_Edge_def Mul_Color_Target_def - -subsubsection {* Correctness of the proof outline of one mutator *} - -lemma Mul_Redirect_Edge: "0\j \ j - \ Mul_Redirect_Edge j n - pre(Mul_Color_Target j n)" -apply (unfold mul_mutator_defs) -apply annhoare -apply(simp_all) -apply clarify -apply(simp add:nth_list_update) -done - -lemma Mul_Color_Target: "0\j \ j - \ Mul_Color_Target j n - .{\Mul_mut_init n \ Z (\Muts!j)}." -apply (unfold mul_mutator_defs) -apply annhoare -apply(simp_all) -apply clarify -apply(simp add:nth_list_update) -done - -lemma Mul_Mutator: "0\j \ j - \ Mul_Mutator j n .{False}." -apply(unfold Mul_Mutator_def) -apply annhoare -apply(simp_all add:Mul_Redirect_Edge Mul_Color_Target) -apply(simp add:mul_mutator_defs Mul_Redirect_Edge_def) -done - -subsubsection {* Interference freedom between mutators *} - -lemma Mul_interfree_Redirect_Edge_Redirect_Edge: - "\0\i; ij; jj\ \ - interfree_aux (Some (Mul_Redirect_Edge i n),{}, Some(Mul_Redirect_Edge j n))" -apply (unfold mul_mutator_defs) -apply interfree_aux -apply safe -apply(simp_all add: nth_list_update) -done - -lemma Mul_interfree_Redirect_Edge_Color_Target: - "\0\i; ij; jj\ \ - interfree_aux (Some(Mul_Redirect_Edge i n),{},Some(Mul_Color_Target j n))" -apply (unfold mul_mutator_defs) -apply interfree_aux -apply safe -apply(simp_all add: nth_list_update) -done - -lemma Mul_interfree_Color_Target_Redirect_Edge: - "\0\i; ij; jj\ \ - interfree_aux (Some(Mul_Color_Target i n),{},Some(Mul_Redirect_Edge j n))" -apply (unfold mul_mutator_defs) -apply interfree_aux -apply safe -apply(simp_all add:nth_list_update) -done - -lemma Mul_interfree_Color_Target_Color_Target: - " \0\i; ij; jj\ \ - interfree_aux (Some(Mul_Color_Target i n),{},Some(Mul_Color_Target j n))" -apply (unfold mul_mutator_defs) -apply interfree_aux -apply safe -apply(simp_all add: nth_list_update) -done - -lemmas mul_mutator_interfree = - Mul_interfree_Redirect_Edge_Redirect_Edge Mul_interfree_Redirect_Edge_Color_Target - Mul_interfree_Color_Target_Redirect_Edge Mul_interfree_Color_Target_Color_Target - -lemma Mul_interfree_Mutator_Mutator: "\i < n; j < n; i \ j\ \ - interfree_aux (Some (Mul_Mutator i n), {}, Some (Mul_Mutator j n))" -apply(unfold Mul_Mutator_def) -apply(interfree_aux) -apply(simp_all add:mul_mutator_interfree) -apply(simp_all add: mul_mutator_defs) -apply(tactic {* TRYALL (interfree_aux_tac) *}) -apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) -apply (simp_all add:nth_list_update) -done - -subsubsection {* Modular Parameterized Mutators *} - -lemma Mul_Parameterized_Mutators: "0 - \- .{\Mul_mut_init n \ (\iMuts!i))}. - COBEGIN - SCHEME [0\ j< n] - Mul_Mutator j n - .{False}. - COEND - .{False}." -apply oghoare -apply(force simp add:Mul_Mutator_def mul_mutator_defs nth_list_update) -apply(erule Mul_Mutator) -apply(simp add:Mul_interfree_Mutator_Mutator) -apply(force simp add:Mul_Mutator_def mul_mutator_defs nth_list_update) -done - -subsection {* The Collector *} - -constdefs - Queue :: "mul_gar_coll_state \ nat" - "Queue \ \ length (filter (\i. \ Z i \ \M!(T i) \ Black) \Muts) \" - -consts M_init :: nodes - -constdefs - Proper_M_init :: "mul_gar_coll_state \ bool" - "Proper_M_init \ \ Blacks M_init=Roots \ length M_init=length \M \" - - Mul_Proper :: "mul_gar_coll_state \ nat \ bool" - "Mul_Proper \ \ \n. Proper_Roots \M \ Proper_Edges (\M, \E) \ \Proper_M_init \ n=length \Muts \" - - Safe :: "mul_gar_coll_state \ bool" - "Safe \ \ Reach \E \ Blacks \M \" - -lemmas mul_collector_defs = Proper_M_init_def Mul_Proper_def Safe_def - -subsubsection {* Blackening Roots *} - -constdefs - Mul_Blacken_Roots :: "nat \ mul_gar_coll_state ann_com" - "Mul_Blacken_Roots n \ - .{\Mul_Proper n}. - \ind:=0;; - .{\Mul_Proper n \ \ind=0}. - WHILE \indM - INV .{\Mul_Proper n \ (\i<\ind. i\Roots \ \M!i=Black) \ \ind\length \M}. - DO .{\Mul_Proper n \ (\i<\ind. i\Roots \ \M!i=Black) \ \indM}. - IF \ind\Roots THEN - .{\Mul_Proper n \ (\i<\ind. i\Roots \ \M!i=Black) \ \indM \ \ind\Roots}. - \M:=\M[\ind:=Black] FI;; - .{\Mul_Proper n \ (\i<\ind+1. i\Roots \ \M!i=Black) \ \indM}. - \ind:=\ind+1 - OD" - -lemma Mul_Blacken_Roots: - "\ Mul_Blacken_Roots n - .{\Mul_Proper n \ Roots \ Blacks \M}." -apply (unfold Mul_Blacken_Roots_def) -apply annhoare -apply(simp_all add:mul_collector_defs Graph_defs) -apply safe -apply(simp_all add:nth_list_update) - apply (erule less_SucE) - apply simp+ - apply force -apply force -done - -subsubsection {* Propagating Black *} - -constdefs - Mul_PBInv :: "mul_gar_coll_state \ bool" - "Mul_PBInv \ \\Safe \ \obc\Blacks \M \ \l<\Queue - \ (\i<\ind. \BtoW(\E!i,\M)) \ \l\\Queue\" - - Mul_Auxk :: "mul_gar_coll_state \ bool" - "Mul_Auxk \ \\l<\Queue \ \M!\k\Black \ \BtoW(\E!\ind, \M) \ \obc\Blacks \M\" - -constdefs - Mul_Propagate_Black :: "nat \ mul_gar_coll_state ann_com" - "Mul_Propagate_Black n \ - .{\Mul_Proper n \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ (\Safe \ \l\\Queue \ \obc\Blacks \M)}. - \ind:=0;; - .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \M \ Blacks \M\Blacks \M \ \bc\Blacks \M - \ (\Safe \ \l\\Queue \ \obc\Blacks \M) \ \ind=0}. - WHILE \indE - INV .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \M \ \bc\Blacks \M - \ \Mul_PBInv \ \ind\length \E}. - DO .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \M \ \bc\Blacks \M - \ \Mul_PBInv \ \indE}. - IF \M!(fst (\E!\ind))=Black THEN - .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \M \ \bc\Blacks \M - \ \Mul_PBInv \ (\M!fst(\E!\ind))=Black \ \indE}. - \k:=snd(\E!\ind);; - .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \M \ \bc\Blacks \M - \ (\Safe \ \obc\Blacks \M \ \l<\Queue \ (\i<\ind. \BtoW(\E!i,\M)) - \ \l\\Queue \ \Mul_Auxk ) \ \kM \ \M!fst(\E!\ind)=Black - \ \indE}. - \\M:=\M[\k:=Black],,\ind:=\ind+1\ - ELSE .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \M \ \bc\Blacks \M - \ \Mul_PBInv \ \indE}. - \IF \M!(fst (\E!\ind))\Black THEN \ind:=\ind+1 FI\ FI - OD" - -lemma Mul_Propagate_Black: - "\ Mul_Propagate_Black n - .{\Mul_Proper n \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M - \ (\Safe \ \obc\Blacks \M \ \l<\Queue \ (\l\\Queue \ \obc\Blacks \M))}." -apply(unfold Mul_Propagate_Black_def) -apply annhoare -apply(simp_all add:Mul_PBInv_def mul_collector_defs Mul_Auxk_def Graph6 Graph7 Graph8 Graph12 mul_collector_defs Queue_def) ---{* 8 subgoals left *} -apply force -apply force -apply force -apply(force simp add:BtoW_def Graph_defs) ---{* 4 subgoals left *} -apply clarify -apply(simp add: mul_collector_defs Graph12 Graph6 Graph7 Graph8) -apply(disjE_tac) - apply(simp_all add:Graph12 Graph13) - apply(case_tac "M x! k x=Black") - apply(simp add: Graph10) - apply(rule disjI2, rule disjI1, erule subset_psubset_trans, erule Graph11, force) -apply(case_tac "M x! k x=Black") - apply(simp add: Graph10 BtoW_def) - apply(rule disjI2, clarify, erule less_SucE, force) - apply(case_tac "M x!snd(E x! ind x)=Black") - apply(force) - apply(force) -apply(rule disjI2, rule disjI1, erule subset_psubset_trans, erule Graph11, force) ---{* 3 subgoals left *} -apply force ---{* 2 subgoals left *} -apply clarify -apply(conjI_tac) -apply(disjE_tac) - apply (simp_all) -apply clarify -apply(erule less_SucE) - apply force -apply (simp add:BtoW_def) ---{* 1 subgoal left *} -apply clarify -apply simp -apply(disjE_tac) -apply (simp_all) -apply(rule disjI1 , rule Graph1) - apply simp_all -done - -subsubsection {* Counting Black Nodes *} - -constdefs - Mul_CountInv :: "mul_gar_coll_state \ nat \ bool" - "Mul_CountInv \ \ \ind. {i. i \Ma!i=Black}\\bc \" - - Mul_Count :: "nat \ mul_gar_coll_state ann_com" - "Mul_Count n \ - .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M - \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M) ) - \ \q \bc={}}. - \ind:=0;; - .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M - \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M) ) - \ \q \bc={} \ \ind=0}. - WHILE \indM - INV .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ \Mul_CountInv \ind - \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) - \ \q \ind\length \M}. - DO .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ \Mul_CountInv \ind - \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) - \ \q \indM}. - IF \M!\ind=Black - THEN .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ \Mul_CountInv \ind - \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) - \ \q \indM \ \M!\ind=Black}. - \bc:=insert \ind \bc - FI;; - .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ \Mul_CountInv (\ind+1) - \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) - \ \q \indM}. - \ind:=\ind+1 - OD" - -lemma Mul_Count: - "\ Mul_Count n - .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ Blacks \Ma\\bc - \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) - \ \q edges \ edges" - -axioms - Append_to_free0: "length (Append_to_free (i, e)) = length e" - Append_to_free1: "Proper_Edges (m, e) - \ Proper_Edges (m, Append_to_free(i, e))" - Append_to_free2: "i \ Reach e - \ n \ Reach (Append_to_free(i, e)) = ( n = i \ n \ Reach e)" - -constdefs - Mul_AppendInv :: "mul_gar_coll_state \ nat \ bool" - "Mul_AppendInv \ \ \ind. (\i. ind\i \ iM \ i\Reach \E \ \M!i=Black)\" - - Mul_Append :: "nat \ mul_gar_coll_state ann_com" - "Mul_Append n \ - .{\Mul_Proper n \ Roots\Blacks \M \ \Safe}. - \ind:=0;; - .{\Mul_Proper n \ Roots\Blacks \M \ \Safe \ \ind=0}. - WHILE \indM - INV .{\Mul_Proper n \ \Mul_AppendInv \ind \ \ind\length \M}. - DO .{\Mul_Proper n \ \Mul_AppendInv \ind \ \indM}. - IF \M!\ind=Black THEN - .{\Mul_Proper n \ \Mul_AppendInv \ind \ \indM \ \M!\ind=Black}. - \M:=\M[\ind:=White] - ELSE - .{\Mul_Proper n \ \Mul_AppendInv \ind \ \indM \ \ind\Reach \E}. - \E:=Append_to_free(\ind,\E) - FI;; - .{\Mul_Proper n \ \Mul_AppendInv (\ind+1) \ \indM}. - \ind:=\ind+1 - OD" - -lemma Mul_Append: - "\ Mul_Append n - .{\Mul_Proper n}." -apply(unfold Mul_Append_def) -apply annhoare -apply(simp_all add: mul_collector_defs Mul_AppendInv_def - Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) -apply(force simp add:Blacks_def) -apply(force simp add:Blacks_def) -apply(force simp add:Blacks_def) -apply(force simp add:Graph_defs) -apply force -apply(force simp add:Append_to_free1 Append_to_free2) -apply force -apply force -done - -subsubsection {* Collector *} - -constdefs - Mul_Collector :: "nat \ mul_gar_coll_state ann_com" - "Mul_Collector n \ -.{\Mul_Proper n}. -WHILE True INV .{\Mul_Proper n}. -DO -Mul_Blacken_Roots n ;; -.{\Mul_Proper n \ Roots\Blacks \M}. - \obc:={};; -.{\Mul_Proper n \ Roots\Blacks \M \ \obc={}}. - \bc:=Roots;; -.{\Mul_Proper n \ Roots\Blacks \M \ \obc={} \ \bc=Roots}. - \l:=0;; -.{\Mul_Proper n \ Roots\Blacks \M \ \obc={} \ \bc=Roots \ \l=0}. - WHILE \lMul_Proper n \ Roots\Blacks \M \ \bc\Blacks \M \ - (\Safe \ (\l\\Queue \ \bc\Blacks \M) \ \lMul_Proper n \ Roots\Blacks \M \ \bc\Blacks \M - \ (\Safe \ \l\\Queue \ \bc\Blacks \M)}. - \obc:=\bc;; - Mul_Propagate_Black n;; - .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \M \ \bc\Blacks \M - \ (\Safe \ \obc\Blacks \M \ \l<\Queue - \ (\l\\Queue \ \obc\Blacks \M))}. - \bc:={};; - .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \M \ \bc\Blacks \M - \ (\Safe \ \obc\Blacks \M \ \l<\Queue - \ (\l\\Queue \ \obc\Blacks \M)) \ \bc={}}. - \ \Ma:=\M,, \q:=\Queue \;; - Mul_Count n;; - .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ Blacks \Ma\\bc - \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) - \ \qobc=\bc THEN - .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ Blacks \Ma\\bc - \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) - \ \q \obc=\bc}. - \l:=\l+1 - ELSE .{\Mul_Proper n \ Roots\Blacks \M - \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M - \ length \Ma=length \M \ Blacks \Ma\\bc - \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) - \ \q \obc\\bc}. - \l:=0 FI - OD;; - Mul_Append n -OD" - -lemmas mul_modules = Mul_Redirect_Edge_def Mul_Color_Target_def - Mul_Blacken_Roots_def Mul_Propagate_Black_def - Mul_Count_def Mul_Append_def - -lemma Mul_Collector: - "\ Mul_Collector n - .{False}." -apply(unfold Mul_Collector_def) -apply annhoare -apply(simp_all only:pre.simps Mul_Blacken_Roots - Mul_Propagate_Black Mul_Count Mul_Append) -apply(simp_all add:mul_modules) -apply(simp_all add:mul_collector_defs Queue_def) -apply force -apply force -apply force -apply (force simp add: less_Suc_eq_le) -apply force -apply (force dest:subset_antisym) -apply force -apply force -apply force -done - -subsection {* Interference Freedom *} - -lemma le_length_filter_update[rule_format]: - "\i. (\P (list!i) \ P j) \ i length(filter P list) \ length(filter P (list[i:=j]))" -apply(induct_tac "list") - apply(simp) -apply(clarify) -apply(case_tac i) - apply(simp) -apply(simp) -done - -lemma less_length_filter_update [rule_format]: - "\i. P j \ \(P (list!i)) \ i length(filter P list) < length(filter P (list[i:=j]))" -apply(induct_tac "list") - apply(simp) -apply(clarify) -apply(case_tac i) - apply(simp) -apply(simp) -done - -lemma Mul_interfree_Blacken_Roots_Redirect_Edge: "\0\j; j \ - interfree_aux (Some(Mul_Blacken_Roots n),{},Some(Mul_Redirect_Edge j n))" -apply (unfold mul_modules) -apply interfree_aux -apply safe -apply(simp_all add:Graph6 Graph9 Graph12 nth_list_update mul_mutator_defs mul_collector_defs) -done - -lemma Mul_interfree_Redirect_Edge_Blacken_Roots: "\0\j; j\ - interfree_aux (Some(Mul_Redirect_Edge j n ),{},Some (Mul_Blacken_Roots n))" -apply (unfold mul_modules) -apply interfree_aux -apply safe -apply(simp_all add:mul_mutator_defs nth_list_update) -done - -lemma Mul_interfree_Blacken_Roots_Color_Target: "\0\j; j\ - interfree_aux (Some(Mul_Blacken_Roots n),{},Some (Mul_Color_Target j n ))" -apply (unfold mul_modules) -apply interfree_aux -apply safe -apply(simp_all add:mul_mutator_defs mul_collector_defs nth_list_update Graph7 Graph8 Graph9 Graph12) -done - -lemma Mul_interfree_Color_Target_Blacken_Roots: "\0\j; j\ - interfree_aux (Some(Mul_Color_Target j n ),{},Some (Mul_Blacken_Roots n ))" -apply (unfold mul_modules) -apply interfree_aux -apply safe -apply(simp_all add:mul_mutator_defs nth_list_update) -done - -lemma Mul_interfree_Propagate_Black_Redirect_Edge: "\0\j; j\ - interfree_aux (Some(Mul_Propagate_Black n),{},Some (Mul_Redirect_Edge j n ))" -apply (unfold mul_modules) -apply interfree_aux -apply(simp_all add:mul_mutator_defs mul_collector_defs Mul_PBInv_def nth_list_update Graph6) ---{* 7 subgoals left *} -apply clarify -apply(disjE_tac) - apply(simp_all add:Graph6) - apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) -apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) ---{* 6 subgoals left *} -apply clarify -apply(disjE_tac) - apply(simp_all add:Graph6) - apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) -apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) ---{* 5 subgoals left *} -apply clarify -apply(disjE_tac) - apply(simp_all add:Graph6) - apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) -apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(erule conjE) -apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule conjI) - apply(rule impI,(rule disjI2)+,rule conjI) - apply clarify - apply(case_tac "R (Muts x! j)=i") - apply (force simp add: nth_list_update BtoW_def) - apply (force simp add: nth_list_update) - apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(rule impI,(rule disjI2)+, erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) - apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) -apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) -apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) ---{* 4 subgoals left *} -apply clarify -apply(disjE_tac) - apply(simp_all add:Graph6) - apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) -apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(erule conjE) -apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule conjI) - apply(rule impI,(rule disjI2)+,rule conjI) - apply clarify - apply(case_tac "R (Muts x! j)=i") - apply (force simp add: nth_list_update BtoW_def) - apply (force simp add: nth_list_update) - apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(rule impI,(rule disjI2)+, erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) - apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) -apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) -apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) ---{* 3 subgoals left *} -apply clarify -apply(disjE_tac) - apply(simp_all add:Graph6) - apply (rule impI) - apply(rule conjI) - apply(rule disjI1,rule subset_trans,erule Graph3,simp,simp) - apply(case_tac "R (Muts x ! j)= ind x") - apply(simp add:nth_list_update) - apply(simp add:nth_list_update) - apply(case_tac "R (Muts x ! j)= ind x") - apply(simp add:nth_list_update) - apply(simp add:nth_list_update) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule conjI) - apply(rule impI) - apply(rule conjI) - apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(case_tac "R (Muts x ! j)= ind x") - apply(simp add:nth_list_update) - apply(simp add:nth_list_update) - apply(rule impI) - apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(rule conjI) - apply(rule impI) - apply(rule conjI) - apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(case_tac "R (Muts x ! j)= ind x") - apply(simp add:nth_list_update) - apply(simp add:nth_list_update) - apply(rule impI) - apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(erule conjE) - apply(rule conjI) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule impI,rule conjI,(rule disjI2)+,rule conjI) - apply clarify - apply(case_tac "R (Muts x! j)=i") - apply (force simp add: nth_list_update BtoW_def) - apply (force simp add: nth_list_update) - apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(case_tac "R (Muts x ! j)= ind x") - apply(simp add:nth_list_update) - apply(simp add:nth_list_update) - apply(rule impI,rule conjI) - apply(rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) - apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) - apply(case_tac "R (Muts x! j)=ind x") - apply (force simp add: nth_list_update) - apply (force simp add: nth_list_update) -apply(rule impI, (rule disjI2)+, erule le_trans) -apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) ---{* 2 subgoals left *} -apply clarify -apply(rule conjI) - apply(disjE_tac) - apply(simp_all add:Mul_Auxk_def Graph6) - apply (rule impI) - apply(rule conjI) - apply(rule disjI1,rule subset_trans,erule Graph3,simp,simp) - apply(case_tac "R (Muts x ! j)= ind x") - apply(simp add:nth_list_update) - apply(simp add:nth_list_update) - apply(case_tac "R (Muts x ! j)= ind x") - apply(simp add:nth_list_update) - apply(simp add:nth_list_update) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule impI) - apply(rule conjI) - apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(case_tac "R (Muts x ! j)= ind x") - apply(simp add:nth_list_update) - apply(simp add:nth_list_update) - apply(rule impI) - apply(rule conjI) - apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(case_tac "R (Muts x ! j)= ind x") - apply(simp add:nth_list_update) - apply(simp add:nth_list_update) -apply(rule impI) -apply(rule conjI) - apply(erule conjE)+ - apply(case_tac "M x!(T (Muts x!j))=Black") - apply((rule disjI2)+,rule conjI) - apply clarify - apply(case_tac "R (Muts x! j)=i") - apply (force simp add: nth_list_update BtoW_def) - apply (force simp add: nth_list_update) - apply(rule conjI) - apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(rule impI) - apply(case_tac "R (Muts x ! j)= ind x") - apply(simp add:nth_list_update BtoW_def) - apply (simp add:nth_list_update) - apply(rule impI) - apply simp - apply(disjE_tac) - apply(rule disjI1, erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply force - apply(rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) - apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) - apply(case_tac "R (Muts x ! j)= ind x") - apply(simp add:nth_list_update) - apply(simp add:nth_list_update) -apply(disjE_tac) -apply simp_all -apply(conjI_tac) - apply(rule impI) - apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(erule conjE)+ -apply(rule impI,(rule disjI2)+,rule conjI) - apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(rule impI)+ -apply simp -apply(disjE_tac) - apply(rule disjI1, erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply force ---{* 1 subgoal left *} -apply clarify -apply(disjE_tac) - apply(simp_all add:Graph6) - apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) -apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(erule conjE) -apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule conjI) - apply(rule impI,(rule disjI2)+,rule conjI) - apply clarify - apply(case_tac "R (Muts x! j)=i") - apply (force simp add: nth_list_update BtoW_def) - apply (force simp add: nth_list_update) - apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(rule impI,(rule disjI2)+, erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) - apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) -apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) -apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) -done - -lemma Mul_interfree_Redirect_Edge_Propagate_Black: "\0\j; j\ - interfree_aux (Some(Mul_Redirect_Edge j n ),{},Some (Mul_Propagate_Black n))" -apply (unfold mul_modules) -apply interfree_aux -apply safe -apply(simp_all add:mul_mutator_defs nth_list_update) -done - -lemma Mul_interfree_Propagate_Black_Color_Target: "\0\j; j\ - interfree_aux (Some(Mul_Propagate_Black n),{},Some (Mul_Color_Target j n ))" -apply (unfold mul_modules) -apply interfree_aux -apply(simp_all add: mul_collector_defs mul_mutator_defs) ---{* 7 subgoals left *} -apply clarify -apply (simp add:Graph7 Graph8 Graph12) -apply(disjE_tac) - apply(simp add:Graph7 Graph8 Graph12) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule disjI2,rule disjI1, erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply((rule disjI2)+,erule subset_psubset_trans, erule Graph11, simp) -apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) ---{* 6 subgoals left *} -apply clarify -apply (simp add:Graph7 Graph8 Graph12) -apply(disjE_tac) - apply(simp add:Graph7 Graph8 Graph12) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule disjI2,rule disjI1, erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply((rule disjI2)+,erule subset_psubset_trans, erule Graph11, simp) -apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) ---{* 5 subgoals left *} -apply clarify -apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12) -apply(disjE_tac) - apply(simp add:Graph7 Graph8 Graph12) - apply(rule disjI2,rule disjI1, erule psubset_subset_trans,simp add:Graph9) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp) -apply(erule conjE) -apply(case_tac "M x!(T (Muts x!j))=Black") - apply((rule disjI2)+) - apply (rule conjI) - apply(simp add:Graph10) - apply(erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) -apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp) ---{* 4 subgoals left *} -apply clarify -apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12) -apply(disjE_tac) - apply(simp add:Graph7 Graph8 Graph12) - apply(rule disjI2,rule disjI1, erule psubset_subset_trans,simp add:Graph9) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp) -apply(erule conjE) -apply(case_tac "M x!(T (Muts x!j))=Black") - apply((rule disjI2)+) - apply (rule conjI) - apply(simp add:Graph10) - apply(erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) -apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp) ---{* 3 subgoals left *} -apply clarify -apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12) -apply(case_tac "M x!(T (Muts x!j))=Black") - apply(simp add:Graph10) - apply(disjE_tac) - apply simp_all - apply(rule disjI2, rule disjI2, rule disjI1,erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply(erule conjE) - apply((rule disjI2)+,erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) -apply(rule conjI) - apply(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11) -apply (force simp add:nth_list_update) ---{* 2 subgoals left *} -apply clarify -apply(simp add:Mul_Auxk_def Graph7 Graph8 Graph12) -apply(case_tac "M x!(T (Muts x!j))=Black") - apply(simp add:Graph10) - apply(disjE_tac) - apply simp_all - apply(rule disjI2, rule disjI2, rule disjI1,erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply(erule conjE)+ - apply((rule disjI2)+,rule conjI, erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply((rule impI)+) - apply simp - apply(erule disjE) - apply(rule disjI1, erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply force -apply(rule conjI) - apply(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11) -apply (force simp add:nth_list_update) ---{* 1 subgoal left *} -apply clarify -apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12) -apply(case_tac "M x!(T (Muts x!j))=Black") - apply(simp add:Graph10) - apply(disjE_tac) - apply simp_all - apply(rule disjI2, rule disjI2, rule disjI1,erule less_le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply(erule conjE) - apply((rule disjI2)+,erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) -apply(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11) -done - -lemma Mul_interfree_Color_Target_Propagate_Black: "\0\j; j\ - interfree_aux (Some(Mul_Color_Target j n),{},Some(Mul_Propagate_Black n ))" -apply (unfold mul_modules) -apply interfree_aux -apply safe -apply(simp_all add:mul_mutator_defs nth_list_update) -done - -lemma Mul_interfree_Count_Redirect_Edge: "\0\j; j\ - interfree_aux (Some(Mul_Count n ),{},Some(Mul_Redirect_Edge j n))" -apply (unfold mul_modules) -apply interfree_aux ---{* 9 subgoals left *} -apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def Graph6) -apply clarify -apply disjE_tac - apply(simp add:Graph6) - apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) - apply(simp add:Graph6) -apply clarify -apply disjE_tac - apply(simp add:Graph6) - apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(simp add:Graph6) ---{* 8 subgoals left *} -apply(simp add:mul_mutator_defs nth_list_update) ---{* 7 subgoals left *} -apply(simp add:mul_mutator_defs mul_collector_defs) -apply clarify -apply disjE_tac - apply(simp add:Graph6) - apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) - apply(simp add:Graph6) -apply clarify -apply disjE_tac - apply(simp add:Graph6) - apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(simp add:Graph6) ---{* 6 subgoals left *} -apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def) -apply clarify -apply disjE_tac - apply(simp add:Graph6 Queue_def) - apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) - apply(simp add:Graph6) -apply clarify -apply disjE_tac - apply(simp add:Graph6) - apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(simp add:Graph6) ---{* 5 subgoals left *} -apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def) -apply clarify -apply disjE_tac - apply(simp add:Graph6) - apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) - apply(simp add:Graph6) -apply clarify -apply disjE_tac - apply(simp add:Graph6) - apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(simp add:Graph6) ---{* 4 subgoals left *} -apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def) -apply clarify -apply disjE_tac - apply(simp add:Graph6) - apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) - apply(simp add:Graph6) -apply clarify -apply disjE_tac - apply(simp add:Graph6) - apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(simp add:Graph6) ---{* 3 subgoals left *} -apply(simp add:mul_mutator_defs nth_list_update) ---{* 2 subgoals left *} -apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def) -apply clarify -apply disjE_tac - apply(simp add:Graph6) - apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) - apply(simp add:Graph6) -apply clarify -apply disjE_tac - apply(simp add:Graph6) - apply(rule conjI) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) - apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(simp add:Graph6) ---{* 1 subgoal left *} -apply(simp add:mul_mutator_defs nth_list_update) -done - -lemma Mul_interfree_Redirect_Edge_Count: "\0\j; j\ - interfree_aux (Some(Mul_Redirect_Edge j n),{},Some(Mul_Count n ))" -apply (unfold mul_modules) -apply interfree_aux -apply safe -apply(simp_all add:mul_mutator_defs nth_list_update) -done - -lemma Mul_interfree_Count_Color_Target: "\0\j; j\ - interfree_aux (Some(Mul_Count n ),{},Some(Mul_Color_Target j n))" -apply (unfold mul_modules) -apply interfree_aux -apply(simp_all add:mul_collector_defs mul_mutator_defs Mul_CountInv_def) ---{* 6 subgoals left *} -apply clarify -apply disjE_tac - apply (simp add: Graph7 Graph8 Graph12) - apply (simp add: Graph7 Graph8 Graph12) -apply clarify -apply disjE_tac - apply (simp add: Graph7 Graph8 Graph12) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) -apply (simp add: Graph7 Graph8 Graph12) -apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) ---{* 5 subgoals left *} -apply clarify -apply disjE_tac - apply (simp add: Graph7 Graph8 Graph12) - apply (simp add: Graph7 Graph8 Graph12) -apply clarify -apply disjE_tac - apply (simp add: Graph7 Graph8 Graph12) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) -apply (simp add: Graph7 Graph8 Graph12) -apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) ---{* 4 subgoals left *} -apply clarify -apply disjE_tac - apply (simp add: Graph7 Graph8 Graph12) - apply (simp add: Graph7 Graph8 Graph12) -apply clarify -apply disjE_tac - apply (simp add: Graph7 Graph8 Graph12) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) -apply (simp add: Graph7 Graph8 Graph12) -apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) ---{* 3 subgoals left *} -apply clarify -apply disjE_tac - apply (simp add: Graph7 Graph8 Graph12) - apply (simp add: Graph7 Graph8 Graph12) -apply clarify -apply disjE_tac - apply (simp add: Graph7 Graph8 Graph12) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) -apply (simp add: Graph7 Graph8 Graph12) -apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) ---{* 2 subgoals left *} -apply clarify -apply disjE_tac - apply (simp add: Graph7 Graph8 Graph12 nth_list_update) - apply (simp add: Graph7 Graph8 Graph12 nth_list_update) -apply clarify -apply disjE_tac - apply (simp add: Graph7 Graph8 Graph12) - apply(rule conjI) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) - apply (simp add: nth_list_update) -apply (simp add: Graph7 Graph8 Graph12) -apply(rule conjI) - apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) -apply (simp add: nth_list_update) ---{* 1 subgoal left *} -apply clarify -apply disjE_tac - apply (simp add: Graph7 Graph8 Graph12) - apply (simp add: Graph7 Graph8 Graph12) -apply clarify -apply disjE_tac - apply (simp add: Graph7 Graph8 Graph12) - apply(case_tac "M x!(T (Muts x!j))=Black") - apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) - apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) - apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) -apply (simp add: Graph7 Graph8 Graph12) -apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) -done - -lemma Mul_interfree_Color_Target_Count: "\0\j; j\ - interfree_aux (Some(Mul_Color_Target j n),{}, Some(Mul_Count n ))" -apply (unfold mul_modules) -apply interfree_aux -apply safe -apply(simp_all add:mul_mutator_defs nth_list_update) -done - -lemma Mul_interfree_Append_Redirect_Edge: "\0\j; j\ - interfree_aux (Some(Mul_Append n),{}, Some(Mul_Redirect_Edge j n))" -apply (unfold mul_modules) -apply interfree_aux -apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) -apply(simp_all add:Graph6 Append_to_free0 Append_to_free1 mul_collector_defs mul_mutator_defs Mul_AppendInv_def) -apply(erule_tac x=j in allE, force dest:Graph3)+ -done - -lemma Mul_interfree_Redirect_Edge_Append: "\0\j; j\ - interfree_aux (Some(Mul_Redirect_Edge j n),{},Some(Mul_Append n))" -apply (unfold mul_modules) -apply interfree_aux -apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) -apply(simp_all add:mul_collector_defs Append_to_free0 Mul_AppendInv_def mul_mutator_defs nth_list_update) -done - -lemma Mul_interfree_Append_Color_Target: "\0\j; j\ - interfree_aux (Some(Mul_Append n),{}, Some(Mul_Color_Target j n))" -apply (unfold mul_modules) -apply interfree_aux -apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) -apply(simp_all add:mul_mutator_defs mul_collector_defs Mul_AppendInv_def Graph7 Graph8 Append_to_free0 Append_to_free1 - Graph12 nth_list_update) -done - -lemma Mul_interfree_Color_Target_Append: "\0\j; j\ - interfree_aux (Some(Mul_Color_Target j n),{}, Some(Mul_Append n))" -apply (unfold mul_modules) -apply interfree_aux -apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) -apply(simp_all add: mul_mutator_defs nth_list_update) -apply(simp add:Mul_AppendInv_def Append_to_free0) -done - -subsubsection {* Interference freedom Collector-Mutator *} - -lemmas mul_collector_mutator_interfree = - Mul_interfree_Blacken_Roots_Redirect_Edge Mul_interfree_Blacken_Roots_Color_Target - Mul_interfree_Propagate_Black_Redirect_Edge Mul_interfree_Propagate_Black_Color_Target - Mul_interfree_Count_Redirect_Edge Mul_interfree_Count_Color_Target - Mul_interfree_Append_Redirect_Edge Mul_interfree_Append_Color_Target - Mul_interfree_Redirect_Edge_Blacken_Roots Mul_interfree_Color_Target_Blacken_Roots - Mul_interfree_Redirect_Edge_Propagate_Black Mul_interfree_Color_Target_Propagate_Black - Mul_interfree_Redirect_Edge_Count Mul_interfree_Color_Target_Count - Mul_interfree_Redirect_Edge_Append Mul_interfree_Color_Target_Append - -lemma Mul_interfree_Collector_Mutator: "j - interfree_aux (Some (Mul_Collector n), {}, Some (Mul_Mutator j n))" -apply(unfold Mul_Collector_def Mul_Mutator_def) -apply interfree_aux -apply(simp_all add:mul_collector_mutator_interfree) -apply(unfold mul_modules mul_collector_defs mul_mutator_defs) -apply(tactic {* TRYALL (interfree_aux_tac) *}) ---{* 42 subgoals left *} -apply (clarify,simp add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)+ ---{* 24 subgoals left *} -apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) ---{* 14 subgoals left *} -apply(tactic {* TRYALL (clarify_tac @{claset}) *}) -apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) -apply(tactic {* TRYALL (rtac conjI) *}) -apply(tactic {* TRYALL (rtac impI) *}) -apply(tactic {* TRYALL (etac disjE) *}) -apply(tactic {* TRYALL (etac conjE) *}) -apply(tactic {* TRYALL (etac disjE) *}) -apply(tactic {* TRYALL (etac disjE) *}) ---{* 72 subgoals left *} -apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) ---{* 35 subgoals left *} -apply(tactic {* TRYALL(EVERY'[rtac disjI1,rtac subset_trans,etac @{thm Graph3},force_tac @{clasimpset}, assume_tac]) *}) ---{* 28 subgoals left *} -apply(tactic {* TRYALL (etac conjE) *}) -apply(tactic {* TRYALL (etac disjE) *}) ---{* 34 subgoals left *} -apply(rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) -apply(case_tac [!] "M x!(T (Muts x ! j))=Black") -apply(simp_all add:Graph10) ---{* 47 subgoals left *} -apply(tactic {* TRYALL(EVERY'[REPEAT o (rtac disjI2),etac (thm "subset_psubset_trans"),etac (thm "Graph11"),force_tac @{clasimpset}]) *}) ---{* 41 subgoals left *} -apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI1, etac @{thm le_trans}, force_tac (@{claset},@{simpset} addsimps [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])]) *}) ---{* 35 subgoals left *} -apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac (thm "psubset_subset_trans"),rtac (thm "Graph9"),force_tac @{clasimpset}]) *}) ---{* 31 subgoals left *} -apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac (thm "subset_psubset_trans"),etac (thm "Graph11"),force_tac @{clasimpset}]) *}) ---{* 29 subgoals left *} -apply(tactic {* TRYALL(EVERY'[REPEAT o (rtac disjI2),etac (thm "subset_psubset_trans"),etac (thm "subset_psubset_trans"),etac (thm "Graph11"),force_tac @{clasimpset}]) *}) ---{* 25 subgoals left *} -apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI2, rtac disjI1, etac @{thm le_trans}, force_tac (@{claset},@{simpset} addsimps [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])]) *}) ---{* 10 subgoals left *} -apply(rule disjI2,rule disjI2,rule conjI,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update, rule disjI1, rule less_imp_le, erule less_le_trans, force simp add:Queue_def less_Suc_eq_le le_length_filter_update)+ -done - -subsubsection {* Interference freedom Mutator-Collector *} - -lemma Mul_interfree_Mutator_Collector: " j < n \ - interfree_aux (Some (Mul_Mutator j n), {}, Some (Mul_Collector n))" -apply(unfold Mul_Collector_def Mul_Mutator_def) -apply interfree_aux -apply(simp_all add:mul_collector_mutator_interfree) -apply(unfold mul_modules mul_collector_defs mul_mutator_defs) -apply(tactic {* TRYALL (interfree_aux_tac) *}) ---{* 76 subgoals left *} -apply (clarify,simp add: nth_list_update)+ ---{* 56 subgoals left *} -apply(clarify,simp add:Mul_AppendInv_def Append_to_free0 nth_list_update)+ -done - -subsubsection {* The Multi-Mutator Garbage Collection Algorithm *} - -text {* The total number of verification conditions is 328 *} - -lemma Mul_Gar_Coll: - "\- .{\Mul_Proper n \ \Mul_mut_init n \ (\iMuts!i))}. - COBEGIN - Mul_Collector n - .{False}. - \ - SCHEME [0\ j< n] - Mul_Mutator j n - .{False}. - COEND - .{False}." -apply oghoare ---{* Strengthening the precondition *} -apply(rule Int_greatest) - apply (case_tac n) - apply(force simp add: Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append) - apply(simp add: Mul_Mutator_def mul_collector_defs mul_mutator_defs nth_append) - apply force -apply clarify -apply(case_tac i) - apply(simp add:Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append) -apply(simp add: Mul_Mutator_def mul_mutator_defs mul_collector_defs nth_append nth_map_upt) ---{* Collector *} -apply(rule Mul_Collector) ---{* Mutator *} -apply(erule Mul_Mutator) ---{* Interference freedom *} -apply(simp add:Mul_interfree_Collector_Mutator) -apply(simp add:Mul_interfree_Mutator_Collector) -apply(simp add:Mul_interfree_Mutator_Mutator) ---{* Weakening of the postcondition *} -apply(case_tac n) - apply(simp add:Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append) -apply(simp add:Mul_Mutator_def mul_mutator_defs mul_collector_defs nth_append) -done - -end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/OG_Com.thy --- a/src/HOL/HoareParallel/OG_Com.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,55 +0,0 @@ - -header {* \chapter{The Owicki-Gries Method} - -\section{Abstract Syntax} *} - -theory OG_Com imports Main begin - -text {* Type abbreviations for boolean expressions and assertions: *} - -types - 'a bexp = "'a set" - 'a assn = "'a set" - -text {* The syntax of commands is defined by two mutually recursive -datatypes: @{text "'a ann_com"} for annotated commands and @{text "'a -com"} for non-annotated commands. *} - -datatype 'a ann_com = - AnnBasic "('a assn)" "('a \ 'a)" - | AnnSeq "('a ann_com)" "('a ann_com)" - | AnnCond1 "('a assn)" "('a bexp)" "('a ann_com)" "('a ann_com)" - | AnnCond2 "('a assn)" "('a bexp)" "('a ann_com)" - | AnnWhile "('a assn)" "('a bexp)" "('a assn)" "('a ann_com)" - | AnnAwait "('a assn)" "('a bexp)" "('a com)" -and 'a com = - Parallel "('a ann_com option \ 'a assn) list" - | Basic "('a \ 'a)" - | Seq "('a com)" "('a com)" - | Cond "('a bexp)" "('a com)" "('a com)" - | While "('a bexp)" "('a assn)" "('a com)" - -text {* The function @{text pre} extracts the precondition of an -annotated command: *} - -consts - pre ::"'a ann_com \ 'a assn" -primrec - "pre (AnnBasic r f) = r" - "pre (AnnSeq c1 c2) = pre c1" - "pre (AnnCond1 r b c1 c2) = r" - "pre (AnnCond2 r b c) = r" - "pre (AnnWhile r b i c) = r" - "pre (AnnAwait r b c) = r" - -text {* Well-formedness predicate for atomic programs: *} - -consts atom_com :: "'a com \ bool" -primrec - "atom_com (Parallel Ts) = False" - "atom_com (Basic f) = True" - "atom_com (Seq c1 c2) = (atom_com c1 \ atom_com c2)" - "atom_com (Cond b c1 c2) = (atom_com c1 \ atom_com c2)" - "atom_com (While b i c) = atom_com c" - -end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/OG_Examples.thy --- a/src/HOL/HoareParallel/OG_Examples.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,549 +0,0 @@ - -header {* \section{Examples} *} - -theory OG_Examples imports OG_Syntax begin - -subsection {* Mutual Exclusion *} - -subsubsection {* Peterson's Algorithm I*} - -text {* Eike Best. "Semantics of Sequential and Parallel Programs", page 217. *} - -record Petersons_mutex_1 = - pr1 :: nat - pr2 :: nat - in1 :: bool - in2 :: bool - hold :: nat - -lemma Petersons_mutex_1: - "\- .{\pr1=0 \ \\in1 \ \pr2=0 \ \\in2 }. - COBEGIN .{\pr1=0 \ \\in1}. - WHILE True INV .{\pr1=0 \ \\in1}. - DO - .{\pr1=0 \ \\in1}. \ \in1:=True,,\pr1:=1 \;; - .{\pr1=1 \ \in1}. \ \hold:=1,,\pr1:=2 \;; - .{\pr1=2 \ \in1 \ (\hold=1 \ \hold=2 \ \pr2=2)}. - AWAIT (\\in2 \ \(\hold=1)) THEN \pr1:=3 END;; - .{\pr1=3 \ \in1 \ (\hold=1 \ \hold=2 \ \pr2=2)}. - \\in1:=False,,\pr1:=0\ - OD .{\pr1=0 \ \\in1}. - \ - .{\pr2=0 \ \\in2}. - WHILE True INV .{\pr2=0 \ \\in2}. - DO - .{\pr2=0 \ \\in2}. \ \in2:=True,,\pr2:=1 \;; - .{\pr2=1 \ \in2}. \ \hold:=2,,\pr2:=2 \;; - .{\pr2=2 \ \in2 \ (\hold=2 \ (\hold=1 \ \pr1=2))}. - AWAIT (\\in1 \ \(\hold=2)) THEN \pr2:=3 END;; - .{\pr2=3 \ \in2 \ (\hold=2 \ (\hold=1 \ \pr1=2))}. - \\in2:=False,,\pr2:=0\ - OD .{\pr2=0 \ \\in2}. - COEND - .{\pr1=0 \ \\in1 \ \pr2=0 \ \\in2}." -apply oghoare ---{* 104 verification conditions. *} -apply auto -done - -subsubsection {*Peterson's Algorithm II: A Busy Wait Solution *} - -text {* Apt and Olderog. "Verification of sequential and concurrent Programs", page 282. *} - -record Busy_wait_mutex = - flag1 :: bool - flag2 :: bool - turn :: nat - after1 :: bool - after2 :: bool - -lemma Busy_wait_mutex: - "\- .{True}. - \flag1:=False,, \flag2:=False,, - COBEGIN .{\\flag1}. - WHILE True - INV .{\\flag1}. - DO .{\\flag1}. \ \flag1:=True,,\after1:=False \;; - .{\flag1 \ \\after1}. \ \turn:=1,,\after1:=True \;; - .{\flag1 \ \after1 \ (\turn=1 \ \turn=2)}. - WHILE \(\flag2 \ \turn=2) - INV .{\flag1 \ \after1 \ (\turn=1 \ \turn=2)}. - DO .{\flag1 \ \after1 \ (\turn=1 \ \turn=2)}. SKIP OD;; - .{\flag1 \ \after1 \ (\flag2 \ \after2 \ \turn=2)}. - \flag1:=False - OD - .{False}. - \ - .{\\flag2}. - WHILE True - INV .{\\flag2}. - DO .{\\flag2}. \ \flag2:=True,,\after2:=False \;; - .{\flag2 \ \\after2}. \ \turn:=2,,\after2:=True \;; - .{\flag2 \ \after2 \ (\turn=1 \ \turn=2)}. - WHILE \(\flag1 \ \turn=1) - INV .{\flag2 \ \after2 \ (\turn=1 \ \turn=2)}. - DO .{\flag2 \ \after2 \ (\turn=1 \ \turn=2)}. SKIP OD;; - .{\flag2 \ \after2 \ (\flag1 \ \after1 \ \turn=1)}. - \flag2:=False - OD - .{False}. - COEND - .{False}." -apply oghoare ---{* 122 vc *} -apply auto -done - -subsubsection {* Peterson's Algorithm III: A Solution using Semaphores *} - -record Semaphores_mutex = - out :: bool - who :: nat - -lemma Semaphores_mutex: - "\- .{i\j}. - \out:=True ,, - COBEGIN .{i\j}. - WHILE True INV .{i\j}. - DO .{i\j}. AWAIT \out THEN \out:=False,, \who:=i END;; - .{\\out \ \who=i \ i\j}. \out:=True OD - .{False}. - \ - .{i\j}. - WHILE True INV .{i\j}. - DO .{i\j}. AWAIT \out THEN \out:=False,,\who:=j END;; - .{\\out \ \who=j \ i\j}. \out:=True OD - .{False}. - COEND - .{False}." -apply oghoare ---{* 38 vc *} -apply auto -done - -subsubsection {* Peterson's Algorithm III: Parameterized version: *} - -lemma Semaphores_parameterized_mutex: - "0 \- .{True}. - \out:=True ,, - COBEGIN - SCHEME [0\ i< n] - .{True}. - WHILE True INV .{True}. - DO .{True}. AWAIT \out THEN \out:=False,, \who:=i END;; - .{\\out \ \who=i}. \out:=True OD - .{False}. - COEND - .{False}." -apply oghoare ---{* 20 vc *} -apply auto -done - -subsubsection{* The Ticket Algorithm *} - -record Ticket_mutex = - num :: nat - nextv :: nat - turn :: "nat list" - index :: nat - -lemma Ticket_mutex: - "\ 0n=length \turn \ 0<\nextv \ (\k l. k l k\l - \ \turn!k < \num \ (\turn!k =0 \ \turn!k\\turn!l))\ \ - \ \- .{n=length \turn}. - \index:= 0,, - WHILE \index < n INV .{n=length \turn \ (\i<\index. \turn!i=0)}. - DO \turn:= \turn[\index:=0],, \index:=\index +1 OD,, - \num:=1 ,, \nextv:=1 ,, - COBEGIN - SCHEME [0\ i< n] - .{\I}. - WHILE True INV .{\I}. - DO .{\I}. \ \turn :=\turn[i:=\num],, \num:=\num+1 \;; - .{\I}. WAIT \turn!i=\nextv END;; - .{\I \ \turn!i=\nextv}. \nextv:=\nextv+1 - OD - .{False}. - COEND - .{False}." -apply oghoare ---{* 35 vc *} -apply simp_all ---{* 21 vc *} -apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) ---{* 11 vc *} -apply simp_all -apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) ---{* 10 subgoals left *} -apply(erule less_SucE) - apply simp -apply simp ---{* 9 subgoals left *} -apply(case_tac "i=k") - apply force -apply simp -apply(case_tac "i=l") - apply force -apply force ---{* 8 subgoals left *} -prefer 8 -apply force -apply force ---{* 6 subgoals left *} -prefer 6 -apply(erule_tac x=i in allE) -apply fastsimp ---{* 5 subgoals left *} -prefer 5 -apply(case_tac [!] "j=k") ---{* 10 subgoals left *} -apply simp_all -apply(erule_tac x=k in allE) -apply force ---{* 9 subgoals left *} -apply(case_tac "j=l") - apply simp - apply(erule_tac x=k in allE) - apply(erule_tac x=k in allE) - apply(erule_tac x=l in allE) - apply force -apply(erule_tac x=k in allE) -apply(erule_tac x=k in allE) -apply(erule_tac x=l in allE) -apply force ---{* 8 subgoals left *} -apply force -apply(case_tac "j=l") - apply simp -apply(erule_tac x=k in allE) -apply(erule_tac x=l in allE) -apply force -apply force -apply force ---{* 5 subgoals left *} -apply(erule_tac x=k in allE) -apply(erule_tac x=l in allE) -apply(case_tac "j=l") - apply force -apply force -apply force ---{* 3 subgoals left *} -apply(erule_tac x=k in allE) -apply(erule_tac x=l in allE) -apply(case_tac "j=l") - apply force -apply force -apply force ---{* 1 subgoals left *} -apply(erule_tac x=k in allE) -apply(erule_tac x=l in allE) -apply(case_tac "j=l") - apply force -apply force -done - -subsection{* Parallel Zero Search *} - -text {* Synchronized Zero Search. Zero-6 *} - -text {*Apt and Olderog. "Verification of sequential and concurrent Programs" page 294: *} - -record Zero_search = - turn :: nat - found :: bool - x :: nat - y :: nat - -lemma Zero_search: - "\I1= \ a\\x \ (\found \ (a<\x \ f(\x)=0) \ (\y\a \ f(\y)=0)) - \ (\\found \ a<\ x \ f(\x)\0) \ ; - I2= \\y\a+1 \ (\found \ (a<\x \ f(\x)=0) \ (\y\a \ f(\y)=0)) - \ (\\found \ \y\a \ f(\y)\0) \ \ \ - \- .{\ u. f(u)=0}. - \turn:=1,, \found:= False,, - \x:=a,, \y:=a+1 ,, - COBEGIN .{\I1}. - WHILE \\found - INV .{\I1}. - DO .{a\\x \ (\found \ \y\a \ f(\y)=0) \ (a<\x \ f(\x)\0)}. - WAIT \turn=1 END;; - .{a\\x \ (\found \ \y\a \ f(\y)=0) \ (a<\x \ f(\x)\0)}. - \turn:=2;; - .{a\\x \ (\found \ \y\a \ f(\y)=0) \ (a<\x \ f(\x)\0)}. - \ \x:=\x+1,, - IF f(\x)=0 THEN \found:=True ELSE SKIP FI\ - OD;; - .{\I1 \ \found}. - \turn:=2 - .{\I1 \ \found}. - \ - .{\I2}. - WHILE \\found - INV .{\I2}. - DO .{\y\a+1 \ (\found \ a<\x \ f(\x)=0) \ (\y\a \ f(\y)\0)}. - WAIT \turn=2 END;; - .{\y\a+1 \ (\found \ a<\x \ f(\x)=0) \ (\y\a \ f(\y)\0)}. - \turn:=1;; - .{\y\a+1 \ (\found \ a<\x \ f(\x)=0) \ (\y\a \ f(\y)\0)}. - \ \y:=(\y - 1),, - IF f(\y)=0 THEN \found:=True ELSE SKIP FI\ - OD;; - .{\I2 \ \found}. - \turn:=1 - .{\I2 \ \found}. - COEND - .{f(\x)=0 \ f(\y)=0}." -apply oghoare ---{* 98 verification conditions *} -apply auto ---{* auto takes about 3 minutes !! *} -done - -text {* Easier Version: without AWAIT. Apt and Olderog. page 256: *} - -lemma Zero_Search_2: -"\I1=\ a\\x \ (\found \ (a<\x \ f(\x)=0) \ (\y\a \ f(\y)=0)) - \ (\\found \ a<\x \ f(\x)\0)\; - I2= \\y\a+1 \ (\found \ (a<\x \ f(\x)=0) \ (\y\a \ f(\y)=0)) - \ (\\found \ \y\a \ f(\y)\0)\\ \ - \- .{\u. f(u)=0}. - \found:= False,, - \x:=a,, \y:=a+1,, - COBEGIN .{\I1}. - WHILE \\found - INV .{\I1}. - DO .{a\\x \ (\found \ \y\a \ f(\y)=0) \ (a<\x \ f(\x)\0)}. - \ \x:=\x+1,,IF f(\x)=0 THEN \found:=True ELSE SKIP FI\ - OD - .{\I1 \ \found}. - \ - .{\I2}. - WHILE \\found - INV .{\I2}. - DO .{\y\a+1 \ (\found \ a<\x \ f(\x)=0) \ (\y\a \ f(\y)\0)}. - \ \y:=(\y - 1),,IF f(\y)=0 THEN \found:=True ELSE SKIP FI\ - OD - .{\I2 \ \found}. - COEND - .{f(\x)=0 \ f(\y)=0}." -apply oghoare ---{* 20 vc *} -apply auto ---{* auto takes aprox. 2 minutes. *} -done - -subsection {* Producer/Consumer *} - -subsubsection {* Previous lemmas *} - -lemma nat_lemma2: "\ b = m*(n::nat) + t; a = s*n + u; t=u; b-a < n \ \ m \ s" -proof - - assume "b = m*(n::nat) + t" "a = s*n + u" "t=u" - hence "(m - s) * n = b - a" by (simp add: diff_mult_distrib) - also assume "\ < n" - finally have "m - s < 1" by simp - thus ?thesis by arith -qed - -lemma mod_lemma: "\ (c::nat) \ a; a < b; b - c < n \ \ b mod n \ a mod n" -apply(subgoal_tac "b=b div n*n + b mod n" ) - prefer 2 apply (simp add: mod_div_equality [symmetric]) -apply(subgoal_tac "a=a div n*n + a mod n") - prefer 2 - apply(simp add: mod_div_equality [symmetric]) -apply(subgoal_tac "b - a \ b - c") - prefer 2 apply arith -apply(drule le_less_trans) -back - apply assumption -apply(frule less_not_refl2) -apply(drule less_imp_le) -apply (drule_tac m = "a" and k = n in div_le_mono) -apply(safe) -apply(frule_tac b = "b" and a = "a" and n = "n" in nat_lemma2, assumption, assumption) -apply assumption -apply(drule order_antisym, assumption) -apply(rotate_tac -3) -apply(simp) -done - - -subsubsection {* Producer/Consumer Algorithm *} - -record Producer_consumer = - ins :: nat - outs :: nat - li :: nat - lj :: nat - vx :: nat - vy :: nat - buffer :: "nat list" - b :: "nat list" - -text {* The whole proof takes aprox. 4 minutes. *} - -lemma Producer_consumer: - "\INIT= \0 0buffer \ length \b=length a\ ; - I= \(\k<\ins. \outs\k \ (a ! k) = \buffer ! (k mod (length \buffer))) \ - \outs\\ins \ \ins-\outs\length \buffer\ ; - I1= \\I \ \li\length a\ ; - p1= \\I1 \ \li=\ins\ ; - I2 = \\I \ (\k<\lj. (a ! k)=(\b ! k)) \ \lj\length a\ ; - p2 = \\I2 \ \lj=\outs\ \ \ - \- .{\INIT}. - \ins:=0,, \outs:=0,, \li:=0,, \lj:=0,, - COBEGIN .{\p1 \ \INIT}. - WHILE \li p1 \ \INIT}. - DO .{\p1 \ \INIT \ \livx:= (a ! \li);; - .{\p1 \ \INIT \ \li \vx=(a ! \li)}. - WAIT \ins-\outs < length \buffer END;; - .{\p1 \ \INIT \ \li \vx=(a ! \li) - \ \ins-\outs < length \buffer}. - \buffer:=(list_update \buffer (\ins mod (length \buffer)) \vx);; - .{\p1 \ \INIT \ \li (a ! \li)=(\buffer ! (\ins mod (length \buffer))) - \ \ins-\outs buffer}. - \ins:=\ins+1;; - .{\I1 \ \INIT \ (\li+1)=\ins \ \lili:=\li+1 - OD - .{\p1 \ \INIT \ \li=length a}. - \ - .{\p2 \ \INIT}. - WHILE \lj < length a - INV .{\p2 \ \INIT}. - DO .{\p2 \ \lj \INIT}. - WAIT \outs<\ins END;; - .{\p2 \ \lj \outs<\ins \ \INIT}. - \vy:=(\buffer ! (\outs mod (length \buffer)));; - .{\p2 \ \lj \outs<\ins \ \vy=(a ! \lj) \ \INIT}. - \outs:=\outs+1;; - .{\I2 \ (\lj+1)=\outs \ \lj \vy=(a ! \lj) \ \INIT}. - \b:=(list_update \b \lj \vy);; - .{\I2 \ (\lj+1)=\outs \ \lj (a ! \lj)=(\b ! \lj) \ \INIT}. - \lj:=\lj+1 - OD - .{\p2 \ \lj=length a \ \INIT}. - COEND - .{ \kb ! k)}." -apply oghoare ---{* 138 vc *} -apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) ---{* 112 subgoals left *} -apply(simp_all (no_asm)) -apply(tactic {*ALLGOALS (conjI_Tac (K all_tac)) *}) ---{* 930 subgoals left *} -apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) -apply(simp_all (asm_lr) only:length_0_conv [THEN sym]) ---{* 44 subgoals left *} -apply (simp_all (asm_lr) del:length_0_conv add: neq0_conv nth_list_update mod_less_divisor mod_lemma) ---{* 32 subgoals left *} -apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) - -apply(tactic {* TRYALL (Lin_Arith.tac @{context}) *}) ---{* 9 subgoals left *} -apply (force simp add:less_Suc_eq) -apply(drule sym) -apply (force simp add:less_Suc_eq)+ -done - -subsection {* Parameterized Examples *} - -subsubsection {* Set Elements of an Array to Zero *} - -record Example1 = - a :: "nat \ nat" - -lemma Example1: - "\- .{True}. - COBEGIN SCHEME [0\ia:=\a (i:=0) .{\a i=0}. COEND - .{\i < n. \a i = 0}." -apply oghoare -apply simp_all -done - -text {* Same example with lists as auxiliary variables. *} -record Example1_list = - A :: "nat list" -lemma Example1_list: - "\- .{n < length \A}. - COBEGIN - SCHEME [0\iA}. \A:=\A[i:=0] .{\A!i=0}. - COEND - .{\i < n. \A!i = 0}." -apply oghoare -apply force+ -done - -subsubsection {* Increment a Variable in Parallel *} - -text {* First some lemmas about summation properties. *} -(* -lemma Example2_lemma1: "!!b. j (\i::nat b j = 0 " -apply(induct n) - apply simp_all -apply(force simp add: less_Suc_eq) -done -*) -lemma Example2_lemma2_aux: "!!b. j - (\i=0..i=0..i=0.. s \ (\i::nat=0..i=0..j \ Suc (\i::nat=0..i=0..i=0..i=0..j") - apply(drule_tac b="b" and t="(Suc 0)" in Example2_lemma2_aux2) -apply(rotate_tac -1) -apply(erule ssubst) -apply simp_all -done - - -record Example2 = - c :: "nat \ nat" - x :: nat - -lemma Example_2: "0 - \- .{\x=0 \ (\i=0..c i)=0}. - COBEGIN - SCHEME [0\ix=(\i=0..c i) \ \c i=0}. - \ \x:=\x+(Suc 0),, \c:=\c (i:=(Suc 0)) \ - .{\x=(\i=0..c i) \ \c i=(Suc 0)}. - COEND - .{\x=n}." -apply oghoare -apply (simp_all cong del: strong_setsum_cong) -apply (tactic {* ALLGOALS (clarify_tac @{claset}) *}) -apply (simp_all cong del: strong_setsum_cong) - apply(erule (1) Example2_lemma2) - apply(erule (1) Example2_lemma2) - apply(erule (1) Example2_lemma2) -apply(simp) -done - -end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/OG_Hoare.thy --- a/src/HOL/HoareParallel/OG_Hoare.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,469 +0,0 @@ - -header {* \section{The Proof System} *} - -theory OG_Hoare imports OG_Tran begin - -consts assertions :: "'a ann_com \ ('a assn) set" -primrec - "assertions (AnnBasic r f) = {r}" - "assertions (AnnSeq c1 c2) = assertions c1 \ assertions c2" - "assertions (AnnCond1 r b c1 c2) = {r} \ assertions c1 \ assertions c2" - "assertions (AnnCond2 r b c) = {r} \ assertions c" - "assertions (AnnWhile r b i c) = {r, i} \ assertions c" - "assertions (AnnAwait r b c) = {r}" - -consts atomics :: "'a ann_com \ ('a assn \ 'a com) set" -primrec - "atomics (AnnBasic r f) = {(r, Basic f)}" - "atomics (AnnSeq c1 c2) = atomics c1 \ atomics c2" - "atomics (AnnCond1 r b c1 c2) = atomics c1 \ atomics c2" - "atomics (AnnCond2 r b c) = atomics c" - "atomics (AnnWhile r b i c) = atomics c" - "atomics (AnnAwait r b c) = {(r \ b, c)}" - -consts com :: "'a ann_triple_op \ 'a ann_com_op" -primrec "com (c, q) = c" - -consts post :: "'a ann_triple_op \ 'a assn" -primrec "post (c, q) = q" - -constdefs interfree_aux :: "('a ann_com_op \ 'a assn \ 'a ann_com_op) \ bool" - "interfree_aux \ \(co, q, co'). co'= None \ - (\(r,a) \ atomics (the co'). \= (q \ r) a q \ - (co = None \ (\p \ assertions (the co). \= (p \ r) a p)))" - -constdefs interfree :: "(('a ann_triple_op) list) \ bool" - "interfree Ts \ \i j. i < length Ts \ j < length Ts \ i \ j \ - interfree_aux (com (Ts!i), post (Ts!i), com (Ts!j)) " - -inductive - oghoare :: "'a assn \ 'a com \ 'a assn \ bool" ("(3\- _//_//_)" [90,55,90] 50) - and ann_hoare :: "'a ann_com \ 'a assn \ bool" ("(2\ _// _)" [60,90] 45) -where - AnnBasic: "r \ {s. f s \ q} \ \ (AnnBasic r f) q" - -| AnnSeq: "\ \ c0 pre c1; \ c1 q \ \ \ (AnnSeq c0 c1) q" - -| AnnCond1: "\ r \ b \ pre c1; \ c1 q; r \ -b \ pre c2; \ c2 q\ - \ \ (AnnCond1 r b c1 c2) q" -| AnnCond2: "\ r \ b \ pre c; \ c q; r \ -b \ q \ \ \ (AnnCond2 r b c) q" - -| AnnWhile: "\ r \ i; i \ b \ pre c; \ c i; i \ -b \ q \ - \ \ (AnnWhile r b i c) q" - -| AnnAwait: "\ atom_com c; \- (r \ b) c q \ \ \ (AnnAwait r b c) q" - -| AnnConseq: "\\ c q; q \ q' \ \ \ c q'" - - -| Parallel: "\ \ic q. Ts!i = (Some c, q) \ \ c q; interfree Ts \ - \ \- (\i\{i. ii\{i. i- {s. f s \q} (Basic f) q" - -| Seq: "\ \- p c1 r; \- r c2 q \ \ \- p (Seq c1 c2) q " - -| Cond: "\ \- (p \ b) c1 q; \- (p \ -b) c2 q \ \ \- p (Cond b c1 c2) q" - -| While: "\ \- (p \ b) c p \ \ \- p (While b i c) (p \ -b)" - -| Conseq: "\ p' \ p; \- p c q ; q \ q' \ \ \- p' c q'" - -section {* Soundness *} -(* In the version Isabelle-10-Sep-1999: HOL: The THEN and ELSE -parts of conditional expressions (if P then x else y) are no longer -simplified. (This allows the simplifier to unfold recursive -functional programs.) To restore the old behaviour, we declare -@{text "lemmas [cong del] = if_weak_cong"}. *) - -lemmas [cong del] = if_weak_cong - -lemmas ann_hoare_induct = oghoare_ann_hoare.induct [THEN conjunct2] -lemmas oghoare_induct = oghoare_ann_hoare.induct [THEN conjunct1] - -lemmas AnnBasic = oghoare_ann_hoare.AnnBasic -lemmas AnnSeq = oghoare_ann_hoare.AnnSeq -lemmas AnnCond1 = oghoare_ann_hoare.AnnCond1 -lemmas AnnCond2 = oghoare_ann_hoare.AnnCond2 -lemmas AnnWhile = oghoare_ann_hoare.AnnWhile -lemmas AnnAwait = oghoare_ann_hoare.AnnAwait -lemmas AnnConseq = oghoare_ann_hoare.AnnConseq - -lemmas Parallel = oghoare_ann_hoare.Parallel -lemmas Basic = oghoare_ann_hoare.Basic -lemmas Seq = oghoare_ann_hoare.Seq -lemmas Cond = oghoare_ann_hoare.Cond -lemmas While = oghoare_ann_hoare.While -lemmas Conseq = oghoare_ann_hoare.Conseq - -subsection {* Soundness of the System for Atomic Programs *} - -lemma Basic_ntran [rule_format]: - "(Basic f, s) -Pn\ (Parallel Ts, t) \ All_None Ts \ t = f s" -apply(induct "n") - apply(simp (no_asm)) -apply(fast dest: rel_pow_Suc_D2 Parallel_empty_lemma elim: transition_cases) -done - -lemma SEM_fwhile: "SEM S (p \ b) \ p \ SEM (fwhile b S k) p \ (p \ -b)" -apply (induct "k") - apply(simp (no_asm) add: L3_5v_lemma3) -apply(simp (no_asm) add: L3_5iv L3_5ii Parallel_empty) -apply(rule conjI) - apply (blast dest: L3_5i) -apply(simp add: SEM_def sem_def id_def) -apply (blast dest: Basic_ntran rtrancl_imp_UN_rel_pow) -done - -lemma atom_hoare_sound [rule_format]: - " \- p c q \ atom_com(c) \ \= p c q" -apply (unfold com_validity_def) -apply(rule oghoare_induct) -apply simp_all ---{*Basic*} - apply(simp add: SEM_def sem_def) - apply(fast dest: rtrancl_imp_UN_rel_pow Basic_ntran) ---{* Seq *} - apply(rule impI) - apply(rule subset_trans) - prefer 2 apply simp - apply(simp add: L3_5ii L3_5i) ---{* Cond *} - apply(simp add: L3_5iv) ---{* While *} - apply (force simp add: L3_5v dest: SEM_fwhile) ---{* Conseq *} -apply(force simp add: SEM_def sem_def) -done - -subsection {* Soundness of the System for Component Programs *} - -inductive_cases ann_transition_cases: - "(None,s) -1\ (c', s')" - "(Some (AnnBasic r f),s) -1\ (c', s')" - "(Some (AnnSeq c1 c2), s) -1\ (c', s')" - "(Some (AnnCond1 r b c1 c2), s) -1\ (c', s')" - "(Some (AnnCond2 r b c), s) -1\ (c', s')" - "(Some (AnnWhile r b I c), s) -1\ (c', s')" - "(Some (AnnAwait r b c),s) -1\ (c', s')" - -text {* Strong Soundness for Component Programs:*} - -lemma ann_hoare_case_analysis [rule_format]: - defines I: "I \ \C q'. - ((\r f. C = AnnBasic r f \ (\q. r \ {s. f s \ q} \ q \ q')) \ - (\c0 c1. C = AnnSeq c0 c1 \ (\q. q \ q' \ \ c0 pre c1 \ \ c1 q)) \ - (\r b c1 c2. C = AnnCond1 r b c1 c2 \ (\q. q \ q' \ - r \ b \ pre c1 \ \ c1 q \ r \ -b \ pre c2 \ \ c2 q)) \ - (\r b c. C = AnnCond2 r b c \ - (\q. q \ q' \ r \ b \ pre c \ \ c q \ r \ -b \ q)) \ - (\r i b c. C = AnnWhile r b i c \ - (\q. q \ q' \ r \ i \ i \ b \ pre c \ \ c i \ i \ -b \ q)) \ - (\r b c. C = AnnAwait r b c \ (\q. q \ q' \ \- (r \ b) c q)))" - shows "\ C q' \ I C q'" -apply(rule ann_hoare_induct) -apply (simp_all add: I) - apply(rule_tac x=q in exI,simp)+ -apply(rule conjI,clarify,simp,clarify,rule_tac x=qa in exI,fast)+ -apply(clarify,simp,clarify,rule_tac x=qa in exI,fast) -done - -lemma Help: "(transition \ {(x,y). True}) = (transition)" -apply force -done - -lemma Strong_Soundness_aux_aux [rule_format]: - "(co, s) -1\ (co', t) \ (\c. co = Some c \ s\ pre c \ - (\q. \ c q \ (if co' = None then t\q else t \ pre(the co') \ \ (the co') q )))" -apply(rule ann_transition_transition.induct [THEN conjunct1]) -apply simp_all ---{* Basic *} - apply clarify - apply(frule ann_hoare_case_analysis) - apply force ---{* Seq *} - apply clarify - apply(frule ann_hoare_case_analysis,simp) - apply(fast intro: AnnConseq) - apply clarify - apply(frule ann_hoare_case_analysis,simp) - apply clarify - apply(rule conjI) - apply force - apply(rule AnnSeq,simp) - apply(fast intro: AnnConseq) ---{* Cond1 *} - apply clarify - apply(frule ann_hoare_case_analysis,simp) - apply(fast intro: AnnConseq) - apply clarify - apply(frule ann_hoare_case_analysis,simp) - apply(fast intro: AnnConseq) ---{* Cond2 *} - apply clarify - apply(frule ann_hoare_case_analysis,simp) - apply(fast intro: AnnConseq) - apply clarify - apply(frule ann_hoare_case_analysis,simp) - apply(fast intro: AnnConseq) ---{* While *} - apply clarify - apply(frule ann_hoare_case_analysis,simp) - apply force - apply clarify - apply(frule ann_hoare_case_analysis,simp) - apply auto - apply(rule AnnSeq) - apply simp - apply(rule AnnWhile) - apply simp_all ---{* Await *} -apply(frule ann_hoare_case_analysis,simp) -apply clarify -apply(drule atom_hoare_sound) - apply simp -apply(simp add: com_validity_def SEM_def sem_def) -apply(simp add: Help All_None_def) -apply force -done - -lemma Strong_Soundness_aux: "\ (Some c, s) -*\ (co, t); s \ pre c; \ c q \ - \ if co = None then t \ q else t \ pre (the co) \ \ (the co) q" -apply(erule rtrancl_induct2) - apply simp -apply(case_tac "a") - apply(fast elim: ann_transition_cases) -apply(erule Strong_Soundness_aux_aux) - apply simp -apply simp_all -done - -lemma Strong_Soundness: "\ (Some c, s)-*\(co, t); s \ pre c; \ c q \ - \ if co = None then t\q else t \ pre (the co)" -apply(force dest:Strong_Soundness_aux) -done - -lemma ann_hoare_sound: "\ c q \ \ c q" -apply (unfold ann_com_validity_def ann_SEM_def ann_sem_def) -apply clarify -apply(drule Strong_Soundness) -apply simp_all -done - -subsection {* Soundness of the System for Parallel Programs *} - -lemma Parallel_length_post_P1: "(Parallel Ts,s) -P1\ (R', t) \ - (\Rs. R' = (Parallel Rs) \ (length Rs) = (length Ts) \ - (\i. i post(Rs ! i) = post(Ts ! i)))" -apply(erule transition_cases) -apply simp -apply clarify -apply(case_tac "i=ia") -apply simp+ -done - -lemma Parallel_length_post_PStar: "(Parallel Ts,s) -P*\ (R',t) \ - (\Rs. R' = (Parallel Rs) \ (length Rs) = (length Ts) \ - (\i. i post(Ts ! i) = post(Rs ! i)))" -apply(erule rtrancl_induct2) - apply(simp_all) -apply clarify -apply simp -apply(drule Parallel_length_post_P1) -apply auto -done - -lemma assertions_lemma: "pre c \ assertions c" -apply(rule ann_com_com.induct [THEN conjunct1]) -apply auto -done - -lemma interfree_aux1 [rule_format]: - "(c,s) -1\ (r,t) \ (interfree_aux(c1, q1, c) \ interfree_aux(c1, q1, r))" -apply (rule ann_transition_transition.induct [THEN conjunct1]) -apply(safe) -prefer 13 -apply (rule TrueI) -apply (simp_all add:interfree_aux_def) -apply force+ -done - -lemma interfree_aux2 [rule_format]: - "(c,s) -1\ (r,t) \ (interfree_aux(c, q, a) \ interfree_aux(r, q, a) )" -apply (rule ann_transition_transition.induct [THEN conjunct1]) -apply(force simp add:interfree_aux_def)+ -done - -lemma interfree_lemma: "\ (Some c, s) -1\ (r, t);interfree Ts ; i \ interfree (Ts[i:= (r, q)])" -apply(simp add: interfree_def) -apply clarify -apply(case_tac "i=j") - apply(drule_tac t = "ia" in not_sym) - apply simp_all -apply(force elim: interfree_aux1) -apply(force elim: interfree_aux2 simp add:nth_list_update) -done - -text {* Strong Soundness Theorem for Parallel Programs:*} - -lemma Parallel_Strong_Soundness_Seq_aux: - "\interfree Ts; i - \ interfree (Ts[i:=(Some c0, pre c1)])" -apply(simp add: interfree_def) -apply clarify -apply(case_tac "i=j") - apply(force simp add: nth_list_update interfree_aux_def) -apply(case_tac "i=ia") - apply(erule_tac x=ia in allE) - apply(force simp add:interfree_aux_def assertions_lemma) -apply simp -done - -lemma Parallel_Strong_Soundness_Seq [rule_format (no_asm)]: - "\ \i post(Ts!i) - else b \ pre(the(com(Ts!i))) \ \ the(com(Ts!i)) post(Ts!i)); - com(Ts ! i) = Some(AnnSeq c0 c1); i \ - (\ia post(Ts[i:=(Some c0, pre c1)]! ia) - else b \ pre(the(com(Ts[i:=(Some c0, pre c1)]! ia))) \ - \ the(com(Ts[i:=(Some c0, pre c1)]! ia)) post(Ts[i:=(Some c0, pre c1)]! ia))) - \ interfree (Ts[i:= (Some c0, pre c1)])" -apply(rule conjI) - apply safe - apply(case_tac "i=ia") - apply simp - apply(force dest: ann_hoare_case_analysis) - apply simp -apply(fast elim: Parallel_Strong_Soundness_Seq_aux) -done - -lemma Parallel_Strong_Soundness_aux_aux [rule_format]: - "(Some c, b) -1\ (co, t) \ - (\Ts. i com(Ts ! i) = Some c \ - (\ipost(Ts!i) - else b\pre(the(com(Ts!i))) \ \ the(com(Ts!i)) post(Ts!i))) \ - interfree Ts \ - (\j. j i\j \ (if com(Ts!j) = None then t\post(Ts!j) - else t\pre(the(com(Ts!j))) \ \ the(com(Ts!j)) post(Ts!j))) )" -apply(rule ann_transition_transition.induct [THEN conjunct1]) -apply safe -prefer 11 -apply(rule TrueI) -apply simp_all ---{* Basic *} - apply(erule_tac x = "i" in all_dupE, erule (1) notE impE) - apply(erule_tac x = "j" in allE , erule (1) notE impE) - apply(simp add: interfree_def) - apply(erule_tac x = "j" in allE,simp) - apply(erule_tac x = "i" in allE,simp) - apply(drule_tac t = "i" in not_sym) - apply(case_tac "com(Ts ! j)=None") - apply(force intro: converse_rtrancl_into_rtrancl - simp add: interfree_aux_def com_validity_def SEM_def sem_def All_None_def) - apply(simp add:interfree_aux_def) - apply clarify - apply simp - apply(erule_tac x="pre y" in ballE) - apply(force intro: converse_rtrancl_into_rtrancl - simp add: com_validity_def SEM_def sem_def All_None_def) - apply(simp add:assertions_lemma) ---{* Seqs *} - apply(erule_tac x = "Ts[i:=(Some c0, pre c1)]" in allE) - apply(drule Parallel_Strong_Soundness_Seq,simp+) - apply(erule_tac x = "Ts[i:=(Some c0, pre c1)]" in allE) - apply(drule Parallel_Strong_Soundness_Seq,simp+) ---{* Await *} -apply(rule_tac x = "i" in allE , assumption , erule (1) notE impE) -apply(erule_tac x = "j" in allE , erule (1) notE impE) -apply(simp add: interfree_def) -apply(erule_tac x = "j" in allE,simp) -apply(erule_tac x = "i" in allE,simp) -apply(drule_tac t = "i" in not_sym) -apply(case_tac "com(Ts ! j)=None") - apply(force intro: converse_rtrancl_into_rtrancl simp add: interfree_aux_def - com_validity_def SEM_def sem_def All_None_def Help) -apply(simp add:interfree_aux_def) -apply clarify -apply simp -apply(erule_tac x="pre y" in ballE) - apply(force intro: converse_rtrancl_into_rtrancl - simp add: com_validity_def SEM_def sem_def All_None_def Help) -apply(simp add:assertions_lemma) -done - -lemma Parallel_Strong_Soundness_aux [rule_format]: - "\(Ts',s) -P*\ (Rs',t); Ts' = (Parallel Ts); interfree Ts; - \i. i (\c q. (Ts ! i) = (Some c, q) \ s\(pre c) \ \ c q ) \ \ - \Rs. Rs' = (Parallel Rs) \ (\j. j - (if com(Rs ! j) = None then t\post(Ts ! j) - else t\pre(the(com(Rs ! j))) \ \ the(com(Rs ! j)) post(Ts ! j))) \ interfree Rs" -apply(erule rtrancl_induct2) - apply clarify ---{* Base *} - apply force ---{* Induction step *} -apply clarify -apply(drule Parallel_length_post_PStar) -apply clarify -apply (ind_cases "(Parallel Ts, s) -P1\ (Parallel Rs, t)" for Ts s Rs t) -apply(rule conjI) - apply clarify - apply(case_tac "i=j") - apply(simp split del:split_if) - apply(erule Strong_Soundness_aux_aux,simp+) - apply force - apply force - apply(simp split del: split_if) - apply(erule Parallel_Strong_Soundness_aux_aux) - apply(simp_all add: split del:split_if) - apply force -apply(rule interfree_lemma) -apply simp_all -done - -lemma Parallel_Strong_Soundness: - "\(Parallel Ts, s) -P*\ (Parallel Rs, t); interfree Ts; ji. i (\c q. Ts ! i = (Some c, q) \ s\pre c \ \ c q) \ \ - if com(Rs ! j) = None then t\post(Ts ! j) else t\pre (the(com(Rs ! j)))" -apply(drule Parallel_Strong_Soundness_aux) -apply simp+ -done - -lemma oghoare_sound [rule_format]: "\- p c q \ \= p c q" -apply (unfold com_validity_def) -apply(rule oghoare_induct) -apply(rule TrueI)+ ---{* Parallel *} - apply(simp add: SEM_def sem_def) - apply clarify - apply(frule Parallel_length_post_PStar) - apply clarify - apply(drule_tac j=xb in Parallel_Strong_Soundness) - apply clarify - apply simp - apply force - apply simp - apply(erule_tac V = "\i. ?P i" in thin_rl) - apply(drule_tac s = "length Rs" in sym) - apply(erule allE, erule impE, assumption) - apply(force dest: nth_mem simp add: All_None_def) ---{* Basic *} - apply(simp add: SEM_def sem_def) - apply(force dest: rtrancl_imp_UN_rel_pow Basic_ntran) ---{* Seq *} - apply(rule subset_trans) - prefer 2 apply assumption - apply(simp add: L3_5ii L3_5i) ---{* Cond *} - apply(simp add: L3_5iv) ---{* While *} - apply(simp add: L3_5v) - apply (blast dest: SEM_fwhile) ---{* Conseq *} -apply(auto simp add: SEM_def sem_def) -done - -end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/OG_Syntax.thy --- a/src/HOL/HoareParallel/OG_Syntax.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,140 +0,0 @@ -theory OG_Syntax -imports OG_Tactics Quote_Antiquote -begin - -text{* Syntax for commands and for assertions and boolean expressions in - commands @{text com} and annotated commands @{text ann_com}. *} - -syntax - "_Assign" :: "idt \ 'b \ 'a com" ("(\_ :=/ _)" [70, 65] 61) - "_AnnAssign" :: "'a assn \ idt \ 'b \ 'a com" ("(_ \_ :=/ _)" [90,70,65] 61) - -translations - "\\x := a" \ "Basic \\\(_update_name x (\_. a))\" - "r \\x := a" \ "AnnBasic r \\\(_update_name x (\_. a))\" - -syntax - "_AnnSkip" :: "'a assn \ 'a ann_com" ("_//SKIP" [90] 63) - "_AnnSeq" :: "'a ann_com \ 'a ann_com \ 'a ann_com" ("_;;/ _" [60,61] 60) - - "_AnnCond1" :: "'a assn \ 'a bexp \ 'a ann_com \ 'a ann_com \ 'a ann_com" - ("_ //IF _ /THEN _ /ELSE _ /FI" [90,0,0,0] 61) - "_AnnCond2" :: "'a assn \ 'a bexp \ 'a ann_com \ 'a ann_com" - ("_ //IF _ /THEN _ /FI" [90,0,0] 61) - "_AnnWhile" :: "'a assn \ 'a bexp \ 'a assn \ 'a ann_com \ 'a ann_com" - ("_ //WHILE _ /INV _ //DO _//OD" [90,0,0,0] 61) - "_AnnAwait" :: "'a assn \ 'a bexp \ 'a com \ 'a ann_com" - ("_ //AWAIT _ /THEN /_ /END" [90,0,0] 61) - "_AnnAtom" :: "'a assn \ 'a com \ 'a ann_com" ("_//\_\" [90,0] 61) - "_AnnWait" :: "'a assn \ 'a bexp \ 'a ann_com" ("_//WAIT _ END" [90,0] 61) - - "_Skip" :: "'a com" ("SKIP" 63) - "_Seq" :: "'a com \ 'a com \ 'a com" ("_,,/ _" [55, 56] 55) - "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" - ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61) - "_Cond2" :: "'a bexp \ 'a com \ 'a com" ("IF _ THEN _ FI" [0,0] 56) - "_While_inv" :: "'a bexp \ 'a assn \ 'a com \ 'a com" - ("(0WHILE _/ INV _ //DO _ /OD)" [0, 0, 0] 61) - "_While" :: "'a bexp \ 'a com \ 'a com" - ("(0WHILE _ //DO _ /OD)" [0, 0] 61) - -translations - "SKIP" \ "Basic id" - "c_1,, c_2" \ "Seq c_1 c_2" - - "IF b THEN c1 ELSE c2 FI" \ "Cond .{b}. c1 c2" - "IF b THEN c FI" \ "IF b THEN c ELSE SKIP FI" - "WHILE b INV i DO c OD" \ "While .{b}. i c" - "WHILE b DO c OD" \ "WHILE b INV CONST undefined DO c OD" - - "r SKIP" \ "AnnBasic r id" - "c_1;; c_2" \ "AnnSeq c_1 c_2" - "r IF b THEN c1 ELSE c2 FI" \ "AnnCond1 r .{b}. c1 c2" - "r IF b THEN c FI" \ "AnnCond2 r .{b}. c" - "r WHILE b INV i DO c OD" \ "AnnWhile r .{b}. i c" - "r AWAIT b THEN c END" \ "AnnAwait r .{b}. c" - "r \c\" \ "r AWAIT True THEN c END" - "r WAIT b END" \ "r AWAIT b THEN SKIP END" - -nonterminals - prgs - -syntax - "_PAR" :: "prgs \ 'a" ("COBEGIN//_//COEND" [57] 56) - "_prg" :: "['a, 'a] \ prgs" ("_//_" [60, 90] 57) - "_prgs" :: "['a, 'a, prgs] \ prgs" ("_//_//\//_" [60,90,57] 57) - - "_prg_scheme" :: "['a, 'a, 'a, 'a, 'a] \ prgs" - ("SCHEME [_ \ _ < _] _// _" [0,0,0,60, 90] 57) - -translations - "_prg c q" \ "[(Some c, q)]" - "_prgs c q ps" \ "(Some c, q) # ps" - "_PAR ps" \ "Parallel ps" - - "_prg_scheme j i k c q" \ "map (\i. (Some c, q)) [j.. (x, if T = dummyT then T else Term.domain_type T) - | NONE => raise Match); - - fun update_name_tr' (Free x) = Free (upd_tr' x) - | update_name_tr' ((c as Const ("_free", _)) $ Free x) = - c $ Free (upd_tr' x) - | update_name_tr' (Const x) = Const (upd_tr' x) - | update_name_tr' _ = raise Match; - - fun K_tr' (Abs (_,_,t)) = if null (loose_bnos t) then t else raise Match - | K_tr' (Abs (_,_,Abs (_,_,t)$Bound 0)) = if null (loose_bnos t) then t else raise Match - | K_tr' _ = raise Match; - - fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) = - quote_tr' (Syntax.const "_Assign" $ update_name_tr' f) - (Abs (x, dummyT, K_tr' k) :: ts) - | assign_tr' _ = raise Match; - - fun annassign_tr' (r :: Abs (x, _, f $ k $ Bound 0) :: ts) = - quote_tr' (Syntax.const "_AnnAssign" $ r $ update_name_tr' f) - (Abs (x, dummyT, K_tr' k) :: ts) - | annassign_tr' _ = raise Match; - - fun Parallel_PAR [(Const ("Cons",_) $ (Const ("Pair",_) $ (Const ("Some",_) $ t1 ) $ t2) $ Const ("Nil",_))] = - (Syntax.const "_prg" $ t1 $ t2) - | Parallel_PAR [(Const ("Cons",_) $ (Const ("Pair",_) $ (Const ("Some",_) $ t1) $ t2) $ ts)] = - (Syntax.const "_prgs" $ t1 $ t2 $ Parallel_PAR [ts]) - | Parallel_PAR _ = raise Match; - -fun Parallel_tr' ts = Syntax.const "_PAR" $ Parallel_PAR ts; - in - [("Collect", assert_tr'), ("Basic", assign_tr'), - ("Cond", bexp_tr' "_Cond"), ("While", bexp_tr' "_While_inv"), - ("AnnBasic", annassign_tr'), - ("AnnWhile", annbexp_tr' "_AnnWhile"), ("AnnAwait", annbexp_tr' "_AnnAwait"), - ("AnnCond1", annbexp_tr' "_AnnCond1"), ("AnnCond2", annbexp_tr' "_AnnCond2")] - - end - -*} - -end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/OG_Tactics.thy --- a/src/HOL/HoareParallel/OG_Tactics.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,496 +0,0 @@ -header {* \section{Generation of Verification Conditions} *} - -theory OG_Tactics -imports OG_Hoare -begin - -lemmas ann_hoare_intros=AnnBasic AnnSeq AnnCond1 AnnCond2 AnnWhile AnnAwait AnnConseq -lemmas oghoare_intros=Parallel Basic Seq Cond While Conseq - -lemma ParallelConseqRule: - "\ p \ (\i\{i. i- (\i\{i. ii\{i. ii\{i. i q \ - \ \- p (Parallel Ts) q" -apply (rule Conseq) -prefer 2 - apply fast -apply assumption+ -done - -lemma SkipRule: "p \ q \ \- p (Basic id) q" -apply(rule oghoare_intros) - prefer 2 apply(rule Basic) - prefer 2 apply(rule subset_refl) -apply(simp add:Id_def) -done - -lemma BasicRule: "p \ {s. (f s)\q} \ \- p (Basic f) q" -apply(rule oghoare_intros) - prefer 2 apply(rule oghoare_intros) - prefer 2 apply(rule subset_refl) -apply assumption -done - -lemma SeqRule: "\ \- p c1 r; \- r c2 q \ \ \- p (Seq c1 c2) q" -apply(rule Seq) -apply fast+ -done - -lemma CondRule: - "\ p \ {s. (s\b \ s\w) \ (s\b \ s\w')}; \- w c1 q; \- w' c2 q \ - \ \- p (Cond b c1 c2) q" -apply(rule Cond) - apply(rule Conseq) - prefer 4 apply(rule Conseq) -apply simp_all -apply force+ -done - -lemma WhileRule: "\ p \ i; \- (i \ b) c i ; (i \ (-b)) \ q \ - \ \- p (While b i c) q" -apply(rule Conseq) - prefer 2 apply(rule While) -apply assumption+ -done - -text {* Three new proof rules for special instances of the @{text -AnnBasic} and the @{text AnnAwait} commands when the transformation -performed on the state is the identity, and for an @{text AnnAwait} -command where the boolean condition is @{text "{s. True}"}: *} - -lemma AnnatomRule: - "\ atom_com(c); \- r c q \ \ \ (AnnAwait r {s. True} c) q" -apply(rule AnnAwait) -apply simp_all -done - -lemma AnnskipRule: - "r \ q \ \ (AnnBasic r id) q" -apply(rule AnnBasic) -apply simp -done - -lemma AnnwaitRule: - "\ (r \ b) \ q \ \ \ (AnnAwait r b (Basic id)) q" -apply(rule AnnAwait) - apply simp -apply(rule BasicRule) -apply simp -done - -text {* Lemmata to avoid using the definition of @{text -map_ann_hoare}, @{text interfree_aux}, @{text interfree_swap} and -@{text interfree} by splitting it into different cases: *} - -lemma interfree_aux_rule1: "interfree_aux(co, q, None)" -by(simp add:interfree_aux_def) - -lemma interfree_aux_rule2: - "\(R,r)\(atomics a). \- (q \ R) r q \ interfree_aux(None, q, Some a)" -apply(simp add:interfree_aux_def) -apply(force elim:oghoare_sound) -done - -lemma interfree_aux_rule3: - "(\(R, r)\(atomics a). \- (q \ R) r q \ (\p\(assertions c). \- (p \ R) r p)) - \ interfree_aux(Some c, q, Some a)" -apply(simp add:interfree_aux_def) -apply(force elim:oghoare_sound) -done - -lemma AnnBasic_assertions: - "\interfree_aux(None, r, Some a); interfree_aux(None, q, Some a)\ \ - interfree_aux(Some (AnnBasic r f), q, Some a)" -apply(simp add: interfree_aux_def) -by force - -lemma AnnSeq_assertions: - "\ interfree_aux(Some c1, q, Some a); interfree_aux(Some c2, q, Some a)\\ - interfree_aux(Some (AnnSeq c1 c2), q, Some a)" -apply(simp add: interfree_aux_def) -by force - -lemma AnnCond1_assertions: - "\ interfree_aux(None, r, Some a); interfree_aux(Some c1, q, Some a); - interfree_aux(Some c2, q, Some a)\\ - interfree_aux(Some(AnnCond1 r b c1 c2), q, Some a)" -apply(simp add: interfree_aux_def) -by force - -lemma AnnCond2_assertions: - "\ interfree_aux(None, r, Some a); interfree_aux(Some c, q, Some a)\\ - interfree_aux(Some (AnnCond2 r b c), q, Some a)" -apply(simp add: interfree_aux_def) -by force - -lemma AnnWhile_assertions: - "\ interfree_aux(None, r, Some a); interfree_aux(None, i, Some a); - interfree_aux(Some c, q, Some a)\\ - interfree_aux(Some (AnnWhile r b i c), q, Some a)" -apply(simp add: interfree_aux_def) -by force - -lemma AnnAwait_assertions: - "\ interfree_aux(None, r, Some a); interfree_aux(None, q, Some a)\\ - interfree_aux(Some (AnnAwait r b c), q, Some a)" -apply(simp add: interfree_aux_def) -by force - -lemma AnnBasic_atomics: - "\- (q \ r) (Basic f) q \ interfree_aux(None, q, Some (AnnBasic r f))" -by(simp add: interfree_aux_def oghoare_sound) - -lemma AnnSeq_atomics: - "\ interfree_aux(Any, q, Some a1); interfree_aux(Any, q, Some a2)\\ - interfree_aux(Any, q, Some (AnnSeq a1 a2))" -apply(simp add: interfree_aux_def) -by force - -lemma AnnCond1_atomics: - "\ interfree_aux(Any, q, Some a1); interfree_aux(Any, q, Some a2)\\ - interfree_aux(Any, q, Some (AnnCond1 r b a1 a2))" -apply(simp add: interfree_aux_def) -by force - -lemma AnnCond2_atomics: - "interfree_aux (Any, q, Some a)\ interfree_aux(Any, q, Some (AnnCond2 r b a))" -by(simp add: interfree_aux_def) - -lemma AnnWhile_atomics: "interfree_aux (Any, q, Some a) - \ interfree_aux(Any, q, Some (AnnWhile r b i a))" -by(simp add: interfree_aux_def) - -lemma Annatom_atomics: - "\- (q \ r) a q \ interfree_aux (None, q, Some (AnnAwait r {x. True} a))" -by(simp add: interfree_aux_def oghoare_sound) - -lemma AnnAwait_atomics: - "\- (q \ (r \ b)) a q \ interfree_aux (None, q, Some (AnnAwait r b a))" -by(simp add: interfree_aux_def oghoare_sound) - -constdefs - interfree_swap :: "('a ann_triple_op * ('a ann_triple_op) list) \ bool" - "interfree_swap == \(x, xs). \y\set xs. interfree_aux (com x, post x, com y) - \ interfree_aux(com y, post y, com x)" - -lemma interfree_swap_Empty: "interfree_swap (x, [])" -by(simp add:interfree_swap_def) - -lemma interfree_swap_List: - "\ interfree_aux (com x, post x, com y); - interfree_aux (com y, post y ,com x); interfree_swap (x, xs) \ - \ interfree_swap (x, y#xs)" -by(simp add:interfree_swap_def) - -lemma interfree_swap_Map: "\k. i\k \ k interfree_aux (com x, post x, c k) - \ interfree_aux (c k, Q k, com x) - \ interfree_swap (x, map (\k. (c k, Q k)) [i.. interfree_swap(x, xs); interfree xs \ \ interfree (x#xs)" -apply(simp add:interfree_def interfree_swap_def) -apply clarify -apply(case_tac i) - apply(case_tac j) - apply simp_all -apply(case_tac j,simp+) -done - -lemma interfree_Map: - "(\i j. a\i \ i a\j \ j i\j \ interfree_aux (c i, Q i, c j)) - \ interfree (map (\k. (c k, Q k)) [a.. bool " ("[\] _" [0] 45) - "[\] Ts == (\ic q. Ts!i=(Some c, q) \ \ c q)" - -lemma MapAnnEmpty: "[\] []" -by(simp add:map_ann_hoare_def) - -lemma MapAnnList: "\ \ c q ; [\] xs \ \ [\] (Some c,q)#xs" -apply(simp add:map_ann_hoare_def) -apply clarify -apply(case_tac i,simp+) -done - -lemma MapAnnMap: - "\k. i\k \ k \ (c k) (Q k) \ [\] map (\k. (Some (c k), Q k)) [i.. [\] Ts ; interfree Ts \ - \ \- (\i\{i. ii\{i. i \k (c k) (Q k); - \k l. k l k\l \ interfree_aux (Some(c k), Q k, Some(c l)) \ - \ \- (\i\{i. iii\{i. i(b x)}" -by fast -lemma list_length: "length []=0 \ length (x#xs) = Suc(length xs)" -by simp -lemma list_lemmas: "length []=0 \ length (x#xs) = Suc(length xs) -\ (x#xs) ! 0=x \ (x#xs) ! Suc n = xs ! n" -by simp -lemma le_Suc_eq_insert: "{i. i a1 \ a2 \ .. \ an"} returning -@{text n} subgoals, one for each conjunct: *} - -ML {* -fun conjI_Tac tac i st = st |> - ( (EVERY [rtac conjI i, - conjI_Tac tac (i+1), - tac i]) ORELSE (tac i) ) -*} - - -subsubsection {* Tactic for the generation of the verification conditions *} - -text {* The tactic basically uses two subtactics: - -\begin{description} - -\item[HoareRuleTac] is called at the level of parallel programs, it - uses the ParallelTac to solve parallel composition of programs. - This verification has two parts, namely, (1) all component programs are - correct and (2) they are interference free. @{text HoareRuleTac} is - also called at the level of atomic regions, i.e. @{text "\ \"} and - @{text "AWAIT b THEN _ END"}, and at each interference freedom test. - -\item[AnnHoareRuleTac] is for component programs which - are annotated programs and so, there are not unknown assertions - (no need to use the parameter precond, see NOTE). - - NOTE: precond(::bool) informs if the subgoal has the form @{text "\- ?p c q"}, - in this case we have precond=False and the generated verification - condition would have the form @{text "?p \ \"} which can be solved by - @{text "rtac subset_refl"}, if True we proceed to simplify it using - the simplification tactics above. - -\end{description} -*} - -ML {* - - fun WlpTac i = (rtac (@{thm SeqRule}) i) THEN (HoareRuleTac false (i+1)) -and HoareRuleTac precond i st = st |> - ( (WlpTac i THEN HoareRuleTac precond i) - ORELSE - (FIRST[rtac (@{thm SkipRule}) i, - rtac (@{thm BasicRule}) i, - EVERY[rtac (@{thm ParallelConseqRule}) i, - ParallelConseq (i+2), - ParallelTac (i+1), - ParallelConseq i], - EVERY[rtac (@{thm CondRule}) i, - HoareRuleTac false (i+2), - HoareRuleTac false (i+1)], - EVERY[rtac (@{thm WhileRule}) i, - HoareRuleTac true (i+1)], - K all_tac i ] - THEN (if precond then (K all_tac i) else (rtac (@{thm subset_refl}) i)))) - -and AnnWlpTac i = (rtac (@{thm AnnSeq}) i) THEN (AnnHoareRuleTac (i+1)) -and AnnHoareRuleTac i st = st |> - ( (AnnWlpTac i THEN AnnHoareRuleTac i ) - ORELSE - (FIRST[(rtac (@{thm AnnskipRule}) i), - EVERY[rtac (@{thm AnnatomRule}) i, - HoareRuleTac true (i+1)], - (rtac (@{thm AnnwaitRule}) i), - rtac (@{thm AnnBasic}) i, - EVERY[rtac (@{thm AnnCond1}) i, - AnnHoareRuleTac (i+3), - AnnHoareRuleTac (i+1)], - EVERY[rtac (@{thm AnnCond2}) i, - AnnHoareRuleTac (i+1)], - EVERY[rtac (@{thm AnnWhile}) i, - AnnHoareRuleTac (i+2)], - EVERY[rtac (@{thm AnnAwait}) i, - HoareRuleTac true (i+1)], - K all_tac i])) - -and ParallelTac i = EVERY[rtac (@{thm ParallelRule}) i, - interfree_Tac (i+1), - MapAnn_Tac i] - -and MapAnn_Tac i st = st |> - (FIRST[rtac (@{thm MapAnnEmpty}) i, - EVERY[rtac (@{thm MapAnnList}) i, - MapAnn_Tac (i+1), - AnnHoareRuleTac i], - EVERY[rtac (@{thm MapAnnMap}) i, - rtac (@{thm allI}) i,rtac (@{thm impI}) i, - AnnHoareRuleTac i]]) - -and interfree_swap_Tac i st = st |> - (FIRST[rtac (@{thm interfree_swap_Empty}) i, - EVERY[rtac (@{thm interfree_swap_List}) i, - interfree_swap_Tac (i+2), - interfree_aux_Tac (i+1), - interfree_aux_Tac i ], - EVERY[rtac (@{thm interfree_swap_Map}) i, - rtac (@{thm allI}) i,rtac (@{thm impI}) i, - conjI_Tac (interfree_aux_Tac) i]]) - -and interfree_Tac i st = st |> - (FIRST[rtac (@{thm interfree_Empty}) i, - EVERY[rtac (@{thm interfree_List}) i, - interfree_Tac (i+1), - interfree_swap_Tac i], - EVERY[rtac (@{thm interfree_Map}) i, - rtac (@{thm allI}) i,rtac (@{thm allI}) i,rtac (@{thm impI}) i, - interfree_aux_Tac i ]]) - -and interfree_aux_Tac i = (before_interfree_simp_tac i ) THEN - (FIRST[rtac (@{thm interfree_aux_rule1}) i, - dest_assertions_Tac i]) - -and dest_assertions_Tac i st = st |> - (FIRST[EVERY[rtac (@{thm AnnBasic_assertions}) i, - dest_atomics_Tac (i+1), - dest_atomics_Tac i], - EVERY[rtac (@{thm AnnSeq_assertions}) i, - dest_assertions_Tac (i+1), - dest_assertions_Tac i], - EVERY[rtac (@{thm AnnCond1_assertions}) i, - dest_assertions_Tac (i+2), - dest_assertions_Tac (i+1), - dest_atomics_Tac i], - EVERY[rtac (@{thm AnnCond2_assertions}) i, - dest_assertions_Tac (i+1), - dest_atomics_Tac i], - EVERY[rtac (@{thm AnnWhile_assertions}) i, - dest_assertions_Tac (i+2), - dest_atomics_Tac (i+1), - dest_atomics_Tac i], - EVERY[rtac (@{thm AnnAwait_assertions}) i, - dest_atomics_Tac (i+1), - dest_atomics_Tac i], - dest_atomics_Tac i]) - -and dest_atomics_Tac i st = st |> - (FIRST[EVERY[rtac (@{thm AnnBasic_atomics}) i, - HoareRuleTac true i], - EVERY[rtac (@{thm AnnSeq_atomics}) i, - dest_atomics_Tac (i+1), - dest_atomics_Tac i], - EVERY[rtac (@{thm AnnCond1_atomics}) i, - dest_atomics_Tac (i+1), - dest_atomics_Tac i], - EVERY[rtac (@{thm AnnCond2_atomics}) i, - dest_atomics_Tac i], - EVERY[rtac (@{thm AnnWhile_atomics}) i, - dest_atomics_Tac i], - EVERY[rtac (@{thm Annatom_atomics}) i, - HoareRuleTac true i], - EVERY[rtac (@{thm AnnAwait_atomics}) i, - HoareRuleTac true i], - K all_tac i]) -*} - - -text {* The final tactic is given the name @{text oghoare}: *} - -ML {* -val oghoare_tac = SUBGOAL (fn (_, i) => - (HoareRuleTac true i)) -*} - -text {* Notice that the tactic for parallel programs @{text -"oghoare_tac"} is initially invoked with the value @{text true} for -the parameter @{text precond}. - -Parts of the tactic can be also individually used to generate the -verification conditions for annotated sequential programs and to -generate verification conditions out of interference freedom tests: *} - -ML {* val annhoare_tac = SUBGOAL (fn (_, i) => - (AnnHoareRuleTac i)) - -val interfree_aux_tac = SUBGOAL (fn (_, i) => - (interfree_aux_Tac i)) -*} - -text {* The so defined ML tactics are then ``exported'' to be used in -Isabelle proofs. *} - -method_setup oghoare = {* - Scan.succeed (K (SIMPLE_METHOD' oghoare_tac)) *} - "verification condition generator for the oghoare logic" - -method_setup annhoare = {* - Scan.succeed (K (SIMPLE_METHOD' annhoare_tac)) *} - "verification condition generator for the ann_hoare logic" - -method_setup interfree_aux = {* - Scan.succeed (K (SIMPLE_METHOD' interfree_aux_tac)) *} - "verification condition generator for interference freedom tests" - -text {* Tactics useful for dealing with the generated verification conditions: *} - -method_setup conjI_tac = {* - Scan.succeed (K (SIMPLE_METHOD' (conjI_Tac (K all_tac)))) *} - "verification condition generator for interference freedom tests" - -ML {* -fun disjE_Tac tac i st = st |> - ( (EVERY [etac disjE i, - disjE_Tac tac (i+1), - tac i]) ORELSE (tac i) ) -*} - -method_setup disjE_tac = {* - Scan.succeed (K (SIMPLE_METHOD' (disjE_Tac (K all_tac)))) *} - "verification condition generator for interference freedom tests" - -end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/OG_Tran.thy --- a/src/HOL/HoareParallel/OG_Tran.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,309 +0,0 @@ - -header {* \section{Operational Semantics} *} - -theory OG_Tran imports OG_Com begin - -types - 'a ann_com_op = "('a ann_com) option" - 'a ann_triple_op = "('a ann_com_op \ 'a assn)" - -consts com :: "'a ann_triple_op \ 'a ann_com_op" -primrec "com (c, q) = c" - -consts post :: "'a ann_triple_op \ 'a assn" -primrec "post (c, q) = q" - -constdefs - All_None :: "'a ann_triple_op list \ bool" - "All_None Ts \ \(c, q) \ set Ts. c = None" - -subsection {* The Transition Relation *} - -inductive_set - ann_transition :: "(('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a)) set" - and transition :: "(('a com \ 'a) \ ('a com \ 'a)) set" - and ann_transition' :: "('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a) \ bool" - ("_ -1\ _"[81,81] 100) - and transition' :: "('a com \ 'a) \ ('a com \ 'a) \ bool" - ("_ -P1\ _"[81,81] 100) - and transitions :: "('a com \ 'a) \ ('a com \ 'a) \ bool" - ("_ -P*\ _"[81,81] 100) -where - "con_0 -1\ con_1 \ (con_0, con_1) \ ann_transition" -| "con_0 -P1\ con_1 \ (con_0, con_1) \ transition" -| "con_0 -P*\ con_1 \ (con_0, con_1) \ transition\<^sup>*" - -| AnnBasic: "(Some (AnnBasic r f), s) -1\ (None, f s)" - -| AnnSeq1: "(Some c0, s) -1\ (None, t) \ - (Some (AnnSeq c0 c1), s) -1\ (Some c1, t)" -| AnnSeq2: "(Some c0, s) -1\ (Some c2, t) \ - (Some (AnnSeq c0 c1), s) -1\ (Some (AnnSeq c2 c1), t)" - -| AnnCond1T: "s \ b \ (Some (AnnCond1 r b c1 c2), s) -1\ (Some c1, s)" -| AnnCond1F: "s \ b \ (Some (AnnCond1 r b c1 c2), s) -1\ (Some c2, s)" - -| AnnCond2T: "s \ b \ (Some (AnnCond2 r b c), s) -1\ (Some c, s)" -| AnnCond2F: "s \ b \ (Some (AnnCond2 r b c), s) -1\ (None, s)" - -| AnnWhileF: "s \ b \ (Some (AnnWhile r b i c), s) -1\ (None, s)" -| AnnWhileT: "s \ b \ (Some (AnnWhile r b i c), s) -1\ - (Some (AnnSeq c (AnnWhile i b i c)), s)" - -| AnnAwait: "\ s \ b; atom_com c; (c, s) -P*\ (Parallel [], t) \ \ - (Some (AnnAwait r b c), s) -1\ (None, t)" - -| Parallel: "\ i (r, t) \ - \ (Parallel Ts, s) -P1\ (Parallel (Ts [i:=(r, q)]), t)" - -| Basic: "(Basic f, s) -P1\ (Parallel [], f s)" - -| Seq1: "All_None Ts \ (Seq (Parallel Ts) c, s) -P1\ (c, s)" -| Seq2: "(c0, s) -P1\ (c2, t) \ (Seq c0 c1, s) -P1\ (Seq c2 c1, t)" - -| CondT: "s \ b \ (Cond b c1 c2, s) -P1\ (c1, s)" -| CondF: "s \ b \ (Cond b c1 c2, s) -P1\ (c2, s)" - -| WhileF: "s \ b \ (While b i c, s) -P1\ (Parallel [], s)" -| WhileT: "s \ b \ (While b i c, s) -P1\ (Seq c (While b i c), s)" - -monos "rtrancl_mono" - -text {* The corresponding syntax translations are: *} - -abbreviation - ann_transition_n :: "('a ann_com_op \ 'a) \ nat \ ('a ann_com_op \ 'a) - \ bool" ("_ -_\ _"[81,81] 100) where - "con_0 -n\ con_1 \ (con_0, con_1) \ ann_transition ^^ n" - -abbreviation - ann_transitions :: "('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a) \ bool" - ("_ -*\ _"[81,81] 100) where - "con_0 -*\ con_1 \ (con_0, con_1) \ ann_transition\<^sup>*" - -abbreviation - transition_n :: "('a com \ 'a) \ nat \ ('a com \ 'a) \ bool" - ("_ -P_\ _"[81,81,81] 100) where - "con_0 -Pn\ con_1 \ (con_0, con_1) \ transition ^^ n" - -subsection {* Definition of Semantics *} - -constdefs - ann_sem :: "'a ann_com \ 'a \ 'a set" - "ann_sem c \ \s. {t. (Some c, s) -*\ (None, t)}" - - ann_SEM :: "'a ann_com \ 'a set \ 'a set" - "ann_SEM c S \ \ann_sem c ` S" - - sem :: "'a com \ 'a \ 'a set" - "sem c \ \s. {t. \Ts. (c, s) -P*\ (Parallel Ts, t) \ All_None Ts}" - - SEM :: "'a com \ 'a set \ 'a set" - "SEM c S \ \sem c ` S " - -syntax "_Omega" :: "'a com" ("\" 63) -translations "\" \ "While CONST UNIV CONST UNIV (Basic id)" - -consts fwhile :: "'a bexp \ 'a com \ nat \ 'a com" -primrec - "fwhile b c 0 = \" - "fwhile b c (Suc n) = Cond b (Seq c (fwhile b c n)) (Basic id)" - -subsubsection {* Proofs *} - -declare ann_transition_transition.intros [intro] -inductive_cases transition_cases: - "(Parallel T,s) -P1\ t" - "(Basic f, s) -P1\ t" - "(Seq c1 c2, s) -P1\ t" - "(Cond b c1 c2, s) -P1\ t" - "(While b i c, s) -P1\ t" - -lemma Parallel_empty_lemma [rule_format (no_asm)]: - "(Parallel [],s) -Pn\ (Parallel Ts,t) \ Ts=[] \ n=0 \ s=t" -apply(induct n) - apply(simp (no_asm)) -apply clarify -apply(drule rel_pow_Suc_D2) -apply(force elim:transition_cases) -done - -lemma Parallel_AllNone_lemma [rule_format (no_asm)]: - "All_None Ss \ (Parallel Ss,s) -Pn\ (Parallel Ts,t) \ Ts=Ss \ n=0 \ s=t" -apply(induct "n") - apply(simp (no_asm)) -apply clarify -apply(drule rel_pow_Suc_D2) -apply clarify -apply(erule transition_cases,simp_all) -apply(force dest:nth_mem simp add:All_None_def) -done - -lemma Parallel_AllNone: "All_None Ts \ (SEM (Parallel Ts) X) = X" -apply (unfold SEM_def sem_def) -apply auto -apply(drule rtrancl_imp_UN_rel_pow) -apply clarify -apply(drule Parallel_AllNone_lemma) -apply auto -done - -lemma Parallel_empty: "Ts=[] \ (SEM (Parallel Ts) X) = X" -apply(rule Parallel_AllNone) -apply(simp add:All_None_def) -done - -text {* Set of lemmas from Apt and Olderog "Verification of sequential -and concurrent programs", page 63. *} - -lemma L3_5i: "X\Y \ SEM c X \ SEM c Y" -apply (unfold SEM_def) -apply force -done - -lemma L3_5ii_lemma1: - "\ (c1, s1) -P*\ (Parallel Ts, s2); All_None Ts; - (c2, s2) -P*\ (Parallel Ss, s3); All_None Ss \ - \ (Seq c1 c2, s1) -P*\ (Parallel Ss, s3)" -apply(erule converse_rtrancl_induct2) -apply(force intro:converse_rtrancl_into_rtrancl)+ -done - -lemma L3_5ii_lemma2 [rule_format (no_asm)]: - "\c1 c2 s t. (Seq c1 c2, s) -Pn\ (Parallel Ts, t) \ - (All_None Ts) \ (\y m Rs. (c1,s) -P*\ (Parallel Rs, y) \ - (All_None Rs) \ (c2, y) -Pm\ (Parallel Ts, t) \ m \ n)" -apply(induct "n") - apply(force) -apply(safe dest!: rel_pow_Suc_D2) -apply(erule transition_cases,simp_all) - apply (fast intro!: le_SucI) -apply (fast intro!: le_SucI elim!: rel_pow_imp_rtrancl converse_rtrancl_into_rtrancl) -done - -lemma L3_5ii_lemma3: - "\(Seq c1 c2,s) -P*\ (Parallel Ts,t); All_None Ts\ \ - (\y Rs. (c1,s) -P*\ (Parallel Rs,y) \ All_None Rs - \ (c2,y) -P*\ (Parallel Ts,t))" -apply(drule rtrancl_imp_UN_rel_pow) -apply(fast dest: L3_5ii_lemma2 rel_pow_imp_rtrancl) -done - -lemma L3_5ii: "SEM (Seq c1 c2) X = SEM c2 (SEM c1 X)" -apply (unfold SEM_def sem_def) -apply auto - apply(fast dest: L3_5ii_lemma3) -apply(fast elim: L3_5ii_lemma1) -done - -lemma L3_5iii: "SEM (Seq (Seq c1 c2) c3) X = SEM (Seq c1 (Seq c2 c3)) X" -apply (simp (no_asm) add: L3_5ii) -done - -lemma L3_5iv: - "SEM (Cond b c1 c2) X = (SEM c1 (X \ b)) Un (SEM c2 (X \ (-b)))" -apply (unfold SEM_def sem_def) -apply auto -apply(erule converse_rtranclE) - prefer 2 - apply (erule transition_cases,simp_all) - apply(fast intro: converse_rtrancl_into_rtrancl elim: transition_cases)+ -done - - -lemma L3_5v_lemma1[rule_format]: - "(S,s) -Pn\ (T,t) \ S=\ \ (\(\Rs. T=(Parallel Rs) \ All_None Rs))" -apply (unfold UNIV_def) -apply(rule nat_less_induct) -apply safe -apply(erule rel_pow_E2) - apply simp_all -apply(erule transition_cases) - apply simp_all -apply(erule rel_pow_E2) - apply(simp add: Id_def) -apply(erule transition_cases,simp_all) -apply clarify -apply(erule transition_cases,simp_all) -apply(erule rel_pow_E2,simp) -apply clarify -apply(erule transition_cases) - apply simp+ - apply clarify - apply(erule transition_cases) -apply simp_all -done - -lemma L3_5v_lemma2: "\(\, s) -P*\ (Parallel Ts, t); All_None Ts \ \ False" -apply(fast dest: rtrancl_imp_UN_rel_pow L3_5v_lemma1) -done - -lemma L3_5v_lemma3: "SEM (\) S = {}" -apply (unfold SEM_def sem_def) -apply(fast dest: L3_5v_lemma2) -done - -lemma L3_5v_lemma4 [rule_format]: - "\s. (While b i c, s) -Pn\ (Parallel Ts, t) \ All_None Ts \ - (\k. (fwhile b c k, s) -P*\ (Parallel Ts, t))" -apply(rule nat_less_induct) -apply safe -apply(erule rel_pow_E2) - apply safe -apply(erule transition_cases,simp_all) - apply (rule_tac x = "1" in exI) - apply(force dest: Parallel_empty_lemma intro: converse_rtrancl_into_rtrancl simp add: Id_def) -apply safe -apply(drule L3_5ii_lemma2) - apply safe -apply(drule le_imp_less_Suc) -apply (erule allE , erule impE,assumption) -apply (erule allE , erule impE, assumption) -apply safe -apply (rule_tac x = "k+1" in exI) -apply(simp (no_asm)) -apply(rule converse_rtrancl_into_rtrancl) - apply fast -apply(fast elim: L3_5ii_lemma1) -done - -lemma L3_5v_lemma5 [rule_format]: - "\s. (fwhile b c k, s) -P*\ (Parallel Ts, t) \ All_None Ts \ - (While b i c, s) -P*\ (Parallel Ts,t)" -apply(induct "k") - apply(force dest: L3_5v_lemma2) -apply safe -apply(erule converse_rtranclE) - apply simp_all -apply(erule transition_cases,simp_all) - apply(rule converse_rtrancl_into_rtrancl) - apply(fast) - apply(fast elim!: L3_5ii_lemma1 dest: L3_5ii_lemma3) -apply(drule rtrancl_imp_UN_rel_pow) -apply clarify -apply(erule rel_pow_E2) - apply simp_all -apply(erule transition_cases,simp_all) -apply(fast dest: Parallel_empty_lemma) -done - -lemma L3_5v: "SEM (While b i c) = (\x. (\k. SEM (fwhile b c k) x))" -apply(rule ext) -apply (simp add: SEM_def sem_def) -apply safe - apply(drule rtrancl_imp_UN_rel_pow,simp) - apply clarify - apply(fast dest:L3_5v_lemma4) -apply(fast intro: L3_5v_lemma5) -done - -section {* Validity of Correctness Formulas *} - -constdefs - com_validity :: "'a assn \ 'a com \ 'a assn \ bool" ("(3\= _// _//_)" [90,55,90] 50) - "\= p c q \ SEM c p \ q" - - ann_com_validity :: "'a ann_com \ 'a assn \ bool" ("\ _ _" [60,90] 45) - "\ c q \ ann_SEM c (pre c) \ q" - -end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/Quote_Antiquote.thy --- a/src/HOL/HoareParallel/Quote_Antiquote.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,24 +0,0 @@ - -header {* \section{Concrete Syntax} *} - -theory Quote_Antiquote imports Main begin - -syntax - "_quote" :: "'b \ ('a \ 'b)" ("(\_\)" [0] 1000) - "_antiquote" :: "('a \ 'b) \ 'b" ("\_" [1000] 1000) - "_Assert" :: "'a \ 'a set" ("(.{_}.)" [0] 1000) - -syntax (xsymbols) - "_Assert" :: "'a \ 'a set" ("(\_\)" [0] 1000) - -translations - ".{b}." \ "Collect \b\" - -parse_translation {* - let - fun quote_tr [t] = Syntax.quote_tr "_antiquote" t - | quote_tr ts = raise TERM ("quote_tr", ts); - in [("_quote", quote_tr)] end -*} - -end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/RG_Com.thy --- a/src/HOL/HoareParallel/RG_Com.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,25 +0,0 @@ - -header {* \chapter{The Rely-Guarantee Method} - -\section {Abstract Syntax} -*} - -theory RG_Com imports Main begin - -text {* Semantics of assertions and boolean expressions (bexp) as sets -of states. Syntax of commands @{text com} and parallel commands -@{text par_com}. *} - -types - 'a bexp = "'a set" - -datatype 'a com = - Basic "'a \'a" - | Seq "'a com" "'a com" - | Cond "'a bexp" "'a com" "'a com" - | While "'a bexp" "'a com" - | Await "'a bexp" "'a com" - -types 'a par_com = "(('a com) option) list" - -end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/RG_Examples.thy --- a/src/HOL/HoareParallel/RG_Examples.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,359 +0,0 @@ -header {* \section{Examples} *} - -theory RG_Examples -imports RG_Syntax -begin - -lemmas definitions [simp]= stable_def Pre_def Rely_def Guar_def Post_def Com_def - -subsection {* Set Elements of an Array to Zero *} - -lemma le_less_trans2: "\(j::nat) j\ \ i (a::nat) < c; b\d \ \ a + b < c + d" -by simp - -record Example1 = - A :: "nat list" - -lemma Example1: - "\ COBEGIN - SCHEME [0 \ i < n] - (\A := \A [i := 0], - \ n < length \A \, - \ length \A = length \A \ \A ! i = \A ! i \, - \ length \A = length \A \ (\j j \ \A ! j = \A ! j) \, - \ \A ! i = 0 \) - COEND - SAT [\ n < length \A \, \ \A = \A \, \ True \, \ \i < n. \A ! i = 0 \]" -apply(rule Parallel) -apply (auto intro!: Basic) -done - -lemma Example1_parameterized: -"k < t \ - \ COBEGIN - SCHEME [k*n\i<(Suc k)*n] (\A:=\A[i:=0], - \t*n < length \A\, - \t*n < length \A \ length \A=length \A \ \A!i = \A!i\, - \t*n < length \A \ length \A=length \A \ (\jA . i\j \ \A!j = \A!j)\, - \\A!i=0\) - COEND - SAT [\t*n < length \A\, - \t*n < length \A \ length \A=length \A \ (\iA!(k*n+i)=\A!(k*n+i))\, - \t*n < length \A \ length \A=length \A \ - (\iA . (i \A!i = \A!i) \ ((Suc k)*n \ i\ \A!i = \A!i))\, - \\iA!(k*n+i) = 0\]" -apply(rule Parallel) - apply auto - apply(erule_tac x="k*n +i" in allE) - apply(subgoal_tac "k*n+i COBEGIN - (\ \x:=\x+1;; \c_0:=\c_0 + 1 \, - \\x=\c_0 + \c_1 \ \c_0=0\, - \\c_0 = \c_0 \ - (\x=\c_0 + \c_1 - \ \x = \c_0 + \c_1)\, - \\c_1 = \c_1 \ - (\x=\c_0 + \c_1 - \ \x =\c_0 + \c_1)\, - \\x=\c_0 + \c_1 \ \c_0=1 \) - \ - (\ \x:=\x+1;; \c_1:=\c_1+1 \, - \\x=\c_0 + \c_1 \ \c_1=0 \, - \\c_1 = \c_1 \ - (\x=\c_0 + \c_1 - \ \x = \c_0 + \c_1)\, - \\c_0 = \c_0 \ - (\x=\c_0 + \c_1 - \ \x =\c_0 + \c_1)\, - \\x=\c_0 + \c_1 \ \c_1=1\) - COEND - SAT [\\x=0 \ \c_0=0 \ \c_1=0\, - \\x=\x \ \c_0= \c_0 \ \c_1=\c_1\, - \True\, - \\x=2\]" -apply(rule Parallel) - apply simp_all - apply clarify - apply(case_tac i) - apply simp - apply(rule conjI) - apply clarify - apply simp - apply clarify - apply simp - apply(case_tac j,simp) - apply simp - apply simp - apply(rule conjI) - apply clarify - apply simp - apply clarify - apply simp - apply(subgoal_tac "j=0") - apply (rotate_tac -1) - apply (simp (asm_lr)) - apply arith - apply clarify - apply(case_tac i,simp,simp) - apply clarify - apply simp - apply(erule_tac x=0 in all_dupE) - apply(erule_tac x=1 in allE,simp) -apply clarify -apply(case_tac i,simp) - apply(rule Await) - apply simp_all - apply(clarify) - apply(rule Seq) - prefer 2 - apply(rule Basic) - apply simp_all - apply(rule subset_refl) - apply(rule Basic) - apply simp_all - apply clarify - apply simp -apply(rule Await) - apply simp_all -apply(clarify) -apply(rule Seq) - prefer 2 - apply(rule Basic) - apply simp_all - apply(rule subset_refl) -apply(auto intro!: Basic) -done - -subsubsection {* Parameterized *} - -lemma Example2_lemma2_aux: "j - (\i=0..i=0..i=0.. s \ (\i::nat=0..i=0..j \ Suc (\i::nat=0..i=0..i=0..i=0..j") - apply(drule_tac b="b" and t="(Suc 0)" in Example2_lemma2_aux2) -apply(rotate_tac -1) -apply(erule ssubst) -apply simp_all -done - -lemma Example2_lemma2_Suc0: "\j \ - Suc (\i::nat=0..< n. b i)=(\i=0..< n. (b (j:=Suc 0)) i)" -by(simp add:Example2_lemma2) - -record Example2_parameterized = - C :: "nat \ nat" - y :: nat - -lemma Example2_parameterized: "0 - \ COBEGIN SCHEME [0\i \y:=\y+1;; \C:=\C (i:=1) \, - \\y=(\i=0..C i) \ \C i=0\, - \\C i = \C i \ - (\y=(\i=0..C i) \ \y =(\i=0..C i))\, - \(\jj \ \C j = \C j) \ - (\y=(\i=0..C i) \ \y =(\i=0..C i))\, - \\y=(\i=0..C i) \ \C i=1\) - COEND - SAT [\\y=0 \ (\i=0..C i)=0 \, \\C=\C \ \y=\y\, \True\, \\y=n\]" -apply(rule Parallel) -apply force -apply force -apply(force) -apply clarify -apply simp -apply(simp cong:setsum_ivl_cong) -apply clarify -apply simp -apply(rule Await) -apply simp_all -apply clarify -apply(rule Seq) -prefer 2 -apply(rule Basic) -apply(rule subset_refl) -apply simp+ -apply(rule Basic) -apply simp -apply clarify -apply simp -apply(simp add:Example2_lemma2_Suc0 cong:if_cong) -apply simp+ -done - -subsection {* Find Least Element *} - -text {* A previous lemma: *} - -lemma mod_aux :"\i < (n::nat); a mod n = i; j < a + n; j mod n = i; a < j\ \ False" -apply(subgoal_tac "a=a div n*n + a mod n" ) - prefer 2 apply (simp (no_asm_use)) -apply(subgoal_tac "j=j div n*n + j mod n") - prefer 2 apply (simp (no_asm_use)) -apply simp -apply(subgoal_tac "a div n*n < j div n*n") -prefer 2 apply arith -apply(subgoal_tac "j div n*n < (a div n + 1)*n") -prefer 2 apply simp -apply (simp only:mult_less_cancel2) -apply arith -done - -record Example3 = - X :: "nat \ nat" - Y :: "nat \ nat" - -lemma Example3: "m mod n=0 \ - \ COBEGIN - SCHEME [0\ijX i < \Y j) DO - IF P(B!(\X i)) THEN \Y:=\Y (i:=\X i) - ELSE \X:= \X (i:=(\X i)+ n) FI - OD, - \(\X i) mod n=i \ (\j<\X i. j mod n=i \ \P(B!j)) \ (\Y i P(B!(\Y i)) \ \Y i\ m+i)\, - \(\jj \ \Y j \ \Y j) \ \X i = \X i \ - \Y i = \Y i\, - \(\jj \ \X j = \X j \ \Y j = \Y j) \ - \Y i \ \Y i\, - \(\X i) mod n=i \ (\j<\X i. j mod n=i \ \P(B!j)) \ (\Y i P(B!(\Y i)) \ \Y i\ m+i) \ (\jY j \ \X i) \) - COEND - SAT [\ \iX i=i \ \Y i=m+i \,\\X=\X \ \Y=\Y\,\True\, - \\iX i) mod n=i \ (\j<\X i. j mod n=i \ \P(B!j)) \ - (\Y i P(B!(\Y i)) \ \Y i\ m+i) \ (\jY j \ \X i)\]" -apply(rule Parallel) ---{*5 subgoals left *} -apply force+ -apply clarify -apply simp -apply(rule While) - apply force - apply force - apply force - apply(rule_tac pre'="\ \X i mod n = i \ (\j. j<\X i \ j mod n = i \ \P(B!j)) \ (\Y i < n * q \ P (B!(\Y i))) \ \X i<\Y i\" in Conseq) - apply force - apply(rule subset_refl)+ - apply(rule Cond) - apply force - apply(rule Basic) - apply force - apply fastsimp - apply force - apply force - apply(rule Basic) - apply simp - apply clarify - apply simp - apply (case_tac "X x (j mod n) \ j") - apply (drule le_imp_less_or_eq) - apply (erule disjE) - apply (drule_tac j=j and n=n and i="j mod n" and a="X x (j mod n)" in mod_aux) - apply auto -done - -text {* Same but with a list as auxiliary variable: *} - -record Example3_list = - X :: "nat list" - Y :: "nat list" - -lemma Example3_list: "m mod n=0 \ \ (COBEGIN SCHEME [0\ijX!i < \Y!j) DO - IF P(B!(\X!i)) THEN \Y:=\Y[i:=\X!i] ELSE \X:= \X[i:=(\X!i)+ n] FI - OD, - \nX \ nY \ (\X!i) mod n=i \ (\j<\X!i. j mod n=i \ \P(B!j)) \ (\Y!i P(B!(\Y!i)) \ \Y!i\ m+i)\, - \(\jj \ \Y!j \ \Y!j) \ \X!i = \X!i \ - \Y!i = \Y!i \ length \X = length \X \ length \Y = length \Y\, - \(\jj \ \X!j = \X!j \ \Y!j = \Y!j) \ - \Y!i \ \Y!i \ length \X = length \X \ length \Y = length \Y\, - \(\X!i) mod n=i \ (\j<\X!i. j mod n=i \ \P(B!j)) \ (\Y!i P(B!(\Y!i)) \ \Y!i\ m+i) \ (\jY!j \ \X!i) \) COEND) - SAT [\nX \ nY \ (\iX!i=i \ \Y!i=m+i) \, - \\X=\X \ \Y=\Y\, - \True\, - \\iX!i) mod n=i \ (\j<\X!i. j mod n=i \ \P(B!j)) \ - (\Y!i P(B!(\Y!i)) \ \Y!i\ m+i) \ (\jY!j \ \X!i)\]" -apply(rule Parallel) ---{* 5 subgoals left *} -apply force+ -apply clarify -apply simp -apply(rule While) - apply force - apply force - apply force - apply(rule_tac pre'="\nX \ nY \ \X ! i mod n = i \ (\j. j < \X ! i \ j mod n = i \ \ P (B ! j)) \ (\Y ! i < n * q \ P (B ! (\Y ! i))) \ \X!i<\Y!i\" in Conseq) - apply force - apply(rule subset_refl)+ - apply(rule Cond) - apply force - apply(rule Basic) - apply force - apply force - apply force - apply force - apply(rule Basic) - apply simp - apply clarify - apply simp - apply(rule allI) - apply(rule impI)+ - apply(case_tac "X x ! i\ j") - apply(drule le_imp_less_or_eq) - apply(erule disjE) - apply(drule_tac j=j and n=n and i=i and a="X x ! i" in mod_aux) - apply auto -done - -end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/RG_Hoare.thy --- a/src/HOL/HoareParallel/RG_Hoare.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1375 +0,0 @@ -header {* \section{The Proof System} *} - -theory RG_Hoare imports RG_Tran begin - -subsection {* Proof System for Component Programs *} - -declare Un_subset_iff [iff del] -declare Cons_eq_map_conv[iff] - -constdefs - stable :: "'a set \ ('a \ 'a) set \ bool" - "stable \ \f g. (\x y. x \ f \ (x, y) \ g \ y \ f)" - -inductive - rghoare :: "['a com, 'a set, ('a \ 'a) set, ('a \ 'a) set, 'a set] \ bool" - ("\ _ sat [_, _, _, _]" [60,0,0,0,0] 45) -where - Basic: "\ pre \ {s. f s \ post}; {(s,t). s \ pre \ (t=f s \ t=s)} \ guar; - stable pre rely; stable post rely \ - \ \ Basic f sat [pre, rely, guar, post]" - -| Seq: "\ \ P sat [pre, rely, guar, mid]; \ Q sat [mid, rely, guar, post] \ - \ \ Seq P Q sat [pre, rely, guar, post]" - -| Cond: "\ stable pre rely; \ P1 sat [pre \ b, rely, guar, post]; - \ P2 sat [pre \ -b, rely, guar, post]; \s. (s,s)\guar \ - \ \ Cond b P1 P2 sat [pre, rely, guar, post]" - -| While: "\ stable pre rely; (pre \ -b) \ post; stable post rely; - \ P sat [pre \ b, rely, guar, pre]; \s. (s,s)\guar \ - \ \ While b P sat [pre, rely, guar, post]" - -| Await: "\ stable pre rely; stable post rely; - \V. \ P sat [pre \ b \ {V}, {(s, t). s = t}, - UNIV, {s. (V, s) \ guar} \ post] \ - \ \ Await b P sat [pre, rely, guar, post]" - -| Conseq: "\ pre \ pre'; rely \ rely'; guar' \ guar; post' \ post; - \ P sat [pre', rely', guar', post'] \ - \ \ P sat [pre, rely, guar, post]" - -constdefs - Pre :: "'a rgformula \ 'a set" - "Pre x \ fst(snd x)" - Post :: "'a rgformula \ 'a set" - "Post x \ snd(snd(snd(snd x)))" - Rely :: "'a rgformula \ ('a \ 'a) set" - "Rely x \ fst(snd(snd x))" - Guar :: "'a rgformula \ ('a \ 'a) set" - "Guar x \ fst(snd(snd(snd x)))" - Com :: "'a rgformula \ 'a com" - "Com x \ fst x" - -subsection {* Proof System for Parallel Programs *} - -types 'a par_rgformula = "('a rgformula) list \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set" - -inductive - par_rghoare :: "('a rgformula) list \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set \ bool" - ("\ _ SAT [_, _, _, _]" [60,0,0,0,0] 45) -where - Parallel: - "\ \i (\j\{j. j j\i}. Guar(xs!j)) \ Rely(xs!i); - (\j\{j. j guar; - pre \ (\i\{i. ii\{i. i post; - \i Com(xs!i) sat [Pre(xs!i),Rely(xs!i),Guar(xs!i),Post(xs!i)] \ - \ \ xs SAT [pre, rely, guar, post]" - -section {* Soundness*} - -subsubsection {* Some previous lemmas *} - -lemma tl_of_assum_in_assum: - "(P, s) # (P, t) # xs \ assum (pre, rely) \ stable pre rely - \ (P, t) # xs \ assum (pre, rely)" -apply(simp add:assum_def) -apply clarify -apply(rule conjI) - apply(erule_tac x=0 in allE) - apply(simp (no_asm_use)only:stable_def) - apply(erule allE,erule allE,erule impE,assumption,erule mp) - apply(simp add:Env) -apply clarify -apply(erule_tac x="Suc i" in allE) -apply simp -done - -lemma etran_in_comm: - "(P, t) # xs \ comm(guar, post) \ (P, s) # (P, t) # xs \ comm(guar, post)" -apply(simp add:comm_def) -apply clarify -apply(case_tac i,simp+) -done - -lemma ctran_in_comm: - "\(s, s) \ guar; (Q, s) # xs \ comm(guar, post)\ - \ (P, s) # (Q, s) # xs \ comm(guar, post)" -apply(simp add:comm_def) -apply clarify -apply(case_tac i,simp+) -done - -lemma takecptn_is_cptn [rule_format, elim!]: - "\j. c \ cptn \ take (Suc j) c \ cptn" -apply(induct "c") - apply(force elim: cptn.cases) -apply clarify -apply(case_tac j) - apply simp - apply(rule CptnOne) -apply simp -apply(force intro:cptn.intros elim:cptn.cases) -done - -lemma dropcptn_is_cptn [rule_format,elim!]: - "\j cptn \ drop j c \ cptn" -apply(induct "c") - apply(force elim: cptn.cases) -apply clarify -apply(case_tac j,simp+) -apply(erule cptn.cases) - apply simp - apply force -apply force -done - -lemma takepar_cptn_is_par_cptn [rule_format,elim]: - "\j. c \ par_cptn \ take (Suc j) c \ par_cptn" -apply(induct "c") - apply(force elim: cptn.cases) -apply clarify -apply(case_tac j,simp) - apply(rule ParCptnOne) -apply(force intro:par_cptn.intros elim:par_cptn.cases) -done - -lemma droppar_cptn_is_par_cptn [rule_format]: - "\j par_cptn \ drop j c \ par_cptn" -apply(induct "c") - apply(force elim: par_cptn.cases) -apply clarify -apply(case_tac j,simp+) -apply(erule par_cptn.cases) - apply simp - apply force -apply force -done - -lemma tl_of_cptn_is_cptn: "\x # xs \ cptn; xs \ []\ \ xs \ cptn" -apply(subgoal_tac "1 < length (x # xs)") - apply(drule dropcptn_is_cptn,simp+) -done - -lemma not_ctran_None [rule_format]: - "\s. (None, s)#xs \ cptn \ (\i xs!i)" -apply(induct xs,simp+) -apply clarify -apply(erule cptn.cases,simp) - apply simp - apply(case_tac i,simp) - apply(rule Env) - apply simp -apply(force elim:ctran.cases) -done - -lemma cptn_not_empty [simp]:"[] \ cptn" -apply(force elim:cptn.cases) -done - -lemma etran_or_ctran [rule_format]: - "\m i. x\cptn \ m \ length x - \ (\i. Suc i < m \ \ x!i -c\ x!Suc i) \ Suc i < m - \ x!i -e\ x!Suc i" -apply(induct x,simp) -apply clarify -apply(erule cptn.cases,simp) - apply(case_tac i,simp) - apply(rule Env) - apply simp - apply(erule_tac x="m - 1" in allE) - apply(case_tac m,simp,simp) - apply(subgoal_tac "(\i. Suc i < nata \ (((P, t) # xs) ! i, xs ! i) \ ctran)") - apply force - apply clarify - apply(erule_tac x="Suc ia" in allE,simp) -apply(erule_tac x="0" and P="\j. ?H j \ (?J j) \ ctran" in allE,simp) -done - -lemma etran_or_ctran2 [rule_format]: - "\i. Suc i x\cptn \ (x!i -c\ x!Suc i \ \ x!i -e\ x!Suc i) - \ (x!i -e\ x!Suc i \ \ x!i -c\ x!Suc i)" -apply(induct x) - apply simp -apply clarify -apply(erule cptn.cases,simp) - apply(case_tac i,simp+) -apply(case_tac i,simp) - apply(force elim:etran.cases) -apply simp -done - -lemma etran_or_ctran2_disjI1: - "\ x\cptn; Suc i x!Suc i\ \ \ x!i -e\ x!Suc i" -by(drule etran_or_ctran2,simp_all) - -lemma etran_or_ctran2_disjI2: - "\ x\cptn; Suc i x!Suc i\ \ \ x!i -c\ x!Suc i" -by(drule etran_or_ctran2,simp_all) - -lemma not_ctran_None2 [rule_format]: - "\ (None, s) # xs \cptn; i \ \ ((None, s) # xs) ! i -c\ xs ! i" -apply(frule not_ctran_None,simp) -apply(case_tac i,simp) - apply(force elim:etranE) -apply simp -apply(rule etran_or_ctran2_disjI2,simp_all) -apply(force intro:tl_of_cptn_is_cptn) -done - -lemma Ex_first_occurrence [rule_format]: "P (n::nat) \ (\m. P m \ (\i P i))"; -apply(rule nat_less_induct) -apply clarify -apply(case_tac "\m. m \ P m") -apply auto -done - -lemma stability [rule_format]: - "\j k. x \ cptn \ stable p rely \ j\k \ k snd(x!j)\p \ - (\i. (Suc i) - (x!i -e\ x!(Suc i)) \ (snd(x!i), snd(x!(Suc i))) \ rely) \ - (\i. j\i \ i x!i -e\ x!Suc i) \ snd(x!k)\p \ fst(x!j)=fst(x!k)" -apply(induct x) - apply clarify - apply(force elim:cptn.cases) -apply clarify -apply(erule cptn.cases,simp) - apply simp - apply(case_tac k,simp,simp) - apply(case_tac j,simp) - apply(erule_tac x=0 in allE) - apply(erule_tac x="nat" and P="\j. (0\j) \ (?J j)" in allE,simp) - apply(subgoal_tac "t\p") - apply(subgoal_tac "(\i. i < length xs \ ((P, t) # xs) ! i -e\ xs ! i \ (snd (((P, t) # xs) ! i), snd (xs ! i)) \ rely)") - apply clarify - apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j)\etran" in allE,simp) - apply clarify - apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j) \ (?T j)\rely" in allE,simp) - apply(erule_tac x=0 and P="\j. (?H j) \ (?J j)\etran \ ?T j" in allE,simp) - apply(simp(no_asm_use) only:stable_def) - apply(erule_tac x=s in allE) - apply(erule_tac x=t in allE) - apply simp - apply(erule mp) - apply(erule mp) - apply(rule Env) - apply simp - apply(erule_tac x="nata" in allE) - apply(erule_tac x="nat" and P="\j. (?s\j) \ (?J j)" in allE,simp) - apply(subgoal_tac "(\i. i < length xs \ ((P, t) # xs) ! i -e\ xs ! i \ (snd (((P, t) # xs) ! i), snd (xs ! i)) \ rely)") - apply clarify - apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j)\etran" in allE,simp) - apply clarify - apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j) \ (?T j)\rely" in allE,simp) -apply(case_tac k,simp,simp) -apply(case_tac j) - apply(erule_tac x=0 and P="\j. (?H j) \ (?J j)\etran" in allE,simp) - apply(erule etran.cases,simp) -apply(erule_tac x="nata" in allE) -apply(erule_tac x="nat" and P="\j. (?s\j) \ (?J j)" in allE,simp) -apply(subgoal_tac "(\i. i < length xs \ ((Q, t) # xs) ! i -e\ xs ! i \ (snd (((Q, t) # xs) ! i), snd (xs ! i)) \ rely)") - apply clarify - apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j)\etran" in allE,simp) -apply clarify -apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j) \ (?T j)\rely" in allE,simp) -done - -subsection {* Soundness of the System for Component Programs *} - -subsubsection {* Soundness of the Basic rule *} - -lemma unique_ctran_Basic [rule_format]: - "\s i. x \ cptn \ x ! 0 = (Some (Basic f), s) \ - Suc i x!i -c\ x!Suc i \ - (\j. Suc j i\j \ x!j -e\ x!Suc j)" -apply(induct x,simp) -apply simp -apply clarify -apply(erule cptn.cases,simp) - apply(case_tac i,simp+) - apply clarify - apply(case_tac j,simp) - apply(rule Env) - apply simp -apply clarify -apply simp -apply(case_tac i) - apply(case_tac j,simp,simp) - apply(erule ctran.cases,simp_all) - apply(force elim: not_ctran_None) -apply(ind_cases "((Some (Basic f), sa), Q, t) \ ctran" for sa Q t) -apply simp -apply(drule_tac i=nat in not_ctran_None,simp) -apply(erule etranE,simp) -done - -lemma exists_ctran_Basic_None [rule_format]: - "\s i. x \ cptn \ x ! 0 = (Some (Basic f), s) - \ i fst(x!i)=None \ (\j x!Suc j)" -apply(induct x,simp) -apply simp -apply clarify -apply(erule cptn.cases,simp) - apply(case_tac i,simp,simp) - apply(erule_tac x=nat in allE,simp) - apply clarify - apply(rule_tac x="Suc j" in exI,simp,simp) -apply clarify -apply(case_tac i,simp,simp) -apply(rule_tac x=0 in exI,simp) -done - -lemma Basic_sound: - " \pre \ {s. f s \ post}; {(s, t). s \ pre \ t = f s} \ guar; - stable pre rely; stable post rely\ - \ \ Basic f sat [pre, rely, guar, post]" -apply(unfold com_validity_def) -apply clarify -apply(simp add:comm_def) -apply(rule conjI) - apply clarify - apply(simp add:cp_def assum_def) - apply clarify - apply(frule_tac j=0 and k=i and p=pre in stability) - apply simp_all - apply(erule_tac x=ia in allE,simp) - apply(erule_tac i=i and f=f in unique_ctran_Basic,simp_all) - apply(erule subsetD,simp) - apply(case_tac "x!i") - apply clarify - apply(drule_tac s="Some (Basic f)" in sym,simp) - apply(thin_tac "\j. ?H j") - apply(force elim:ctran.cases) -apply clarify -apply(simp add:cp_def) -apply clarify -apply(frule_tac i="length x - 1" and f=f in exists_ctran_Basic_None,simp+) - apply(case_tac x,simp+) - apply(rule last_fst_esp,simp add:last_length) - apply (case_tac x,simp+) -apply(simp add:assum_def) -apply clarify -apply(frule_tac j=0 and k="j" and p=pre in stability) - apply simp_all - apply(erule_tac x=i in allE,simp) - apply(erule_tac i=j and f=f in unique_ctran_Basic,simp_all) -apply(case_tac "x!j") -apply clarify -apply simp -apply(drule_tac s="Some (Basic f)" in sym,simp) -apply(case_tac "x!Suc j",simp) -apply(rule ctran.cases,simp) -apply(simp_all) -apply(drule_tac c=sa in subsetD,simp) -apply clarify -apply(frule_tac j="Suc j" and k="length x - 1" and p=post in stability,simp_all) - apply(case_tac x,simp+) - apply(erule_tac x=i in allE) -apply(erule_tac i=j and f=f in unique_ctran_Basic,simp_all) - apply arith+ -apply(case_tac x) -apply(simp add:last_length)+ -done - -subsubsection{* Soundness of the Await rule *} - -lemma unique_ctran_Await [rule_format]: - "\s i. x \ cptn \ x ! 0 = (Some (Await b c), s) \ - Suc i x!i -c\ x!Suc i \ - (\j. Suc j i\j \ x!j -e\ x!Suc j)" -apply(induct x,simp+) -apply clarify -apply(erule cptn.cases,simp) - apply(case_tac i,simp+) - apply clarify - apply(case_tac j,simp) - apply(rule Env) - apply simp -apply clarify -apply simp -apply(case_tac i) - apply(case_tac j,simp,simp) - apply(erule ctran.cases,simp_all) - apply(force elim: not_ctran_None) -apply(ind_cases "((Some (Await b c), sa), Q, t) \ ctran" for sa Q t,simp) -apply(drule_tac i=nat in not_ctran_None,simp) -apply(erule etranE,simp) -done - -lemma exists_ctran_Await_None [rule_format]: - "\s i. x \ cptn \ x ! 0 = (Some (Await b c), s) - \ i fst(x!i)=None \ (\j x!Suc j)" -apply(induct x,simp+) -apply clarify -apply(erule cptn.cases,simp) - apply(case_tac i,simp+) - apply(erule_tac x=nat in allE,simp) - apply clarify - apply(rule_tac x="Suc j" in exI,simp,simp) -apply clarify -apply(case_tac i,simp,simp) -apply(rule_tac x=0 in exI,simp) -done - -lemma Star_imp_cptn: - "(P, s) -c*\ (R, t) \ \l \ cp P s. (last l)=(R, t) - \ (\i. Suc i l!i -c\ l!Suc i)" -apply (erule converse_rtrancl_induct2) - apply(rule_tac x="[(R,t)]" in bexI) - apply simp - apply(simp add:cp_def) - apply(rule CptnOne) -apply clarify -apply(rule_tac x="(a, b)#l" in bexI) - apply (rule conjI) - apply(case_tac l,simp add:cp_def) - apply(simp add:last_length) - apply clarify -apply(case_tac i,simp) -apply(simp add:cp_def) -apply force -apply(simp add:cp_def) - apply(case_tac l) - apply(force elim:cptn.cases) -apply simp -apply(erule CptnComp) -apply clarify -done - -lemma Await_sound: - "\stable pre rely; stable post rely; - \V. \ P sat [pre \ b \ {s. s = V}, {(s, t). s = t}, - UNIV, {s. (V, s) \ guar} \ post] \ - \ P sat [pre \ b \ {s. s = V}, {(s, t). s = t}, - UNIV, {s. (V, s) \ guar} \ post] \ - \ \ Await b P sat [pre, rely, guar, post]" -apply(unfold com_validity_def) -apply clarify -apply(simp add:comm_def) -apply(rule conjI) - apply clarify - apply(simp add:cp_def assum_def) - apply clarify - apply(frule_tac j=0 and k=i and p=pre in stability,simp_all) - apply(erule_tac x=ia in allE,simp) - apply(subgoal_tac "x\ cp (Some(Await b P)) s") - apply(erule_tac i=i in unique_ctran_Await,force,simp_all) - apply(simp add:cp_def) ---{* here starts the different part. *} - apply(erule ctran.cases,simp_all) - apply(drule Star_imp_cptn) - apply clarify - apply(erule_tac x=sa in allE) - apply clarify - apply(erule_tac x=sa in allE) - apply(drule_tac c=l in subsetD) - apply (simp add:cp_def) - apply clarify - apply(erule_tac x=ia and P="\i. ?H i \ (?J i,?I i)\ctran" in allE,simp) - apply(erule etranE,simp) - apply simp -apply clarify -apply(simp add:cp_def) -apply clarify -apply(frule_tac i="length x - 1" in exists_ctran_Await_None,force) - apply (case_tac x,simp+) - apply(rule last_fst_esp,simp add:last_length) - apply(case_tac x, (simp add:cptn_not_empty)+) -apply clarify -apply(simp add:assum_def) -apply clarify -apply(frule_tac j=0 and k="j" and p=pre in stability,simp_all) - apply(erule_tac x=i in allE,simp) - apply(erule_tac i=j in unique_ctran_Await,force,simp_all) -apply(case_tac "x!j") -apply clarify -apply simp -apply(drule_tac s="Some (Await b P)" in sym,simp) -apply(case_tac "x!Suc j",simp) -apply(rule ctran.cases,simp) -apply(simp_all) -apply(drule Star_imp_cptn) -apply clarify -apply(erule_tac x=sa in allE) -apply clarify -apply(erule_tac x=sa in allE) -apply(drule_tac c=l in subsetD) - apply (simp add:cp_def) - apply clarify - apply(erule_tac x=i and P="\i. ?H i \ (?J i,?I i)\ctran" in allE,simp) - apply(erule etranE,simp) -apply simp -apply clarify -apply(frule_tac j="Suc j" and k="length x - 1" and p=post in stability,simp_all) - apply(case_tac x,simp+) - apply(erule_tac x=i in allE) -apply(erule_tac i=j in unique_ctran_Await,force,simp_all) - apply arith+ -apply(case_tac x) -apply(simp add:last_length)+ -done - -subsubsection{* Soundness of the Conditional rule *} - -lemma Cond_sound: - "\ stable pre rely; \ P1 sat [pre \ b, rely, guar, post]; - \ P2 sat [pre \ - b, rely, guar, post]; \s. (s,s)\guar\ - \ \ (Cond b P1 P2) sat [pre, rely, guar, post]" -apply(unfold com_validity_def) -apply clarify -apply(simp add:cp_def comm_def) -apply(case_tac "\i. Suc i x!i -c\ x!Suc i") - prefer 2 - apply simp - apply clarify - apply(frule_tac j="0" and k="length x - 1" and p=pre in stability,simp+) - apply(case_tac x,simp+) - apply(simp add:assum_def) - apply(simp add:assum_def) - apply(erule_tac m="length x" in etran_or_ctran,simp+) - apply(case_tac x, (simp add:last_length)+) -apply(erule exE) -apply(drule_tac n=i and P="\i. ?H i \ (?J i,?I i)\ ctran" in Ex_first_occurrence) -apply clarify -apply (simp add:assum_def) -apply(frule_tac j=0 and k="m" and p=pre in stability,simp+) - apply(erule_tac m="Suc m" in etran_or_ctran,simp+) -apply(erule ctran.cases,simp_all) - apply(erule_tac x="sa" in allE) - apply(drule_tac c="drop (Suc m) x" in subsetD) - apply simp - apply clarify - apply simp - apply clarify - apply(case_tac "i\m") - apply(drule le_imp_less_or_eq) - apply(erule disjE) - apply(erule_tac x=i in allE, erule impE, assumption) - apply simp+ - apply(erule_tac x="i - (Suc m)" and P="\j. ?H j \ ?J j \ (?I j)\guar" in allE) - apply(subgoal_tac "(Suc m)+(i - Suc m) \ length x") - apply(subgoal_tac "(Suc m)+Suc (i - Suc m) \ length x") - apply(rotate_tac -2) - apply simp - apply arith - apply arith -apply(case_tac "length (drop (Suc m) x)",simp) -apply(erule_tac x="sa" in allE) -back -apply(drule_tac c="drop (Suc m) x" in subsetD,simp) - apply clarify -apply simp -apply clarify -apply(case_tac "i\m") - apply(drule le_imp_less_or_eq) - apply(erule disjE) - apply(erule_tac x=i in allE, erule impE, assumption) - apply simp - apply simp -apply(erule_tac x="i - (Suc m)" and P="\j. ?H j \ ?J j \ (?I j)\guar" in allE) -apply(subgoal_tac "(Suc m)+(i - Suc m) \ length x") - apply(subgoal_tac "(Suc m)+Suc (i - Suc m) \ length x") - apply(rotate_tac -2) - apply simp - apply arith -apply arith -done - -subsubsection{* Soundness of the Sequential rule *} - -inductive_cases Seq_cases [elim!]: "(Some (Seq P Q), s) -c\ t" - -lemma last_lift_not_None: "fst ((lift Q) ((x#xs)!(length xs))) \ None" -apply(subgoal_tac "length xs cptn_mod \ \s P. x !0=(Some (Seq P Q), s) \ - (\iSome Q) \ - (\xs\ cp (Some P) s. x=map (lift Q) xs)" -apply(erule cptn_mod.induct) -apply(unfold cp_def) -apply safe -apply simp_all - apply(simp add:lift_def) - apply(rule_tac x="[(Some Pa, sa)]" in exI,simp add:CptnOne) - apply(subgoal_tac "(\i < Suc (length xs). fst (((Some (Seq Pa Q), t) # xs) ! i) \ Some Q)") - apply clarify - apply(rule_tac x="(Some Pa, sa) #(Some Pa, t) # zs" in exI,simp) - apply(rule conjI,erule CptnEnv) - apply(simp (no_asm_use) add:lift_def) - apply clarify - apply(erule_tac x="Suc i" in allE, simp) - apply(ind_cases "((Some (Seq Pa Q), sa), None, t) \ ctran" for Pa sa t) - apply(rule_tac x="(Some P, sa) # xs" in exI, simp add:cptn_iff_cptn_mod lift_def) -apply(erule_tac x="length xs" in allE, simp) -apply(simp only:Cons_lift_append) -apply(subgoal_tac "length xs < length ((Some P, sa) # xs)") - apply(simp only :nth_append length_map last_length nth_map) - apply(case_tac "last((Some P, sa) # xs)") - apply(simp add:lift_def) -apply simp -done -declare map_eq_Cons_conv [simp del] Cons_eq_map_conv [simp del] - -lemma Seq_sound2 [rule_format]: - "x \ cptn \ \s P i. x!0=(Some (Seq P Q), s) \ i fst(x!i)=Some Q \ - (\j(Some Q)) \ - (\xs ys. xs \ cp (Some P) s \ length xs=Suc i - \ ys \ cp (Some Q) (snd(xs !i)) \ x=(map (lift Q) xs)@tl ys)" -apply(erule cptn.induct) -apply(unfold cp_def) -apply safe -apply simp_all - apply(case_tac i,simp+) - apply(erule allE,erule impE,assumption,simp) - apply clarify - apply(subgoal_tac "(\j < nat. fst (((Some (Seq Pa Q), t) # xs) ! j) \ Some Q)",clarify) - prefer 2 - apply force - apply(case_tac xsa,simp,simp) - apply(rule_tac x="(Some Pa, sa) #(Some Pa, t) # list" in exI,simp) - apply(rule conjI,erule CptnEnv) - apply(simp (no_asm_use) add:lift_def) - apply(rule_tac x=ys in exI,simp) -apply(ind_cases "((Some (Seq Pa Q), sa), t) \ ctran" for Pa sa t) - apply simp - apply(rule_tac x="(Some Pa, sa)#[(None, ta)]" in exI,simp) - apply(rule conjI) - apply(drule_tac xs="[]" in CptnComp,force simp add:CptnOne,simp) - apply(case_tac i, simp+) - apply(case_tac nat,simp+) - apply(rule_tac x="(Some Q,ta)#xs" in exI,simp add:lift_def) - apply(case_tac nat,simp+) - apply(force) -apply(case_tac i, simp+) -apply(case_tac nat,simp+) -apply(erule_tac x="Suc nata" in allE,simp) -apply clarify -apply(subgoal_tac "(\j Some Q)",clarify) - prefer 2 - apply clarify - apply force -apply(rule_tac x="(Some Pa, sa)#(Some P2, ta)#(tl xsa)" in exI,simp) -apply(rule conjI,erule CptnComp) -apply(rule nth_tl_if,force,simp+) -apply(rule_tac x=ys in exI,simp) -apply(rule conjI) -apply(rule nth_tl_if,force,simp+) - apply(rule tl_zero,simp+) - apply force -apply(rule conjI,simp add:lift_def) -apply(subgoal_tac "lift Q (Some P2, ta) =(Some (Seq P2 Q), ta)") - apply(simp add:Cons_lift del:map.simps) - apply(rule nth_tl_if) - apply force - apply simp+ -apply(simp add:lift_def) -done -(* -lemma last_lift_not_None3: "fst (last (map (lift Q) (x#xs))) \ None" -apply(simp only:last_length [THEN sym]) -apply(subgoal_tac "length xs None" -apply(simp only:last_length [THEN sym]) -apply(subgoal_tac "length xs\ P sat [pre, rely, guar, mid]; \ Q sat [mid, rely, guar, post]\ - \ \ Seq P Q sat [pre, rely, guar, post]" -apply(unfold com_validity_def) -apply clarify -apply(case_tac "\i[]") - apply(drule last_conv_nth) - apply (simp del:map.simps) - apply(simp only:last_lift_not_None) - apply simp ---{* @{text "\i[]") - apply(drule last_conv_nth,simp) - apply(rule conjI) - apply(erule mp) - apply(case_tac "xs!m") - apply(case_tac "fst(xs!m)",simp) - apply(simp add:lift_def nth_append) - apply clarify - apply(erule_tac x="m+i" in allE) - back - back - apply(case_tac ys,(simp add:nth_append)+) - apply (case_tac i, (simp add:snd_lift)+) - apply(erule mp) - apply(case_tac "xs!m") - apply(force elim:etran.cases intro:Env simp add:lift_def) - apply simp -apply simp -apply clarify -apply(rule conjI,clarify) - apply(case_tac "i[]") - apply(drule last_conv_nth) - apply(simp add: snd_lift nth_append) - apply(rule conjI,clarify) - apply(case_tac ys,simp+) - apply clarify - apply(case_tac ys,simp+) -done - -subsubsection{* Soundness of the While rule *} - -lemma last_append[rule_format]: - "\xs. ys\[] \ ((xs@ys)!(length (xs@ys) - (Suc 0)))=(ys!(length ys - (Suc 0)))" -apply(induct ys) - apply simp -apply clarify -apply (simp add:nth_append length_append) -done - -lemma assum_after_body: - "\ \ P sat [pre \ b, rely, guar, pre]; - (Some P, s) # xs \ cptn_mod; fst (last ((Some P, s) # xs)) = None; s \ b; - (Some (While b P), s) # (Some (Seq P (While b P)), s) # - map (lift (While b P)) xs @ ys \ assum (pre, rely)\ - \ (Some (While b P), snd (last ((Some P, s) # xs))) # ys \ assum (pre, rely)" -apply(simp add:assum_def com_validity_def cp_def cptn_iff_cptn_mod) -apply clarify -apply(erule_tac x=s in allE) -apply(drule_tac c="(Some P, s) # xs" in subsetD,simp) - apply clarify - apply(erule_tac x="Suc i" in allE) - apply simp - apply(simp add:Cons_lift_append nth_append snd_lift del:map.simps) - apply(erule mp) - apply(erule etranE,simp) - apply(case_tac "fst(((Some P, s) # xs) ! i)") - apply(force intro:Env simp add:lift_def) - apply(force intro:Env simp add:lift_def) -apply(rule conjI) - apply clarify - apply(simp add:comm_def last_length) -apply clarify -apply(rule conjI) - apply(simp add:comm_def) -apply clarify -apply(erule_tac x="Suc(length xs + i)" in allE,simp) -apply(case_tac i, simp add:nth_append Cons_lift_append snd_lift del:map.simps) - apply(simp add:last_length) - apply(erule mp) - apply(case_tac "last xs") - apply(simp add:lift_def) -apply(simp add:Cons_lift_append nth_append snd_lift del:map.simps) -done - -lemma While_sound_aux [rule_format]: - "\ pre \ - b \ post; \ P sat [pre \ b, rely, guar, pre]; \s. (s, s) \ guar; - stable pre rely; stable post rely; x \ cptn_mod \ - \ \s xs. x=(Some(While b P),s)#xs \ x\assum(pre, rely) \ x \ comm (guar, post)" -apply(erule cptn_mod.induct) -apply safe -apply (simp_all del:last.simps) ---{* 5 subgoals left *} -apply(simp add:comm_def) ---{* 4 subgoals left *} -apply(rule etran_in_comm) -apply(erule mp) -apply(erule tl_of_assum_in_assum,simp) ---{* While-None *} -apply(ind_cases "((Some (While b P), s), None, t) \ ctran" for s t) -apply(simp add:comm_def) -apply(simp add:cptn_iff_cptn_mod [THEN sym]) -apply(rule conjI,clarify) - apply(force simp add:assum_def) -apply clarify -apply(rule conjI, clarify) - apply(case_tac i,simp,simp) - apply(force simp add:not_ctran_None2) -apply(subgoal_tac "\i. Suc i < length ((None, t) # xs) \ (((None, t) # xs) ! i, ((None, t) # xs) ! Suc i)\ etran") - prefer 2 - apply clarify - apply(rule_tac m="length ((None, s) # xs)" in etran_or_ctran,simp+) - apply(erule not_ctran_None2,simp) - apply simp+ -apply(frule_tac j="0" and k="length ((None, s) # xs) - 1" and p=post in stability,simp+) - apply(force simp add:assum_def subsetD) - apply(simp add:assum_def) - apply clarify - apply(erule_tac x="i" in allE,simp) - apply(erule_tac x="Suc i" in allE,simp) - apply simp -apply clarify -apply (simp add:last_length) ---{* WhileOne *} -apply(thin_tac "P = While b P \ ?Q") -apply(rule ctran_in_comm,simp) -apply(simp add:Cons_lift del:map.simps) -apply(simp add:comm_def del:map.simps) -apply(rule conjI) - apply clarify - apply(case_tac "fst(((Some P, sa) # xs) ! i)") - apply(case_tac "((Some P, sa) # xs) ! i") - apply (simp add:lift_def) - apply(ind_cases "(Some (While b P), ba) -c\ t" for ba t) - apply simp - apply simp - apply(simp add:snd_lift del:map.simps) - apply(simp only:com_validity_def cp_def cptn_iff_cptn_mod) - apply(erule_tac x=sa in allE) - apply(drule_tac c="(Some P, sa) # xs" in subsetD) - apply (simp add:assum_def del:map.simps) - apply clarify - apply(erule_tac x="Suc ia" in allE,simp add:snd_lift del:map.simps) - apply(erule mp) - apply(case_tac "fst(((Some P, sa) # xs) ! ia)") - apply(erule etranE,simp add:lift_def) - apply(rule Env) - apply(erule etranE,simp add:lift_def) - apply(rule Env) - apply (simp add:comm_def del:map.simps) - apply clarify - apply(erule allE,erule impE,assumption) - apply(erule mp) - apply(case_tac "((Some P, sa) # xs) ! i") - apply(case_tac "xs!i") - apply(simp add:lift_def) - apply(case_tac "fst(xs!i)") - apply force - apply force ---{* last=None *} -apply clarify -apply(subgoal_tac "(map (lift (While b P)) ((Some P, sa) # xs))\[]") - apply(drule last_conv_nth) - apply (simp del:map.simps) - apply(simp only:last_lift_not_None) -apply simp ---{* WhileMore *} -apply(thin_tac "P = While b P \ ?Q") -apply(rule ctran_in_comm,simp del:last.simps) ---{* metiendo la hipotesis antes de dividir la conclusion. *} -apply(subgoal_tac "(Some (While b P), snd (last ((Some P, sa) # xs))) # ys \ assum (pre, rely)") - apply (simp del:last.simps) - prefer 2 - apply(erule assum_after_body) - apply (simp del:last.simps)+ ---{* lo de antes. *} -apply(simp add:comm_def del:map.simps last.simps) -apply(rule conjI) - apply clarify - apply(simp only:Cons_lift_append) - apply(case_tac "i t" for ba t) - apply simp - apply simp - apply(simp add:snd_lift del:map.simps last.simps) - apply(thin_tac " \i. i < length ys \ ?P i") - apply(simp only:com_validity_def cp_def cptn_iff_cptn_mod) - apply(erule_tac x=sa in allE) - apply(drule_tac c="(Some P, sa) # xs" in subsetD) - apply (simp add:assum_def del:map.simps last.simps) - apply clarify - apply(erule_tac x="Suc ia" in allE,simp add:nth_append snd_lift del:map.simps last.simps, erule mp) - apply(case_tac "fst(((Some P, sa) # xs) ! ia)") - apply(erule etranE,simp add:lift_def) - apply(rule Env) - apply(erule etranE,simp add:lift_def) - apply(rule Env) - apply (simp add:comm_def del:map.simps) - apply clarify - apply(erule allE,erule impE,assumption) - apply(erule mp) - apply(case_tac "((Some P, sa) # xs) ! i") - apply(case_tac "xs!i") - apply(simp add:lift_def) - apply(case_tac "fst(xs!i)") - apply force - apply force ---{* @{text "i \ length xs"} *} -apply(subgoal_tac "i-length xs length xs"} *} -apply(case_tac "i-length xs") - apply arith -apply(simp add:nth_append del:map.simps last.simps) -apply(rotate_tac -3) -apply(subgoal_tac "i- Suc (length xs)=nat") - prefer 2 - apply arith -apply simp ---{* last=None *} -apply clarify -apply(case_tac ys) - apply(simp add:Cons_lift del:map.simps last.simps) - apply(subgoal_tac "(map (lift (While b P)) ((Some P, sa) # xs))\[]") - apply(drule last_conv_nth) - apply (simp del:map.simps) - apply(simp only:last_lift_not_None) - apply simp -apply(subgoal_tac "((Some (Seq P (While b P)), sa) # map (lift (While b P)) xs @ ys)\[]") - apply(drule last_conv_nth) - apply (simp del:map.simps last.simps) - apply(simp add:nth_append del:last.simps) - apply(subgoal_tac "((Some (While b P), snd (last ((Some P, sa) # xs))) # a # list)\[]") - apply(drule last_conv_nth) - apply (simp del:map.simps last.simps) - apply simp -apply simp -done - -lemma While_sound: - "\stable pre rely; pre \ - b \ post; stable post rely; - \ P sat [pre \ b, rely, guar, pre]; \s. (s,s)\guar\ - \ \ While b P sat [pre, rely, guar, post]" -apply(unfold com_validity_def) -apply clarify -apply(erule_tac xs="tl x" in While_sound_aux) - apply(simp add:com_validity_def) - apply force - apply simp_all -apply(simp add:cptn_iff_cptn_mod cp_def) -apply(simp add:cp_def) -apply clarify -apply(rule nth_equalityI) - apply simp_all - apply(case_tac x,simp+) -apply clarify -apply(case_tac i,simp+) -apply(case_tac x,simp+) -done - -subsubsection{* Soundness of the Rule of Consequence *} - -lemma Conseq_sound: - "\pre \ pre'; rely \ rely'; guar' \ guar; post' \ post; - \ P sat [pre', rely', guar', post']\ - \ \ P sat [pre, rely, guar, post]" -apply(simp add:com_validity_def assum_def comm_def) -apply clarify -apply(erule_tac x=s in allE) -apply(drule_tac c=x in subsetD) - apply force -apply force -done - -subsubsection {* Soundness of the system for sequential component programs *} - -theorem rgsound: - "\ P sat [pre, rely, guar, post] \ \ P sat [pre, rely, guar, post]" -apply(erule rghoare.induct) - apply(force elim:Basic_sound) - apply(force elim:Seq_sound) - apply(force elim:Cond_sound) - apply(force elim:While_sound) - apply(force elim:Await_sound) -apply(erule Conseq_sound,simp+) -done - -subsection {* Soundness of the System for Parallel Programs *} - -constdefs - ParallelCom :: "('a rgformula) list \ 'a par_com" - "ParallelCom Ps \ map (Some \ fst) Ps" - -lemma two: - "\ \i (\j\{j. j < length xs \ j \ i}. Guar (xs ! j)) - \ Rely (xs ! i); - pre \ (\i\{i. i < length xs}. Pre (xs ! i)); - \i Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)]; - length xs=length clist; x \ par_cp (ParallelCom xs) s; x\par_assum(pre, rely); - \icp (Some(Com(xs!i))) s; x \ clist \ - \ \j i. i Suc j (clist!i!j) -c\ (clist!i!Suc j) - \ (snd(clist!i!j), snd(clist!i!Suc j)) \ Guar(xs!i)" -apply(unfold par_cp_def) -apply (rule ccontr) ---{* By contradiction: *} -apply (simp del: Un_subset_iff) -apply(erule exE) ---{* the first c-tran that does not satisfy the guarantee-condition is from @{text "\_i"} at step @{text "m"}. *} -apply(drule_tac n=j and P="\j. \i. ?H i j" in Ex_first_occurrence) -apply(erule exE) -apply clarify ---{* @{text "\_i \ A(pre, rely_1)"} *} -apply(subgoal_tac "take (Suc (Suc m)) (clist!i) \ assum(Pre(xs!i), Rely(xs!i))") ---{* but this contradicts @{text "\ \_i sat [pre_i,rely_i,guar_i,post_i]"} *} - apply(erule_tac x=i and P="\i. ?H i \ \ (?J i) sat [?I i,?K i,?M i,?N i]" in allE,erule impE,assumption) - apply(simp add:com_validity_def) - apply(erule_tac x=s in allE) - apply(simp add:cp_def comm_def) - apply(drule_tac c="take (Suc (Suc m)) (clist ! i)" in subsetD) - apply simp - apply (blast intro: takecptn_is_cptn) - apply simp - apply clarify - apply(erule_tac x=m and P="\j. ?I j \ ?J j \ ?H j" in allE) - apply (simp add:conjoin_def same_length_def) -apply(simp add:assum_def del: Un_subset_iff) -apply(rule conjI) - apply(erule_tac x=i and P="\j. ?H j \ ?I j \cp (?K j) (?J j)" in allE) - apply(simp add:cp_def par_assum_def) - apply(drule_tac c="s" in subsetD,simp) - apply simp -apply clarify -apply(erule_tac x=i and P="\j. ?H j \ ?M \ UNION (?S j) (?T j) \ (?L j)" in allE) -apply(simp del: Un_subset_iff) -apply(erule subsetD) -apply simp -apply(simp add:conjoin_def compat_label_def) -apply clarify -apply(erule_tac x=ia and P="\j. ?H j \ (?P j) \ ?Q j" in allE,simp) ---{* each etran in @{text "\_1[0\m]"} corresponds to *} -apply(erule disjE) ---{* a c-tran in some @{text "\_{ib}"} *} - apply clarify - apply(case_tac "i=ib",simp) - apply(erule etranE,simp) - apply(erule_tac x="ib" and P="\i. ?H i \ (?I i) \ (?J i)" in allE) - apply (erule etranE) - apply(case_tac "ia=m",simp) - apply simp - apply(erule_tac x=ia and P="\j. ?H j \ (\ i. ?P i j)" in allE) - apply(subgoal_tac "ia"}, -therefore it satisfies @{text "rely \ guar_{ib}"} *} -apply (force simp add:par_assum_def same_state_def) -done - - -lemma three [rule_format]: - "\ xs\[]; \i (\j\{j. j < length xs \ j \ i}. Guar (xs ! j)) - \ Rely (xs ! i); - pre \ (\i\{i. i < length xs}. Pre (xs ! i)); - \i Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)]; - length xs=length clist; x \ par_cp (ParallelCom xs) s; x \ par_assum(pre, rely); - \icp (Some(Com(xs!i))) s; x \ clist \ - \ \j i. i Suc j (clist!i!j) -e\ (clist!i!Suc j) - \ (snd(clist!i!j), snd(clist!i!Suc j)) \ rely \ (\j\{j. j < length xs \ j \ i}. Guar (xs ! j))" -apply(drule two) - apply simp_all -apply clarify -apply(simp add:conjoin_def compat_label_def) -apply clarify -apply(erule_tac x=j and P="\j. ?H j \ (?J j \ (\i. ?P i j)) \ ?I j" in allE,simp) -apply(erule disjE) - prefer 2 - apply(force simp add:same_state_def par_assum_def) -apply clarify -apply(case_tac "i=ia",simp) - apply(erule etranE,simp) -apply(erule_tac x="ia" and P="\i. ?H i \ (?I i) \ (?J i)" in allE,simp) -apply(erule_tac x=j and P="\j. \i. ?S j i \ (?I j i, ?H j i)\ ctran \ (?P i j)" in allE) -apply(erule_tac x=ia and P="\j. ?S j \ (?I j, ?H j)\ ctran \ (?P j)" in allE) -apply(simp add:same_state_def) -apply(erule_tac x=i and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in all_dupE) -apply(erule_tac x=ia and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in allE,simp) -done - -lemma four: - "\xs\[]; \i < length xs. rely \ (\j\{j. j < length xs \ j \ i}. Guar (xs ! j)) - \ Rely (xs ! i); - (\j\{j. j < length xs}. Guar (xs ! j)) \ guar; - pre \ (\i\{i. i < length xs}. Pre (xs ! i)); - \i < length xs. - \ Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)]; - x \ par_cp (ParallelCom xs) s; x \ par_assum (pre, rely); Suc i < length x; - x ! i -pc\ x ! Suc i\ - \ (snd (x ! i), snd (x ! Suc i)) \ guar" -apply(simp add: ParallelCom_def del: Un_subset_iff) -apply(subgoal_tac "(map (Some \ fst) xs)\[]") - prefer 2 - apply simp -apply(frule rev_subsetD) - apply(erule one [THEN equalityD1]) -apply(erule subsetD) -apply (simp del: Un_subset_iff) -apply clarify -apply(drule_tac pre=pre and rely=rely and x=x and s=s and xs=xs and clist=clist in two) -apply(assumption+) - apply(erule sym) - apply(simp add:ParallelCom_def) - apply assumption - apply(simp add:Com_def) - apply assumption -apply(simp add:conjoin_def same_program_def) -apply clarify -apply(erule_tac x=i and P="\j. ?H j \ fst(?I j)=(?J j)" in all_dupE) -apply(erule_tac x="Suc i" and P="\j. ?H j \ fst(?I j)=(?J j)" in allE) -apply(erule par_ctranE,simp) -apply(erule_tac x=i and P="\j. \i. ?S j i \ (?I j i, ?H j i)\ ctran \ (?P i j)" in allE) -apply(erule_tac x=ia and P="\j. ?S j \ (?I j, ?H j)\ ctran \ (?P j)" in allE) -apply(rule_tac x=ia in exI) -apply(simp add:same_state_def) -apply(erule_tac x=ia and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in all_dupE,simp) -apply(erule_tac x=ia and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in allE,simp) -apply(erule_tac x=i and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in all_dupE) -apply(erule_tac x=i and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in all_dupE,simp) -apply(erule_tac x="Suc i" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) -apply(erule mp) -apply(subgoal_tac "r=fst(clist ! ia ! Suc i)",simp) -apply(drule_tac i=ia in list_eq_if) -back -apply simp_all -done - -lemma parcptn_not_empty [simp]:"[] \ par_cptn" -apply(force elim:par_cptn.cases) -done - -lemma five: - "\xs\[]; \i (\j\{j. j < length xs \ j \ i}. Guar (xs ! j)) - \ Rely (xs ! i); - pre \ (\i\{i. i < length xs}. Pre (xs ! i)); - (\i\{i. i < length xs}. Post (xs ! i)) \ post; - \i < length xs. - \ Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)]; - x \ par_cp (ParallelCom xs) s; x \ par_assum (pre, rely); - All_None (fst (last x)) \ \ snd (last x) \ post" -apply(simp add: ParallelCom_def del: Un_subset_iff) -apply(subgoal_tac "(map (Some \ fst) xs)\[]") - prefer 2 - apply simp -apply(frule rev_subsetD) - apply(erule one [THEN equalityD1]) -apply(erule subsetD) -apply(simp del: Un_subset_iff) -apply clarify -apply(subgoal_tac "\iassum(Pre(xs!i), Rely(xs!i))") - apply(erule_tac x=i and P="\i. ?H i \ \ (?J i) sat [?I i,?K i,?M i,?N i]" in allE,erule impE,assumption) - apply(simp add:com_validity_def) - apply(erule_tac x=s in allE) - apply(erule_tac x=i and P="\j. ?H j \ (?I j) \ cp (?J j) s" in allE,simp) - apply(drule_tac c="clist!i" in subsetD) - apply (force simp add:Com_def) - apply(simp add:comm_def conjoin_def same_program_def del:last.simps) - apply clarify - apply(erule_tac x="length x - 1" and P="\j. ?H j \ fst(?I j)=(?J j)" in allE) - apply (simp add:All_None_def same_length_def) - apply(erule_tac x=i and P="\j. ?H j \ length(?J j)=(?K j)" in allE) - apply(subgoal_tac "length x - 1 < length x",simp) - apply(case_tac "x\[]") - apply(simp add: last_conv_nth) - apply(erule_tac x="clist!i" in ballE) - apply(simp add:same_state_def) - apply(subgoal_tac "clist!i\[]") - apply(simp add: last_conv_nth) - apply(case_tac x) - apply (force simp add:par_cp_def) - apply (force simp add:par_cp_def) - apply force - apply (force simp add:par_cp_def) - apply(case_tac x) - apply (force simp add:par_cp_def) - apply (force simp add:par_cp_def) -apply clarify -apply(simp add:assum_def) -apply(rule conjI) - apply(simp add:conjoin_def same_state_def par_cp_def) - apply clarify - apply(erule_tac x=ia and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in allE,simp) - apply(erule_tac x=0 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) - apply(case_tac x,simp+) - apply (simp add:par_assum_def) - apply clarify - apply(drule_tac c="snd (clist ! ia ! 0)" in subsetD) - apply assumption - apply simp -apply clarify -apply(erule_tac x=ia in all_dupE) -apply(rule subsetD, erule mp, assumption) -apply(erule_tac pre=pre and rely=rely and x=x and s=s in three) - apply(erule_tac x=ic in allE,erule mp) - apply simp_all - apply(simp add:ParallelCom_def) - apply(force simp add:Com_def) -apply(simp add:conjoin_def same_length_def) -done - -lemma ParallelEmpty [rule_format]: - "\i s. x \ par_cp (ParallelCom []) s \ - Suc i < length x \ (x ! i, x ! Suc i) \ par_ctran" -apply(induct_tac x) - apply(simp add:par_cp_def ParallelCom_def) -apply clarify -apply(case_tac list,simp,simp) -apply(case_tac i) - apply(simp add:par_cp_def ParallelCom_def) - apply(erule par_ctranE,simp) -apply(simp add:par_cp_def ParallelCom_def) -apply clarify -apply(erule par_cptn.cases,simp) - apply simp -apply(erule par_ctranE) -back -apply simp -done - -theorem par_rgsound: - "\ c SAT [pre, rely, guar, post] \ - \ (ParallelCom c) SAT [pre, rely, guar, post]" -apply(erule par_rghoare.induct) -apply(case_tac xs,simp) - apply(simp add:par_com_validity_def par_comm_def) - apply clarify - apply(case_tac "post=UNIV",simp) - apply clarify - apply(drule ParallelEmpty) - apply assumption - apply simp - apply clarify - apply simp -apply(subgoal_tac "xs\[]") - prefer 2 - apply simp -apply(thin_tac "xs = a # list") -apply(simp add:par_com_validity_def par_comm_def) -apply clarify -apply(rule conjI) - apply clarify - apply(erule_tac pre=pre and rely=rely and guar=guar and x=x and s=s and xs=xs in four) - apply(assumption+) - apply clarify - apply (erule allE, erule impE, assumption,erule rgsound) - apply(assumption+) -apply clarify -apply(erule_tac pre=pre and rely=rely and post=post and x=x and s=s and xs=xs in five) - apply(assumption+) - apply clarify - apply (erule allE, erule impE, assumption,erule rgsound) - apply(assumption+) -done - -end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/RG_Syntax.thy --- a/src/HOL/HoareParallel/RG_Syntax.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,95 +0,0 @@ -header {* \section{Concrete Syntax} *} - -theory RG_Syntax -imports RG_Hoare Quote_Antiquote -begin - -syntax - "_Assign" :: "idt \ 'b \ 'a com" ("(\_ :=/ _)" [70, 65] 61) - "_skip" :: "'a com" ("SKIP") - "_Seq" :: "'a com \ 'a com \ 'a com" ("(_;;/ _)" [60,61] 60) - "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" ("(0IF _/ THEN _/ ELSE _/FI)" [0, 0, 0] 61) - "_Cond2" :: "'a bexp \ 'a com \ 'a com" ("(0IF _ THEN _ FI)" [0,0] 56) - "_While" :: "'a bexp \ 'a com \ 'a com" ("(0WHILE _ /DO _ /OD)" [0, 0] 61) - "_Await" :: "'a bexp \ 'a com \ 'a com" ("(0AWAIT _ /THEN /_ /END)" [0,0] 61) - "_Atom" :: "'a com \ 'a com" ("(\_\)" 61) - "_Wait" :: "'a bexp \ 'a com" ("(0WAIT _ END)" 61) - -translations - "\\x := a" \ "Basic \\\(_update_name x (\_. a))\" - "SKIP" \ "Basic id" - "c1;; c2" \ "Seq c1 c2" - "IF b THEN c1 ELSE c2 FI" \ "Cond .{b}. c1 c2" - "IF b THEN c FI" \ "IF b THEN c ELSE SKIP FI" - "WHILE b DO c OD" \ "While .{b}. c" - "AWAIT b THEN c END" \ "Await .{b}. c" - "\c\" \ "AWAIT True THEN c END" - "WAIT b END" \ "AWAIT b THEN SKIP END" - -nonterminals - prgs - -syntax - "_PAR" :: "prgs \ 'a" ("COBEGIN//_//COEND" 60) - "_prg" :: "'a \ prgs" ("_" 57) - "_prgs" :: "['a, prgs] \ prgs" ("_//\//_" [60,57] 57) - -translations - "_prg a" \ "[a]" - "_prgs a ps" \ "a # ps" - "_PAR ps" \ "ps" - -syntax - "_prg_scheme" :: "['a, 'a, 'a, 'a] \ prgs" ("SCHEME [_ \ _ < _] _" [0,0,0,60] 57) - -translations - "_prg_scheme j i k c" \ "(map (\i. c) [j.. 'a" ("\_") - "_after" :: "id \ 'a" ("\_") - -translations - "\x" \ "x \fst" - "\x" \ "x \snd" - -print_translation {* - let - fun quote_tr' f (t :: ts) = - Term.list_comb (f $ Syntax.quote_tr' "_antiquote" t, ts) - | quote_tr' _ _ = raise Match; - - val assert_tr' = quote_tr' (Syntax.const "_Assert"); - - fun bexp_tr' name ((Const ("Collect", _) $ t) :: ts) = - quote_tr' (Syntax.const name) (t :: ts) - | bexp_tr' _ _ = raise Match; - - fun upd_tr' (x_upd, T) = - (case try (unsuffix Record.updateN) x_upd of - SOME x => (x, if T = dummyT then T else Term.domain_type T) - | NONE => raise Match); - - fun update_name_tr' (Free x) = Free (upd_tr' x) - | update_name_tr' ((c as Const ("_free", _)) $ Free x) = - c $ Free (upd_tr' x) - | update_name_tr' (Const x) = Const (upd_tr' x) - | update_name_tr' _ = raise Match; - - fun K_tr' (Abs (_,_,t)) = if null (loose_bnos t) then t else raise Match - | K_tr' (Abs (_,_,Abs (_,_,t)$Bound 0)) = if null (loose_bnos t) then t else raise Match - | K_tr' _ = raise Match; - - fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) = - quote_tr' (Syntax.const "_Assign" $ update_name_tr' f) - (Abs (x, dummyT, K_tr' k) :: ts) - | assign_tr' _ = raise Match; - in - [("Collect", assert_tr'), ("Basic", assign_tr'), - ("Cond", bexp_tr' "_Cond"), ("While", bexp_tr' "_While_inv")] - end -*} - -end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/RG_Tran.thy --- a/src/HOL/HoareParallel/RG_Tran.thy Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,1075 +0,0 @@ -header {* \section{Operational Semantics} *} - -theory RG_Tran -imports RG_Com -begin - -subsection {* Semantics of Component Programs *} - -subsubsection {* Environment transitions *} - -types 'a conf = "(('a com) option) \ 'a" - -inductive_set - etran :: "('a conf \ 'a conf) set" - and etran' :: "'a conf \ 'a conf \ bool" ("_ -e\ _" [81,81] 80) -where - "P -e\ Q \ (P,Q) \ etran" -| Env: "(P, s) -e\ (P, t)" - -lemma etranE: "c -e\ c' \ (\P s t. c = (P, s) \ c' = (P, t) \ Q) \ Q" - by (induct c, induct c', erule etran.cases, blast) - -subsubsection {* Component transitions *} - -inductive_set - ctran :: "('a conf \ 'a conf) set" - and ctran' :: "'a conf \ 'a conf \ bool" ("_ -c\ _" [81,81] 80) - and ctrans :: "'a conf \ 'a conf \ bool" ("_ -c*\ _" [81,81] 80) -where - "P -c\ Q \ (P,Q) \ ctran" -| "P -c*\ Q \ (P,Q) \ ctran^*" - -| Basic: "(Some(Basic f), s) -c\ (None, f s)" - -| Seq1: "(Some P0, s) -c\ (None, t) \ (Some(Seq P0 P1), s) -c\ (Some P1, t)" - -| Seq2: "(Some P0, s) -c\ (Some P2, t) \ (Some(Seq P0 P1), s) -c\ (Some(Seq P2 P1), t)" - -| CondT: "s\b \ (Some(Cond b P1 P2), s) -c\ (Some P1, s)" -| CondF: "s\b \ (Some(Cond b P1 P2), s) -c\ (Some P2, s)" - -| WhileF: "s\b \ (Some(While b P), s) -c\ (None, s)" -| WhileT: "s\b \ (Some(While b P), s) -c\ (Some(Seq P (While b P)), s)" - -| Await: "\s\b; (Some P, s) -c*\ (None, t)\ \ (Some(Await b P), s) -c\ (None, t)" - -monos "rtrancl_mono" - -subsection {* Semantics of Parallel Programs *} - -types 'a par_conf = "('a par_com) \ 'a" - -inductive_set - par_etran :: "('a par_conf \ 'a par_conf) set" - and par_etran' :: "['a par_conf,'a par_conf] \ bool" ("_ -pe\ _" [81,81] 80) -where - "P -pe\ Q \ (P,Q) \ par_etran" -| ParEnv: "(Ps, s) -pe\ (Ps, t)" - -inductive_set - par_ctran :: "('a par_conf \ 'a par_conf) set" - and par_ctran' :: "['a par_conf,'a par_conf] \ bool" ("_ -pc\ _" [81,81] 80) -where - "P -pc\ Q \ (P,Q) \ par_ctran" -| ParComp: "\i (r, t)\ \ (Ps, s) -pc\ (Ps[i:=r], t)" - -lemma par_ctranE: "c -pc\ c' \ - (\i Ps s r t. c = (Ps, s) \ c' = (Ps[i := r], t) \ i < length Ps \ - (Ps ! i, s) -c\ (r, t) \ P) \ P" - by (induct c, induct c', erule par_ctran.cases, blast) - -subsection {* Computations *} - -subsubsection {* Sequential computations *} - -types 'a confs = "('a conf) list" - -inductive_set cptn :: "('a confs) set" -where - CptnOne: "[(P,s)] \ cptn" -| CptnEnv: "(P, t)#xs \ cptn \ (P,s)#(P,t)#xs \ cptn" -| CptnComp: "\(P,s) -c\ (Q,t); (Q, t)#xs \ cptn \ \ (P,s)#(Q,t)#xs \ cptn" - -constdefs - cp :: "('a com) option \ 'a \ ('a confs) set" - "cp P s \ {l. l!0=(P,s) \ l \ cptn}" - -subsubsection {* Parallel computations *} - -types 'a par_confs = "('a par_conf) list" - -inductive_set par_cptn :: "('a par_confs) set" -where - ParCptnOne: "[(P,s)] \ par_cptn" -| ParCptnEnv: "(P,t)#xs \ par_cptn \ (P,s)#(P,t)#xs \ par_cptn" -| ParCptnComp: "\ (P,s) -pc\ (Q,t); (Q,t)#xs \ par_cptn \ \ (P,s)#(Q,t)#xs \ par_cptn" - -constdefs - par_cp :: "'a par_com \ 'a \ ('a par_confs) set" - "par_cp P s \ {l. l!0=(P,s) \ l \ par_cptn}" - -subsection{* Modular Definition of Computation *} - -constdefs - lift :: "'a com \ 'a conf \ 'a conf" - "lift Q \ \(P, s). (if P=None then (Some Q,s) else (Some(Seq (the P) Q), s))" - -inductive_set cptn_mod :: "('a confs) set" -where - CptnModOne: "[(P, s)] \ cptn_mod" -| CptnModEnv: "(P, t)#xs \ cptn_mod \ (P, s)#(P, t)#xs \ cptn_mod" -| CptnModNone: "\(Some P, s) -c\ (None, t); (None, t)#xs \ cptn_mod \ \ (Some P,s)#(None, t)#xs \cptn_mod" -| CptnModCondT: "\(Some P0, s)#ys \ cptn_mod; s \ b \ \ (Some(Cond b P0 P1), s)#(Some P0, s)#ys \ cptn_mod" -| CptnModCondF: "\(Some P1, s)#ys \ cptn_mod; s \ b \ \ (Some(Cond b P0 P1), s)#(Some P1, s)#ys \ cptn_mod" -| CptnModSeq1: "\(Some P0, s)#xs \ cptn_mod; zs=map (lift P1) xs \ - \ (Some(Seq P0 P1), s)#zs \ cptn_mod" -| CptnModSeq2: - "\(Some P0, s)#xs \ cptn_mod; fst(last ((Some P0, s)#xs)) = None; - (Some P1, snd(last ((Some P0, s)#xs)))#ys \ cptn_mod; - zs=(map (lift P1) xs)@ys \ \ (Some(Seq P0 P1), s)#zs \ cptn_mod" - -| CptnModWhile1: - "\ (Some P, s)#xs \ cptn_mod; s \ b; zs=map (lift (While b P)) xs \ - \ (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \ cptn_mod" -| CptnModWhile2: - "\ (Some P, s)#xs \ cptn_mod; fst(last ((Some P, s)#xs))=None; s \ b; - zs=(map (lift (While b P)) xs)@ys; - (Some(While b P), snd(last ((Some P, s)#xs)))#ys \ cptn_mod\ - \ (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \ cptn_mod" - -subsection {* Equivalence of Both Definitions.*} - -lemma last_length: "((a#xs)!(length xs))=last (a#xs)" -apply simp -apply(induct xs,simp+) -apply(case_tac xs) -apply simp_all -done - -lemma div_seq [rule_format]: "list \ cptn_mod \ - (\s P Q zs. list=(Some (Seq P Q), s)#zs \ - (\xs. (Some P, s)#xs \ cptn_mod \ (zs=(map (lift Q) xs) \ - ( fst(((Some P, s)#xs)!length xs)=None \ - (\ys. (Some Q, snd(((Some P, s)#xs)!length xs))#ys \ cptn_mod - \ zs=(map (lift (Q)) xs)@ys)))))" -apply(erule cptn_mod.induct) -apply simp_all - apply clarify - apply(force intro:CptnModOne) - apply clarify - apply(erule_tac x=Pa in allE) - apply(erule_tac x=Q in allE) - apply simp - apply clarify - apply(erule disjE) - apply(rule_tac x="(Some Pa,t)#xsa" in exI) - apply(rule conjI) - apply clarify - apply(erule CptnModEnv) - apply(rule disjI1) - apply(simp add:lift_def) - apply clarify - apply(rule_tac x="(Some Pa,t)#xsa" in exI) - apply(rule conjI) - apply(erule CptnModEnv) - apply(rule disjI2) - apply(rule conjI) - apply(case_tac xsa,simp,simp) - apply(rule_tac x="ys" in exI) - apply(rule conjI) - apply simp - apply(simp add:lift_def) - apply clarify - apply(erule ctran.cases,simp_all) - apply clarify - apply(rule_tac x="xs" in exI) - apply simp - apply clarify -apply(rule_tac x="xs" in exI) -apply(simp add: last_length) -done - -lemma cptn_onlyif_cptn_mod_aux [rule_format]: - "\s Q t xs.((Some a, s), Q, t) \ ctran \ (Q, t) # xs \ cptn_mod - \ (Some a, s) # (Q, t) # xs \ cptn_mod" -apply(induct a) -apply simp_all ---{* basic *} -apply clarify -apply(erule ctran.cases,simp_all) -apply(rule CptnModNone,rule Basic,simp) -apply clarify -apply(erule ctran.cases,simp_all) ---{* Seq1 *} -apply(rule_tac xs="[(None,ta)]" in CptnModSeq2) - apply(erule CptnModNone) - apply(rule CptnModOne) - apply simp -apply simp -apply(simp add:lift_def) ---{* Seq2 *} -apply(erule_tac x=sa in allE) -apply(erule_tac x="Some P2" in allE) -apply(erule allE,erule impE, assumption) -apply(drule div_seq,simp) -apply force -apply clarify -apply(erule disjE) - apply clarify - apply(erule allE,erule impE, assumption) - apply(erule_tac CptnModSeq1) - apply(simp add:lift_def) -apply clarify -apply(erule allE,erule impE, assumption) -apply(erule_tac CptnModSeq2) - apply (simp add:last_length) - apply (simp add:last_length) -apply(simp add:lift_def) ---{* Cond *} -apply clarify -apply(erule ctran.cases,simp_all) -apply(force elim: CptnModCondT) -apply(force elim: CptnModCondF) ---{* While *} -apply clarify -apply(erule ctran.cases,simp_all) -apply(rule CptnModNone,erule WhileF,simp) -apply(drule div_seq,force) -apply clarify -apply (erule disjE) - apply(force elim:CptnModWhile1) -apply clarify -apply(force simp add:last_length elim:CptnModWhile2) ---{* await *} -apply clarify -apply(erule ctran.cases,simp_all) -apply(rule CptnModNone,erule Await,simp+) -done - -lemma cptn_onlyif_cptn_mod [rule_format]: "c \ cptn \ c \ cptn_mod" -apply(erule cptn.induct) - apply(rule CptnModOne) - apply(erule CptnModEnv) -apply(case_tac P) - apply simp - apply(erule ctran.cases,simp_all) -apply(force elim:cptn_onlyif_cptn_mod_aux) -done - -lemma lift_is_cptn: "c\cptn \ map (lift P) c \ cptn" -apply(erule cptn.induct) - apply(force simp add:lift_def CptnOne) - apply(force intro:CptnEnv simp add:lift_def) -apply(force simp add:lift_def intro:CptnComp Seq2 Seq1 elim:ctran.cases) -done - -lemma cptn_append_is_cptn [rule_format]: - "\b a. b#c1\cptn \ a#c2\cptn \ (b#c1)!length c1=a \ b#c1@c2\cptn" -apply(induct c1) - apply simp -apply clarify -apply(erule cptn.cases,simp_all) - apply(force intro:CptnEnv) -apply(force elim:CptnComp) -done - -lemma last_lift: "\xs\[]; fst(xs!(length xs - (Suc 0)))=None\ - \ fst((map (lift P) xs)!(length (map (lift P) xs)- (Suc 0)))=(Some P)" -apply(case_tac "(xs ! (length xs - (Suc 0)))") -apply (simp add:lift_def) -done - -lemma last_fst [rule_format]: "P((a#x)!length x) \ \P a \ P (x!(length x - (Suc 0)))" -apply(induct x,simp+) -done - -lemma last_fst_esp: - "fst(((Some a,s)#xs)!(length xs))=None \ fst(xs!(length xs - (Suc 0)))=None" -apply(erule last_fst) -apply simp -done - -lemma last_snd: "xs\[] \ - snd(((map (lift P) xs))!(length (map (lift P) xs) - (Suc 0)))=snd(xs!(length xs - (Suc 0)))" -apply(case_tac "(xs ! (length xs - (Suc 0)))",simp) -apply (simp add:lift_def) -done - -lemma Cons_lift: "(Some (Seq P Q), s) # (map (lift Q) xs) = map (lift Q) ((Some P, s) # xs)" -by(simp add:lift_def) - -lemma Cons_lift_append: - "(Some (Seq P Q), s) # (map (lift Q) xs) @ ys = map (lift Q) ((Some P, s) # xs)@ ys " -by(simp add:lift_def) - -lemma lift_nth: "i map (lift Q) xs ! i = lift Q (xs! i)" -by (simp add:lift_def) - -lemma snd_lift: "i< length xs \ snd(lift Q (xs ! i))= snd (xs ! i)" -apply(case_tac "xs!i") -apply(simp add:lift_def) -done - -lemma cptn_if_cptn_mod: "c \ cptn_mod \ c \ cptn" -apply(erule cptn_mod.induct) - apply(rule CptnOne) - apply(erule CptnEnv) - apply(erule CptnComp,simp) - apply(rule CptnComp) - apply(erule CondT,simp) - apply(rule CptnComp) - apply(erule CondF,simp) ---{* Seq1 *} -apply(erule cptn.cases,simp_all) - apply(rule CptnOne) - apply clarify - apply(drule_tac P=P1 in lift_is_cptn) - apply(simp add:lift_def) - apply(rule CptnEnv,simp) -apply clarify -apply(simp add:lift_def) -apply(rule conjI) - apply clarify - apply(rule CptnComp) - apply(rule Seq1,simp) - apply(drule_tac P=P1 in lift_is_cptn) - apply(simp add:lift_def) -apply clarify -apply(rule CptnComp) - apply(rule Seq2,simp) -apply(drule_tac P=P1 in lift_is_cptn) -apply(simp add:lift_def) ---{* Seq2 *} -apply(rule cptn_append_is_cptn) - apply(drule_tac P=P1 in lift_is_cptn) - apply(simp add:lift_def) - apply simp -apply(case_tac "xs\[]") - apply(drule_tac P=P1 in last_lift) - apply(rule last_fst_esp) - apply (simp add:last_length) - apply(simp add:Cons_lift del:map.simps) - apply(rule conjI, clarify, simp) - apply(case_tac "(((Some P0, s) # xs) ! length xs)") - apply clarify - apply (simp add:lift_def last_length) -apply (simp add:last_length) ---{* While1 *} -apply(rule CptnComp) -apply(rule WhileT,simp) -apply(drule_tac P="While b P" in lift_is_cptn) -apply(simp add:lift_def) ---{* While2 *} -apply(rule CptnComp) -apply(rule WhileT,simp) -apply(rule cptn_append_is_cptn) -apply(drule_tac P="While b P" in lift_is_cptn) - apply(simp add:lift_def) - apply simp -apply(case_tac "xs\[]") - apply(drule_tac P="While b P" in last_lift) - apply(rule last_fst_esp,simp add:last_length) - apply(simp add:Cons_lift del:map.simps) - apply(rule conjI, clarify, simp) - apply(case_tac "(((Some P, s) # xs) ! length xs)") - apply clarify - apply (simp add:last_length lift_def) -apply simp -done - -theorem cptn_iff_cptn_mod: "(c \ cptn) = (c \ cptn_mod)" -apply(rule iffI) - apply(erule cptn_onlyif_cptn_mod) -apply(erule cptn_if_cptn_mod) -done - -section {* Validity of Correctness Formulas*} - -subsection {* Validity for Component Programs. *} - -types 'a rgformula = "'a com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set" - -constdefs - assum :: "('a set \ ('a \ 'a) set) \ ('a confs) set" - "assum \ \(pre, rely). {c. snd(c!0) \ pre \ (\i. Suc i - c!i -e\ c!(Suc i) \ (snd(c!i), snd(c!Suc i)) \ rely)}" - - comm :: "(('a \ 'a) set \ 'a set) \ ('a confs) set" - "comm \ \(guar, post). {c. (\i. Suc i - c!i -c\ c!(Suc i) \ (snd(c!i), snd(c!Suc i)) \ guar) \ - (fst (last c) = None \ snd (last c) \ post)}" - - com_validity :: "'a com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set \ bool" - ("\ _ sat [_, _, _, _]" [60,0,0,0,0] 45) - "\ P sat [pre, rely, guar, post] \ - \s. cp (Some P) s \ assum(pre, rely) \ comm(guar, post)" - -subsection {* Validity for Parallel Programs. *} - -constdefs - All_None :: "(('a com) option) list \ bool" - "All_None xs \ \c\set xs. c=None" - - par_assum :: "('a set \ ('a \ 'a) set) \ ('a par_confs) set" - "par_assum \ \(pre, rely). {c. snd(c!0) \ pre \ (\i. Suc i - c!i -pe\ c!Suc i \ (snd(c!i), snd(c!Suc i)) \ rely)}" - - par_comm :: "(('a \ 'a) set \ 'a set) \ ('a par_confs) set" - "par_comm \ \(guar, post). {c. (\i. Suc i - c!i -pc\ c!Suc i \ (snd(c!i), snd(c!Suc i)) \ guar) \ - (All_None (fst (last c)) \ snd( last c) \ post)}" - - par_com_validity :: "'a par_com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set -\ 'a set \ bool" ("\ _ SAT [_, _, _, _]" [60,0,0,0,0] 45) - "\ Ps SAT [pre, rely, guar, post] \ - \s. par_cp Ps s \ par_assum(pre, rely) \ par_comm(guar, post)" - -subsection {* Compositionality of the Semantics *} - -subsubsection {* Definition of the conjoin operator *} - -constdefs - same_length :: "'a par_confs \ ('a confs) list \ bool" - "same_length c clist \ (\i ('a confs) list \ bool" - "same_state c clist \ (\i j ('a confs) list \ bool" - "same_program c clist \ (\jx. fst(nth x j)) clist)" - - compat_label :: "'a par_confs \ ('a confs) list \ bool" - "compat_label c clist \ (\j. Suc j - (c!j -pc\ c!Suc j \ (\i (clist!i)! Suc j \ - (\li \ (clist!l)!j -e\ (clist!l)! Suc j))) \ - (c!j -pe\ c!Suc j \ (\i (clist!i)! Suc j)))" - - conjoin :: "'a par_confs \ ('a confs) list \ bool" ("_ \ _" [65,65] 64) - "c \ clist \ (same_length c clist) \ (same_state c clist) \ (same_program c clist) \ (compat_label c clist)" - -subsubsection {* Some previous lemmas *} - -lemma list_eq_if [rule_format]: - "\ys. xs=ys \ (length xs = length ys) \ (\i (\i ys!0=a; ys\[] \ \ ys=(a#(tl ys))" -apply(case_tac ys) - apply simp+ -done - -lemma nth_tl_if [rule_format]: "ys\[] \ ys!0=a \ P ys \ P (a#(tl ys))" -apply(induct ys) - apply simp+ -done - -lemma nth_tl_onlyif [rule_format]: "ys\[] \ ys!0=a \ P (a#(tl ys)) \ P ys" -apply(induct ys) - apply simp+ -done - -lemma seq_not_eq1: "Seq c1 c2\c1" -apply(rule com.induct) -apply simp_all -apply clarify -done - -lemma seq_not_eq2: "Seq c1 c2\c2" -apply(rule com.induct) -apply simp_all -apply clarify -done - -lemma if_not_eq1: "Cond b c1 c2 \c1" -apply(rule com.induct) -apply simp_all -apply clarify -done - -lemma if_not_eq2: "Cond b c1 c2\c2" -apply(rule com.induct) -apply simp_all -apply clarify -done - -lemmas seq_and_if_not_eq [simp] = seq_not_eq1 seq_not_eq2 -seq_not_eq1 [THEN not_sym] seq_not_eq2 [THEN not_sym] -if_not_eq1 if_not_eq2 if_not_eq1 [THEN not_sym] if_not_eq2 [THEN not_sym] - -lemma prog_not_eq_in_ctran_aux: - assumes c: "(P,s) -c\ (Q,t)" - shows "P\Q" using c - by (induct x1 \ "(P,s)" x2 \ "(Q,t)" arbitrary: P s Q t) auto - -lemma prog_not_eq_in_ctran [simp]: "\ (P,s) -c\ (P,t)" -apply clarify -apply(drule prog_not_eq_in_ctran_aux) -apply simp -done - -lemma prog_not_eq_in_par_ctran_aux [rule_format]: "(P,s) -pc\ (Q,t) \ (P\Q)" -apply(erule par_ctran.induct) -apply(drule prog_not_eq_in_ctran_aux) -apply clarify -apply(drule list_eq_if) - apply simp_all -apply force -done - -lemma prog_not_eq_in_par_ctran [simp]: "\ (P,s) -pc\ (P,t)" -apply clarify -apply(drule prog_not_eq_in_par_ctran_aux) -apply simp -done - -lemma tl_in_cptn: "\ a#xs \cptn; xs\[] \ \ xs\cptn" -apply(force elim:cptn.cases) -done - -lemma tl_zero[rule_format]: - "P (ys!Suc j) \ Suc j ys\[] \ P (tl(ys)!j)" -apply(induct ys) - apply simp_all -done - -subsection {* The Semantics is Compositional *} - -lemma aux_if [rule_format]: - "\xs s clist. (length clist = length xs \ (\i cptn) - \ ((xs, s)#ys \ map (\i. (fst i,s)#snd i) (zip xs clist)) - \ (xs, s)#ys \ par_cptn)" -apply(induct ys) - apply(clarify) - apply(rule ParCptnOne) -apply(clarify) -apply(simp add:conjoin_def compat_label_def) -apply clarify -apply(erule_tac x="0" and P="\j. ?H j \ (?P j \ ?Q j)" in all_dupE,simp) -apply(erule disjE) ---{* first step is a Component step *} - apply clarify - apply simp - apply(subgoal_tac "a=(xs[i:=(fst(clist!i!0))])") - apply(subgoal_tac "b=snd(clist!i!0)",simp) - prefer 2 - apply(simp add: same_state_def) - apply(erule_tac x=i in allE,erule impE,assumption, - erule_tac x=1 and P="\j. (?H j) \ (snd (?d j))=(snd (?e j))" in allE,simp) - prefer 2 - apply(simp add:same_program_def) - apply(erule_tac x=1 and P="\j. ?H j \ (fst (?s j))=(?t j)" in allE,simp) - apply(rule nth_equalityI,simp) - apply clarify - apply(case_tac "i=ia",simp,simp) - apply(erule_tac x=ia and P="\j. ?H j \ ?I j \ ?J j" in allE) - apply(drule_tac t=i in not_sym,simp) - apply(erule etranE,simp) - apply(rule ParCptnComp) - apply(erule ParComp,simp) ---{* applying the induction hypothesis *} - apply(erule_tac x="xs[i := fst (clist ! i ! 0)]" in allE) - apply(erule_tac x="snd (clist ! i ! 0)" in allE) - apply(erule mp) - apply(rule_tac x="map tl clist" in exI,simp) - apply(rule conjI,clarify) - apply(case_tac "i=ia",simp) - apply(rule nth_tl_if) - apply(force simp add:same_length_def length_Suc_conv) - apply simp - apply(erule allE,erule impE,assumption,erule tl_in_cptn) - apply(force simp add:same_length_def length_Suc_conv) - apply(rule nth_tl_if) - apply(force simp add:same_length_def length_Suc_conv) - apply(simp add:same_state_def) - apply(erule_tac x=ia in allE, erule impE, assumption, - erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) - apply(erule_tac x=ia and P="\j. ?H j \ ?I j \ ?J j" in allE) - apply(drule_tac t=i in not_sym,simp) - apply(erule etranE,simp) - apply(erule allE,erule impE,assumption,erule tl_in_cptn) - apply(force simp add:same_length_def length_Suc_conv) - apply(simp add:same_length_def same_state_def) - apply(rule conjI) - apply clarify - apply(case_tac j,simp,simp) - apply(erule_tac x=ia in allE, erule impE, assumption, - erule_tac x="Suc(Suc nat)" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) - apply(force simp add:same_length_def length_Suc_conv) - apply(rule conjI) - apply(simp add:same_program_def) - apply clarify - apply(case_tac j,simp) - apply(rule nth_equalityI,simp) - apply clarify - apply(case_tac "i=ia",simp,simp) - apply(erule_tac x="Suc(Suc nat)" and P="\j. ?H j \ (fst (?s j))=(?t j)" in allE,simp) - apply(rule nth_equalityI,simp,simp) - apply(force simp add:length_Suc_conv) - apply(rule allI,rule impI) - apply(erule_tac x="Suc j" and P="\j. ?H j \ (?I j \ ?J j)" in allE,simp) - apply(erule disjE) - apply clarify - apply(rule_tac x=ia in exI,simp) - apply(case_tac "i=ia",simp) - apply(rule conjI) - apply(force simp add: length_Suc_conv) - apply clarify - apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE,assumption) - apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE,assumption) - apply simp - apply(case_tac j,simp) - apply(rule tl_zero) - apply(erule_tac x=l in allE, erule impE, assumption, - erule_tac x=1 and P="\j. (?H j) \ (snd (?d j))=(snd (?e j))" in allE,simp) - apply(force elim:etranE intro:Env) - apply force - apply force - apply simp - apply(rule tl_zero) - apply(erule tl_zero) - apply force - apply force - apply force - apply force - apply(rule conjI,simp) - apply(rule nth_tl_if) - apply force - apply(erule_tac x=ia in allE, erule impE, assumption, - erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) - apply(erule_tac x=ia and P="\j. ?H j \ ?I j \ ?J j" in allE) - apply(drule_tac t=i in not_sym,simp) - apply(erule etranE,simp) - apply(erule tl_zero) - apply force - apply force - apply clarify - apply(case_tac "i=l",simp) - apply(rule nth_tl_if) - apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply simp - apply(erule_tac P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE,assumption,erule impE,assumption) - apply(erule tl_zero,force) - apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply(rule nth_tl_if) - apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply(erule_tac x=l in allE, erule impE, assumption, - erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) - apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE, assumption,simp) - apply(erule etranE,simp) - apply(rule tl_zero) - apply force - apply force - apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply(rule disjI2) - apply(case_tac j,simp) - apply clarify - apply(rule tl_zero) - apply(erule_tac x=ia and P="\j. ?H j \ ?I j\etran" in allE,erule impE, assumption) - apply(case_tac "i=ia",simp,simp) - apply(erule_tac x=ia in allE, erule impE, assumption, - erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) - apply(erule_tac x=ia and P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE, assumption,simp) - apply(force elim:etranE intro:Env) - apply force - apply(erule_tac x=ia and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply simp - apply clarify - apply(rule tl_zero) - apply(rule tl_zero,force) - apply force - apply(erule_tac x=ia and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply force - apply(erule_tac x=ia and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) ---{* first step is an environmental step *} -apply clarify -apply(erule par_etran.cases) -apply simp -apply(rule ParCptnEnv) -apply(erule_tac x="Ps" in allE) -apply(erule_tac x="t" in allE) -apply(erule mp) -apply(rule_tac x="map tl clist" in exI,simp) -apply(rule conjI) - apply clarify - apply(erule_tac x=i and P="\j. ?H j \ (?I ?s j) \ cptn" in allE,simp) - apply(erule cptn.cases) - apply(simp add:same_length_def) - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply(simp add:same_state_def) - apply(erule_tac x=i in allE, erule impE, assumption, - erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) - apply(erule_tac x=i and P="\j. ?H j \ ?J j \etran" in allE,simp) - apply(erule etranE,simp) -apply(simp add:same_state_def same_length_def) -apply(rule conjI,clarify) - apply(case_tac j,simp,simp) - apply(erule_tac x=i in allE, erule impE, assumption, - erule_tac x="Suc(Suc nat)" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) - apply(rule tl_zero) - apply(simp) - apply force - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) -apply(rule conjI) - apply(simp add:same_program_def) - apply clarify - apply(case_tac j,simp) - apply(rule nth_equalityI,simp) - apply clarify - apply simp - apply(erule_tac x="Suc(Suc nat)" and P="\j. ?H j \ (fst (?s j))=(?t j)" in allE,simp) - apply(rule nth_equalityI,simp,simp) - apply(force simp add:length_Suc_conv) -apply(rule allI,rule impI) -apply(erule_tac x="Suc j" and P="\j. ?H j \ (?I j \ ?J j)" in allE,simp) -apply(erule disjE) - apply clarify - apply(rule_tac x=i in exI,simp) - apply(rule conjI) - apply(erule_tac x=i and P="\i. ?H i \ ?J i \etran" in allE, erule impE, assumption) - apply(erule etranE,simp) - apply(erule_tac x=i in allE, erule impE, assumption, - erule_tac x=1 and P="\j. (?H j) \ (snd (?d j))=(snd (?e j))" in allE,simp) - apply(rule nth_tl_if) - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply simp - apply(erule tl_zero,force) - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply clarify - apply(erule_tac x=l and P="\i. ?H i \ ?J i \etran" in allE, erule impE, assumption) - apply(erule etranE,simp) - apply(erule_tac x=l in allE, erule impE, assumption, - erule_tac x=1 and P="\j. (?H j) \ (snd (?d j))=(snd (?e j))" in allE,simp) - apply(rule nth_tl_if) - apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply simp - apply(rule tl_zero,force) - apply force - apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) -apply(rule disjI2) -apply simp -apply clarify -apply(case_tac j,simp) - apply(rule tl_zero) - apply(erule_tac x=i and P="\i. ?H i \ ?J i \etran" in allE, erule impE, assumption) - apply(erule_tac x=i and P="\i. ?H i \ ?J i \etran" in allE, erule impE, assumption) - apply(force elim:etranE intro:Env) - apply force - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) -apply simp -apply(rule tl_zero) - apply(rule tl_zero,force) - apply force - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply force -apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) -done - -lemma less_Suc_0 [iff]: "(n < Suc 0) = (n = 0)" -by auto - -lemma aux_onlyif [rule_format]: "\xs s. (xs, s)#ys \ par_cptn \ - (\clist. (length clist = length xs) \ - (xs, s)#ys \ map (\i. (fst i,s)#(snd i)) (zip xs clist) \ - (\i cptn))" -apply(induct ys) - apply(clarify) - apply(rule_tac x="map (\i. []) [0..[] \ (\ys. ((xs, s)#ys \ par_cptn) = - (\clist. length clist= length xs \ - ((xs, s)#ys \ map (\i. (fst i,s)#(snd i)) (zip xs clist)) \ - (\i cptn))) = - (par_cp (xs) s = {c. \clist. (length clist)=(length xs) \ - (\i cp(xs!i) s) \ c \ clist})" -apply (rule iffI) - apply(rule subset_antisym) - apply(rule subsetI) - apply(clarify) - apply(simp add:par_cp_def cp_def) - apply(case_tac x) - apply(force elim:par_cptn.cases) - apply simp - apply(erule_tac x="list" in allE) - apply clarify - apply simp - apply(rule_tac x="map (\i. (fst i, s) # snd i) (zip xs clist)" in exI,simp) - apply(rule subsetI) - apply(clarify) - apply(case_tac x) - apply(erule_tac x=0 in allE) - apply(simp add:cp_def conjoin_def same_length_def same_program_def same_state_def compat_label_def) - apply clarify - apply(erule cptn.cases,force,force,force) - apply(simp add:par_cp_def conjoin_def same_length_def same_program_def same_state_def compat_label_def) - apply clarify - apply(erule_tac x=0 and P="\j. ?H j \ (length (?s j) = ?t)" in all_dupE) - apply(subgoal_tac "a = xs") - apply(subgoal_tac "b = s",simp) - prefer 3 - apply(erule_tac x=0 and P="\j. ?H j \ (fst (?s j))=((?t j))" in allE) - apply (simp add:cp_def) - apply(rule nth_equalityI,simp,simp) - prefer 2 - apply(erule_tac x=0 in allE) - apply (simp add:cp_def) - apply(erule_tac x=0 and P="\j. ?H j \ (\i. ?T i \ (snd (?d j i))=(snd (?e j i)))" in allE,simp) - apply(erule_tac x=0 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) - apply(erule_tac x=list in allE) - apply(rule_tac x="map tl clist" in exI,simp) - apply(rule conjI) - apply clarify - apply(case_tac j,simp) - apply(erule_tac x=i in allE, erule impE, assumption, - erule_tac x="0" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) - apply(erule_tac x=i in allE, erule impE, assumption, - erule_tac x="Suc nat" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!i",simp,simp) - apply(rule conjI) - apply clarify - apply(rule nth_equalityI,simp,simp) - apply(case_tac j) - apply clarify - apply(erule_tac x=i in allE) - apply(simp add:cp_def) - apply clarify - apply simp - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!i",simp,simp) - apply(thin_tac "?H = (\i. ?J i)") - apply(rule conjI) - apply clarify - apply(erule_tac x=j in allE,erule impE, assumption,erule disjE) - apply clarify - apply(rule_tac x=i in exI,simp) - apply(case_tac j,simp) - apply(rule conjI) - apply(erule_tac x=i in allE) - apply(simp add:cp_def) - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!i",simp,simp) - apply clarify - apply(erule_tac x=l in allE) - apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE) - apply clarify - apply(simp add:cp_def) - apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!l",simp,simp) - apply simp - apply(rule conjI) - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!i",simp,simp) - apply clarify - apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE) - apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!l",simp,simp) - apply clarify - apply(erule_tac x=i in allE) - apply(simp add:cp_def) - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!i",simp) - apply(rule nth_tl_if,simp,simp) - apply(erule_tac x=i and P="\j. ?H j \ (?P j)\etran" in allE, erule impE, assumption,simp) - apply(simp add:cp_def) - apply clarify - apply(rule nth_tl_if) - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!i",simp,simp) - apply force - apply force -apply clarify -apply(rule iffI) - apply(simp add:par_cp_def) - apply(erule_tac c="(xs, s) # ys" in equalityCE) - apply simp - apply clarify - apply(rule_tac x="map tl clist" in exI) - apply simp - apply (rule conjI) - apply(simp add:conjoin_def cp_def) - apply(rule conjI) - apply clarify - apply(unfold same_length_def) - apply clarify - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,simp) - apply(rule conjI) - apply(simp add:same_state_def) - apply clarify - apply(erule_tac x=i in allE, erule impE, assumption, - erule_tac x=j and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) - apply(case_tac j,simp) - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!i",simp,simp) - apply(rule conjI) - apply(simp add:same_program_def) - apply clarify - apply(rule nth_equalityI,simp,simp) - apply(case_tac j,simp) - apply clarify - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!i",simp,simp) - apply clarify - apply(simp add:compat_label_def) - apply(rule allI,rule impI) - apply(erule_tac x=j in allE,erule impE, assumption) - apply(erule disjE) - apply clarify - apply(rule_tac x=i in exI,simp) - apply(rule conjI) - apply(erule_tac x=i in allE) - apply(case_tac j,simp) - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!i",simp,simp) - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!i",simp,simp) - apply clarify - apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE) - apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE) - apply(case_tac "clist!l",simp,simp) - apply(erule_tac x=l in allE,simp) - apply(rule disjI2) - apply clarify - apply(rule tl_zero) - apply(case_tac j,simp,simp) - apply(rule tl_zero,force) - apply force - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply force - apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) - apply clarify - apply(erule_tac x=i in allE) - apply(simp add:cp_def) - apply(rule nth_tl_if) - apply(simp add:conjoin_def) - apply clarify - apply(simp add:same_length_def) - apply(erule_tac x=i in allE,simp) - apply simp - apply simp - apply simp -apply clarify -apply(erule_tac c="(xs, s) # ys" in equalityCE) - apply(simp add:par_cp_def) -apply simp -apply(erule_tac x="map (\i. (fst i, s) # snd i) (zip xs clist)" in allE) -apply simp -apply clarify -apply(simp add:cp_def) -done - -theorem one: "xs\[] \ - par_cp xs s = {c. \clist. (length clist)=(length xs) \ - (\i cp(xs!i) s) \ c \ clist}" -apply(frule one_iff_aux) -apply(drule sym) -apply(erule iffD2) -apply clarify -apply(rule iffI) - apply(erule aux_onlyif) -apply clarify -apply(force intro:aux_if) -done - -end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/ROOT.ML --- a/src/HOL/HoareParallel/ROOT.ML Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,3 +0,0 @@ -(* $Id$ *) - -use_thys ["OG_Examples", "Gar_Coll", "Mul_Gar_Coll", "RG_Examples"]; diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/document/root.bib --- a/src/HOL/HoareParallel/document/root.bib Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,28 +0,0 @@ -@inproceedings{NipkowP-FASE99,author={Tobias Nipkow and Prensa Nieto, Leonor}, -title={{Owicki/Gries} in {Isabelle/HOL}}, -booktitle={Fundamental Approaches to Software Engineering (FASE'99)}, -editor={J.-P. Finance},publisher="Springer",series="LNCS",volume=1577, -pages={188--203},year=1999} - -@InProceedings{PrenEsp00, - author = {Prensa Nieto, Leonor and Javier Esparza}, - title = {Verifying Single and Multi-mutator Garbage Collectors - with {Owicki/Gries} in {Isabelle/HOL}}, - booktitle = {Mathematical Foundations of Computer Science (MFCS 2000)}, - editor = {M. Nielsen and B. Rovan}, - publisher = {Springer-Verlag}, - series = {LNCS}, - volume = 1893, - pages = {619--628}, - year = 2000 -} - -@PhdThesis{Prensa-PhD,author={Leonor Prensa Nieto}, -title={Verification of Parallel Programs with the Owicki-Gries and -Rely-Guarantee Methods in Isabelle/HOL}, -school={Technische Universit{\"a}t M{\"u}nchen},year=2002} - -@inproceedings{Prensa-ESOP03,author={Prensa Nieto, Leonor}, -title={The {Rely-Guarantee} Method in {Isabelle/HOL}}, -booktitle={European Symposium on Programming (ESOP'03)},editor={P. Degano}, -publisher=Springer,series=LNCS,volume=2618,pages={348--362},year=2003} diff -r c876bcb601fc -r 5b65449d7669 src/HOL/HoareParallel/document/root.tex --- a/src/HOL/HoareParallel/document/root.tex Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,62 +0,0 @@ - -% $Id$ - -\documentclass[11pt,a4paper]{report} -\usepackage{graphicx} -\usepackage[english]{babel} -\usepackage{isabelle,isabellesym} -\usepackage{pdfsetup} - -\urlstyle{rm} -\isabellestyle{it} - -\renewcommand{\isamarkupheader}[1]{#1} - -\begin{document} - -\title{Hoare Logic for Parallel Programs} -\author{Leonor Prensa Nieto} -\maketitle - -\begin{abstract}\noindent - In the following theories a formalization of the Owicki-Gries and - the rely-guarantee methods is presented. These methods are widely - used for correctness proofs of parallel imperative programs with - shared variables. We define syntax, semantics and proof rules in - Isabelle/HOL. The proof rules also provide for programs - parameterized in the number of parallel components. Their - correctness w.r.t.\ the semantics is proven. Completeness proofs - for both methods are extended to the new case of parameterized - programs. (These proofs have not been formalized in Isabelle. They - can be found in~\cite{Prensa-PhD}.) Using this formalizations we - verify several non-trivial examples for parameterized and - non-parameterized programs. For the automatic generation of - verification conditions with the Owicki-Gries method we define a - tactic based on the proof rules. The most involved examples are the - verification of two garbage-collection algorithms, the second one - parameterized in the number of mutators. - -For excellent descriptions of this work see -\cite{NipkowP-FASE99,PrenEsp00,Prensa-PhD,Prensa-ESOP03}. - -\end{abstract} - -\pagestyle{plain} -\thispagestyle{empty} -\tableofcontents - -\clearpage - -\begin{center} - \includegraphics[scale=0.7]{session_graph} -\end{center} - -\newpage - -\parindent 0pt\parskip 0.5ex -\input{session} - -\bibliographystyle{plain} -\bibliography{root} - -\end{document} diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/Gar_Coll.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/Gar_Coll.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,846 @@ + +header {* \section{The Single Mutator Case} *} + +theory Gar_Coll imports Graph OG_Syntax begin + +declare psubsetE [rule del] + +text {* Declaration of variables: *} + +record gar_coll_state = + M :: nodes + E :: edges + bc :: "nat set" + obc :: "nat set" + Ma :: nodes + ind :: nat + k :: nat + z :: bool + +subsection {* The Mutator *} + +text {* The mutator first redirects an arbitrary edge @{text "R"} from +an arbitrary accessible node towards an arbitrary accessible node +@{text "T"}. It then colors the new target @{text "T"} black. + +We declare the arbitrarily selected node and edge as constants:*} + +consts R :: nat T :: nat + +text {* \noindent The following predicate states, given a list of +nodes @{text "m"} and a list of edges @{text "e"}, the conditions +under which the selected edge @{text "R"} and node @{text "T"} are +valid: *} + +constdefs + Mut_init :: "gar_coll_state \ bool" + "Mut_init \ \ T \ Reach \E \ R < length \E \ T < length \M \" + +text {* \noindent For the mutator we +consider two modules, one for each action. An auxiliary variable +@{text "\z"} is set to false if the mutator has already redirected an +edge but has not yet colored the new target. *} + +constdefs + Redirect_Edge :: "gar_coll_state ann_com" + "Redirect_Edge \ .{\Mut_init \ \z}. \\E:=\E[R:=(fst(\E!R), T)],, \z:= (\\z)\" + + Color_Target :: "gar_coll_state ann_com" + "Color_Target \ .{\Mut_init \ \\z}. \\M:=\M[T:=Black],, \z:= (\\z)\" + + Mutator :: "gar_coll_state ann_com" + "Mutator \ + .{\Mut_init \ \z}. + WHILE True INV .{\Mut_init \ \z}. + DO Redirect_Edge ;; Color_Target OD" + +subsubsection {* Correctness of the mutator *} + +lemmas mutator_defs = Mut_init_def Redirect_Edge_def Color_Target_def + +lemma Redirect_Edge: + "\ Redirect_Edge pre(Color_Target)" +apply (unfold mutator_defs) +apply annhoare +apply(simp_all) +apply(force elim:Graph2) +done + +lemma Color_Target: + "\ Color_Target .{\Mut_init \ \z}." +apply (unfold mutator_defs) +apply annhoare +apply(simp_all) +done + +lemma Mutator: + "\ Mutator .{False}." +apply(unfold Mutator_def) +apply annhoare +apply(simp_all add:Redirect_Edge Color_Target) +apply(simp add:mutator_defs Redirect_Edge_def) +done + +subsection {* The Collector *} + +text {* \noindent A constant @{text "M_init"} is used to give @{text "\Ma"} a +suitable first value, defined as a list of nodes where only the @{text +"Roots"} are black. *} + +consts M_init :: nodes + +constdefs + Proper_M_init :: "gar_coll_state \ bool" + "Proper_M_init \ \ Blacks M_init=Roots \ length M_init=length \M \" + + Proper :: "gar_coll_state \ bool" + "Proper \ \ Proper_Roots \M \ Proper_Edges(\M, \E) \ \Proper_M_init \" + + Safe :: "gar_coll_state \ bool" + "Safe \ \ Reach \E \ Blacks \M \" + +lemmas collector_defs = Proper_M_init_def Proper_def Safe_def + +subsubsection {* Blackening the roots *} + +constdefs + Blacken_Roots :: " gar_coll_state ann_com" + "Blacken_Roots \ + .{\Proper}. + \ind:=0;; + .{\Proper \ \ind=0}. + WHILE \indM + INV .{\Proper \ (\i<\ind. i \ Roots \ \M!i=Black) \ \ind\length \M}. + DO .{\Proper \ (\i<\ind. i \ Roots \ \M!i=Black) \ \indM}. + IF \ind\Roots THEN + .{\Proper \ (\i<\ind. i \ Roots \ \M!i=Black) \ \indM \ \ind\Roots}. + \M:=\M[\ind:=Black] FI;; + .{\Proper \ (\i<\ind+1. i \ Roots \ \M!i=Black) \ \indM}. + \ind:=\ind+1 + OD" + +lemma Blacken_Roots: + "\ Blacken_Roots .{\Proper \ Roots\Blacks \M}." +apply (unfold Blacken_Roots_def) +apply annhoare +apply(simp_all add:collector_defs Graph_defs) +apply safe +apply(simp_all add:nth_list_update) + apply (erule less_SucE) + apply simp+ + apply force +apply force +done + +subsubsection {* Propagating black *} + +constdefs + PBInv :: "gar_coll_state \ nat \ bool" + "PBInv \ \ \ind. \obc < Blacks \M \ (\i BtoW (\E!i, \M) \ + (\\z \ i=R \ (snd(\E!R)) = T \ (\r. ind \ r \ r < length \E \ BtoW(\E!r,\M))))\" + +constdefs + Propagate_Black_aux :: "gar_coll_state ann_com" + "Propagate_Black_aux \ + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M}. + \ind:=0;; + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M \ \ind=0}. + WHILE \indE + INV .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \ind\length \E}. + DO .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \indE}. + IF \M!(fst (\E!\ind)) = Black THEN + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \indE \ \M!fst(\E!\ind)=Black}. + \M:=\M[snd(\E!\ind):=Black];; + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv (\ind + 1) \ \indE}. + \ind:=\ind+1 + FI + OD" + +lemma Propagate_Black_aux: + "\ Propagate_Black_aux + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ ( \obc < Blacks \M \ \Safe)}." +apply (unfold Propagate_Black_aux_def PBInv_def collector_defs) +apply annhoare +apply(simp_all add:Graph6 Graph7 Graph8 Graph12) + apply force + apply force + apply force +--{* 4 subgoals left *} +apply clarify +apply(simp add:Proper_Edges_def Proper_Roots_def Graph6 Graph7 Graph8 Graph12) +apply (erule disjE) + apply(rule disjI1) + apply(erule Graph13) + apply force +apply (case_tac "M x ! snd (E x ! ind x)=Black") + apply (simp add: Graph10 BtoW_def) + apply (rule disjI2) + apply clarify + apply (erule less_SucE) + apply (erule_tac x=i in allE , erule (1) notE impE) + apply simp + apply clarify + apply (drule_tac y = r in le_imp_less_or_eq) + apply (erule disjE) + apply (subgoal_tac "Suc (ind x)\r") + apply fast + apply arith + apply fast + apply fast +apply(rule disjI1) +apply(erule subset_psubset_trans) +apply(erule Graph11) +apply fast +--{* 3 subgoals left *} +apply force +apply force +--{* last *} +apply clarify +apply simp +apply(subgoal_tac "ind x = length (E x)") + apply (rotate_tac -1) + apply (simp (asm_lr)) + apply(drule Graph1) + apply simp + apply clarify + apply(erule allE, erule impE, assumption) + apply force + apply force +apply arith +done + +subsubsection {* Refining propagating black *} + +constdefs + Auxk :: "gar_coll_state \ bool" + "Auxk \ \\kM \ (\M!\k\Black \ \BtoW(\E!\ind, \M) \ + \obcM \ (\\z \ \ind=R \ snd(\E!R)=T + \ (\r. \ind rE \ BtoW(\E!r, \M))))\" + +constdefs + Propagate_Black :: " gar_coll_state ann_com" + "Propagate_Black \ + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M}. + \ind:=0;; + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M \ \ind=0}. + WHILE \indE + INV .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \ind\length \E}. + DO .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \indE}. + IF (\M!(fst (\E!\ind)))=Black THEN + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \indE \ (\M!fst(\E!\ind))=Black}. + \k:=(snd(\E!\ind));; + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \indE \ (\M!fst(\E!\ind))=Black + \ \Auxk}. + \\M:=\M[\k:=Black],, \ind:=\ind+1\ + ELSE .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ \PBInv \ind \ \indE}. + \IF (\M!(fst (\E!\ind)))\Black THEN \ind:=\ind+1 FI\ + FI + OD" + +lemma Propagate_Black: + "\ Propagate_Black + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ ( \obc < Blacks \M \ \Safe)}." +apply (unfold Propagate_Black_def PBInv_def Auxk_def collector_defs) +apply annhoare +apply(simp_all add:Graph6 Graph7 Graph8 Graph12) + apply force + apply force + apply force +--{* 5 subgoals left *} +apply clarify +apply(simp add:BtoW_def Proper_Edges_def) +--{* 4 subgoals left *} +apply clarify +apply(simp add:Proper_Edges_def Graph6 Graph7 Graph8 Graph12) +apply (erule disjE) + apply (rule disjI1) + apply (erule psubset_subset_trans) + apply (erule Graph9) +apply (case_tac "M x!k x=Black") + apply (case_tac "M x ! snd (E x ! ind x)=Black") + apply (simp add: Graph10 BtoW_def) + apply (rule disjI2) + apply clarify + apply (erule less_SucE) + apply (erule_tac x=i in allE , erule (1) notE impE) + apply simp + apply clarify + apply (drule_tac y = r in le_imp_less_or_eq) + apply (erule disjE) + apply (subgoal_tac "Suc (ind x)\r") + apply fast + apply arith + apply fast + apply fast + apply (simp add: Graph10 BtoW_def) + apply (erule disjE) + apply (erule disjI1) + apply clarify + apply (erule less_SucE) + apply force + apply simp + apply (subgoal_tac "Suc R\r") + apply fast + apply arith +apply(rule disjI1) +apply(erule subset_psubset_trans) +apply(erule Graph11) +apply fast +--{* 3 subgoals left *} +apply force +--{* 2 subgoals left *} +apply clarify +apply(simp add:Proper_Edges_def Graph6 Graph7 Graph8 Graph12) +apply (erule disjE) + apply fast +apply clarify +apply (erule less_SucE) + apply (erule_tac x=i in allE , erule (1) notE impE) + apply simp + apply clarify + apply (drule_tac y = r in le_imp_less_or_eq) + apply (erule disjE) + apply (subgoal_tac "Suc (ind x)\r") + apply fast + apply arith + apply (simp add: BtoW_def) +apply (simp add: BtoW_def) +--{* last *} +apply clarify +apply simp +apply(subgoal_tac "ind x = length (E x)") + apply (rotate_tac -1) + apply (simp (asm_lr)) + apply(drule Graph1) + apply simp + apply clarify + apply(erule allE, erule impE, assumption) + apply force + apply force +apply arith +done + +subsubsection {* Counting black nodes *} + +constdefs + CountInv :: "gar_coll_state \ nat \ bool" + "CountInv \ \ \ind. {i. i \Ma!i=Black}\\bc \" + +constdefs + Count :: " gar_coll_state ann_com" + "Count \ + .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ (\obc < Blacks \Ma \ \Safe) \ \bc={}}. + \ind:=0;; + .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ (\obc < Blacks \Ma \ \Safe) \ \bc={} + \ \ind=0}. + WHILE \indM + INV .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \CountInv \ind + \ ( \obc < Blacks \Ma \ \Safe) \ \ind\length \M}. + DO .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \CountInv \ind + \ ( \obc < Blacks \Ma \ \Safe) \ \indM}. + IF \M!\ind=Black + THEN .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \CountInv \ind + \ ( \obc < Blacks \Ma \ \Safe) \ \indM \ \M!\ind=Black}. + \bc:=insert \ind \bc + FI;; + .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \CountInv (\ind+1) + \ ( \obc < Blacks \Ma \ \Safe) \ \indM}. + \ind:=\ind+1 + OD" + +lemma Count: + "\ Count + .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\\bc \ \bc\Blacks \M \ length \Ma=length \M + \ (\obc < Blacks \Ma \ \Safe)}." +apply(unfold Count_def) +apply annhoare +apply(simp_all add:CountInv_def Graph6 Graph7 Graph8 Graph12 Blacks_def collector_defs) + apply force + apply force + apply force + apply clarify + apply simp + apply(fast elim:less_SucE) + apply clarify + apply simp + apply(fast elim:less_SucE) + apply force +apply force +done + +subsubsection {* Appending garbage nodes to the free list *} + +consts Append_to_free :: "nat \ edges \ edges" + +axioms + Append_to_free0: "length (Append_to_free (i, e)) = length e" + Append_to_free1: "Proper_Edges (m, e) + \ Proper_Edges (m, Append_to_free(i, e))" + Append_to_free2: "i \ Reach e + \ n \ Reach (Append_to_free(i, e)) = ( n = i \ n \ Reach e)" + +constdefs + AppendInv :: "gar_coll_state \ nat \ bool" + "AppendInv \ \\ind. \iM. ind\i \ i\Reach \E \ \M!i=Black\" + +constdefs + Append :: " gar_coll_state ann_com" + "Append \ + .{\Proper \ Roots\Blacks \M \ \Safe}. + \ind:=0;; + .{\Proper \ Roots\Blacks \M \ \Safe \ \ind=0}. + WHILE \indM + INV .{\Proper \ \AppendInv \ind \ \ind\length \M}. + DO .{\Proper \ \AppendInv \ind \ \indM}. + IF \M!\ind=Black THEN + .{\Proper \ \AppendInv \ind \ \indM \ \M!\ind=Black}. + \M:=\M[\ind:=White] + ELSE .{\Proper \ \AppendInv \ind \ \indM \ \ind\Reach \E}. + \E:=Append_to_free(\ind,\E) + FI;; + .{\Proper \ \AppendInv (\ind+1) \ \indM}. + \ind:=\ind+1 + OD" + +lemma Append: + "\ Append .{\Proper}." +apply(unfold Append_def AppendInv_def) +apply annhoare +apply(simp_all add:collector_defs Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) + apply(force simp:Blacks_def nth_list_update) + apply force + apply force + apply(force simp add:Graph_defs) + apply force + apply clarify + apply simp + apply(rule conjI) + apply (erule Append_to_free1) + apply clarify + apply (drule_tac n = "i" in Append_to_free2) + apply force + apply force +apply force +done + +subsubsection {* Correctness of the Collector *} + +constdefs + Collector :: " gar_coll_state ann_com" + "Collector \ +.{\Proper}. + WHILE True INV .{\Proper}. + DO + Blacken_Roots;; + .{\Proper \ Roots\Blacks \M}. + \obc:={};; + .{\Proper \ Roots\Blacks \M \ \obc={}}. + \bc:=Roots;; + .{\Proper \ Roots\Blacks \M \ \obc={} \ \bc=Roots}. + \Ma:=M_init;; + .{\Proper \ Roots\Blacks \M \ \obc={} \ \bc=Roots \ \Ma=M_init}. + WHILE \obc\\bc + INV .{\Proper \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\\bc \ \bc\Blacks \M + \ length \Ma=length \M \ (\obc < Blacks \Ma \ \Safe)}. + DO .{\Proper \ Roots\Blacks \M \ \bc\Blacks \M}. + \obc:=\bc;; + Propagate_Black;; + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ (\obc < Blacks \M \ \Safe)}. + \Ma:=\M;; + .{\Proper \ Roots\Blacks \M \ \obc\Blacks \Ma + \ Blacks \Ma\Blacks \M \ \bc\Blacks \M \ length \Ma=length \M + \ ( \obc < Blacks \Ma \ \Safe)}. + \bc:={};; + Count + OD;; + Append + OD" + +lemma Collector: + "\ Collector .{False}." +apply(unfold Collector_def) +apply annhoare +apply(simp_all add: Blacken_Roots Propagate_Black Count Append) +apply(simp_all add:Blacken_Roots_def Propagate_Black_def Count_def Append_def collector_defs) + apply (force simp add: Proper_Roots_def) + apply force + apply force +apply clarify +apply (erule disjE) +apply(simp add:psubsetI) + apply(force dest:subset_antisym) +done + +subsection {* Interference Freedom *} + +lemmas modules = Redirect_Edge_def Color_Target_def Blacken_Roots_def + Propagate_Black_def Count_def Append_def +lemmas Invariants = PBInv_def Auxk_def CountInv_def AppendInv_def +lemmas abbrev = collector_defs mutator_defs Invariants + +lemma interfree_Blacken_Roots_Redirect_Edge: + "interfree_aux (Some Blacken_Roots, {}, Some Redirect_Edge)" +apply (unfold modules) +apply interfree_aux +apply safe +apply (simp_all add:Graph6 Graph12 abbrev) +done + +lemma interfree_Redirect_Edge_Blacken_Roots: + "interfree_aux (Some Redirect_Edge, {}, Some Blacken_Roots)" +apply (unfold modules) +apply interfree_aux +apply safe +apply(simp add:abbrev)+ +done + +lemma interfree_Blacken_Roots_Color_Target: + "interfree_aux (Some Blacken_Roots, {}, Some Color_Target)" +apply (unfold modules) +apply interfree_aux +apply safe +apply(simp_all add:Graph7 Graph8 nth_list_update abbrev) +done + +lemma interfree_Color_Target_Blacken_Roots: + "interfree_aux (Some Color_Target, {}, Some Blacken_Roots)" +apply (unfold modules ) +apply interfree_aux +apply safe +apply(simp add:abbrev)+ +done + +lemma interfree_Propagate_Black_Redirect_Edge: + "interfree_aux (Some Propagate_Black, {}, Some Redirect_Edge)" +apply (unfold modules ) +apply interfree_aux +--{* 11 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(erule conjE)+ +apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) + apply(erule Graph4) + apply(simp)+ + apply (simp add:BtoW_def) + apply (simp add:BtoW_def) +apply(rule conjI) + apply (force simp add:BtoW_def) +apply(erule Graph4) + apply simp+ +--{* 7 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(erule conjE)+ +apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) + apply(erule Graph4) + apply(simp)+ + apply (simp add:BtoW_def) + apply (simp add:BtoW_def) +apply(rule conjI) + apply (force simp add:BtoW_def) +apply(erule Graph4) + apply simp+ +--{* 6 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(erule conjE)+ +apply(rule conjI) + apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) + apply(erule Graph4) + apply(simp)+ + apply (simp add:BtoW_def) + apply (simp add:BtoW_def) + apply(rule conjI) + apply (force simp add:BtoW_def) + apply(erule Graph4) + apply simp+ +apply(simp add:BtoW_def nth_list_update) +apply force +--{* 5 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +--{* 4 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(rule conjI) + apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) + apply(erule Graph4) + apply(simp)+ + apply (simp add:BtoW_def) + apply (simp add:BtoW_def) + apply(rule conjI) + apply (force simp add:BtoW_def) + apply(erule Graph4) + apply simp+ +apply(rule conjI) + apply(simp add:nth_list_update) + apply force +apply(rule impI, rule impI, erule disjE, erule disjI1, case_tac "R = (ind x)" ,case_tac "M x ! T = Black") + apply(force simp add:BtoW_def) + apply(case_tac "M x !snd (E x ! ind x)=Black") + apply(rule disjI2) + apply simp + apply (erule Graph5) + apply simp+ + apply(force simp add:BtoW_def) +apply(force simp add:BtoW_def) +--{* 3 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +--{* 2 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12) +apply(erule disjE, erule disjI1, rule disjI2, rule allI, (rule impI)+, case_tac "R=i", rule conjI, erule sym) + apply clarify + apply(erule Graph4) + apply(simp)+ + apply (simp add:BtoW_def) + apply (simp add:BtoW_def) +apply(rule conjI) + apply (force simp add:BtoW_def) +apply(erule Graph4) + apply simp+ +done + +lemma interfree_Redirect_Edge_Propagate_Black: + "interfree_aux (Some Redirect_Edge, {}, Some Propagate_Black)" +apply (unfold modules ) +apply interfree_aux +apply(clarify, simp add:abbrev)+ +done + +lemma interfree_Propagate_Black_Color_Target: + "interfree_aux (Some Propagate_Black, {}, Some Color_Target)" +apply (unfold modules ) +apply interfree_aux +--{* 11 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12)+ +apply(erule conjE)+ +apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, + case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, + erule allE, erule impE, assumption, erule impE, assumption, + simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +--{* 7 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +apply(erule conjE)+ +apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, + case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, + erule allE, erule impE, assumption, erule impE, assumption, + simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +--{* 6 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +apply clarify +apply (rule conjI) + apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, + case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, + erule allE, erule impE, assumption, erule impE, assumption, + simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +apply(simp add:nth_list_update) +--{* 5 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +--{* 4 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +apply (rule conjI) + apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, + case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, + erule allE, erule impE, assumption, erule impE, assumption, + simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +apply(rule conjI) +apply(simp add:nth_list_update) +apply(rule impI,rule impI, case_tac "M x!T=Black",rotate_tac -1, force simp add: BtoW_def Graph10, + erule subset_psubset_trans, erule Graph11, force) +--{* 3 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +--{* 2 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +apply(erule disjE,rule disjI1,erule psubset_subset_trans,erule Graph9, + case_tac "M x!T=Black", rule disjI2,rotate_tac -1, simp add: Graph10, clarify, + erule allE, erule impE, assumption, erule impE, assumption, + simp add:BtoW_def, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +--{* 3 subgoals left *} +apply(simp add:abbrev) +done + +lemma interfree_Color_Target_Propagate_Black: + "interfree_aux (Some Color_Target, {}, Some Propagate_Black)" +apply (unfold modules ) +apply interfree_aux +apply(clarify, simp add:abbrev)+ +done + +lemma interfree_Count_Redirect_Edge: + "interfree_aux (Some Count, {}, Some Redirect_Edge)" +apply (unfold modules) +apply interfree_aux +--{* 9 subgoals left *} +apply(simp_all add:abbrev Graph6 Graph12) +--{* 6 subgoals left *} +apply(clarify, simp add:abbrev Graph6 Graph12, + erule disjE,erule disjI1,rule disjI2,rule subset_trans, erule Graph3,force,force)+ +done + +lemma interfree_Redirect_Edge_Count: + "interfree_aux (Some Redirect_Edge, {}, Some Count)" +apply (unfold modules ) +apply interfree_aux +apply(clarify,simp add:abbrev)+ +apply(simp add:abbrev) +done + +lemma interfree_Count_Color_Target: + "interfree_aux (Some Count, {}, Some Color_Target)" +apply (unfold modules ) +apply interfree_aux +--{* 9 subgoals left *} +apply(simp_all add:abbrev Graph7 Graph8 Graph12) +--{* 6 subgoals left *} +apply(clarify,simp add:abbrev Graph7 Graph8 Graph12, + erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9)+ +--{* 2 subgoals left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12) +apply(rule conjI) + apply(erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9) +apply(simp add:nth_list_update) +--{* 1 subgoal left *} +apply(clarify, simp add:abbrev Graph7 Graph8 Graph12, + erule disjE, erule disjI1, rule disjI2,erule subset_trans, erule Graph9) +done + +lemma interfree_Color_Target_Count: + "interfree_aux (Some Color_Target, {}, Some Count)" +apply (unfold modules ) +apply interfree_aux +apply(clarify, simp add:abbrev)+ +apply(simp add:abbrev) +done + +lemma interfree_Append_Redirect_Edge: + "interfree_aux (Some Append, {}, Some Redirect_Edge)" +apply (unfold modules ) +apply interfree_aux +apply( simp_all add:abbrev Graph6 Append_to_free0 Append_to_free1 Graph12) +apply(clarify, simp add:abbrev Graph6 Append_to_free0 Append_to_free1 Graph12, force dest:Graph3)+ +done + +lemma interfree_Redirect_Edge_Append: + "interfree_aux (Some Redirect_Edge, {}, Some Append)" +apply (unfold modules ) +apply interfree_aux +apply(clarify, simp add:abbrev Append_to_free0)+ +apply (force simp add: Append_to_free2) +apply(clarify, simp add:abbrev Append_to_free0)+ +done + +lemma interfree_Append_Color_Target: + "interfree_aux (Some Append, {}, Some Color_Target)" +apply (unfold modules ) +apply interfree_aux +apply(clarify, simp add:abbrev Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12 nth_list_update)+ +apply(simp add:abbrev Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12 nth_list_update) +done + +lemma interfree_Color_Target_Append: + "interfree_aux (Some Color_Target, {}, Some Append)" +apply (unfold modules ) +apply interfree_aux +apply(clarify, simp add:abbrev Append_to_free0)+ +apply (force simp add: Append_to_free2) +apply(clarify,simp add:abbrev Append_to_free0)+ +done + +lemmas collector_mutator_interfree = + interfree_Blacken_Roots_Redirect_Edge interfree_Blacken_Roots_Color_Target + interfree_Propagate_Black_Redirect_Edge interfree_Propagate_Black_Color_Target + interfree_Count_Redirect_Edge interfree_Count_Color_Target + interfree_Append_Redirect_Edge interfree_Append_Color_Target + interfree_Redirect_Edge_Blacken_Roots interfree_Color_Target_Blacken_Roots + interfree_Redirect_Edge_Propagate_Black interfree_Color_Target_Propagate_Black + interfree_Redirect_Edge_Count interfree_Color_Target_Count + interfree_Redirect_Edge_Append interfree_Color_Target_Append + +subsubsection {* Interference freedom Collector-Mutator *} + +lemma interfree_Collector_Mutator: + "interfree_aux (Some Collector, {}, Some Mutator)" +apply(unfold Collector_def Mutator_def) +apply interfree_aux +apply(simp_all add:collector_mutator_interfree) +apply(unfold modules collector_defs mutator_defs) +apply(tactic {* TRYALL (interfree_aux_tac) *}) +--{* 32 subgoals left *} +apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) +--{* 20 subgoals left *} +apply(tactic{* TRYALL (clarify_tac @{claset}) *}) +apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) +apply(tactic {* TRYALL (etac disjE) *}) +apply simp_all +apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac subset_trans,etac @{thm Graph3},force_tac @{clasimpset}, assume_tac]) *}) +apply(tactic {* TRYALL(EVERY'[rtac disjI2,etac subset_trans,rtac @{thm Graph9},force_tac @{clasimpset}]) *}) +apply(tactic {* TRYALL(EVERY'[rtac disjI1,etac @{thm psubset_subset_trans},rtac @{thm Graph9},force_tac @{clasimpset}]) *}) +done + +subsubsection {* Interference freedom Mutator-Collector *} + +lemma interfree_Mutator_Collector: + "interfree_aux (Some Mutator, {}, Some Collector)" +apply(unfold Collector_def Mutator_def) +apply interfree_aux +apply(simp_all add:collector_mutator_interfree) +apply(unfold modules collector_defs mutator_defs) +apply(tactic {* TRYALL (interfree_aux_tac) *}) +--{* 64 subgoals left *} +apply(simp_all add:nth_list_update Invariants Append_to_free0)+ +apply(tactic{* TRYALL (clarify_tac @{claset}) *}) +--{* 4 subgoals left *} +apply force +apply(simp add:Append_to_free2) +apply force +apply(simp add:Append_to_free2) +done + +subsubsection {* The Garbage Collection algorithm *} + +text {* In total there are 289 verification conditions. *} + +lemma Gar_Coll: + "\- .{\Proper \ \Mut_init \ \z}. + COBEGIN + Collector + .{False}. + \ + Mutator + .{False}. + COEND + .{False}." +apply oghoare +apply(force simp add: Mutator_def Collector_def modules) +apply(rule Collector) +apply(rule Mutator) +apply(simp add:interfree_Collector_Mutator) +apply(simp add:interfree_Mutator_Collector) +apply force +done + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/Graph.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/Graph.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,410 @@ +header {* \chapter{Case Study: Single and Multi-Mutator Garbage Collection Algorithms} + +\section {Formalization of the Memory} *} + +theory Graph imports Main begin + +datatype node = Black | White + +types + nodes = "node list" + edge = "nat \ nat" + edges = "edge list" + +consts Roots :: "nat set" + +constdefs + Proper_Roots :: "nodes \ bool" + "Proper_Roots M \ Roots\{} \ Roots \ {i. i edges) \ bool" + "Proper_Edges \ (\(M,E). \i snd(E!i) nodes) \ bool" + "BtoW \ (\(e,M). (M!fst e)=Black \ (M!snd e)\Black)" + + Blacks :: "nodes \ nat set" + "Blacks M \ {i. i M!i=Black}" + + Reach :: "edges \ nat set" + "Reach E \ {x. (\path. 1 path!(length path - 1)\Roots \ x=path!0 + \ (\ij x\Roots}" + +text{* Reach: the set of reachable nodes is the set of Roots together with the +nodes reachable from some Root by a path represented by a list of + nodes (at least two since we traverse at least one edge), where two +consecutive nodes correspond to an edge in E. *} + +subsection {* Proofs about Graphs *} + +lemmas Graph_defs= Blacks_def Proper_Roots_def Proper_Edges_def BtoW_def +declare Graph_defs [simp] + +subsubsection{* Graph 1 *} + +lemma Graph1_aux [rule_format]: + "\ Roots\Blacks M; \iBtoW(E!i,M)\ + \ 1< length path \ (path!(length path - 1))\Roots \ + (\ij. j < length E \ E!j=(path!(Suc i), path!i))) + \ M!(path!0) = Black" +apply(induct_tac "path") + apply force +apply clarify +apply simp +apply(case_tac "list") + apply force +apply simp +apply(rotate_tac -2) +apply(erule_tac x = "0" in all_dupE) +apply simp +apply clarify +apply(erule allE , erule (1) notE impE) +apply simp +apply(erule mp) +apply(case_tac "lista") + apply force +apply simp +apply(erule mp) +apply clarify +apply(erule_tac x = "Suc i" in allE) +apply force +done + +lemma Graph1: + "\Roots\Blacks M; Proper_Edges(M, E); \iBtoW(E!i,M) \ + \ Reach E\Blacks M" +apply (unfold Reach_def) +apply simp +apply clarify +apply(erule disjE) + apply clarify + apply(rule conjI) + apply(subgoal_tac "0< length path - Suc 0") + apply(erule allE , erule (1) notE impE) + apply force + apply simp + apply(rule Graph1_aux) +apply auto +done + +subsubsection{* Graph 2 *} + +lemma Ex_first_occurrence [rule_format]: + "P (n::nat) \ (\m. P m \ (\i. i \ P i))"; +apply(rule nat_less_induct) +apply clarify +apply(case_tac "\m. m \ P m") +apply auto +done + +lemma Compl_lemma: "(n::nat)\l \ (\m. m\l \ n=l - m)" +apply(rule_tac x = "l - n" in exI) +apply arith +done + +lemma Ex_last_occurrence: + "\P (n::nat); n\l\ \ (\m. P (l - m) \ (\i. i \P (l - i)))" +apply(drule Compl_lemma) +apply clarify +apply(erule Ex_first_occurrence) +done + +lemma Graph2: + "\T \ Reach E; R \ T \ Reach (E[R:=(fst(E!R), T)])" +apply (unfold Reach_def) +apply clarify +apply simp +apply(case_tac "\zpath!z") + apply(rule_tac x = "path" in exI) + apply simp + apply clarify + apply(erule allE , erule (1) notE impE) + apply clarify + apply(rule_tac x = "j" in exI) + apply(case_tac "j=R") + apply(erule_tac x = "Suc i" in allE) + apply simp + apply (force simp add:nth_list_update) +apply simp +apply(erule exE) +apply(subgoal_tac "z \ length path - Suc 0") + prefer 2 apply arith +apply(drule_tac P = "\m. m fst(E!R)=path!m" in Ex_last_occurrence) + apply assumption +apply clarify +apply simp +apply(rule_tac x = "(path!0)#(drop (length path - Suc m) path)" in exI) +apply simp +apply(case_tac "length path - (length path - Suc m)") + apply arith +apply simp +apply(subgoal_tac "(length path - Suc m) + nat \ length path") + prefer 2 apply arith +apply(drule nth_drop) +apply simp +apply(subgoal_tac "length path - Suc m + nat = length path - Suc 0") + prefer 2 apply arith +apply simp +apply clarify +apply(case_tac "i") + apply(force simp add: nth_list_update) +apply simp +apply(subgoal_tac "(length path - Suc m) + nata \ length path") + prefer 2 apply arith +apply(subgoal_tac "(length path - Suc m) + (Suc nata) \ length path") + prefer 2 apply arith +apply simp +apply(erule_tac x = "length path - Suc m + nata" in allE) +apply simp +apply clarify +apply(rule_tac x = "j" in exI) +apply(case_tac "R=j") + prefer 2 apply force +apply simp +apply(drule_tac t = "path ! (length path - Suc m)" in sym) +apply simp +apply(case_tac " length path - Suc 0 < m") + apply(subgoal_tac "(length path - Suc m)=0") + prefer 2 apply arith + apply(simp del: diff_is_0_eq) + apply(subgoal_tac "Suc nata\nat") + prefer 2 apply arith + apply(drule_tac n = "Suc nata" in Compl_lemma) + apply clarify + using [[linarith_split_limit = 0]] + apply force + using [[linarith_split_limit = 9]] +apply(drule leI) +apply(subgoal_tac "Suc (length path - Suc m + nata)=(length path - Suc 0) - (m - Suc nata)") + apply(erule_tac x = "m - (Suc nata)" in allE) + apply(case_tac "m") + apply simp + apply simp +apply simp +done + + +subsubsection{* Graph 3 *} + +lemma Graph3: + "\ T\Reach E; R \ Reach(E[R:=(fst(E!R),T)]) \ Reach E" +apply (unfold Reach_def) +apply clarify +apply simp +apply(case_tac "\i(length path\m)") + prefer 2 apply arith + apply(simp) + apply(rule conjI) + apply(subgoal_tac "\(m + length patha - 1 < m)") + prefer 2 apply arith + apply(simp add: nth_append) + apply(rule conjI) + apply(case_tac "m") + apply force + apply(case_tac "path") + apply force + apply force + apply clarify + apply(case_tac "Suc i\m") + apply(erule_tac x = "i" in allE) + apply simp + apply clarify + apply(rule_tac x = "j" in exI) + apply(case_tac "Suc i(length path\Suc m)" ) + prefer 2 apply arith + apply clarsimp + apply(erule_tac x = "i" in allE) + apply simp + apply clarify + apply(case_tac "R=j") + apply(force simp add: nth_list_update) + apply(force simp add: nth_list_update) +--{* the changed edge is not part of the path *} +apply(rule_tac x = "path" in exI) +apply simp +apply clarify +apply(erule_tac x = "i" in allE) +apply clarify +apply(case_tac "R=j") + apply(erule_tac x = "i" in allE) + apply simp +apply(force simp add: nth_list_update) +done + +subsubsection{* Graph 4 *} + +lemma Graph4: + "\T \ Reach E; Roots\Blacks M; I\length E; TiBtoW(E!i,M); RBlack\ \ + (\r. I\r \ r BtoW(E[R:=(fst(E!R),T)]!r,M))" +apply (unfold Reach_def) +apply simp +apply(erule disjE) + prefer 2 apply force +apply clarify +--{* there exist a black node in the path to T *} +apply(case_tac "\m T \ Reach E ; Roots \ Blacks M; \iBtoW(E!i,M); T Black\ + \ (\r. R r BtoW(E[R:=(fst(E!R),T)]!r,M))" +apply (unfold Reach_def) +apply simp +apply(erule disjE) + prefer 2 apply force +apply clarify +--{* there exist a black node in the path to T*} +apply(case_tac "\mR") + apply(drule le_imp_less_or_eq [of _ R]) + apply(erule disjE) + apply(erule allE , erule (1) notE impE) + apply force + apply force + apply(rule_tac x = "j" in exI) + apply(force simp add: nth_list_update) +apply simp +apply(rotate_tac -1) +apply(erule_tac x = "length path - 1" in allE) +apply(case_tac "length path") + apply force +apply force +done + +subsubsection {* Other lemmas about graphs *} + +lemma Graph6: + "\Proper_Edges(M,E); R \ Proper_Edges(M,E[R:=(fst(E!R),T)])" +apply (unfold Proper_Edges_def) + apply(force simp add: nth_list_update) +done + +lemma Graph7: + "\Proper_Edges(M,E)\ \ Proper_Edges(M[T:=a],E)" +apply (unfold Proper_Edges_def) +apply force +done + +lemma Graph8: + "\Proper_Roots(M)\ \ Proper_Roots(M[T:=a])" +apply (unfold Proper_Roots_def) +apply force +done + +text{* Some specific lemmata for the verification of garbage collection algorithms. *} + +lemma Graph9: "j Blacks M\Blacks (M[j := Black])" +apply (unfold Blacks_def) + apply(force simp add: nth_list_update) +done + +lemma Graph10 [rule_format (no_asm)]: "\i. M!i=a \M[i:=a]=M" +apply(induct_tac "M") +apply auto +apply(case_tac "i") +apply auto +done + +lemma Graph11 [rule_format (no_asm)]: + "\ M!j\Black;j \ Blacks M \ Blacks (M[j := Black])" +apply (unfold Blacks_def) +apply(rule psubsetI) + apply(force simp add: nth_list_update) +apply safe +apply(erule_tac c = "j" in equalityCE) +apply auto +done + +lemma Graph12: "\a\Blacks M;j \ a\Blacks (M[j := Black])" +apply (unfold Blacks_def) +apply(force simp add: nth_list_update) +done + +lemma Graph13: "\a\ Blacks M;j \ a \ Blacks (M[j := Black])" +apply (unfold Blacks_def) +apply(erule psubset_subset_trans) +apply(force simp add: nth_list_update) +done + +declare Graph_defs [simp del] + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/Hoare_Parallel.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/Hoare_Parallel.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,5 @@ +theory Hoare_Parallel +imports OG_Examples Gar_Coll Mul_Gar_Coll RG_Examples +begin + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/Mul_Gar_Coll.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,1283 @@ + +header {* \section{The Multi-Mutator Case} *} + +theory Mul_Gar_Coll imports Graph OG_Syntax begin + +text {* The full theory takes aprox. 18 minutes. *} + +record mut = + Z :: bool + R :: nat + T :: nat + +text {* Declaration of variables: *} + +record mul_gar_coll_state = + M :: nodes + E :: edges + bc :: "nat set" + obc :: "nat set" + Ma :: nodes + ind :: nat + k :: nat + q :: nat + l :: nat + Muts :: "mut list" + +subsection {* The Mutators *} + +constdefs + Mul_mut_init :: "mul_gar_coll_state \ nat \ bool" + "Mul_mut_init \ \ \n. n=length \Muts \ (\iMuts!i)E + \ T (\Muts!i)M) \" + + Mul_Redirect_Edge :: "nat \ nat \ mul_gar_coll_state ann_com" + "Mul_Redirect_Edge j n \ + .{\Mul_mut_init n \ Z (\Muts!j)}. + \IF T(\Muts!j) \ Reach \E THEN + \E:= \E[R (\Muts!j):= (fst (\E!R(\Muts!j)), T (\Muts!j))] FI,, + \Muts:= \Muts[j:= (\Muts!j) \Z:=False\]\" + + Mul_Color_Target :: "nat \ nat \ mul_gar_coll_state ann_com" + "Mul_Color_Target j n \ + .{\Mul_mut_init n \ \ Z (\Muts!j)}. + \\M:=\M[T (\Muts!j):=Black],, \Muts:=\Muts[j:= (\Muts!j) \Z:=True\]\" + + Mul_Mutator :: "nat \ nat \ mul_gar_coll_state ann_com" + "Mul_Mutator j n \ + .{\Mul_mut_init n \ Z (\Muts!j)}. + WHILE True + INV .{\Mul_mut_init n \ Z (\Muts!j)}. + DO Mul_Redirect_Edge j n ;; + Mul_Color_Target j n + OD" + +lemmas mul_mutator_defs = Mul_mut_init_def Mul_Redirect_Edge_def Mul_Color_Target_def + +subsubsection {* Correctness of the proof outline of one mutator *} + +lemma Mul_Redirect_Edge: "0\j \ j + \ Mul_Redirect_Edge j n + pre(Mul_Color_Target j n)" +apply (unfold mul_mutator_defs) +apply annhoare +apply(simp_all) +apply clarify +apply(simp add:nth_list_update) +done + +lemma Mul_Color_Target: "0\j \ j + \ Mul_Color_Target j n + .{\Mul_mut_init n \ Z (\Muts!j)}." +apply (unfold mul_mutator_defs) +apply annhoare +apply(simp_all) +apply clarify +apply(simp add:nth_list_update) +done + +lemma Mul_Mutator: "0\j \ j + \ Mul_Mutator j n .{False}." +apply(unfold Mul_Mutator_def) +apply annhoare +apply(simp_all add:Mul_Redirect_Edge Mul_Color_Target) +apply(simp add:mul_mutator_defs Mul_Redirect_Edge_def) +done + +subsubsection {* Interference freedom between mutators *} + +lemma Mul_interfree_Redirect_Edge_Redirect_Edge: + "\0\i; ij; jj\ \ + interfree_aux (Some (Mul_Redirect_Edge i n),{}, Some(Mul_Redirect_Edge j n))" +apply (unfold mul_mutator_defs) +apply interfree_aux +apply safe +apply(simp_all add: nth_list_update) +done + +lemma Mul_interfree_Redirect_Edge_Color_Target: + "\0\i; ij; jj\ \ + interfree_aux (Some(Mul_Redirect_Edge i n),{},Some(Mul_Color_Target j n))" +apply (unfold mul_mutator_defs) +apply interfree_aux +apply safe +apply(simp_all add: nth_list_update) +done + +lemma Mul_interfree_Color_Target_Redirect_Edge: + "\0\i; ij; jj\ \ + interfree_aux (Some(Mul_Color_Target i n),{},Some(Mul_Redirect_Edge j n))" +apply (unfold mul_mutator_defs) +apply interfree_aux +apply safe +apply(simp_all add:nth_list_update) +done + +lemma Mul_interfree_Color_Target_Color_Target: + " \0\i; ij; jj\ \ + interfree_aux (Some(Mul_Color_Target i n),{},Some(Mul_Color_Target j n))" +apply (unfold mul_mutator_defs) +apply interfree_aux +apply safe +apply(simp_all add: nth_list_update) +done + +lemmas mul_mutator_interfree = + Mul_interfree_Redirect_Edge_Redirect_Edge Mul_interfree_Redirect_Edge_Color_Target + Mul_interfree_Color_Target_Redirect_Edge Mul_interfree_Color_Target_Color_Target + +lemma Mul_interfree_Mutator_Mutator: "\i < n; j < n; i \ j\ \ + interfree_aux (Some (Mul_Mutator i n), {}, Some (Mul_Mutator j n))" +apply(unfold Mul_Mutator_def) +apply(interfree_aux) +apply(simp_all add:mul_mutator_interfree) +apply(simp_all add: mul_mutator_defs) +apply(tactic {* TRYALL (interfree_aux_tac) *}) +apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) +apply (simp_all add:nth_list_update) +done + +subsubsection {* Modular Parameterized Mutators *} + +lemma Mul_Parameterized_Mutators: "0 + \- .{\Mul_mut_init n \ (\iMuts!i))}. + COBEGIN + SCHEME [0\ j< n] + Mul_Mutator j n + .{False}. + COEND + .{False}." +apply oghoare +apply(force simp add:Mul_Mutator_def mul_mutator_defs nth_list_update) +apply(erule Mul_Mutator) +apply(simp add:Mul_interfree_Mutator_Mutator) +apply(force simp add:Mul_Mutator_def mul_mutator_defs nth_list_update) +done + +subsection {* The Collector *} + +constdefs + Queue :: "mul_gar_coll_state \ nat" + "Queue \ \ length (filter (\i. \ Z i \ \M!(T i) \ Black) \Muts) \" + +consts M_init :: nodes + +constdefs + Proper_M_init :: "mul_gar_coll_state \ bool" + "Proper_M_init \ \ Blacks M_init=Roots \ length M_init=length \M \" + + Mul_Proper :: "mul_gar_coll_state \ nat \ bool" + "Mul_Proper \ \ \n. Proper_Roots \M \ Proper_Edges (\M, \E) \ \Proper_M_init \ n=length \Muts \" + + Safe :: "mul_gar_coll_state \ bool" + "Safe \ \ Reach \E \ Blacks \M \" + +lemmas mul_collector_defs = Proper_M_init_def Mul_Proper_def Safe_def + +subsubsection {* Blackening Roots *} + +constdefs + Mul_Blacken_Roots :: "nat \ mul_gar_coll_state ann_com" + "Mul_Blacken_Roots n \ + .{\Mul_Proper n}. + \ind:=0;; + .{\Mul_Proper n \ \ind=0}. + WHILE \indM + INV .{\Mul_Proper n \ (\i<\ind. i\Roots \ \M!i=Black) \ \ind\length \M}. + DO .{\Mul_Proper n \ (\i<\ind. i\Roots \ \M!i=Black) \ \indM}. + IF \ind\Roots THEN + .{\Mul_Proper n \ (\i<\ind. i\Roots \ \M!i=Black) \ \indM \ \ind\Roots}. + \M:=\M[\ind:=Black] FI;; + .{\Mul_Proper n \ (\i<\ind+1. i\Roots \ \M!i=Black) \ \indM}. + \ind:=\ind+1 + OD" + +lemma Mul_Blacken_Roots: + "\ Mul_Blacken_Roots n + .{\Mul_Proper n \ Roots \ Blacks \M}." +apply (unfold Mul_Blacken_Roots_def) +apply annhoare +apply(simp_all add:mul_collector_defs Graph_defs) +apply safe +apply(simp_all add:nth_list_update) + apply (erule less_SucE) + apply simp+ + apply force +apply force +done + +subsubsection {* Propagating Black *} + +constdefs + Mul_PBInv :: "mul_gar_coll_state \ bool" + "Mul_PBInv \ \\Safe \ \obc\Blacks \M \ \l<\Queue + \ (\i<\ind. \BtoW(\E!i,\M)) \ \l\\Queue\" + + Mul_Auxk :: "mul_gar_coll_state \ bool" + "Mul_Auxk \ \\l<\Queue \ \M!\k\Black \ \BtoW(\E!\ind, \M) \ \obc\Blacks \M\" + +constdefs + Mul_Propagate_Black :: "nat \ mul_gar_coll_state ann_com" + "Mul_Propagate_Black n \ + .{\Mul_Proper n \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \l\\Queue \ \obc\Blacks \M)}. + \ind:=0;; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ Blacks \M\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \l\\Queue \ \obc\Blacks \M) \ \ind=0}. + WHILE \indE + INV .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ \Mul_PBInv \ \ind\length \E}. + DO .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ \Mul_PBInv \ \indE}. + IF \M!(fst (\E!\ind))=Black THEN + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ \Mul_PBInv \ (\M!fst(\E!\ind))=Black \ \indE}. + \k:=snd(\E!\ind);; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \obc\Blacks \M \ \l<\Queue \ (\i<\ind. \BtoW(\E!i,\M)) + \ \l\\Queue \ \Mul_Auxk ) \ \kM \ \M!fst(\E!\ind)=Black + \ \indE}. + \\M:=\M[\k:=Black],,\ind:=\ind+1\ + ELSE .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ \Mul_PBInv \ \indE}. + \IF \M!(fst (\E!\ind))\Black THEN \ind:=\ind+1 FI\ FI + OD" + +lemma Mul_Propagate_Black: + "\ Mul_Propagate_Black n + .{\Mul_Proper n \ Roots\Blacks \M \ \obc\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \obc\Blacks \M \ \l<\Queue \ (\l\\Queue \ \obc\Blacks \M))}." +apply(unfold Mul_Propagate_Black_def) +apply annhoare +apply(simp_all add:Mul_PBInv_def mul_collector_defs Mul_Auxk_def Graph6 Graph7 Graph8 Graph12 mul_collector_defs Queue_def) +--{* 8 subgoals left *} +apply force +apply force +apply force +apply(force simp add:BtoW_def Graph_defs) +--{* 4 subgoals left *} +apply clarify +apply(simp add: mul_collector_defs Graph12 Graph6 Graph7 Graph8) +apply(disjE_tac) + apply(simp_all add:Graph12 Graph13) + apply(case_tac "M x! k x=Black") + apply(simp add: Graph10) + apply(rule disjI2, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +apply(case_tac "M x! k x=Black") + apply(simp add: Graph10 BtoW_def) + apply(rule disjI2, clarify, erule less_SucE, force) + apply(case_tac "M x!snd(E x! ind x)=Black") + apply(force) + apply(force) +apply(rule disjI2, rule disjI1, erule subset_psubset_trans, erule Graph11, force) +--{* 3 subgoals left *} +apply force +--{* 2 subgoals left *} +apply clarify +apply(conjI_tac) +apply(disjE_tac) + apply (simp_all) +apply clarify +apply(erule less_SucE) + apply force +apply (simp add:BtoW_def) +--{* 1 subgoal left *} +apply clarify +apply simp +apply(disjE_tac) +apply (simp_all) +apply(rule disjI1 , rule Graph1) + apply simp_all +done + +subsubsection {* Counting Black Nodes *} + +constdefs + Mul_CountInv :: "mul_gar_coll_state \ nat \ bool" + "Mul_CountInv \ \ \ind. {i. i \Ma!i=Black}\\bc \" + + Mul_Count :: "nat \ mul_gar_coll_state ann_com" + "Mul_Count n \ + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M) ) + \ \q \bc={}}. + \ind:=0;; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M) ) + \ \q \bc={} \ \ind=0}. + WHILE \indM + INV .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \Mul_CountInv \ind + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q \ind\length \M}. + DO .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \Mul_CountInv \ind + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q \indM}. + IF \M!\ind=Black + THEN .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \Mul_CountInv \ind + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q \indM \ \M!\ind=Black}. + \bc:=insert \ind \bc + FI;; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ \Mul_CountInv (\ind+1) + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q \indM}. + \ind:=\ind+1 + OD" + +lemma Mul_Count: + "\ Mul_Count n + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ Blacks \Ma\\bc + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q edges \ edges" + +axioms + Append_to_free0: "length (Append_to_free (i, e)) = length e" + Append_to_free1: "Proper_Edges (m, e) + \ Proper_Edges (m, Append_to_free(i, e))" + Append_to_free2: "i \ Reach e + \ n \ Reach (Append_to_free(i, e)) = ( n = i \ n \ Reach e)" + +constdefs + Mul_AppendInv :: "mul_gar_coll_state \ nat \ bool" + "Mul_AppendInv \ \ \ind. (\i. ind\i \ iM \ i\Reach \E \ \M!i=Black)\" + + Mul_Append :: "nat \ mul_gar_coll_state ann_com" + "Mul_Append n \ + .{\Mul_Proper n \ Roots\Blacks \M \ \Safe}. + \ind:=0;; + .{\Mul_Proper n \ Roots\Blacks \M \ \Safe \ \ind=0}. + WHILE \indM + INV .{\Mul_Proper n \ \Mul_AppendInv \ind \ \ind\length \M}. + DO .{\Mul_Proper n \ \Mul_AppendInv \ind \ \indM}. + IF \M!\ind=Black THEN + .{\Mul_Proper n \ \Mul_AppendInv \ind \ \indM \ \M!\ind=Black}. + \M:=\M[\ind:=White] + ELSE + .{\Mul_Proper n \ \Mul_AppendInv \ind \ \indM \ \ind\Reach \E}. + \E:=Append_to_free(\ind,\E) + FI;; + .{\Mul_Proper n \ \Mul_AppendInv (\ind+1) \ \indM}. + \ind:=\ind+1 + OD" + +lemma Mul_Append: + "\ Mul_Append n + .{\Mul_Proper n}." +apply(unfold Mul_Append_def) +apply annhoare +apply(simp_all add: mul_collector_defs Mul_AppendInv_def + Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) +apply(force simp add:Blacks_def) +apply(force simp add:Blacks_def) +apply(force simp add:Blacks_def) +apply(force simp add:Graph_defs) +apply force +apply(force simp add:Append_to_free1 Append_to_free2) +apply force +apply force +done + +subsubsection {* Collector *} + +constdefs + Mul_Collector :: "nat \ mul_gar_coll_state ann_com" + "Mul_Collector n \ +.{\Mul_Proper n}. +WHILE True INV .{\Mul_Proper n}. +DO +Mul_Blacken_Roots n ;; +.{\Mul_Proper n \ Roots\Blacks \M}. + \obc:={};; +.{\Mul_Proper n \ Roots\Blacks \M \ \obc={}}. + \bc:=Roots;; +.{\Mul_Proper n \ Roots\Blacks \M \ \obc={} \ \bc=Roots}. + \l:=0;; +.{\Mul_Proper n \ Roots\Blacks \M \ \obc={} \ \bc=Roots \ \l=0}. + WHILE \lMul_Proper n \ Roots\Blacks \M \ \bc\Blacks \M \ + (\Safe \ (\l\\Queue \ \bc\Blacks \M) \ \lMul_Proper n \ Roots\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \l\\Queue \ \bc\Blacks \M)}. + \obc:=\bc;; + Mul_Propagate_Black n;; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \obc\Blacks \M \ \l<\Queue + \ (\l\\Queue \ \obc\Blacks \M))}. + \bc:={};; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \M \ \bc\Blacks \M + \ (\Safe \ \obc\Blacks \M \ \l<\Queue + \ (\l\\Queue \ \obc\Blacks \M)) \ \bc={}}. + \ \Ma:=\M,, \q:=\Queue \;; + Mul_Count n;; + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ Blacks \Ma\\bc + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \qobc=\bc THEN + .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ Blacks \Ma\\bc + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q \obc=\bc}. + \l:=\l+1 + ELSE .{\Mul_Proper n \ Roots\Blacks \M + \ \obc\Blacks \Ma \ Blacks \Ma\Blacks \M \ \bc\Blacks \M + \ length \Ma=length \M \ Blacks \Ma\\bc + \ (\Safe \ \obc\Blacks \Ma \ \l<\q \ (\q\\Queue \ \obc\Blacks \M)) + \ \q \obc\\bc}. + \l:=0 FI + OD;; + Mul_Append n +OD" + +lemmas mul_modules = Mul_Redirect_Edge_def Mul_Color_Target_def + Mul_Blacken_Roots_def Mul_Propagate_Black_def + Mul_Count_def Mul_Append_def + +lemma Mul_Collector: + "\ Mul_Collector n + .{False}." +apply(unfold Mul_Collector_def) +apply annhoare +apply(simp_all only:pre.simps Mul_Blacken_Roots + Mul_Propagate_Black Mul_Count Mul_Append) +apply(simp_all add:mul_modules) +apply(simp_all add:mul_collector_defs Queue_def) +apply force +apply force +apply force +apply (force simp add: less_Suc_eq_le) +apply force +apply (force dest:subset_antisym) +apply force +apply force +apply force +done + +subsection {* Interference Freedom *} + +lemma le_length_filter_update[rule_format]: + "\i. (\P (list!i) \ P j) \ i length(filter P list) \ length(filter P (list[i:=j]))" +apply(induct_tac "list") + apply(simp) +apply(clarify) +apply(case_tac i) + apply(simp) +apply(simp) +done + +lemma less_length_filter_update [rule_format]: + "\i. P j \ \(P (list!i)) \ i length(filter P list) < length(filter P (list[i:=j]))" +apply(induct_tac "list") + apply(simp) +apply(clarify) +apply(case_tac i) + apply(simp) +apply(simp) +done + +lemma Mul_interfree_Blacken_Roots_Redirect_Edge: "\0\j; j \ + interfree_aux (Some(Mul_Blacken_Roots n),{},Some(Mul_Redirect_Edge j n))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:Graph6 Graph9 Graph12 nth_list_update mul_mutator_defs mul_collector_defs) +done + +lemma Mul_interfree_Redirect_Edge_Blacken_Roots: "\0\j; j\ + interfree_aux (Some(Mul_Redirect_Edge j n ),{},Some (Mul_Blacken_Roots n))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Blacken_Roots_Color_Target: "\0\j; j\ + interfree_aux (Some(Mul_Blacken_Roots n),{},Some (Mul_Color_Target j n ))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs mul_collector_defs nth_list_update Graph7 Graph8 Graph9 Graph12) +done + +lemma Mul_interfree_Color_Target_Blacken_Roots: "\0\j; j\ + interfree_aux (Some(Mul_Color_Target j n ),{},Some (Mul_Blacken_Roots n ))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Propagate_Black_Redirect_Edge: "\0\j; j\ + interfree_aux (Some(Mul_Propagate_Black n),{},Some (Mul_Redirect_Edge j n ))" +apply (unfold mul_modules) +apply interfree_aux +apply(simp_all add:mul_mutator_defs mul_collector_defs Mul_PBInv_def nth_list_update Graph6) +--{* 7 subgoals left *} +apply clarify +apply(disjE_tac) + apply(simp_all add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +--{* 6 subgoals left *} +apply clarify +apply(disjE_tac) + apply(simp_all add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule impI,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +--{* 5 subgoals left *} +apply clarify +apply(disjE_tac) + apply(simp_all add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(erule conjE) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule conjI) + apply(rule impI,(rule disjI2)+,rule conjI) + apply clarify + apply(case_tac "R (Muts x! j)=i") + apply (force simp add: nth_list_update BtoW_def) + apply (force simp add: nth_list_update) + apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,(rule disjI2)+, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) + apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) +apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) +apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) +--{* 4 subgoals left *} +apply clarify +apply(disjE_tac) + apply(simp_all add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(erule conjE) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule conjI) + apply(rule impI,(rule disjI2)+,rule conjI) + apply clarify + apply(case_tac "R (Muts x! j)=i") + apply (force simp add: nth_list_update BtoW_def) + apply (force simp add: nth_list_update) + apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,(rule disjI2)+, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) + apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) +apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) +apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) +--{* 3 subgoals left *} +apply clarify +apply(disjE_tac) + apply(simp_all add:Graph6) + apply (rule impI) + apply(rule conjI) + apply(rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule conjI) + apply(rule impI) + apply(rule conjI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(rule impI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule conjI) + apply(rule impI) + apply(rule conjI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(rule impI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(erule conjE) + apply(rule conjI) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule impI,rule conjI,(rule disjI2)+,rule conjI) + apply clarify + apply(case_tac "R (Muts x! j)=i") + apply (force simp add: nth_list_update BtoW_def) + apply (force simp add: nth_list_update) + apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(rule impI,rule conjI) + apply(rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) + apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) + apply(case_tac "R (Muts x! j)=ind x") + apply (force simp add: nth_list_update) + apply (force simp add: nth_list_update) +apply(rule impI, (rule disjI2)+, erule le_trans) +apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +--{* 2 subgoals left *} +apply clarify +apply(rule conjI) + apply(disjE_tac) + apply(simp_all add:Mul_Auxk_def Graph6) + apply (rule impI) + apply(rule conjI) + apply(rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule impI) + apply(rule conjI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) + apply(rule impI) + apply(rule conjI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) +apply(rule impI) +apply(rule conjI) + apply(erule conjE)+ + apply(case_tac "M x!(T (Muts x!j))=Black") + apply((rule disjI2)+,rule conjI) + apply clarify + apply(case_tac "R (Muts x! j)=i") + apply (force simp add: nth_list_update BtoW_def) + apply (force simp add: nth_list_update) + apply(rule conjI) + apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update BtoW_def) + apply (simp add:nth_list_update) + apply(rule impI) + apply simp + apply(disjE_tac) + apply(rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply force + apply(rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) + apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) + apply(case_tac "R (Muts x ! j)= ind x") + apply(simp add:nth_list_update) + apply(simp add:nth_list_update) +apply(disjE_tac) +apply simp_all +apply(conjI_tac) + apply(rule impI) + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(erule conjE)+ +apply(rule impI,(rule disjI2)+,rule conjI) + apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule impI)+ +apply simp +apply(disjE_tac) + apply(rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply force +--{* 1 subgoal left *} +apply clarify +apply(disjE_tac) + apply(simp_all add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(erule conjE) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule conjI) + apply(rule impI,(rule disjI2)+,rule conjI) + apply clarify + apply(case_tac "R (Muts x! j)=i") + apply (force simp add: nth_list_update BtoW_def) + apply (force simp add: nth_list_update) + apply(erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,(rule disjI2)+, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) + apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) +apply(rule impI,rule disjI2,rule disjI2,rule disjI1, erule le_less_trans) +apply(force simp add:Queue_def less_Suc_eq_le less_length_filter_update) +done + +lemma Mul_interfree_Redirect_Edge_Propagate_Black: "\0\j; j\ + interfree_aux (Some(Mul_Redirect_Edge j n ),{},Some (Mul_Propagate_Black n))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Propagate_Black_Color_Target: "\0\j; j\ + interfree_aux (Some(Mul_Propagate_Black n),{},Some (Mul_Color_Target j n ))" +apply (unfold mul_modules) +apply interfree_aux +apply(simp_all add: mul_collector_defs mul_mutator_defs) +--{* 7 subgoals left *} +apply clarify +apply (simp add:Graph7 Graph8 Graph12) +apply(disjE_tac) + apply(simp add:Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,erule subset_psubset_trans, erule Graph11, simp) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +--{* 6 subgoals left *} +apply clarify +apply (simp add:Graph7 Graph8 Graph12) +apply(disjE_tac) + apply(simp add:Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,erule subset_psubset_trans, erule Graph11, simp) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +--{* 5 subgoals left *} +apply clarify +apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12) +apply(disjE_tac) + apply(simp add:Graph7 Graph8 Graph12) + apply(rule disjI2,rule disjI1, erule psubset_subset_trans,simp add:Graph9) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp) +apply(erule conjE) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply((rule disjI2)+) + apply (rule conjI) + apply(simp add:Graph10) + apply(erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) +apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp) +--{* 4 subgoals left *} +apply clarify +apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12) +apply(disjE_tac) + apply(simp add:Graph7 Graph8 Graph12) + apply(rule disjI2,rule disjI1, erule psubset_subset_trans,simp add:Graph9) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2,rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp) +apply(erule conjE) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply((rule disjI2)+) + apply (rule conjI) + apply(simp add:Graph10) + apply(erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) +apply(rule disjI2,rule disjI1,erule subset_psubset_trans, erule Graph11, simp) +--{* 3 subgoals left *} +apply clarify +apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply(simp add:Graph10) + apply(disjE_tac) + apply simp_all + apply(rule disjI2, rule disjI2, rule disjI1,erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply(erule conjE) + apply((rule disjI2)+,erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) +apply(rule conjI) + apply(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11) +apply (force simp add:nth_list_update) +--{* 2 subgoals left *} +apply clarify +apply(simp add:Mul_Auxk_def Graph7 Graph8 Graph12) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply(simp add:Graph10) + apply(disjE_tac) + apply simp_all + apply(rule disjI2, rule disjI2, rule disjI1,erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply(erule conjE)+ + apply((rule disjI2)+,rule conjI, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule impI)+) + apply simp + apply(erule disjE) + apply(rule disjI1, erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply force +apply(rule conjI) + apply(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11) +apply (force simp add:nth_list_update) +--{* 1 subgoal left *} +apply clarify +apply (simp add:mul_collector_defs Mul_PBInv_def Graph7 Graph8 Graph12) +apply(case_tac "M x!(T (Muts x!j))=Black") + apply(simp add:Graph10) + apply(disjE_tac) + apply simp_all + apply(rule disjI2, rule disjI2, rule disjI1,erule less_le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply(erule conjE) + apply((rule disjI2)+,erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) +apply(rule disjI2,rule disjI1, erule subset_psubset_trans,simp add:Graph11) +done + +lemma Mul_interfree_Color_Target_Propagate_Black: "\0\j; j\ + interfree_aux (Some(Mul_Color_Target j n),{},Some(Mul_Propagate_Black n ))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Count_Redirect_Edge: "\0\j; j\ + interfree_aux (Some(Mul_Count n ),{},Some(Mul_Redirect_Edge j n))" +apply (unfold mul_modules) +apply interfree_aux +--{* 9 subgoals left *} +apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(simp add:Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(simp add:Graph6) +--{* 8 subgoals left *} +apply(simp add:mul_mutator_defs nth_list_update) +--{* 7 subgoals left *} +apply(simp add:mul_mutator_defs mul_collector_defs) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(simp add:Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(simp add:Graph6) +--{* 6 subgoals left *} +apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def) +apply clarify +apply disjE_tac + apply(simp add:Graph6 Queue_def) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(simp add:Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(simp add:Graph6) +--{* 5 subgoals left *} +apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(simp add:Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(simp add:Graph6) +--{* 4 subgoals left *} +apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(simp add:Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(simp add:Graph6) +--{* 3 subgoals left *} +apply(simp add:mul_mutator_defs nth_list_update) +--{* 2 subgoals left *} +apply(simp add:mul_mutator_defs mul_collector_defs Mul_CountInv_def) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule impI,rule disjI1,rule subset_trans,erule Graph3,simp,simp) + apply(simp add:Graph6) +apply clarify +apply disjE_tac + apply(simp add:Graph6) + apply(rule conjI) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) + apply(rule impI,rule disjI2,rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(simp add:Graph6) +--{* 1 subgoal left *} +apply(simp add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Redirect_Edge_Count: "\0\j; j\ + interfree_aux (Some(Mul_Redirect_Edge j n),{},Some(Mul_Count n ))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Count_Color_Target: "\0\j; j\ + interfree_aux (Some(Mul_Count n ),{},Some(Mul_Color_Target j n))" +apply (unfold mul_modules) +apply interfree_aux +apply(simp_all add:mul_collector_defs mul_mutator_defs Mul_CountInv_def) +--{* 6 subgoals left *} +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply (simp add: Graph7 Graph8 Graph12) +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) +apply (simp add: Graph7 Graph8 Graph12) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +--{* 5 subgoals left *} +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply (simp add: Graph7 Graph8 Graph12) +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) +apply (simp add: Graph7 Graph8 Graph12) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +--{* 4 subgoals left *} +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply (simp add: Graph7 Graph8 Graph12) +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) +apply (simp add: Graph7 Graph8 Graph12) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +--{* 3 subgoals left *} +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply (simp add: Graph7 Graph8 Graph12) +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) +apply (simp add: Graph7 Graph8 Graph12) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +--{* 2 subgoals left *} +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12 nth_list_update) + apply (simp add: Graph7 Graph8 Graph12 nth_list_update) +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply(rule conjI) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) + apply (simp add: nth_list_update) +apply (simp add: Graph7 Graph8 Graph12) +apply(rule conjI) + apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +apply (simp add: nth_list_update) +--{* 1 subgoal left *} +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply (simp add: Graph7 Graph8 Graph12) +apply clarify +apply disjE_tac + apply (simp add: Graph7 Graph8 Graph12) + apply(case_tac "M x!(T (Muts x!j))=Black") + apply(rule disjI2,rule disjI2, rule disjI1, erule le_trans) + apply(force simp add:Queue_def less_Suc_eq_le le_length_filter_update Graph10) + apply((rule disjI2)+,(erule subset_psubset_trans)+, simp add: Graph11) +apply (simp add: Graph7 Graph8 Graph12) +apply((rule disjI2)+,erule psubset_subset_trans, simp add: Graph9) +done + +lemma Mul_interfree_Color_Target_Count: "\0\j; j\ + interfree_aux (Some(Mul_Color_Target j n),{}, Some(Mul_Count n ))" +apply (unfold mul_modules) +apply interfree_aux +apply safe +apply(simp_all add:mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Append_Redirect_Edge: "\0\j; j\ + interfree_aux (Some(Mul_Append n),{}, Some(Mul_Redirect_Edge j n))" +apply (unfold mul_modules) +apply interfree_aux +apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) +apply(simp_all add:Graph6 Append_to_free0 Append_to_free1 mul_collector_defs mul_mutator_defs Mul_AppendInv_def) +apply(erule_tac x=j in allE, force dest:Graph3)+ +done + +lemma Mul_interfree_Redirect_Edge_Append: "\0\j; j\ + interfree_aux (Some(Mul_Redirect_Edge j n),{},Some(Mul_Append n))" +apply (unfold mul_modules) +apply interfree_aux +apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) +apply(simp_all add:mul_collector_defs Append_to_free0 Mul_AppendInv_def mul_mutator_defs nth_list_update) +done + +lemma Mul_interfree_Append_Color_Target: "\0\j; j\ + interfree_aux (Some(Mul_Append n),{}, Some(Mul_Color_Target j n))" +apply (unfold mul_modules) +apply interfree_aux +apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) +apply(simp_all add:mul_mutator_defs mul_collector_defs Mul_AppendInv_def Graph7 Graph8 Append_to_free0 Append_to_free1 + Graph12 nth_list_update) +done + +lemma Mul_interfree_Color_Target_Append: "\0\j; j\ + interfree_aux (Some(Mul_Color_Target j n),{}, Some(Mul_Append n))" +apply (unfold mul_modules) +apply interfree_aux +apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) +apply(simp_all add: mul_mutator_defs nth_list_update) +apply(simp add:Mul_AppendInv_def Append_to_free0) +done + +subsubsection {* Interference freedom Collector-Mutator *} + +lemmas mul_collector_mutator_interfree = + Mul_interfree_Blacken_Roots_Redirect_Edge Mul_interfree_Blacken_Roots_Color_Target + Mul_interfree_Propagate_Black_Redirect_Edge Mul_interfree_Propagate_Black_Color_Target + Mul_interfree_Count_Redirect_Edge Mul_interfree_Count_Color_Target + Mul_interfree_Append_Redirect_Edge Mul_interfree_Append_Color_Target + Mul_interfree_Redirect_Edge_Blacken_Roots Mul_interfree_Color_Target_Blacken_Roots + Mul_interfree_Redirect_Edge_Propagate_Black Mul_interfree_Color_Target_Propagate_Black + Mul_interfree_Redirect_Edge_Count Mul_interfree_Color_Target_Count + Mul_interfree_Redirect_Edge_Append Mul_interfree_Color_Target_Append + +lemma Mul_interfree_Collector_Mutator: "j + interfree_aux (Some (Mul_Collector n), {}, Some (Mul_Mutator j n))" +apply(unfold Mul_Collector_def Mul_Mutator_def) +apply interfree_aux +apply(simp_all add:mul_collector_mutator_interfree) +apply(unfold mul_modules mul_collector_defs mul_mutator_defs) +apply(tactic {* TRYALL (interfree_aux_tac) *}) +--{* 42 subgoals left *} +apply (clarify,simp add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12)+ +--{* 24 subgoals left *} +apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) +--{* 14 subgoals left *} +apply(tactic {* TRYALL (clarify_tac @{claset}) *}) +apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) +apply(tactic {* TRYALL (rtac conjI) *}) +apply(tactic {* TRYALL (rtac impI) *}) +apply(tactic {* TRYALL (etac disjE) *}) +apply(tactic {* TRYALL (etac conjE) *}) +apply(tactic {* TRYALL (etac disjE) *}) +apply(tactic {* TRYALL (etac disjE) *}) +--{* 72 subgoals left *} +apply(simp_all add:Graph6 Graph7 Graph8 Append_to_free0 Append_to_free1 Graph12) +--{* 35 subgoals left *} +apply(tactic {* TRYALL(EVERY'[rtac disjI1,rtac subset_trans,etac @{thm Graph3},force_tac @{clasimpset}, assume_tac]) *}) +--{* 28 subgoals left *} +apply(tactic {* TRYALL (etac conjE) *}) +apply(tactic {* TRYALL (etac disjE) *}) +--{* 34 subgoals left *} +apply(rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(rule disjI2,rule disjI1,erule le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update) +apply(case_tac [!] "M x!(T (Muts x ! j))=Black") +apply(simp_all add:Graph10) +--{* 47 subgoals left *} +apply(tactic {* TRYALL(EVERY'[REPEAT o (rtac disjI2),etac (thm "subset_psubset_trans"),etac (thm "Graph11"),force_tac @{clasimpset}]) *}) +--{* 41 subgoals left *} +apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI1, etac @{thm le_trans}, force_tac (@{claset},@{simpset} addsimps [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])]) *}) +--{* 35 subgoals left *} +apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac (thm "psubset_subset_trans"),rtac (thm "Graph9"),force_tac @{clasimpset}]) *}) +--{* 31 subgoals left *} +apply(tactic {* TRYALL(EVERY'[rtac disjI2,rtac disjI1,etac (thm "subset_psubset_trans"),etac (thm "Graph11"),force_tac @{clasimpset}]) *}) +--{* 29 subgoals left *} +apply(tactic {* TRYALL(EVERY'[REPEAT o (rtac disjI2),etac (thm "subset_psubset_trans"),etac (thm "subset_psubset_trans"),etac (thm "Graph11"),force_tac @{clasimpset}]) *}) +--{* 25 subgoals left *} +apply(tactic {* TRYALL(EVERY'[rtac disjI2, rtac disjI2, rtac disjI1, etac @{thm le_trans}, force_tac (@{claset},@{simpset} addsimps [@{thm Queue_def}, @{thm less_Suc_eq_le}, @{thm le_length_filter_update}])]) *}) +--{* 10 subgoals left *} +apply(rule disjI2,rule disjI2,rule conjI,erule less_le_trans,force simp add:Queue_def less_Suc_eq_le le_length_filter_update, rule disjI1, rule less_imp_le, erule less_le_trans, force simp add:Queue_def less_Suc_eq_le le_length_filter_update)+ +done + +subsubsection {* Interference freedom Mutator-Collector *} + +lemma Mul_interfree_Mutator_Collector: " j < n \ + interfree_aux (Some (Mul_Mutator j n), {}, Some (Mul_Collector n))" +apply(unfold Mul_Collector_def Mul_Mutator_def) +apply interfree_aux +apply(simp_all add:mul_collector_mutator_interfree) +apply(unfold mul_modules mul_collector_defs mul_mutator_defs) +apply(tactic {* TRYALL (interfree_aux_tac) *}) +--{* 76 subgoals left *} +apply (clarify,simp add: nth_list_update)+ +--{* 56 subgoals left *} +apply(clarify,simp add:Mul_AppendInv_def Append_to_free0 nth_list_update)+ +done + +subsubsection {* The Multi-Mutator Garbage Collection Algorithm *} + +text {* The total number of verification conditions is 328 *} + +lemma Mul_Gar_Coll: + "\- .{\Mul_Proper n \ \Mul_mut_init n \ (\iMuts!i))}. + COBEGIN + Mul_Collector n + .{False}. + \ + SCHEME [0\ j< n] + Mul_Mutator j n + .{False}. + COEND + .{False}." +apply oghoare +--{* Strengthening the precondition *} +apply(rule Int_greatest) + apply (case_tac n) + apply(force simp add: Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append) + apply(simp add: Mul_Mutator_def mul_collector_defs mul_mutator_defs nth_append) + apply force +apply clarify +apply(case_tac i) + apply(simp add:Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append) +apply(simp add: Mul_Mutator_def mul_mutator_defs mul_collector_defs nth_append nth_map_upt) +--{* Collector *} +apply(rule Mul_Collector) +--{* Mutator *} +apply(erule Mul_Mutator) +--{* Interference freedom *} +apply(simp add:Mul_interfree_Collector_Mutator) +apply(simp add:Mul_interfree_Mutator_Collector) +apply(simp add:Mul_interfree_Mutator_Mutator) +--{* Weakening of the postcondition *} +apply(case_tac n) + apply(simp add:Mul_Collector_def mul_mutator_defs mul_collector_defs nth_append) +apply(simp add:Mul_Mutator_def mul_mutator_defs mul_collector_defs nth_append) +done + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/OG_Com.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/OG_Com.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,55 @@ + +header {* \chapter{The Owicki-Gries Method} + +\section{Abstract Syntax} *} + +theory OG_Com imports Main begin + +text {* Type abbreviations for boolean expressions and assertions: *} + +types + 'a bexp = "'a set" + 'a assn = "'a set" + +text {* The syntax of commands is defined by two mutually recursive +datatypes: @{text "'a ann_com"} for annotated commands and @{text "'a +com"} for non-annotated commands. *} + +datatype 'a ann_com = + AnnBasic "('a assn)" "('a \ 'a)" + | AnnSeq "('a ann_com)" "('a ann_com)" + | AnnCond1 "('a assn)" "('a bexp)" "('a ann_com)" "('a ann_com)" + | AnnCond2 "('a assn)" "('a bexp)" "('a ann_com)" + | AnnWhile "('a assn)" "('a bexp)" "('a assn)" "('a ann_com)" + | AnnAwait "('a assn)" "('a bexp)" "('a com)" +and 'a com = + Parallel "('a ann_com option \ 'a assn) list" + | Basic "('a \ 'a)" + | Seq "('a com)" "('a com)" + | Cond "('a bexp)" "('a com)" "('a com)" + | While "('a bexp)" "('a assn)" "('a com)" + +text {* The function @{text pre} extracts the precondition of an +annotated command: *} + +consts + pre ::"'a ann_com \ 'a assn" +primrec + "pre (AnnBasic r f) = r" + "pre (AnnSeq c1 c2) = pre c1" + "pre (AnnCond1 r b c1 c2) = r" + "pre (AnnCond2 r b c) = r" + "pre (AnnWhile r b i c) = r" + "pre (AnnAwait r b c) = r" + +text {* Well-formedness predicate for atomic programs: *} + +consts atom_com :: "'a com \ bool" +primrec + "atom_com (Parallel Ts) = False" + "atom_com (Basic f) = True" + "atom_com (Seq c1 c2) = (atom_com c1 \ atom_com c2)" + "atom_com (Cond b c1 c2) = (atom_com c1 \ atom_com c2)" + "atom_com (While b i c) = atom_com c" + +end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/OG_Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/OG_Examples.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,549 @@ + +header {* \section{Examples} *} + +theory OG_Examples imports OG_Syntax begin + +subsection {* Mutual Exclusion *} + +subsubsection {* Peterson's Algorithm I*} + +text {* Eike Best. "Semantics of Sequential and Parallel Programs", page 217. *} + +record Petersons_mutex_1 = + pr1 :: nat + pr2 :: nat + in1 :: bool + in2 :: bool + hold :: nat + +lemma Petersons_mutex_1: + "\- .{\pr1=0 \ \\in1 \ \pr2=0 \ \\in2 }. + COBEGIN .{\pr1=0 \ \\in1}. + WHILE True INV .{\pr1=0 \ \\in1}. + DO + .{\pr1=0 \ \\in1}. \ \in1:=True,,\pr1:=1 \;; + .{\pr1=1 \ \in1}. \ \hold:=1,,\pr1:=2 \;; + .{\pr1=2 \ \in1 \ (\hold=1 \ \hold=2 \ \pr2=2)}. + AWAIT (\\in2 \ \(\hold=1)) THEN \pr1:=3 END;; + .{\pr1=3 \ \in1 \ (\hold=1 \ \hold=2 \ \pr2=2)}. + \\in1:=False,,\pr1:=0\ + OD .{\pr1=0 \ \\in1}. + \ + .{\pr2=0 \ \\in2}. + WHILE True INV .{\pr2=0 \ \\in2}. + DO + .{\pr2=0 \ \\in2}. \ \in2:=True,,\pr2:=1 \;; + .{\pr2=1 \ \in2}. \ \hold:=2,,\pr2:=2 \;; + .{\pr2=2 \ \in2 \ (\hold=2 \ (\hold=1 \ \pr1=2))}. + AWAIT (\\in1 \ \(\hold=2)) THEN \pr2:=3 END;; + .{\pr2=3 \ \in2 \ (\hold=2 \ (\hold=1 \ \pr1=2))}. + \\in2:=False,,\pr2:=0\ + OD .{\pr2=0 \ \\in2}. + COEND + .{\pr1=0 \ \\in1 \ \pr2=0 \ \\in2}." +apply oghoare +--{* 104 verification conditions. *} +apply auto +done + +subsubsection {*Peterson's Algorithm II: A Busy Wait Solution *} + +text {* Apt and Olderog. "Verification of sequential and concurrent Programs", page 282. *} + +record Busy_wait_mutex = + flag1 :: bool + flag2 :: bool + turn :: nat + after1 :: bool + after2 :: bool + +lemma Busy_wait_mutex: + "\- .{True}. + \flag1:=False,, \flag2:=False,, + COBEGIN .{\\flag1}. + WHILE True + INV .{\\flag1}. + DO .{\\flag1}. \ \flag1:=True,,\after1:=False \;; + .{\flag1 \ \\after1}. \ \turn:=1,,\after1:=True \;; + .{\flag1 \ \after1 \ (\turn=1 \ \turn=2)}. + WHILE \(\flag2 \ \turn=2) + INV .{\flag1 \ \after1 \ (\turn=1 \ \turn=2)}. + DO .{\flag1 \ \after1 \ (\turn=1 \ \turn=2)}. SKIP OD;; + .{\flag1 \ \after1 \ (\flag2 \ \after2 \ \turn=2)}. + \flag1:=False + OD + .{False}. + \ + .{\\flag2}. + WHILE True + INV .{\\flag2}. + DO .{\\flag2}. \ \flag2:=True,,\after2:=False \;; + .{\flag2 \ \\after2}. \ \turn:=2,,\after2:=True \;; + .{\flag2 \ \after2 \ (\turn=1 \ \turn=2)}. + WHILE \(\flag1 \ \turn=1) + INV .{\flag2 \ \after2 \ (\turn=1 \ \turn=2)}. + DO .{\flag2 \ \after2 \ (\turn=1 \ \turn=2)}. SKIP OD;; + .{\flag2 \ \after2 \ (\flag1 \ \after1 \ \turn=1)}. + \flag2:=False + OD + .{False}. + COEND + .{False}." +apply oghoare +--{* 122 vc *} +apply auto +done + +subsubsection {* Peterson's Algorithm III: A Solution using Semaphores *} + +record Semaphores_mutex = + out :: bool + who :: nat + +lemma Semaphores_mutex: + "\- .{i\j}. + \out:=True ,, + COBEGIN .{i\j}. + WHILE True INV .{i\j}. + DO .{i\j}. AWAIT \out THEN \out:=False,, \who:=i END;; + .{\\out \ \who=i \ i\j}. \out:=True OD + .{False}. + \ + .{i\j}. + WHILE True INV .{i\j}. + DO .{i\j}. AWAIT \out THEN \out:=False,,\who:=j END;; + .{\\out \ \who=j \ i\j}. \out:=True OD + .{False}. + COEND + .{False}." +apply oghoare +--{* 38 vc *} +apply auto +done + +subsubsection {* Peterson's Algorithm III: Parameterized version: *} + +lemma Semaphores_parameterized_mutex: + "0 \- .{True}. + \out:=True ,, + COBEGIN + SCHEME [0\ i< n] + .{True}. + WHILE True INV .{True}. + DO .{True}. AWAIT \out THEN \out:=False,, \who:=i END;; + .{\\out \ \who=i}. \out:=True OD + .{False}. + COEND + .{False}." +apply oghoare +--{* 20 vc *} +apply auto +done + +subsubsection{* The Ticket Algorithm *} + +record Ticket_mutex = + num :: nat + nextv :: nat + turn :: "nat list" + index :: nat + +lemma Ticket_mutex: + "\ 0n=length \turn \ 0<\nextv \ (\k l. k l k\l + \ \turn!k < \num \ (\turn!k =0 \ \turn!k\\turn!l))\ \ + \ \- .{n=length \turn}. + \index:= 0,, + WHILE \index < n INV .{n=length \turn \ (\i<\index. \turn!i=0)}. + DO \turn:= \turn[\index:=0],, \index:=\index +1 OD,, + \num:=1 ,, \nextv:=1 ,, + COBEGIN + SCHEME [0\ i< n] + .{\I}. + WHILE True INV .{\I}. + DO .{\I}. \ \turn :=\turn[i:=\num],, \num:=\num+1 \;; + .{\I}. WAIT \turn!i=\nextv END;; + .{\I \ \turn!i=\nextv}. \nextv:=\nextv+1 + OD + .{False}. + COEND + .{False}." +apply oghoare +--{* 35 vc *} +apply simp_all +--{* 21 vc *} +apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) +--{* 11 vc *} +apply simp_all +apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) +--{* 10 subgoals left *} +apply(erule less_SucE) + apply simp +apply simp +--{* 9 subgoals left *} +apply(case_tac "i=k") + apply force +apply simp +apply(case_tac "i=l") + apply force +apply force +--{* 8 subgoals left *} +prefer 8 +apply force +apply force +--{* 6 subgoals left *} +prefer 6 +apply(erule_tac x=i in allE) +apply fastsimp +--{* 5 subgoals left *} +prefer 5 +apply(case_tac [!] "j=k") +--{* 10 subgoals left *} +apply simp_all +apply(erule_tac x=k in allE) +apply force +--{* 9 subgoals left *} +apply(case_tac "j=l") + apply simp + apply(erule_tac x=k in allE) + apply(erule_tac x=k in allE) + apply(erule_tac x=l in allE) + apply force +apply(erule_tac x=k in allE) +apply(erule_tac x=k in allE) +apply(erule_tac x=l in allE) +apply force +--{* 8 subgoals left *} +apply force +apply(case_tac "j=l") + apply simp +apply(erule_tac x=k in allE) +apply(erule_tac x=l in allE) +apply force +apply force +apply force +--{* 5 subgoals left *} +apply(erule_tac x=k in allE) +apply(erule_tac x=l in allE) +apply(case_tac "j=l") + apply force +apply force +apply force +--{* 3 subgoals left *} +apply(erule_tac x=k in allE) +apply(erule_tac x=l in allE) +apply(case_tac "j=l") + apply force +apply force +apply force +--{* 1 subgoals left *} +apply(erule_tac x=k in allE) +apply(erule_tac x=l in allE) +apply(case_tac "j=l") + apply force +apply force +done + +subsection{* Parallel Zero Search *} + +text {* Synchronized Zero Search. Zero-6 *} + +text {*Apt and Olderog. "Verification of sequential and concurrent Programs" page 294: *} + +record Zero_search = + turn :: nat + found :: bool + x :: nat + y :: nat + +lemma Zero_search: + "\I1= \ a\\x \ (\found \ (a<\x \ f(\x)=0) \ (\y\a \ f(\y)=0)) + \ (\\found \ a<\ x \ f(\x)\0) \ ; + I2= \\y\a+1 \ (\found \ (a<\x \ f(\x)=0) \ (\y\a \ f(\y)=0)) + \ (\\found \ \y\a \ f(\y)\0) \ \ \ + \- .{\ u. f(u)=0}. + \turn:=1,, \found:= False,, + \x:=a,, \y:=a+1 ,, + COBEGIN .{\I1}. + WHILE \\found + INV .{\I1}. + DO .{a\\x \ (\found \ \y\a \ f(\y)=0) \ (a<\x \ f(\x)\0)}. + WAIT \turn=1 END;; + .{a\\x \ (\found \ \y\a \ f(\y)=0) \ (a<\x \ f(\x)\0)}. + \turn:=2;; + .{a\\x \ (\found \ \y\a \ f(\y)=0) \ (a<\x \ f(\x)\0)}. + \ \x:=\x+1,, + IF f(\x)=0 THEN \found:=True ELSE SKIP FI\ + OD;; + .{\I1 \ \found}. + \turn:=2 + .{\I1 \ \found}. + \ + .{\I2}. + WHILE \\found + INV .{\I2}. + DO .{\y\a+1 \ (\found \ a<\x \ f(\x)=0) \ (\y\a \ f(\y)\0)}. + WAIT \turn=2 END;; + .{\y\a+1 \ (\found \ a<\x \ f(\x)=0) \ (\y\a \ f(\y)\0)}. + \turn:=1;; + .{\y\a+1 \ (\found \ a<\x \ f(\x)=0) \ (\y\a \ f(\y)\0)}. + \ \y:=(\y - 1),, + IF f(\y)=0 THEN \found:=True ELSE SKIP FI\ + OD;; + .{\I2 \ \found}. + \turn:=1 + .{\I2 \ \found}. + COEND + .{f(\x)=0 \ f(\y)=0}." +apply oghoare +--{* 98 verification conditions *} +apply auto +--{* auto takes about 3 minutes !! *} +done + +text {* Easier Version: without AWAIT. Apt and Olderog. page 256: *} + +lemma Zero_Search_2: +"\I1=\ a\\x \ (\found \ (a<\x \ f(\x)=0) \ (\y\a \ f(\y)=0)) + \ (\\found \ a<\x \ f(\x)\0)\; + I2= \\y\a+1 \ (\found \ (a<\x \ f(\x)=0) \ (\y\a \ f(\y)=0)) + \ (\\found \ \y\a \ f(\y)\0)\\ \ + \- .{\u. f(u)=0}. + \found:= False,, + \x:=a,, \y:=a+1,, + COBEGIN .{\I1}. + WHILE \\found + INV .{\I1}. + DO .{a\\x \ (\found \ \y\a \ f(\y)=0) \ (a<\x \ f(\x)\0)}. + \ \x:=\x+1,,IF f(\x)=0 THEN \found:=True ELSE SKIP FI\ + OD + .{\I1 \ \found}. + \ + .{\I2}. + WHILE \\found + INV .{\I2}. + DO .{\y\a+1 \ (\found \ a<\x \ f(\x)=0) \ (\y\a \ f(\y)\0)}. + \ \y:=(\y - 1),,IF f(\y)=0 THEN \found:=True ELSE SKIP FI\ + OD + .{\I2 \ \found}. + COEND + .{f(\x)=0 \ f(\y)=0}." +apply oghoare +--{* 20 vc *} +apply auto +--{* auto takes aprox. 2 minutes. *} +done + +subsection {* Producer/Consumer *} + +subsubsection {* Previous lemmas *} + +lemma nat_lemma2: "\ b = m*(n::nat) + t; a = s*n + u; t=u; b-a < n \ \ m \ s" +proof - + assume "b = m*(n::nat) + t" "a = s*n + u" "t=u" + hence "(m - s) * n = b - a" by (simp add: diff_mult_distrib) + also assume "\ < n" + finally have "m - s < 1" by simp + thus ?thesis by arith +qed + +lemma mod_lemma: "\ (c::nat) \ a; a < b; b - c < n \ \ b mod n \ a mod n" +apply(subgoal_tac "b=b div n*n + b mod n" ) + prefer 2 apply (simp add: mod_div_equality [symmetric]) +apply(subgoal_tac "a=a div n*n + a mod n") + prefer 2 + apply(simp add: mod_div_equality [symmetric]) +apply(subgoal_tac "b - a \ b - c") + prefer 2 apply arith +apply(drule le_less_trans) +back + apply assumption +apply(frule less_not_refl2) +apply(drule less_imp_le) +apply (drule_tac m = "a" and k = n in div_le_mono) +apply(safe) +apply(frule_tac b = "b" and a = "a" and n = "n" in nat_lemma2, assumption, assumption) +apply assumption +apply(drule order_antisym, assumption) +apply(rotate_tac -3) +apply(simp) +done + + +subsubsection {* Producer/Consumer Algorithm *} + +record Producer_consumer = + ins :: nat + outs :: nat + li :: nat + lj :: nat + vx :: nat + vy :: nat + buffer :: "nat list" + b :: "nat list" + +text {* The whole proof takes aprox. 4 minutes. *} + +lemma Producer_consumer: + "\INIT= \0 0buffer \ length \b=length a\ ; + I= \(\k<\ins. \outs\k \ (a ! k) = \buffer ! (k mod (length \buffer))) \ + \outs\\ins \ \ins-\outs\length \buffer\ ; + I1= \\I \ \li\length a\ ; + p1= \\I1 \ \li=\ins\ ; + I2 = \\I \ (\k<\lj. (a ! k)=(\b ! k)) \ \lj\length a\ ; + p2 = \\I2 \ \lj=\outs\ \ \ + \- .{\INIT}. + \ins:=0,, \outs:=0,, \li:=0,, \lj:=0,, + COBEGIN .{\p1 \ \INIT}. + WHILE \li p1 \ \INIT}. + DO .{\p1 \ \INIT \ \livx:= (a ! \li);; + .{\p1 \ \INIT \ \li \vx=(a ! \li)}. + WAIT \ins-\outs < length \buffer END;; + .{\p1 \ \INIT \ \li \vx=(a ! \li) + \ \ins-\outs < length \buffer}. + \buffer:=(list_update \buffer (\ins mod (length \buffer)) \vx);; + .{\p1 \ \INIT \ \li (a ! \li)=(\buffer ! (\ins mod (length \buffer))) + \ \ins-\outs buffer}. + \ins:=\ins+1;; + .{\I1 \ \INIT \ (\li+1)=\ins \ \lili:=\li+1 + OD + .{\p1 \ \INIT \ \li=length a}. + \ + .{\p2 \ \INIT}. + WHILE \lj < length a + INV .{\p2 \ \INIT}. + DO .{\p2 \ \lj \INIT}. + WAIT \outs<\ins END;; + .{\p2 \ \lj \outs<\ins \ \INIT}. + \vy:=(\buffer ! (\outs mod (length \buffer)));; + .{\p2 \ \lj \outs<\ins \ \vy=(a ! \lj) \ \INIT}. + \outs:=\outs+1;; + .{\I2 \ (\lj+1)=\outs \ \lj \vy=(a ! \lj) \ \INIT}. + \b:=(list_update \b \lj \vy);; + .{\I2 \ (\lj+1)=\outs \ \lj (a ! \lj)=(\b ! \lj) \ \INIT}. + \lj:=\lj+1 + OD + .{\p2 \ \lj=length a \ \INIT}. + COEND + .{ \kb ! k)}." +apply oghoare +--{* 138 vc *} +apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) +--{* 112 subgoals left *} +apply(simp_all (no_asm)) +apply(tactic {*ALLGOALS (conjI_Tac (K all_tac)) *}) +--{* 930 subgoals left *} +apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) +apply(simp_all (asm_lr) only:length_0_conv [THEN sym]) +--{* 44 subgoals left *} +apply (simp_all (asm_lr) del:length_0_conv add: neq0_conv nth_list_update mod_less_divisor mod_lemma) +--{* 32 subgoals left *} +apply(tactic {* ALLGOALS (clarify_tac @{claset}) *}) + +apply(tactic {* TRYALL (Lin_Arith.tac @{context}) *}) +--{* 9 subgoals left *} +apply (force simp add:less_Suc_eq) +apply(drule sym) +apply (force simp add:less_Suc_eq)+ +done + +subsection {* Parameterized Examples *} + +subsubsection {* Set Elements of an Array to Zero *} + +record Example1 = + a :: "nat \ nat" + +lemma Example1: + "\- .{True}. + COBEGIN SCHEME [0\ia:=\a (i:=0) .{\a i=0}. COEND + .{\i < n. \a i = 0}." +apply oghoare +apply simp_all +done + +text {* Same example with lists as auxiliary variables. *} +record Example1_list = + A :: "nat list" +lemma Example1_list: + "\- .{n < length \A}. + COBEGIN + SCHEME [0\iA}. \A:=\A[i:=0] .{\A!i=0}. + COEND + .{\i < n. \A!i = 0}." +apply oghoare +apply force+ +done + +subsubsection {* Increment a Variable in Parallel *} + +text {* First some lemmas about summation properties. *} +(* +lemma Example2_lemma1: "!!b. j (\i::nat b j = 0 " +apply(induct n) + apply simp_all +apply(force simp add: less_Suc_eq) +done +*) +lemma Example2_lemma2_aux: "!!b. j + (\i=0..i=0..i=0.. s \ (\i::nat=0..i=0..j \ Suc (\i::nat=0..i=0..i=0..i=0..j") + apply(drule_tac b="b" and t="(Suc 0)" in Example2_lemma2_aux2) +apply(rotate_tac -1) +apply(erule ssubst) +apply simp_all +done + + +record Example2 = + c :: "nat \ nat" + x :: nat + +lemma Example_2: "0 + \- .{\x=0 \ (\i=0..c i)=0}. + COBEGIN + SCHEME [0\ix=(\i=0..c i) \ \c i=0}. + \ \x:=\x+(Suc 0),, \c:=\c (i:=(Suc 0)) \ + .{\x=(\i=0..c i) \ \c i=(Suc 0)}. + COEND + .{\x=n}." +apply oghoare +apply (simp_all cong del: strong_setsum_cong) +apply (tactic {* ALLGOALS (clarify_tac @{claset}) *}) +apply (simp_all cong del: strong_setsum_cong) + apply(erule (1) Example2_lemma2) + apply(erule (1) Example2_lemma2) + apply(erule (1) Example2_lemma2) +apply(simp) +done + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/OG_Hoare.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/OG_Hoare.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,469 @@ + +header {* \section{The Proof System} *} + +theory OG_Hoare imports OG_Tran begin + +consts assertions :: "'a ann_com \ ('a assn) set" +primrec + "assertions (AnnBasic r f) = {r}" + "assertions (AnnSeq c1 c2) = assertions c1 \ assertions c2" + "assertions (AnnCond1 r b c1 c2) = {r} \ assertions c1 \ assertions c2" + "assertions (AnnCond2 r b c) = {r} \ assertions c" + "assertions (AnnWhile r b i c) = {r, i} \ assertions c" + "assertions (AnnAwait r b c) = {r}" + +consts atomics :: "'a ann_com \ ('a assn \ 'a com) set" +primrec + "atomics (AnnBasic r f) = {(r, Basic f)}" + "atomics (AnnSeq c1 c2) = atomics c1 \ atomics c2" + "atomics (AnnCond1 r b c1 c2) = atomics c1 \ atomics c2" + "atomics (AnnCond2 r b c) = atomics c" + "atomics (AnnWhile r b i c) = atomics c" + "atomics (AnnAwait r b c) = {(r \ b, c)}" + +consts com :: "'a ann_triple_op \ 'a ann_com_op" +primrec "com (c, q) = c" + +consts post :: "'a ann_triple_op \ 'a assn" +primrec "post (c, q) = q" + +constdefs interfree_aux :: "('a ann_com_op \ 'a assn \ 'a ann_com_op) \ bool" + "interfree_aux \ \(co, q, co'). co'= None \ + (\(r,a) \ atomics (the co'). \= (q \ r) a q \ + (co = None \ (\p \ assertions (the co). \= (p \ r) a p)))" + +constdefs interfree :: "(('a ann_triple_op) list) \ bool" + "interfree Ts \ \i j. i < length Ts \ j < length Ts \ i \ j \ + interfree_aux (com (Ts!i), post (Ts!i), com (Ts!j)) " + +inductive + oghoare :: "'a assn \ 'a com \ 'a assn \ bool" ("(3\- _//_//_)" [90,55,90] 50) + and ann_hoare :: "'a ann_com \ 'a assn \ bool" ("(2\ _// _)" [60,90] 45) +where + AnnBasic: "r \ {s. f s \ q} \ \ (AnnBasic r f) q" + +| AnnSeq: "\ \ c0 pre c1; \ c1 q \ \ \ (AnnSeq c0 c1) q" + +| AnnCond1: "\ r \ b \ pre c1; \ c1 q; r \ -b \ pre c2; \ c2 q\ + \ \ (AnnCond1 r b c1 c2) q" +| AnnCond2: "\ r \ b \ pre c; \ c q; r \ -b \ q \ \ \ (AnnCond2 r b c) q" + +| AnnWhile: "\ r \ i; i \ b \ pre c; \ c i; i \ -b \ q \ + \ \ (AnnWhile r b i c) q" + +| AnnAwait: "\ atom_com c; \- (r \ b) c q \ \ \ (AnnAwait r b c) q" + +| AnnConseq: "\\ c q; q \ q' \ \ \ c q'" + + +| Parallel: "\ \ic q. Ts!i = (Some c, q) \ \ c q; interfree Ts \ + \ \- (\i\{i. ii\{i. i- {s. f s \q} (Basic f) q" + +| Seq: "\ \- p c1 r; \- r c2 q \ \ \- p (Seq c1 c2) q " + +| Cond: "\ \- (p \ b) c1 q; \- (p \ -b) c2 q \ \ \- p (Cond b c1 c2) q" + +| While: "\ \- (p \ b) c p \ \ \- p (While b i c) (p \ -b)" + +| Conseq: "\ p' \ p; \- p c q ; q \ q' \ \ \- p' c q'" + +section {* Soundness *} +(* In the version Isabelle-10-Sep-1999: HOL: The THEN and ELSE +parts of conditional expressions (if P then x else y) are no longer +simplified. (This allows the simplifier to unfold recursive +functional programs.) To restore the old behaviour, we declare +@{text "lemmas [cong del] = if_weak_cong"}. *) + +lemmas [cong del] = if_weak_cong + +lemmas ann_hoare_induct = oghoare_ann_hoare.induct [THEN conjunct2] +lemmas oghoare_induct = oghoare_ann_hoare.induct [THEN conjunct1] + +lemmas AnnBasic = oghoare_ann_hoare.AnnBasic +lemmas AnnSeq = oghoare_ann_hoare.AnnSeq +lemmas AnnCond1 = oghoare_ann_hoare.AnnCond1 +lemmas AnnCond2 = oghoare_ann_hoare.AnnCond2 +lemmas AnnWhile = oghoare_ann_hoare.AnnWhile +lemmas AnnAwait = oghoare_ann_hoare.AnnAwait +lemmas AnnConseq = oghoare_ann_hoare.AnnConseq + +lemmas Parallel = oghoare_ann_hoare.Parallel +lemmas Basic = oghoare_ann_hoare.Basic +lemmas Seq = oghoare_ann_hoare.Seq +lemmas Cond = oghoare_ann_hoare.Cond +lemmas While = oghoare_ann_hoare.While +lemmas Conseq = oghoare_ann_hoare.Conseq + +subsection {* Soundness of the System for Atomic Programs *} + +lemma Basic_ntran [rule_format]: + "(Basic f, s) -Pn\ (Parallel Ts, t) \ All_None Ts \ t = f s" +apply(induct "n") + apply(simp (no_asm)) +apply(fast dest: rel_pow_Suc_D2 Parallel_empty_lemma elim: transition_cases) +done + +lemma SEM_fwhile: "SEM S (p \ b) \ p \ SEM (fwhile b S k) p \ (p \ -b)" +apply (induct "k") + apply(simp (no_asm) add: L3_5v_lemma3) +apply(simp (no_asm) add: L3_5iv L3_5ii Parallel_empty) +apply(rule conjI) + apply (blast dest: L3_5i) +apply(simp add: SEM_def sem_def id_def) +apply (blast dest: Basic_ntran rtrancl_imp_UN_rel_pow) +done + +lemma atom_hoare_sound [rule_format]: + " \- p c q \ atom_com(c) \ \= p c q" +apply (unfold com_validity_def) +apply(rule oghoare_induct) +apply simp_all +--{*Basic*} + apply(simp add: SEM_def sem_def) + apply(fast dest: rtrancl_imp_UN_rel_pow Basic_ntran) +--{* Seq *} + apply(rule impI) + apply(rule subset_trans) + prefer 2 apply simp + apply(simp add: L3_5ii L3_5i) +--{* Cond *} + apply(simp add: L3_5iv) +--{* While *} + apply (force simp add: L3_5v dest: SEM_fwhile) +--{* Conseq *} +apply(force simp add: SEM_def sem_def) +done + +subsection {* Soundness of the System for Component Programs *} + +inductive_cases ann_transition_cases: + "(None,s) -1\ (c', s')" + "(Some (AnnBasic r f),s) -1\ (c', s')" + "(Some (AnnSeq c1 c2), s) -1\ (c', s')" + "(Some (AnnCond1 r b c1 c2), s) -1\ (c', s')" + "(Some (AnnCond2 r b c), s) -1\ (c', s')" + "(Some (AnnWhile r b I c), s) -1\ (c', s')" + "(Some (AnnAwait r b c),s) -1\ (c', s')" + +text {* Strong Soundness for Component Programs:*} + +lemma ann_hoare_case_analysis [rule_format]: + defines I: "I \ \C q'. + ((\r f. C = AnnBasic r f \ (\q. r \ {s. f s \ q} \ q \ q')) \ + (\c0 c1. C = AnnSeq c0 c1 \ (\q. q \ q' \ \ c0 pre c1 \ \ c1 q)) \ + (\r b c1 c2. C = AnnCond1 r b c1 c2 \ (\q. q \ q' \ + r \ b \ pre c1 \ \ c1 q \ r \ -b \ pre c2 \ \ c2 q)) \ + (\r b c. C = AnnCond2 r b c \ + (\q. q \ q' \ r \ b \ pre c \ \ c q \ r \ -b \ q)) \ + (\r i b c. C = AnnWhile r b i c \ + (\q. q \ q' \ r \ i \ i \ b \ pre c \ \ c i \ i \ -b \ q)) \ + (\r b c. C = AnnAwait r b c \ (\q. q \ q' \ \- (r \ b) c q)))" + shows "\ C q' \ I C q'" +apply(rule ann_hoare_induct) +apply (simp_all add: I) + apply(rule_tac x=q in exI,simp)+ +apply(rule conjI,clarify,simp,clarify,rule_tac x=qa in exI,fast)+ +apply(clarify,simp,clarify,rule_tac x=qa in exI,fast) +done + +lemma Help: "(transition \ {(x,y). True}) = (transition)" +apply force +done + +lemma Strong_Soundness_aux_aux [rule_format]: + "(co, s) -1\ (co', t) \ (\c. co = Some c \ s\ pre c \ + (\q. \ c q \ (if co' = None then t\q else t \ pre(the co') \ \ (the co') q )))" +apply(rule ann_transition_transition.induct [THEN conjunct1]) +apply simp_all +--{* Basic *} + apply clarify + apply(frule ann_hoare_case_analysis) + apply force +--{* Seq *} + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply(fast intro: AnnConseq) + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply clarify + apply(rule conjI) + apply force + apply(rule AnnSeq,simp) + apply(fast intro: AnnConseq) +--{* Cond1 *} + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply(fast intro: AnnConseq) + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply(fast intro: AnnConseq) +--{* Cond2 *} + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply(fast intro: AnnConseq) + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply(fast intro: AnnConseq) +--{* While *} + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply force + apply clarify + apply(frule ann_hoare_case_analysis,simp) + apply auto + apply(rule AnnSeq) + apply simp + apply(rule AnnWhile) + apply simp_all +--{* Await *} +apply(frule ann_hoare_case_analysis,simp) +apply clarify +apply(drule atom_hoare_sound) + apply simp +apply(simp add: com_validity_def SEM_def sem_def) +apply(simp add: Help All_None_def) +apply force +done + +lemma Strong_Soundness_aux: "\ (Some c, s) -*\ (co, t); s \ pre c; \ c q \ + \ if co = None then t \ q else t \ pre (the co) \ \ (the co) q" +apply(erule rtrancl_induct2) + apply simp +apply(case_tac "a") + apply(fast elim: ann_transition_cases) +apply(erule Strong_Soundness_aux_aux) + apply simp +apply simp_all +done + +lemma Strong_Soundness: "\ (Some c, s)-*\(co, t); s \ pre c; \ c q \ + \ if co = None then t\q else t \ pre (the co)" +apply(force dest:Strong_Soundness_aux) +done + +lemma ann_hoare_sound: "\ c q \ \ c q" +apply (unfold ann_com_validity_def ann_SEM_def ann_sem_def) +apply clarify +apply(drule Strong_Soundness) +apply simp_all +done + +subsection {* Soundness of the System for Parallel Programs *} + +lemma Parallel_length_post_P1: "(Parallel Ts,s) -P1\ (R', t) \ + (\Rs. R' = (Parallel Rs) \ (length Rs) = (length Ts) \ + (\i. i post(Rs ! i) = post(Ts ! i)))" +apply(erule transition_cases) +apply simp +apply clarify +apply(case_tac "i=ia") +apply simp+ +done + +lemma Parallel_length_post_PStar: "(Parallel Ts,s) -P*\ (R',t) \ + (\Rs. R' = (Parallel Rs) \ (length Rs) = (length Ts) \ + (\i. i post(Ts ! i) = post(Rs ! i)))" +apply(erule rtrancl_induct2) + apply(simp_all) +apply clarify +apply simp +apply(drule Parallel_length_post_P1) +apply auto +done + +lemma assertions_lemma: "pre c \ assertions c" +apply(rule ann_com_com.induct [THEN conjunct1]) +apply auto +done + +lemma interfree_aux1 [rule_format]: + "(c,s) -1\ (r,t) \ (interfree_aux(c1, q1, c) \ interfree_aux(c1, q1, r))" +apply (rule ann_transition_transition.induct [THEN conjunct1]) +apply(safe) +prefer 13 +apply (rule TrueI) +apply (simp_all add:interfree_aux_def) +apply force+ +done + +lemma interfree_aux2 [rule_format]: + "(c,s) -1\ (r,t) \ (interfree_aux(c, q, a) \ interfree_aux(r, q, a) )" +apply (rule ann_transition_transition.induct [THEN conjunct1]) +apply(force simp add:interfree_aux_def)+ +done + +lemma interfree_lemma: "\ (Some c, s) -1\ (r, t);interfree Ts ; i \ interfree (Ts[i:= (r, q)])" +apply(simp add: interfree_def) +apply clarify +apply(case_tac "i=j") + apply(drule_tac t = "ia" in not_sym) + apply simp_all +apply(force elim: interfree_aux1) +apply(force elim: interfree_aux2 simp add:nth_list_update) +done + +text {* Strong Soundness Theorem for Parallel Programs:*} + +lemma Parallel_Strong_Soundness_Seq_aux: + "\interfree Ts; i + \ interfree (Ts[i:=(Some c0, pre c1)])" +apply(simp add: interfree_def) +apply clarify +apply(case_tac "i=j") + apply(force simp add: nth_list_update interfree_aux_def) +apply(case_tac "i=ia") + apply(erule_tac x=ia in allE) + apply(force simp add:interfree_aux_def assertions_lemma) +apply simp +done + +lemma Parallel_Strong_Soundness_Seq [rule_format (no_asm)]: + "\ \i post(Ts!i) + else b \ pre(the(com(Ts!i))) \ \ the(com(Ts!i)) post(Ts!i)); + com(Ts ! i) = Some(AnnSeq c0 c1); i \ + (\ia post(Ts[i:=(Some c0, pre c1)]! ia) + else b \ pre(the(com(Ts[i:=(Some c0, pre c1)]! ia))) \ + \ the(com(Ts[i:=(Some c0, pre c1)]! ia)) post(Ts[i:=(Some c0, pre c1)]! ia))) + \ interfree (Ts[i:= (Some c0, pre c1)])" +apply(rule conjI) + apply safe + apply(case_tac "i=ia") + apply simp + apply(force dest: ann_hoare_case_analysis) + apply simp +apply(fast elim: Parallel_Strong_Soundness_Seq_aux) +done + +lemma Parallel_Strong_Soundness_aux_aux [rule_format]: + "(Some c, b) -1\ (co, t) \ + (\Ts. i com(Ts ! i) = Some c \ + (\ipost(Ts!i) + else b\pre(the(com(Ts!i))) \ \ the(com(Ts!i)) post(Ts!i))) \ + interfree Ts \ + (\j. j i\j \ (if com(Ts!j) = None then t\post(Ts!j) + else t\pre(the(com(Ts!j))) \ \ the(com(Ts!j)) post(Ts!j))) )" +apply(rule ann_transition_transition.induct [THEN conjunct1]) +apply safe +prefer 11 +apply(rule TrueI) +apply simp_all +--{* Basic *} + apply(erule_tac x = "i" in all_dupE, erule (1) notE impE) + apply(erule_tac x = "j" in allE , erule (1) notE impE) + apply(simp add: interfree_def) + apply(erule_tac x = "j" in allE,simp) + apply(erule_tac x = "i" in allE,simp) + apply(drule_tac t = "i" in not_sym) + apply(case_tac "com(Ts ! j)=None") + apply(force intro: converse_rtrancl_into_rtrancl + simp add: interfree_aux_def com_validity_def SEM_def sem_def All_None_def) + apply(simp add:interfree_aux_def) + apply clarify + apply simp + apply(erule_tac x="pre y" in ballE) + apply(force intro: converse_rtrancl_into_rtrancl + simp add: com_validity_def SEM_def sem_def All_None_def) + apply(simp add:assertions_lemma) +--{* Seqs *} + apply(erule_tac x = "Ts[i:=(Some c0, pre c1)]" in allE) + apply(drule Parallel_Strong_Soundness_Seq,simp+) + apply(erule_tac x = "Ts[i:=(Some c0, pre c1)]" in allE) + apply(drule Parallel_Strong_Soundness_Seq,simp+) +--{* Await *} +apply(rule_tac x = "i" in allE , assumption , erule (1) notE impE) +apply(erule_tac x = "j" in allE , erule (1) notE impE) +apply(simp add: interfree_def) +apply(erule_tac x = "j" in allE,simp) +apply(erule_tac x = "i" in allE,simp) +apply(drule_tac t = "i" in not_sym) +apply(case_tac "com(Ts ! j)=None") + apply(force intro: converse_rtrancl_into_rtrancl simp add: interfree_aux_def + com_validity_def SEM_def sem_def All_None_def Help) +apply(simp add:interfree_aux_def) +apply clarify +apply simp +apply(erule_tac x="pre y" in ballE) + apply(force intro: converse_rtrancl_into_rtrancl + simp add: com_validity_def SEM_def sem_def All_None_def Help) +apply(simp add:assertions_lemma) +done + +lemma Parallel_Strong_Soundness_aux [rule_format]: + "\(Ts',s) -P*\ (Rs',t); Ts' = (Parallel Ts); interfree Ts; + \i. i (\c q. (Ts ! i) = (Some c, q) \ s\(pre c) \ \ c q ) \ \ + \Rs. Rs' = (Parallel Rs) \ (\j. j + (if com(Rs ! j) = None then t\post(Ts ! j) + else t\pre(the(com(Rs ! j))) \ \ the(com(Rs ! j)) post(Ts ! j))) \ interfree Rs" +apply(erule rtrancl_induct2) + apply clarify +--{* Base *} + apply force +--{* Induction step *} +apply clarify +apply(drule Parallel_length_post_PStar) +apply clarify +apply (ind_cases "(Parallel Ts, s) -P1\ (Parallel Rs, t)" for Ts s Rs t) +apply(rule conjI) + apply clarify + apply(case_tac "i=j") + apply(simp split del:split_if) + apply(erule Strong_Soundness_aux_aux,simp+) + apply force + apply force + apply(simp split del: split_if) + apply(erule Parallel_Strong_Soundness_aux_aux) + apply(simp_all add: split del:split_if) + apply force +apply(rule interfree_lemma) +apply simp_all +done + +lemma Parallel_Strong_Soundness: + "\(Parallel Ts, s) -P*\ (Parallel Rs, t); interfree Ts; ji. i (\c q. Ts ! i = (Some c, q) \ s\pre c \ \ c q) \ \ + if com(Rs ! j) = None then t\post(Ts ! j) else t\pre (the(com(Rs ! j)))" +apply(drule Parallel_Strong_Soundness_aux) +apply simp+ +done + +lemma oghoare_sound [rule_format]: "\- p c q \ \= p c q" +apply (unfold com_validity_def) +apply(rule oghoare_induct) +apply(rule TrueI)+ +--{* Parallel *} + apply(simp add: SEM_def sem_def) + apply clarify + apply(frule Parallel_length_post_PStar) + apply clarify + apply(drule_tac j=xb in Parallel_Strong_Soundness) + apply clarify + apply simp + apply force + apply simp + apply(erule_tac V = "\i. ?P i" in thin_rl) + apply(drule_tac s = "length Rs" in sym) + apply(erule allE, erule impE, assumption) + apply(force dest: nth_mem simp add: All_None_def) +--{* Basic *} + apply(simp add: SEM_def sem_def) + apply(force dest: rtrancl_imp_UN_rel_pow Basic_ntran) +--{* Seq *} + apply(rule subset_trans) + prefer 2 apply assumption + apply(simp add: L3_5ii L3_5i) +--{* Cond *} + apply(simp add: L3_5iv) +--{* While *} + apply(simp add: L3_5v) + apply (blast dest: SEM_fwhile) +--{* Conseq *} +apply(auto simp add: SEM_def sem_def) +done + +end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/OG_Syntax.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/OG_Syntax.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,140 @@ +theory OG_Syntax +imports OG_Tactics Quote_Antiquote +begin + +text{* Syntax for commands and for assertions and boolean expressions in + commands @{text com} and annotated commands @{text ann_com}. *} + +syntax + "_Assign" :: "idt \ 'b \ 'a com" ("(\_ :=/ _)" [70, 65] 61) + "_AnnAssign" :: "'a assn \ idt \ 'b \ 'a com" ("(_ \_ :=/ _)" [90,70,65] 61) + +translations + "\\x := a" \ "Basic \\\(_update_name x (\_. a))\" + "r \\x := a" \ "AnnBasic r \\\(_update_name x (\_. a))\" + +syntax + "_AnnSkip" :: "'a assn \ 'a ann_com" ("_//SKIP" [90] 63) + "_AnnSeq" :: "'a ann_com \ 'a ann_com \ 'a ann_com" ("_;;/ _" [60,61] 60) + + "_AnnCond1" :: "'a assn \ 'a bexp \ 'a ann_com \ 'a ann_com \ 'a ann_com" + ("_ //IF _ /THEN _ /ELSE _ /FI" [90,0,0,0] 61) + "_AnnCond2" :: "'a assn \ 'a bexp \ 'a ann_com \ 'a ann_com" + ("_ //IF _ /THEN _ /FI" [90,0,0] 61) + "_AnnWhile" :: "'a assn \ 'a bexp \ 'a assn \ 'a ann_com \ 'a ann_com" + ("_ //WHILE _ /INV _ //DO _//OD" [90,0,0,0] 61) + "_AnnAwait" :: "'a assn \ 'a bexp \ 'a com \ 'a ann_com" + ("_ //AWAIT _ /THEN /_ /END" [90,0,0] 61) + "_AnnAtom" :: "'a assn \ 'a com \ 'a ann_com" ("_//\_\" [90,0] 61) + "_AnnWait" :: "'a assn \ 'a bexp \ 'a ann_com" ("_//WAIT _ END" [90,0] 61) + + "_Skip" :: "'a com" ("SKIP" 63) + "_Seq" :: "'a com \ 'a com \ 'a com" ("_,,/ _" [55, 56] 55) + "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" + ("(0IF _/ THEN _/ ELSE _/ FI)" [0, 0, 0] 61) + "_Cond2" :: "'a bexp \ 'a com \ 'a com" ("IF _ THEN _ FI" [0,0] 56) + "_While_inv" :: "'a bexp \ 'a assn \ 'a com \ 'a com" + ("(0WHILE _/ INV _ //DO _ /OD)" [0, 0, 0] 61) + "_While" :: "'a bexp \ 'a com \ 'a com" + ("(0WHILE _ //DO _ /OD)" [0, 0] 61) + +translations + "SKIP" \ "Basic id" + "c_1,, c_2" \ "Seq c_1 c_2" + + "IF b THEN c1 ELSE c2 FI" \ "Cond .{b}. c1 c2" + "IF b THEN c FI" \ "IF b THEN c ELSE SKIP FI" + "WHILE b INV i DO c OD" \ "While .{b}. i c" + "WHILE b DO c OD" \ "WHILE b INV CONST undefined DO c OD" + + "r SKIP" \ "AnnBasic r id" + "c_1;; c_2" \ "AnnSeq c_1 c_2" + "r IF b THEN c1 ELSE c2 FI" \ "AnnCond1 r .{b}. c1 c2" + "r IF b THEN c FI" \ "AnnCond2 r .{b}. c" + "r WHILE b INV i DO c OD" \ "AnnWhile r .{b}. i c" + "r AWAIT b THEN c END" \ "AnnAwait r .{b}. c" + "r \c\" \ "r AWAIT True THEN c END" + "r WAIT b END" \ "r AWAIT b THEN SKIP END" + +nonterminals + prgs + +syntax + "_PAR" :: "prgs \ 'a" ("COBEGIN//_//COEND" [57] 56) + "_prg" :: "['a, 'a] \ prgs" ("_//_" [60, 90] 57) + "_prgs" :: "['a, 'a, prgs] \ prgs" ("_//_//\//_" [60,90,57] 57) + + "_prg_scheme" :: "['a, 'a, 'a, 'a, 'a] \ prgs" + ("SCHEME [_ \ _ < _] _// _" [0,0,0,60, 90] 57) + +translations + "_prg c q" \ "[(Some c, q)]" + "_prgs c q ps" \ "(Some c, q) # ps" + "_PAR ps" \ "Parallel ps" + + "_prg_scheme j i k c q" \ "map (\i. (Some c, q)) [j.. (x, if T = dummyT then T else Term.domain_type T) + | NONE => raise Match); + + fun update_name_tr' (Free x) = Free (upd_tr' x) + | update_name_tr' ((c as Const ("_free", _)) $ Free x) = + c $ Free (upd_tr' x) + | update_name_tr' (Const x) = Const (upd_tr' x) + | update_name_tr' _ = raise Match; + + fun K_tr' (Abs (_,_,t)) = if null (loose_bnos t) then t else raise Match + | K_tr' (Abs (_,_,Abs (_,_,t)$Bound 0)) = if null (loose_bnos t) then t else raise Match + | K_tr' _ = raise Match; + + fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) = + quote_tr' (Syntax.const "_Assign" $ update_name_tr' f) + (Abs (x, dummyT, K_tr' k) :: ts) + | assign_tr' _ = raise Match; + + fun annassign_tr' (r :: Abs (x, _, f $ k $ Bound 0) :: ts) = + quote_tr' (Syntax.const "_AnnAssign" $ r $ update_name_tr' f) + (Abs (x, dummyT, K_tr' k) :: ts) + | annassign_tr' _ = raise Match; + + fun Parallel_PAR [(Const ("Cons",_) $ (Const ("Pair",_) $ (Const ("Some",_) $ t1 ) $ t2) $ Const ("Nil",_))] = + (Syntax.const "_prg" $ t1 $ t2) + | Parallel_PAR [(Const ("Cons",_) $ (Const ("Pair",_) $ (Const ("Some",_) $ t1) $ t2) $ ts)] = + (Syntax.const "_prgs" $ t1 $ t2 $ Parallel_PAR [ts]) + | Parallel_PAR _ = raise Match; + +fun Parallel_tr' ts = Syntax.const "_PAR" $ Parallel_PAR ts; + in + [("Collect", assert_tr'), ("Basic", assign_tr'), + ("Cond", bexp_tr' "_Cond"), ("While", bexp_tr' "_While_inv"), + ("AnnBasic", annassign_tr'), + ("AnnWhile", annbexp_tr' "_AnnWhile"), ("AnnAwait", annbexp_tr' "_AnnAwait"), + ("AnnCond1", annbexp_tr' "_AnnCond1"), ("AnnCond2", annbexp_tr' "_AnnCond2")] + + end + +*} + +end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/OG_Tactics.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/OG_Tactics.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,496 @@ +header {* \section{Generation of Verification Conditions} *} + +theory OG_Tactics +imports OG_Hoare +begin + +lemmas ann_hoare_intros=AnnBasic AnnSeq AnnCond1 AnnCond2 AnnWhile AnnAwait AnnConseq +lemmas oghoare_intros=Parallel Basic Seq Cond While Conseq + +lemma ParallelConseqRule: + "\ p \ (\i\{i. i- (\i\{i. ii\{i. ii\{i. i q \ + \ \- p (Parallel Ts) q" +apply (rule Conseq) +prefer 2 + apply fast +apply assumption+ +done + +lemma SkipRule: "p \ q \ \- p (Basic id) q" +apply(rule oghoare_intros) + prefer 2 apply(rule Basic) + prefer 2 apply(rule subset_refl) +apply(simp add:Id_def) +done + +lemma BasicRule: "p \ {s. (f s)\q} \ \- p (Basic f) q" +apply(rule oghoare_intros) + prefer 2 apply(rule oghoare_intros) + prefer 2 apply(rule subset_refl) +apply assumption +done + +lemma SeqRule: "\ \- p c1 r; \- r c2 q \ \ \- p (Seq c1 c2) q" +apply(rule Seq) +apply fast+ +done + +lemma CondRule: + "\ p \ {s. (s\b \ s\w) \ (s\b \ s\w')}; \- w c1 q; \- w' c2 q \ + \ \- p (Cond b c1 c2) q" +apply(rule Cond) + apply(rule Conseq) + prefer 4 apply(rule Conseq) +apply simp_all +apply force+ +done + +lemma WhileRule: "\ p \ i; \- (i \ b) c i ; (i \ (-b)) \ q \ + \ \- p (While b i c) q" +apply(rule Conseq) + prefer 2 apply(rule While) +apply assumption+ +done + +text {* Three new proof rules for special instances of the @{text +AnnBasic} and the @{text AnnAwait} commands when the transformation +performed on the state is the identity, and for an @{text AnnAwait} +command where the boolean condition is @{text "{s. True}"}: *} + +lemma AnnatomRule: + "\ atom_com(c); \- r c q \ \ \ (AnnAwait r {s. True} c) q" +apply(rule AnnAwait) +apply simp_all +done + +lemma AnnskipRule: + "r \ q \ \ (AnnBasic r id) q" +apply(rule AnnBasic) +apply simp +done + +lemma AnnwaitRule: + "\ (r \ b) \ q \ \ \ (AnnAwait r b (Basic id)) q" +apply(rule AnnAwait) + apply simp +apply(rule BasicRule) +apply simp +done + +text {* Lemmata to avoid using the definition of @{text +map_ann_hoare}, @{text interfree_aux}, @{text interfree_swap} and +@{text interfree} by splitting it into different cases: *} + +lemma interfree_aux_rule1: "interfree_aux(co, q, None)" +by(simp add:interfree_aux_def) + +lemma interfree_aux_rule2: + "\(R,r)\(atomics a). \- (q \ R) r q \ interfree_aux(None, q, Some a)" +apply(simp add:interfree_aux_def) +apply(force elim:oghoare_sound) +done + +lemma interfree_aux_rule3: + "(\(R, r)\(atomics a). \- (q \ R) r q \ (\p\(assertions c). \- (p \ R) r p)) + \ interfree_aux(Some c, q, Some a)" +apply(simp add:interfree_aux_def) +apply(force elim:oghoare_sound) +done + +lemma AnnBasic_assertions: + "\interfree_aux(None, r, Some a); interfree_aux(None, q, Some a)\ \ + interfree_aux(Some (AnnBasic r f), q, Some a)" +apply(simp add: interfree_aux_def) +by force + +lemma AnnSeq_assertions: + "\ interfree_aux(Some c1, q, Some a); interfree_aux(Some c2, q, Some a)\\ + interfree_aux(Some (AnnSeq c1 c2), q, Some a)" +apply(simp add: interfree_aux_def) +by force + +lemma AnnCond1_assertions: + "\ interfree_aux(None, r, Some a); interfree_aux(Some c1, q, Some a); + interfree_aux(Some c2, q, Some a)\\ + interfree_aux(Some(AnnCond1 r b c1 c2), q, Some a)" +apply(simp add: interfree_aux_def) +by force + +lemma AnnCond2_assertions: + "\ interfree_aux(None, r, Some a); interfree_aux(Some c, q, Some a)\\ + interfree_aux(Some (AnnCond2 r b c), q, Some a)" +apply(simp add: interfree_aux_def) +by force + +lemma AnnWhile_assertions: + "\ interfree_aux(None, r, Some a); interfree_aux(None, i, Some a); + interfree_aux(Some c, q, Some a)\\ + interfree_aux(Some (AnnWhile r b i c), q, Some a)" +apply(simp add: interfree_aux_def) +by force + +lemma AnnAwait_assertions: + "\ interfree_aux(None, r, Some a); interfree_aux(None, q, Some a)\\ + interfree_aux(Some (AnnAwait r b c), q, Some a)" +apply(simp add: interfree_aux_def) +by force + +lemma AnnBasic_atomics: + "\- (q \ r) (Basic f) q \ interfree_aux(None, q, Some (AnnBasic r f))" +by(simp add: interfree_aux_def oghoare_sound) + +lemma AnnSeq_atomics: + "\ interfree_aux(Any, q, Some a1); interfree_aux(Any, q, Some a2)\\ + interfree_aux(Any, q, Some (AnnSeq a1 a2))" +apply(simp add: interfree_aux_def) +by force + +lemma AnnCond1_atomics: + "\ interfree_aux(Any, q, Some a1); interfree_aux(Any, q, Some a2)\\ + interfree_aux(Any, q, Some (AnnCond1 r b a1 a2))" +apply(simp add: interfree_aux_def) +by force + +lemma AnnCond2_atomics: + "interfree_aux (Any, q, Some a)\ interfree_aux(Any, q, Some (AnnCond2 r b a))" +by(simp add: interfree_aux_def) + +lemma AnnWhile_atomics: "interfree_aux (Any, q, Some a) + \ interfree_aux(Any, q, Some (AnnWhile r b i a))" +by(simp add: interfree_aux_def) + +lemma Annatom_atomics: + "\- (q \ r) a q \ interfree_aux (None, q, Some (AnnAwait r {x. True} a))" +by(simp add: interfree_aux_def oghoare_sound) + +lemma AnnAwait_atomics: + "\- (q \ (r \ b)) a q \ interfree_aux (None, q, Some (AnnAwait r b a))" +by(simp add: interfree_aux_def oghoare_sound) + +constdefs + interfree_swap :: "('a ann_triple_op * ('a ann_triple_op) list) \ bool" + "interfree_swap == \(x, xs). \y\set xs. interfree_aux (com x, post x, com y) + \ interfree_aux(com y, post y, com x)" + +lemma interfree_swap_Empty: "interfree_swap (x, [])" +by(simp add:interfree_swap_def) + +lemma interfree_swap_List: + "\ interfree_aux (com x, post x, com y); + interfree_aux (com y, post y ,com x); interfree_swap (x, xs) \ + \ interfree_swap (x, y#xs)" +by(simp add:interfree_swap_def) + +lemma interfree_swap_Map: "\k. i\k \ k interfree_aux (com x, post x, c k) + \ interfree_aux (c k, Q k, com x) + \ interfree_swap (x, map (\k. (c k, Q k)) [i.. interfree_swap(x, xs); interfree xs \ \ interfree (x#xs)" +apply(simp add:interfree_def interfree_swap_def) +apply clarify +apply(case_tac i) + apply(case_tac j) + apply simp_all +apply(case_tac j,simp+) +done + +lemma interfree_Map: + "(\i j. a\i \ i a\j \ j i\j \ interfree_aux (c i, Q i, c j)) + \ interfree (map (\k. (c k, Q k)) [a.. bool " ("[\] _" [0] 45) + "[\] Ts == (\ic q. Ts!i=(Some c, q) \ \ c q)" + +lemma MapAnnEmpty: "[\] []" +by(simp add:map_ann_hoare_def) + +lemma MapAnnList: "\ \ c q ; [\] xs \ \ [\] (Some c,q)#xs" +apply(simp add:map_ann_hoare_def) +apply clarify +apply(case_tac i,simp+) +done + +lemma MapAnnMap: + "\k. i\k \ k \ (c k) (Q k) \ [\] map (\k. (Some (c k), Q k)) [i.. [\] Ts ; interfree Ts \ + \ \- (\i\{i. ii\{i. i \k (c k) (Q k); + \k l. k l k\l \ interfree_aux (Some(c k), Q k, Some(c l)) \ + \ \- (\i\{i. iii\{i. i(b x)}" +by fast +lemma list_length: "length []=0 \ length (x#xs) = Suc(length xs)" +by simp +lemma list_lemmas: "length []=0 \ length (x#xs) = Suc(length xs) +\ (x#xs) ! 0=x \ (x#xs) ! Suc n = xs ! n" +by simp +lemma le_Suc_eq_insert: "{i. i a1 \ a2 \ .. \ an"} returning +@{text n} subgoals, one for each conjunct: *} + +ML {* +fun conjI_Tac tac i st = st |> + ( (EVERY [rtac conjI i, + conjI_Tac tac (i+1), + tac i]) ORELSE (tac i) ) +*} + + +subsubsection {* Tactic for the generation of the verification conditions *} + +text {* The tactic basically uses two subtactics: + +\begin{description} + +\item[HoareRuleTac] is called at the level of parallel programs, it + uses the ParallelTac to solve parallel composition of programs. + This verification has two parts, namely, (1) all component programs are + correct and (2) they are interference free. @{text HoareRuleTac} is + also called at the level of atomic regions, i.e. @{text "\ \"} and + @{text "AWAIT b THEN _ END"}, and at each interference freedom test. + +\item[AnnHoareRuleTac] is for component programs which + are annotated programs and so, there are not unknown assertions + (no need to use the parameter precond, see NOTE). + + NOTE: precond(::bool) informs if the subgoal has the form @{text "\- ?p c q"}, + in this case we have precond=False and the generated verification + condition would have the form @{text "?p \ \"} which can be solved by + @{text "rtac subset_refl"}, if True we proceed to simplify it using + the simplification tactics above. + +\end{description} +*} + +ML {* + + fun WlpTac i = (rtac (@{thm SeqRule}) i) THEN (HoareRuleTac false (i+1)) +and HoareRuleTac precond i st = st |> + ( (WlpTac i THEN HoareRuleTac precond i) + ORELSE + (FIRST[rtac (@{thm SkipRule}) i, + rtac (@{thm BasicRule}) i, + EVERY[rtac (@{thm ParallelConseqRule}) i, + ParallelConseq (i+2), + ParallelTac (i+1), + ParallelConseq i], + EVERY[rtac (@{thm CondRule}) i, + HoareRuleTac false (i+2), + HoareRuleTac false (i+1)], + EVERY[rtac (@{thm WhileRule}) i, + HoareRuleTac true (i+1)], + K all_tac i ] + THEN (if precond then (K all_tac i) else (rtac (@{thm subset_refl}) i)))) + +and AnnWlpTac i = (rtac (@{thm AnnSeq}) i) THEN (AnnHoareRuleTac (i+1)) +and AnnHoareRuleTac i st = st |> + ( (AnnWlpTac i THEN AnnHoareRuleTac i ) + ORELSE + (FIRST[(rtac (@{thm AnnskipRule}) i), + EVERY[rtac (@{thm AnnatomRule}) i, + HoareRuleTac true (i+1)], + (rtac (@{thm AnnwaitRule}) i), + rtac (@{thm AnnBasic}) i, + EVERY[rtac (@{thm AnnCond1}) i, + AnnHoareRuleTac (i+3), + AnnHoareRuleTac (i+1)], + EVERY[rtac (@{thm AnnCond2}) i, + AnnHoareRuleTac (i+1)], + EVERY[rtac (@{thm AnnWhile}) i, + AnnHoareRuleTac (i+2)], + EVERY[rtac (@{thm AnnAwait}) i, + HoareRuleTac true (i+1)], + K all_tac i])) + +and ParallelTac i = EVERY[rtac (@{thm ParallelRule}) i, + interfree_Tac (i+1), + MapAnn_Tac i] + +and MapAnn_Tac i st = st |> + (FIRST[rtac (@{thm MapAnnEmpty}) i, + EVERY[rtac (@{thm MapAnnList}) i, + MapAnn_Tac (i+1), + AnnHoareRuleTac i], + EVERY[rtac (@{thm MapAnnMap}) i, + rtac (@{thm allI}) i,rtac (@{thm impI}) i, + AnnHoareRuleTac i]]) + +and interfree_swap_Tac i st = st |> + (FIRST[rtac (@{thm interfree_swap_Empty}) i, + EVERY[rtac (@{thm interfree_swap_List}) i, + interfree_swap_Tac (i+2), + interfree_aux_Tac (i+1), + interfree_aux_Tac i ], + EVERY[rtac (@{thm interfree_swap_Map}) i, + rtac (@{thm allI}) i,rtac (@{thm impI}) i, + conjI_Tac (interfree_aux_Tac) i]]) + +and interfree_Tac i st = st |> + (FIRST[rtac (@{thm interfree_Empty}) i, + EVERY[rtac (@{thm interfree_List}) i, + interfree_Tac (i+1), + interfree_swap_Tac i], + EVERY[rtac (@{thm interfree_Map}) i, + rtac (@{thm allI}) i,rtac (@{thm allI}) i,rtac (@{thm impI}) i, + interfree_aux_Tac i ]]) + +and interfree_aux_Tac i = (before_interfree_simp_tac i ) THEN + (FIRST[rtac (@{thm interfree_aux_rule1}) i, + dest_assertions_Tac i]) + +and dest_assertions_Tac i st = st |> + (FIRST[EVERY[rtac (@{thm AnnBasic_assertions}) i, + dest_atomics_Tac (i+1), + dest_atomics_Tac i], + EVERY[rtac (@{thm AnnSeq_assertions}) i, + dest_assertions_Tac (i+1), + dest_assertions_Tac i], + EVERY[rtac (@{thm AnnCond1_assertions}) i, + dest_assertions_Tac (i+2), + dest_assertions_Tac (i+1), + dest_atomics_Tac i], + EVERY[rtac (@{thm AnnCond2_assertions}) i, + dest_assertions_Tac (i+1), + dest_atomics_Tac i], + EVERY[rtac (@{thm AnnWhile_assertions}) i, + dest_assertions_Tac (i+2), + dest_atomics_Tac (i+1), + dest_atomics_Tac i], + EVERY[rtac (@{thm AnnAwait_assertions}) i, + dest_atomics_Tac (i+1), + dest_atomics_Tac i], + dest_atomics_Tac i]) + +and dest_atomics_Tac i st = st |> + (FIRST[EVERY[rtac (@{thm AnnBasic_atomics}) i, + HoareRuleTac true i], + EVERY[rtac (@{thm AnnSeq_atomics}) i, + dest_atomics_Tac (i+1), + dest_atomics_Tac i], + EVERY[rtac (@{thm AnnCond1_atomics}) i, + dest_atomics_Tac (i+1), + dest_atomics_Tac i], + EVERY[rtac (@{thm AnnCond2_atomics}) i, + dest_atomics_Tac i], + EVERY[rtac (@{thm AnnWhile_atomics}) i, + dest_atomics_Tac i], + EVERY[rtac (@{thm Annatom_atomics}) i, + HoareRuleTac true i], + EVERY[rtac (@{thm AnnAwait_atomics}) i, + HoareRuleTac true i], + K all_tac i]) +*} + + +text {* The final tactic is given the name @{text oghoare}: *} + +ML {* +val oghoare_tac = SUBGOAL (fn (_, i) => + (HoareRuleTac true i)) +*} + +text {* Notice that the tactic for parallel programs @{text +"oghoare_tac"} is initially invoked with the value @{text true} for +the parameter @{text precond}. + +Parts of the tactic can be also individually used to generate the +verification conditions for annotated sequential programs and to +generate verification conditions out of interference freedom tests: *} + +ML {* val annhoare_tac = SUBGOAL (fn (_, i) => + (AnnHoareRuleTac i)) + +val interfree_aux_tac = SUBGOAL (fn (_, i) => + (interfree_aux_Tac i)) +*} + +text {* The so defined ML tactics are then ``exported'' to be used in +Isabelle proofs. *} + +method_setup oghoare = {* + Scan.succeed (K (SIMPLE_METHOD' oghoare_tac)) *} + "verification condition generator for the oghoare logic" + +method_setup annhoare = {* + Scan.succeed (K (SIMPLE_METHOD' annhoare_tac)) *} + "verification condition generator for the ann_hoare logic" + +method_setup interfree_aux = {* + Scan.succeed (K (SIMPLE_METHOD' interfree_aux_tac)) *} + "verification condition generator for interference freedom tests" + +text {* Tactics useful for dealing with the generated verification conditions: *} + +method_setup conjI_tac = {* + Scan.succeed (K (SIMPLE_METHOD' (conjI_Tac (K all_tac)))) *} + "verification condition generator for interference freedom tests" + +ML {* +fun disjE_Tac tac i st = st |> + ( (EVERY [etac disjE i, + disjE_Tac tac (i+1), + tac i]) ORELSE (tac i) ) +*} + +method_setup disjE_tac = {* + Scan.succeed (K (SIMPLE_METHOD' (disjE_Tac (K all_tac)))) *} + "verification condition generator for interference freedom tests" + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/OG_Tran.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/OG_Tran.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,309 @@ + +header {* \section{Operational Semantics} *} + +theory OG_Tran imports OG_Com begin + +types + 'a ann_com_op = "('a ann_com) option" + 'a ann_triple_op = "('a ann_com_op \ 'a assn)" + +consts com :: "'a ann_triple_op \ 'a ann_com_op" +primrec "com (c, q) = c" + +consts post :: "'a ann_triple_op \ 'a assn" +primrec "post (c, q) = q" + +constdefs + All_None :: "'a ann_triple_op list \ bool" + "All_None Ts \ \(c, q) \ set Ts. c = None" + +subsection {* The Transition Relation *} + +inductive_set + ann_transition :: "(('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a)) set" + and transition :: "(('a com \ 'a) \ ('a com \ 'a)) set" + and ann_transition' :: "('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a) \ bool" + ("_ -1\ _"[81,81] 100) + and transition' :: "('a com \ 'a) \ ('a com \ 'a) \ bool" + ("_ -P1\ _"[81,81] 100) + and transitions :: "('a com \ 'a) \ ('a com \ 'a) \ bool" + ("_ -P*\ _"[81,81] 100) +where + "con_0 -1\ con_1 \ (con_0, con_1) \ ann_transition" +| "con_0 -P1\ con_1 \ (con_0, con_1) \ transition" +| "con_0 -P*\ con_1 \ (con_0, con_1) \ transition\<^sup>*" + +| AnnBasic: "(Some (AnnBasic r f), s) -1\ (None, f s)" + +| AnnSeq1: "(Some c0, s) -1\ (None, t) \ + (Some (AnnSeq c0 c1), s) -1\ (Some c1, t)" +| AnnSeq2: "(Some c0, s) -1\ (Some c2, t) \ + (Some (AnnSeq c0 c1), s) -1\ (Some (AnnSeq c2 c1), t)" + +| AnnCond1T: "s \ b \ (Some (AnnCond1 r b c1 c2), s) -1\ (Some c1, s)" +| AnnCond1F: "s \ b \ (Some (AnnCond1 r b c1 c2), s) -1\ (Some c2, s)" + +| AnnCond2T: "s \ b \ (Some (AnnCond2 r b c), s) -1\ (Some c, s)" +| AnnCond2F: "s \ b \ (Some (AnnCond2 r b c), s) -1\ (None, s)" + +| AnnWhileF: "s \ b \ (Some (AnnWhile r b i c), s) -1\ (None, s)" +| AnnWhileT: "s \ b \ (Some (AnnWhile r b i c), s) -1\ + (Some (AnnSeq c (AnnWhile i b i c)), s)" + +| AnnAwait: "\ s \ b; atom_com c; (c, s) -P*\ (Parallel [], t) \ \ + (Some (AnnAwait r b c), s) -1\ (None, t)" + +| Parallel: "\ i (r, t) \ + \ (Parallel Ts, s) -P1\ (Parallel (Ts [i:=(r, q)]), t)" + +| Basic: "(Basic f, s) -P1\ (Parallel [], f s)" + +| Seq1: "All_None Ts \ (Seq (Parallel Ts) c, s) -P1\ (c, s)" +| Seq2: "(c0, s) -P1\ (c2, t) \ (Seq c0 c1, s) -P1\ (Seq c2 c1, t)" + +| CondT: "s \ b \ (Cond b c1 c2, s) -P1\ (c1, s)" +| CondF: "s \ b \ (Cond b c1 c2, s) -P1\ (c2, s)" + +| WhileF: "s \ b \ (While b i c, s) -P1\ (Parallel [], s)" +| WhileT: "s \ b \ (While b i c, s) -P1\ (Seq c (While b i c), s)" + +monos "rtrancl_mono" + +text {* The corresponding syntax translations are: *} + +abbreviation + ann_transition_n :: "('a ann_com_op \ 'a) \ nat \ ('a ann_com_op \ 'a) + \ bool" ("_ -_\ _"[81,81] 100) where + "con_0 -n\ con_1 \ (con_0, con_1) \ ann_transition ^^ n" + +abbreviation + ann_transitions :: "('a ann_com_op \ 'a) \ ('a ann_com_op \ 'a) \ bool" + ("_ -*\ _"[81,81] 100) where + "con_0 -*\ con_1 \ (con_0, con_1) \ ann_transition\<^sup>*" + +abbreviation + transition_n :: "('a com \ 'a) \ nat \ ('a com \ 'a) \ bool" + ("_ -P_\ _"[81,81,81] 100) where + "con_0 -Pn\ con_1 \ (con_0, con_1) \ transition ^^ n" + +subsection {* Definition of Semantics *} + +constdefs + ann_sem :: "'a ann_com \ 'a \ 'a set" + "ann_sem c \ \s. {t. (Some c, s) -*\ (None, t)}" + + ann_SEM :: "'a ann_com \ 'a set \ 'a set" + "ann_SEM c S \ \ann_sem c ` S" + + sem :: "'a com \ 'a \ 'a set" + "sem c \ \s. {t. \Ts. (c, s) -P*\ (Parallel Ts, t) \ All_None Ts}" + + SEM :: "'a com \ 'a set \ 'a set" + "SEM c S \ \sem c ` S " + +syntax "_Omega" :: "'a com" ("\" 63) +translations "\" \ "While CONST UNIV CONST UNIV (Basic id)" + +consts fwhile :: "'a bexp \ 'a com \ nat \ 'a com" +primrec + "fwhile b c 0 = \" + "fwhile b c (Suc n) = Cond b (Seq c (fwhile b c n)) (Basic id)" + +subsubsection {* Proofs *} + +declare ann_transition_transition.intros [intro] +inductive_cases transition_cases: + "(Parallel T,s) -P1\ t" + "(Basic f, s) -P1\ t" + "(Seq c1 c2, s) -P1\ t" + "(Cond b c1 c2, s) -P1\ t" + "(While b i c, s) -P1\ t" + +lemma Parallel_empty_lemma [rule_format (no_asm)]: + "(Parallel [],s) -Pn\ (Parallel Ts,t) \ Ts=[] \ n=0 \ s=t" +apply(induct n) + apply(simp (no_asm)) +apply clarify +apply(drule rel_pow_Suc_D2) +apply(force elim:transition_cases) +done + +lemma Parallel_AllNone_lemma [rule_format (no_asm)]: + "All_None Ss \ (Parallel Ss,s) -Pn\ (Parallel Ts,t) \ Ts=Ss \ n=0 \ s=t" +apply(induct "n") + apply(simp (no_asm)) +apply clarify +apply(drule rel_pow_Suc_D2) +apply clarify +apply(erule transition_cases,simp_all) +apply(force dest:nth_mem simp add:All_None_def) +done + +lemma Parallel_AllNone: "All_None Ts \ (SEM (Parallel Ts) X) = X" +apply (unfold SEM_def sem_def) +apply auto +apply(drule rtrancl_imp_UN_rel_pow) +apply clarify +apply(drule Parallel_AllNone_lemma) +apply auto +done + +lemma Parallel_empty: "Ts=[] \ (SEM (Parallel Ts) X) = X" +apply(rule Parallel_AllNone) +apply(simp add:All_None_def) +done + +text {* Set of lemmas from Apt and Olderog "Verification of sequential +and concurrent programs", page 63. *} + +lemma L3_5i: "X\Y \ SEM c X \ SEM c Y" +apply (unfold SEM_def) +apply force +done + +lemma L3_5ii_lemma1: + "\ (c1, s1) -P*\ (Parallel Ts, s2); All_None Ts; + (c2, s2) -P*\ (Parallel Ss, s3); All_None Ss \ + \ (Seq c1 c2, s1) -P*\ (Parallel Ss, s3)" +apply(erule converse_rtrancl_induct2) +apply(force intro:converse_rtrancl_into_rtrancl)+ +done + +lemma L3_5ii_lemma2 [rule_format (no_asm)]: + "\c1 c2 s t. (Seq c1 c2, s) -Pn\ (Parallel Ts, t) \ + (All_None Ts) \ (\y m Rs. (c1,s) -P*\ (Parallel Rs, y) \ + (All_None Rs) \ (c2, y) -Pm\ (Parallel Ts, t) \ m \ n)" +apply(induct "n") + apply(force) +apply(safe dest!: rel_pow_Suc_D2) +apply(erule transition_cases,simp_all) + apply (fast intro!: le_SucI) +apply (fast intro!: le_SucI elim!: rel_pow_imp_rtrancl converse_rtrancl_into_rtrancl) +done + +lemma L3_5ii_lemma3: + "\(Seq c1 c2,s) -P*\ (Parallel Ts,t); All_None Ts\ \ + (\y Rs. (c1,s) -P*\ (Parallel Rs,y) \ All_None Rs + \ (c2,y) -P*\ (Parallel Ts,t))" +apply(drule rtrancl_imp_UN_rel_pow) +apply(fast dest: L3_5ii_lemma2 rel_pow_imp_rtrancl) +done + +lemma L3_5ii: "SEM (Seq c1 c2) X = SEM c2 (SEM c1 X)" +apply (unfold SEM_def sem_def) +apply auto + apply(fast dest: L3_5ii_lemma3) +apply(fast elim: L3_5ii_lemma1) +done + +lemma L3_5iii: "SEM (Seq (Seq c1 c2) c3) X = SEM (Seq c1 (Seq c2 c3)) X" +apply (simp (no_asm) add: L3_5ii) +done + +lemma L3_5iv: + "SEM (Cond b c1 c2) X = (SEM c1 (X \ b)) Un (SEM c2 (X \ (-b)))" +apply (unfold SEM_def sem_def) +apply auto +apply(erule converse_rtranclE) + prefer 2 + apply (erule transition_cases,simp_all) + apply(fast intro: converse_rtrancl_into_rtrancl elim: transition_cases)+ +done + + +lemma L3_5v_lemma1[rule_format]: + "(S,s) -Pn\ (T,t) \ S=\ \ (\(\Rs. T=(Parallel Rs) \ All_None Rs))" +apply (unfold UNIV_def) +apply(rule nat_less_induct) +apply safe +apply(erule rel_pow_E2) + apply simp_all +apply(erule transition_cases) + apply simp_all +apply(erule rel_pow_E2) + apply(simp add: Id_def) +apply(erule transition_cases,simp_all) +apply clarify +apply(erule transition_cases,simp_all) +apply(erule rel_pow_E2,simp) +apply clarify +apply(erule transition_cases) + apply simp+ + apply clarify + apply(erule transition_cases) +apply simp_all +done + +lemma L3_5v_lemma2: "\(\, s) -P*\ (Parallel Ts, t); All_None Ts \ \ False" +apply(fast dest: rtrancl_imp_UN_rel_pow L3_5v_lemma1) +done + +lemma L3_5v_lemma3: "SEM (\) S = {}" +apply (unfold SEM_def sem_def) +apply(fast dest: L3_5v_lemma2) +done + +lemma L3_5v_lemma4 [rule_format]: + "\s. (While b i c, s) -Pn\ (Parallel Ts, t) \ All_None Ts \ + (\k. (fwhile b c k, s) -P*\ (Parallel Ts, t))" +apply(rule nat_less_induct) +apply safe +apply(erule rel_pow_E2) + apply safe +apply(erule transition_cases,simp_all) + apply (rule_tac x = "1" in exI) + apply(force dest: Parallel_empty_lemma intro: converse_rtrancl_into_rtrancl simp add: Id_def) +apply safe +apply(drule L3_5ii_lemma2) + apply safe +apply(drule le_imp_less_Suc) +apply (erule allE , erule impE,assumption) +apply (erule allE , erule impE, assumption) +apply safe +apply (rule_tac x = "k+1" in exI) +apply(simp (no_asm)) +apply(rule converse_rtrancl_into_rtrancl) + apply fast +apply(fast elim: L3_5ii_lemma1) +done + +lemma L3_5v_lemma5 [rule_format]: + "\s. (fwhile b c k, s) -P*\ (Parallel Ts, t) \ All_None Ts \ + (While b i c, s) -P*\ (Parallel Ts,t)" +apply(induct "k") + apply(force dest: L3_5v_lemma2) +apply safe +apply(erule converse_rtranclE) + apply simp_all +apply(erule transition_cases,simp_all) + apply(rule converse_rtrancl_into_rtrancl) + apply(fast) + apply(fast elim!: L3_5ii_lemma1 dest: L3_5ii_lemma3) +apply(drule rtrancl_imp_UN_rel_pow) +apply clarify +apply(erule rel_pow_E2) + apply simp_all +apply(erule transition_cases,simp_all) +apply(fast dest: Parallel_empty_lemma) +done + +lemma L3_5v: "SEM (While b i c) = (\x. (\k. SEM (fwhile b c k) x))" +apply(rule ext) +apply (simp add: SEM_def sem_def) +apply safe + apply(drule rtrancl_imp_UN_rel_pow,simp) + apply clarify + apply(fast dest:L3_5v_lemma4) +apply(fast intro: L3_5v_lemma5) +done + +section {* Validity of Correctness Formulas *} + +constdefs + com_validity :: "'a assn \ 'a com \ 'a assn \ bool" ("(3\= _// _//_)" [90,55,90] 50) + "\= p c q \ SEM c p \ q" + + ann_com_validity :: "'a ann_com \ 'a assn \ bool" ("\ _ _" [60,90] 45) + "\ c q \ ann_SEM c (pre c) \ q" + +end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/Quote_Antiquote.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/Quote_Antiquote.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,24 @@ + +header {* \section{Concrete Syntax} *} + +theory Quote_Antiquote imports Main begin + +syntax + "_quote" :: "'b \ ('a \ 'b)" ("(\_\)" [0] 1000) + "_antiquote" :: "('a \ 'b) \ 'b" ("\_" [1000] 1000) + "_Assert" :: "'a \ 'a set" ("(.{_}.)" [0] 1000) + +syntax (xsymbols) + "_Assert" :: "'a \ 'a set" ("(\_\)" [0] 1000) + +translations + ".{b}." \ "Collect \b\" + +parse_translation {* + let + fun quote_tr [t] = Syntax.quote_tr "_antiquote" t + | quote_tr ts = raise TERM ("quote_tr", ts); + in [("_quote", quote_tr)] end +*} + +end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/RG_Com.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/RG_Com.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,25 @@ + +header {* \chapter{The Rely-Guarantee Method} + +\section {Abstract Syntax} +*} + +theory RG_Com imports Main begin + +text {* Semantics of assertions and boolean expressions (bexp) as sets +of states. Syntax of commands @{text com} and parallel commands +@{text par_com}. *} + +types + 'a bexp = "'a set" + +datatype 'a com = + Basic "'a \'a" + | Seq "'a com" "'a com" + | Cond "'a bexp" "'a com" "'a com" + | While "'a bexp" "'a com" + | Await "'a bexp" "'a com" + +types 'a par_com = "(('a com) option) list" + +end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/RG_Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/RG_Examples.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,359 @@ +header {* \section{Examples} *} + +theory RG_Examples +imports RG_Syntax +begin + +lemmas definitions [simp]= stable_def Pre_def Rely_def Guar_def Post_def Com_def + +subsection {* Set Elements of an Array to Zero *} + +lemma le_less_trans2: "\(j::nat) j\ \ i (a::nat) < c; b\d \ \ a + b < c + d" +by simp + +record Example1 = + A :: "nat list" + +lemma Example1: + "\ COBEGIN + SCHEME [0 \ i < n] + (\A := \A [i := 0], + \ n < length \A \, + \ length \A = length \A \ \A ! i = \A ! i \, + \ length \A = length \A \ (\j j \ \A ! j = \A ! j) \, + \ \A ! i = 0 \) + COEND + SAT [\ n < length \A \, \ \A = \A \, \ True \, \ \i < n. \A ! i = 0 \]" +apply(rule Parallel) +apply (auto intro!: Basic) +done + +lemma Example1_parameterized: +"k < t \ + \ COBEGIN + SCHEME [k*n\i<(Suc k)*n] (\A:=\A[i:=0], + \t*n < length \A\, + \t*n < length \A \ length \A=length \A \ \A!i = \A!i\, + \t*n < length \A \ length \A=length \A \ (\jA . i\j \ \A!j = \A!j)\, + \\A!i=0\) + COEND + SAT [\t*n < length \A\, + \t*n < length \A \ length \A=length \A \ (\iA!(k*n+i)=\A!(k*n+i))\, + \t*n < length \A \ length \A=length \A \ + (\iA . (i \A!i = \A!i) \ ((Suc k)*n \ i\ \A!i = \A!i))\, + \\iA!(k*n+i) = 0\]" +apply(rule Parallel) + apply auto + apply(erule_tac x="k*n +i" in allE) + apply(subgoal_tac "k*n+i COBEGIN + (\ \x:=\x+1;; \c_0:=\c_0 + 1 \, + \\x=\c_0 + \c_1 \ \c_0=0\, + \\c_0 = \c_0 \ + (\x=\c_0 + \c_1 + \ \x = \c_0 + \c_1)\, + \\c_1 = \c_1 \ + (\x=\c_0 + \c_1 + \ \x =\c_0 + \c_1)\, + \\x=\c_0 + \c_1 \ \c_0=1 \) + \ + (\ \x:=\x+1;; \c_1:=\c_1+1 \, + \\x=\c_0 + \c_1 \ \c_1=0 \, + \\c_1 = \c_1 \ + (\x=\c_0 + \c_1 + \ \x = \c_0 + \c_1)\, + \\c_0 = \c_0 \ + (\x=\c_0 + \c_1 + \ \x =\c_0 + \c_1)\, + \\x=\c_0 + \c_1 \ \c_1=1\) + COEND + SAT [\\x=0 \ \c_0=0 \ \c_1=0\, + \\x=\x \ \c_0= \c_0 \ \c_1=\c_1\, + \True\, + \\x=2\]" +apply(rule Parallel) + apply simp_all + apply clarify + apply(case_tac i) + apply simp + apply(rule conjI) + apply clarify + apply simp + apply clarify + apply simp + apply(case_tac j,simp) + apply simp + apply simp + apply(rule conjI) + apply clarify + apply simp + apply clarify + apply simp + apply(subgoal_tac "j=0") + apply (rotate_tac -1) + apply (simp (asm_lr)) + apply arith + apply clarify + apply(case_tac i,simp,simp) + apply clarify + apply simp + apply(erule_tac x=0 in all_dupE) + apply(erule_tac x=1 in allE,simp) +apply clarify +apply(case_tac i,simp) + apply(rule Await) + apply simp_all + apply(clarify) + apply(rule Seq) + prefer 2 + apply(rule Basic) + apply simp_all + apply(rule subset_refl) + apply(rule Basic) + apply simp_all + apply clarify + apply simp +apply(rule Await) + apply simp_all +apply(clarify) +apply(rule Seq) + prefer 2 + apply(rule Basic) + apply simp_all + apply(rule subset_refl) +apply(auto intro!: Basic) +done + +subsubsection {* Parameterized *} + +lemma Example2_lemma2_aux: "j + (\i=0..i=0..i=0.. s \ (\i::nat=0..i=0..j \ Suc (\i::nat=0..i=0..i=0..i=0..j") + apply(drule_tac b="b" and t="(Suc 0)" in Example2_lemma2_aux2) +apply(rotate_tac -1) +apply(erule ssubst) +apply simp_all +done + +lemma Example2_lemma2_Suc0: "\j \ + Suc (\i::nat=0..< n. b i)=(\i=0..< n. (b (j:=Suc 0)) i)" +by(simp add:Example2_lemma2) + +record Example2_parameterized = + C :: "nat \ nat" + y :: nat + +lemma Example2_parameterized: "0 + \ COBEGIN SCHEME [0\i \y:=\y+1;; \C:=\C (i:=1) \, + \\y=(\i=0..C i) \ \C i=0\, + \\C i = \C i \ + (\y=(\i=0..C i) \ \y =(\i=0..C i))\, + \(\jj \ \C j = \C j) \ + (\y=(\i=0..C i) \ \y =(\i=0..C i))\, + \\y=(\i=0..C i) \ \C i=1\) + COEND + SAT [\\y=0 \ (\i=0..C i)=0 \, \\C=\C \ \y=\y\, \True\, \\y=n\]" +apply(rule Parallel) +apply force +apply force +apply(force) +apply clarify +apply simp +apply(simp cong:setsum_ivl_cong) +apply clarify +apply simp +apply(rule Await) +apply simp_all +apply clarify +apply(rule Seq) +prefer 2 +apply(rule Basic) +apply(rule subset_refl) +apply simp+ +apply(rule Basic) +apply simp +apply clarify +apply simp +apply(simp add:Example2_lemma2_Suc0 cong:if_cong) +apply simp+ +done + +subsection {* Find Least Element *} + +text {* A previous lemma: *} + +lemma mod_aux :"\i < (n::nat); a mod n = i; j < a + n; j mod n = i; a < j\ \ False" +apply(subgoal_tac "a=a div n*n + a mod n" ) + prefer 2 apply (simp (no_asm_use)) +apply(subgoal_tac "j=j div n*n + j mod n") + prefer 2 apply (simp (no_asm_use)) +apply simp +apply(subgoal_tac "a div n*n < j div n*n") +prefer 2 apply arith +apply(subgoal_tac "j div n*n < (a div n + 1)*n") +prefer 2 apply simp +apply (simp only:mult_less_cancel2) +apply arith +done + +record Example3 = + X :: "nat \ nat" + Y :: "nat \ nat" + +lemma Example3: "m mod n=0 \ + \ COBEGIN + SCHEME [0\ijX i < \Y j) DO + IF P(B!(\X i)) THEN \Y:=\Y (i:=\X i) + ELSE \X:= \X (i:=(\X i)+ n) FI + OD, + \(\X i) mod n=i \ (\j<\X i. j mod n=i \ \P(B!j)) \ (\Y i P(B!(\Y i)) \ \Y i\ m+i)\, + \(\jj \ \Y j \ \Y j) \ \X i = \X i \ + \Y i = \Y i\, + \(\jj \ \X j = \X j \ \Y j = \Y j) \ + \Y i \ \Y i\, + \(\X i) mod n=i \ (\j<\X i. j mod n=i \ \P(B!j)) \ (\Y i P(B!(\Y i)) \ \Y i\ m+i) \ (\jY j \ \X i) \) + COEND + SAT [\ \iX i=i \ \Y i=m+i \,\\X=\X \ \Y=\Y\,\True\, + \\iX i) mod n=i \ (\j<\X i. j mod n=i \ \P(B!j)) \ + (\Y i P(B!(\Y i)) \ \Y i\ m+i) \ (\jY j \ \X i)\]" +apply(rule Parallel) +--{*5 subgoals left *} +apply force+ +apply clarify +apply simp +apply(rule While) + apply force + apply force + apply force + apply(rule_tac pre'="\ \X i mod n = i \ (\j. j<\X i \ j mod n = i \ \P(B!j)) \ (\Y i < n * q \ P (B!(\Y i))) \ \X i<\Y i\" in Conseq) + apply force + apply(rule subset_refl)+ + apply(rule Cond) + apply force + apply(rule Basic) + apply force + apply fastsimp + apply force + apply force + apply(rule Basic) + apply simp + apply clarify + apply simp + apply (case_tac "X x (j mod n) \ j") + apply (drule le_imp_less_or_eq) + apply (erule disjE) + apply (drule_tac j=j and n=n and i="j mod n" and a="X x (j mod n)" in mod_aux) + apply auto +done + +text {* Same but with a list as auxiliary variable: *} + +record Example3_list = + X :: "nat list" + Y :: "nat list" + +lemma Example3_list: "m mod n=0 \ \ (COBEGIN SCHEME [0\ijX!i < \Y!j) DO + IF P(B!(\X!i)) THEN \Y:=\Y[i:=\X!i] ELSE \X:= \X[i:=(\X!i)+ n] FI + OD, + \nX \ nY \ (\X!i) mod n=i \ (\j<\X!i. j mod n=i \ \P(B!j)) \ (\Y!i P(B!(\Y!i)) \ \Y!i\ m+i)\, + \(\jj \ \Y!j \ \Y!j) \ \X!i = \X!i \ + \Y!i = \Y!i \ length \X = length \X \ length \Y = length \Y\, + \(\jj \ \X!j = \X!j \ \Y!j = \Y!j) \ + \Y!i \ \Y!i \ length \X = length \X \ length \Y = length \Y\, + \(\X!i) mod n=i \ (\j<\X!i. j mod n=i \ \P(B!j)) \ (\Y!i P(B!(\Y!i)) \ \Y!i\ m+i) \ (\jY!j \ \X!i) \) COEND) + SAT [\nX \ nY \ (\iX!i=i \ \Y!i=m+i) \, + \\X=\X \ \Y=\Y\, + \True\, + \\iX!i) mod n=i \ (\j<\X!i. j mod n=i \ \P(B!j)) \ + (\Y!i P(B!(\Y!i)) \ \Y!i\ m+i) \ (\jY!j \ \X!i)\]" +apply(rule Parallel) +--{* 5 subgoals left *} +apply force+ +apply clarify +apply simp +apply(rule While) + apply force + apply force + apply force + apply(rule_tac pre'="\nX \ nY \ \X ! i mod n = i \ (\j. j < \X ! i \ j mod n = i \ \ P (B ! j)) \ (\Y ! i < n * q \ P (B ! (\Y ! i))) \ \X!i<\Y!i\" in Conseq) + apply force + apply(rule subset_refl)+ + apply(rule Cond) + apply force + apply(rule Basic) + apply force + apply force + apply force + apply force + apply(rule Basic) + apply simp + apply clarify + apply simp + apply(rule allI) + apply(rule impI)+ + apply(case_tac "X x ! i\ j") + apply(drule le_imp_less_or_eq) + apply(erule disjE) + apply(drule_tac j=j and n=n and i=i and a="X x ! i" in mod_aux) + apply auto +done + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/RG_Hoare.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/RG_Hoare.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,1375 @@ +header {* \section{The Proof System} *} + +theory RG_Hoare imports RG_Tran begin + +subsection {* Proof System for Component Programs *} + +declare Un_subset_iff [iff del] +declare Cons_eq_map_conv[iff] + +constdefs + stable :: "'a set \ ('a \ 'a) set \ bool" + "stable \ \f g. (\x y. x \ f \ (x, y) \ g \ y \ f)" + +inductive + rghoare :: "['a com, 'a set, ('a \ 'a) set, ('a \ 'a) set, 'a set] \ bool" + ("\ _ sat [_, _, _, _]" [60,0,0,0,0] 45) +where + Basic: "\ pre \ {s. f s \ post}; {(s,t). s \ pre \ (t=f s \ t=s)} \ guar; + stable pre rely; stable post rely \ + \ \ Basic f sat [pre, rely, guar, post]" + +| Seq: "\ \ P sat [pre, rely, guar, mid]; \ Q sat [mid, rely, guar, post] \ + \ \ Seq P Q sat [pre, rely, guar, post]" + +| Cond: "\ stable pre rely; \ P1 sat [pre \ b, rely, guar, post]; + \ P2 sat [pre \ -b, rely, guar, post]; \s. (s,s)\guar \ + \ \ Cond b P1 P2 sat [pre, rely, guar, post]" + +| While: "\ stable pre rely; (pre \ -b) \ post; stable post rely; + \ P sat [pre \ b, rely, guar, pre]; \s. (s,s)\guar \ + \ \ While b P sat [pre, rely, guar, post]" + +| Await: "\ stable pre rely; stable post rely; + \V. \ P sat [pre \ b \ {V}, {(s, t). s = t}, + UNIV, {s. (V, s) \ guar} \ post] \ + \ \ Await b P sat [pre, rely, guar, post]" + +| Conseq: "\ pre \ pre'; rely \ rely'; guar' \ guar; post' \ post; + \ P sat [pre', rely', guar', post'] \ + \ \ P sat [pre, rely, guar, post]" + +constdefs + Pre :: "'a rgformula \ 'a set" + "Pre x \ fst(snd x)" + Post :: "'a rgformula \ 'a set" + "Post x \ snd(snd(snd(snd x)))" + Rely :: "'a rgformula \ ('a \ 'a) set" + "Rely x \ fst(snd(snd x))" + Guar :: "'a rgformula \ ('a \ 'a) set" + "Guar x \ fst(snd(snd(snd x)))" + Com :: "'a rgformula \ 'a com" + "Com x \ fst x" + +subsection {* Proof System for Parallel Programs *} + +types 'a par_rgformula = "('a rgformula) list \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set" + +inductive + par_rghoare :: "('a rgformula) list \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set \ bool" + ("\ _ SAT [_, _, _, _]" [60,0,0,0,0] 45) +where + Parallel: + "\ \i (\j\{j. j j\i}. Guar(xs!j)) \ Rely(xs!i); + (\j\{j. j guar; + pre \ (\i\{i. ii\{i. i post; + \i Com(xs!i) sat [Pre(xs!i),Rely(xs!i),Guar(xs!i),Post(xs!i)] \ + \ \ xs SAT [pre, rely, guar, post]" + +section {* Soundness*} + +subsubsection {* Some previous lemmas *} + +lemma tl_of_assum_in_assum: + "(P, s) # (P, t) # xs \ assum (pre, rely) \ stable pre rely + \ (P, t) # xs \ assum (pre, rely)" +apply(simp add:assum_def) +apply clarify +apply(rule conjI) + apply(erule_tac x=0 in allE) + apply(simp (no_asm_use)only:stable_def) + apply(erule allE,erule allE,erule impE,assumption,erule mp) + apply(simp add:Env) +apply clarify +apply(erule_tac x="Suc i" in allE) +apply simp +done + +lemma etran_in_comm: + "(P, t) # xs \ comm(guar, post) \ (P, s) # (P, t) # xs \ comm(guar, post)" +apply(simp add:comm_def) +apply clarify +apply(case_tac i,simp+) +done + +lemma ctran_in_comm: + "\(s, s) \ guar; (Q, s) # xs \ comm(guar, post)\ + \ (P, s) # (Q, s) # xs \ comm(guar, post)" +apply(simp add:comm_def) +apply clarify +apply(case_tac i,simp+) +done + +lemma takecptn_is_cptn [rule_format, elim!]: + "\j. c \ cptn \ take (Suc j) c \ cptn" +apply(induct "c") + apply(force elim: cptn.cases) +apply clarify +apply(case_tac j) + apply simp + apply(rule CptnOne) +apply simp +apply(force intro:cptn.intros elim:cptn.cases) +done + +lemma dropcptn_is_cptn [rule_format,elim!]: + "\j cptn \ drop j c \ cptn" +apply(induct "c") + apply(force elim: cptn.cases) +apply clarify +apply(case_tac j,simp+) +apply(erule cptn.cases) + apply simp + apply force +apply force +done + +lemma takepar_cptn_is_par_cptn [rule_format,elim]: + "\j. c \ par_cptn \ take (Suc j) c \ par_cptn" +apply(induct "c") + apply(force elim: cptn.cases) +apply clarify +apply(case_tac j,simp) + apply(rule ParCptnOne) +apply(force intro:par_cptn.intros elim:par_cptn.cases) +done + +lemma droppar_cptn_is_par_cptn [rule_format]: + "\j par_cptn \ drop j c \ par_cptn" +apply(induct "c") + apply(force elim: par_cptn.cases) +apply clarify +apply(case_tac j,simp+) +apply(erule par_cptn.cases) + apply simp + apply force +apply force +done + +lemma tl_of_cptn_is_cptn: "\x # xs \ cptn; xs \ []\ \ xs \ cptn" +apply(subgoal_tac "1 < length (x # xs)") + apply(drule dropcptn_is_cptn,simp+) +done + +lemma not_ctran_None [rule_format]: + "\s. (None, s)#xs \ cptn \ (\i xs!i)" +apply(induct xs,simp+) +apply clarify +apply(erule cptn.cases,simp) + apply simp + apply(case_tac i,simp) + apply(rule Env) + apply simp +apply(force elim:ctran.cases) +done + +lemma cptn_not_empty [simp]:"[] \ cptn" +apply(force elim:cptn.cases) +done + +lemma etran_or_ctran [rule_format]: + "\m i. x\cptn \ m \ length x + \ (\i. Suc i < m \ \ x!i -c\ x!Suc i) \ Suc i < m + \ x!i -e\ x!Suc i" +apply(induct x,simp) +apply clarify +apply(erule cptn.cases,simp) + apply(case_tac i,simp) + apply(rule Env) + apply simp + apply(erule_tac x="m - 1" in allE) + apply(case_tac m,simp,simp) + apply(subgoal_tac "(\i. Suc i < nata \ (((P, t) # xs) ! i, xs ! i) \ ctran)") + apply force + apply clarify + apply(erule_tac x="Suc ia" in allE,simp) +apply(erule_tac x="0" and P="\j. ?H j \ (?J j) \ ctran" in allE,simp) +done + +lemma etran_or_ctran2 [rule_format]: + "\i. Suc i x\cptn \ (x!i -c\ x!Suc i \ \ x!i -e\ x!Suc i) + \ (x!i -e\ x!Suc i \ \ x!i -c\ x!Suc i)" +apply(induct x) + apply simp +apply clarify +apply(erule cptn.cases,simp) + apply(case_tac i,simp+) +apply(case_tac i,simp) + apply(force elim:etran.cases) +apply simp +done + +lemma etran_or_ctran2_disjI1: + "\ x\cptn; Suc i x!Suc i\ \ \ x!i -e\ x!Suc i" +by(drule etran_or_ctran2,simp_all) + +lemma etran_or_ctran2_disjI2: + "\ x\cptn; Suc i x!Suc i\ \ \ x!i -c\ x!Suc i" +by(drule etran_or_ctran2,simp_all) + +lemma not_ctran_None2 [rule_format]: + "\ (None, s) # xs \cptn; i \ \ ((None, s) # xs) ! i -c\ xs ! i" +apply(frule not_ctran_None,simp) +apply(case_tac i,simp) + apply(force elim:etranE) +apply simp +apply(rule etran_or_ctran2_disjI2,simp_all) +apply(force intro:tl_of_cptn_is_cptn) +done + +lemma Ex_first_occurrence [rule_format]: "P (n::nat) \ (\m. P m \ (\i P i))"; +apply(rule nat_less_induct) +apply clarify +apply(case_tac "\m. m \ P m") +apply auto +done + +lemma stability [rule_format]: + "\j k. x \ cptn \ stable p rely \ j\k \ k snd(x!j)\p \ + (\i. (Suc i) + (x!i -e\ x!(Suc i)) \ (snd(x!i), snd(x!(Suc i))) \ rely) \ + (\i. j\i \ i x!i -e\ x!Suc i) \ snd(x!k)\p \ fst(x!j)=fst(x!k)" +apply(induct x) + apply clarify + apply(force elim:cptn.cases) +apply clarify +apply(erule cptn.cases,simp) + apply simp + apply(case_tac k,simp,simp) + apply(case_tac j,simp) + apply(erule_tac x=0 in allE) + apply(erule_tac x="nat" and P="\j. (0\j) \ (?J j)" in allE,simp) + apply(subgoal_tac "t\p") + apply(subgoal_tac "(\i. i < length xs \ ((P, t) # xs) ! i -e\ xs ! i \ (snd (((P, t) # xs) ! i), snd (xs ! i)) \ rely)") + apply clarify + apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j)\etran" in allE,simp) + apply clarify + apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j) \ (?T j)\rely" in allE,simp) + apply(erule_tac x=0 and P="\j. (?H j) \ (?J j)\etran \ ?T j" in allE,simp) + apply(simp(no_asm_use) only:stable_def) + apply(erule_tac x=s in allE) + apply(erule_tac x=t in allE) + apply simp + apply(erule mp) + apply(erule mp) + apply(rule Env) + apply simp + apply(erule_tac x="nata" in allE) + apply(erule_tac x="nat" and P="\j. (?s\j) \ (?J j)" in allE,simp) + apply(subgoal_tac "(\i. i < length xs \ ((P, t) # xs) ! i -e\ xs ! i \ (snd (((P, t) # xs) ! i), snd (xs ! i)) \ rely)") + apply clarify + apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j)\etran" in allE,simp) + apply clarify + apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j) \ (?T j)\rely" in allE,simp) +apply(case_tac k,simp,simp) +apply(case_tac j) + apply(erule_tac x=0 and P="\j. (?H j) \ (?J j)\etran" in allE,simp) + apply(erule etran.cases,simp) +apply(erule_tac x="nata" in allE) +apply(erule_tac x="nat" and P="\j. (?s\j) \ (?J j)" in allE,simp) +apply(subgoal_tac "(\i. i < length xs \ ((Q, t) # xs) ! i -e\ xs ! i \ (snd (((Q, t) # xs) ! i), snd (xs ! i)) \ rely)") + apply clarify + apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j)\etran" in allE,simp) +apply clarify +apply(erule_tac x="Suc i" and P="\j. (?H j) \ (?J j) \ (?T j)\rely" in allE,simp) +done + +subsection {* Soundness of the System for Component Programs *} + +subsubsection {* Soundness of the Basic rule *} + +lemma unique_ctran_Basic [rule_format]: + "\s i. x \ cptn \ x ! 0 = (Some (Basic f), s) \ + Suc i x!i -c\ x!Suc i \ + (\j. Suc j i\j \ x!j -e\ x!Suc j)" +apply(induct x,simp) +apply simp +apply clarify +apply(erule cptn.cases,simp) + apply(case_tac i,simp+) + apply clarify + apply(case_tac j,simp) + apply(rule Env) + apply simp +apply clarify +apply simp +apply(case_tac i) + apply(case_tac j,simp,simp) + apply(erule ctran.cases,simp_all) + apply(force elim: not_ctran_None) +apply(ind_cases "((Some (Basic f), sa), Q, t) \ ctran" for sa Q t) +apply simp +apply(drule_tac i=nat in not_ctran_None,simp) +apply(erule etranE,simp) +done + +lemma exists_ctran_Basic_None [rule_format]: + "\s i. x \ cptn \ x ! 0 = (Some (Basic f), s) + \ i fst(x!i)=None \ (\j x!Suc j)" +apply(induct x,simp) +apply simp +apply clarify +apply(erule cptn.cases,simp) + apply(case_tac i,simp,simp) + apply(erule_tac x=nat in allE,simp) + apply clarify + apply(rule_tac x="Suc j" in exI,simp,simp) +apply clarify +apply(case_tac i,simp,simp) +apply(rule_tac x=0 in exI,simp) +done + +lemma Basic_sound: + " \pre \ {s. f s \ post}; {(s, t). s \ pre \ t = f s} \ guar; + stable pre rely; stable post rely\ + \ \ Basic f sat [pre, rely, guar, post]" +apply(unfold com_validity_def) +apply clarify +apply(simp add:comm_def) +apply(rule conjI) + apply clarify + apply(simp add:cp_def assum_def) + apply clarify + apply(frule_tac j=0 and k=i and p=pre in stability) + apply simp_all + apply(erule_tac x=ia in allE,simp) + apply(erule_tac i=i and f=f in unique_ctran_Basic,simp_all) + apply(erule subsetD,simp) + apply(case_tac "x!i") + apply clarify + apply(drule_tac s="Some (Basic f)" in sym,simp) + apply(thin_tac "\j. ?H j") + apply(force elim:ctran.cases) +apply clarify +apply(simp add:cp_def) +apply clarify +apply(frule_tac i="length x - 1" and f=f in exists_ctran_Basic_None,simp+) + apply(case_tac x,simp+) + apply(rule last_fst_esp,simp add:last_length) + apply (case_tac x,simp+) +apply(simp add:assum_def) +apply clarify +apply(frule_tac j=0 and k="j" and p=pre in stability) + apply simp_all + apply(erule_tac x=i in allE,simp) + apply(erule_tac i=j and f=f in unique_ctran_Basic,simp_all) +apply(case_tac "x!j") +apply clarify +apply simp +apply(drule_tac s="Some (Basic f)" in sym,simp) +apply(case_tac "x!Suc j",simp) +apply(rule ctran.cases,simp) +apply(simp_all) +apply(drule_tac c=sa in subsetD,simp) +apply clarify +apply(frule_tac j="Suc j" and k="length x - 1" and p=post in stability,simp_all) + apply(case_tac x,simp+) + apply(erule_tac x=i in allE) +apply(erule_tac i=j and f=f in unique_ctran_Basic,simp_all) + apply arith+ +apply(case_tac x) +apply(simp add:last_length)+ +done + +subsubsection{* Soundness of the Await rule *} + +lemma unique_ctran_Await [rule_format]: + "\s i. x \ cptn \ x ! 0 = (Some (Await b c), s) \ + Suc i x!i -c\ x!Suc i \ + (\j. Suc j i\j \ x!j -e\ x!Suc j)" +apply(induct x,simp+) +apply clarify +apply(erule cptn.cases,simp) + apply(case_tac i,simp+) + apply clarify + apply(case_tac j,simp) + apply(rule Env) + apply simp +apply clarify +apply simp +apply(case_tac i) + apply(case_tac j,simp,simp) + apply(erule ctran.cases,simp_all) + apply(force elim: not_ctran_None) +apply(ind_cases "((Some (Await b c), sa), Q, t) \ ctran" for sa Q t,simp) +apply(drule_tac i=nat in not_ctran_None,simp) +apply(erule etranE,simp) +done + +lemma exists_ctran_Await_None [rule_format]: + "\s i. x \ cptn \ x ! 0 = (Some (Await b c), s) + \ i fst(x!i)=None \ (\j x!Suc j)" +apply(induct x,simp+) +apply clarify +apply(erule cptn.cases,simp) + apply(case_tac i,simp+) + apply(erule_tac x=nat in allE,simp) + apply clarify + apply(rule_tac x="Suc j" in exI,simp,simp) +apply clarify +apply(case_tac i,simp,simp) +apply(rule_tac x=0 in exI,simp) +done + +lemma Star_imp_cptn: + "(P, s) -c*\ (R, t) \ \l \ cp P s. (last l)=(R, t) + \ (\i. Suc i l!i -c\ l!Suc i)" +apply (erule converse_rtrancl_induct2) + apply(rule_tac x="[(R,t)]" in bexI) + apply simp + apply(simp add:cp_def) + apply(rule CptnOne) +apply clarify +apply(rule_tac x="(a, b)#l" in bexI) + apply (rule conjI) + apply(case_tac l,simp add:cp_def) + apply(simp add:last_length) + apply clarify +apply(case_tac i,simp) +apply(simp add:cp_def) +apply force +apply(simp add:cp_def) + apply(case_tac l) + apply(force elim:cptn.cases) +apply simp +apply(erule CptnComp) +apply clarify +done + +lemma Await_sound: + "\stable pre rely; stable post rely; + \V. \ P sat [pre \ b \ {s. s = V}, {(s, t). s = t}, + UNIV, {s. (V, s) \ guar} \ post] \ + \ P sat [pre \ b \ {s. s = V}, {(s, t). s = t}, + UNIV, {s. (V, s) \ guar} \ post] \ + \ \ Await b P sat [pre, rely, guar, post]" +apply(unfold com_validity_def) +apply clarify +apply(simp add:comm_def) +apply(rule conjI) + apply clarify + apply(simp add:cp_def assum_def) + apply clarify + apply(frule_tac j=0 and k=i and p=pre in stability,simp_all) + apply(erule_tac x=ia in allE,simp) + apply(subgoal_tac "x\ cp (Some(Await b P)) s") + apply(erule_tac i=i in unique_ctran_Await,force,simp_all) + apply(simp add:cp_def) +--{* here starts the different part. *} + apply(erule ctran.cases,simp_all) + apply(drule Star_imp_cptn) + apply clarify + apply(erule_tac x=sa in allE) + apply clarify + apply(erule_tac x=sa in allE) + apply(drule_tac c=l in subsetD) + apply (simp add:cp_def) + apply clarify + apply(erule_tac x=ia and P="\i. ?H i \ (?J i,?I i)\ctran" in allE,simp) + apply(erule etranE,simp) + apply simp +apply clarify +apply(simp add:cp_def) +apply clarify +apply(frule_tac i="length x - 1" in exists_ctran_Await_None,force) + apply (case_tac x,simp+) + apply(rule last_fst_esp,simp add:last_length) + apply(case_tac x, (simp add:cptn_not_empty)+) +apply clarify +apply(simp add:assum_def) +apply clarify +apply(frule_tac j=0 and k="j" and p=pre in stability,simp_all) + apply(erule_tac x=i in allE,simp) + apply(erule_tac i=j in unique_ctran_Await,force,simp_all) +apply(case_tac "x!j") +apply clarify +apply simp +apply(drule_tac s="Some (Await b P)" in sym,simp) +apply(case_tac "x!Suc j",simp) +apply(rule ctran.cases,simp) +apply(simp_all) +apply(drule Star_imp_cptn) +apply clarify +apply(erule_tac x=sa in allE) +apply clarify +apply(erule_tac x=sa in allE) +apply(drule_tac c=l in subsetD) + apply (simp add:cp_def) + apply clarify + apply(erule_tac x=i and P="\i. ?H i \ (?J i,?I i)\ctran" in allE,simp) + apply(erule etranE,simp) +apply simp +apply clarify +apply(frule_tac j="Suc j" and k="length x - 1" and p=post in stability,simp_all) + apply(case_tac x,simp+) + apply(erule_tac x=i in allE) +apply(erule_tac i=j in unique_ctran_Await,force,simp_all) + apply arith+ +apply(case_tac x) +apply(simp add:last_length)+ +done + +subsubsection{* Soundness of the Conditional rule *} + +lemma Cond_sound: + "\ stable pre rely; \ P1 sat [pre \ b, rely, guar, post]; + \ P2 sat [pre \ - b, rely, guar, post]; \s. (s,s)\guar\ + \ \ (Cond b P1 P2) sat [pre, rely, guar, post]" +apply(unfold com_validity_def) +apply clarify +apply(simp add:cp_def comm_def) +apply(case_tac "\i. Suc i x!i -c\ x!Suc i") + prefer 2 + apply simp + apply clarify + apply(frule_tac j="0" and k="length x - 1" and p=pre in stability,simp+) + apply(case_tac x,simp+) + apply(simp add:assum_def) + apply(simp add:assum_def) + apply(erule_tac m="length x" in etran_or_ctran,simp+) + apply(case_tac x, (simp add:last_length)+) +apply(erule exE) +apply(drule_tac n=i and P="\i. ?H i \ (?J i,?I i)\ ctran" in Ex_first_occurrence) +apply clarify +apply (simp add:assum_def) +apply(frule_tac j=0 and k="m" and p=pre in stability,simp+) + apply(erule_tac m="Suc m" in etran_or_ctran,simp+) +apply(erule ctran.cases,simp_all) + apply(erule_tac x="sa" in allE) + apply(drule_tac c="drop (Suc m) x" in subsetD) + apply simp + apply clarify + apply simp + apply clarify + apply(case_tac "i\m") + apply(drule le_imp_less_or_eq) + apply(erule disjE) + apply(erule_tac x=i in allE, erule impE, assumption) + apply simp+ + apply(erule_tac x="i - (Suc m)" and P="\j. ?H j \ ?J j \ (?I j)\guar" in allE) + apply(subgoal_tac "(Suc m)+(i - Suc m) \ length x") + apply(subgoal_tac "(Suc m)+Suc (i - Suc m) \ length x") + apply(rotate_tac -2) + apply simp + apply arith + apply arith +apply(case_tac "length (drop (Suc m) x)",simp) +apply(erule_tac x="sa" in allE) +back +apply(drule_tac c="drop (Suc m) x" in subsetD,simp) + apply clarify +apply simp +apply clarify +apply(case_tac "i\m") + apply(drule le_imp_less_or_eq) + apply(erule disjE) + apply(erule_tac x=i in allE, erule impE, assumption) + apply simp + apply simp +apply(erule_tac x="i - (Suc m)" and P="\j. ?H j \ ?J j \ (?I j)\guar" in allE) +apply(subgoal_tac "(Suc m)+(i - Suc m) \ length x") + apply(subgoal_tac "(Suc m)+Suc (i - Suc m) \ length x") + apply(rotate_tac -2) + apply simp + apply arith +apply arith +done + +subsubsection{* Soundness of the Sequential rule *} + +inductive_cases Seq_cases [elim!]: "(Some (Seq P Q), s) -c\ t" + +lemma last_lift_not_None: "fst ((lift Q) ((x#xs)!(length xs))) \ None" +apply(subgoal_tac "length xs cptn_mod \ \s P. x !0=(Some (Seq P Q), s) \ + (\iSome Q) \ + (\xs\ cp (Some P) s. x=map (lift Q) xs)" +apply(erule cptn_mod.induct) +apply(unfold cp_def) +apply safe +apply simp_all + apply(simp add:lift_def) + apply(rule_tac x="[(Some Pa, sa)]" in exI,simp add:CptnOne) + apply(subgoal_tac "(\i < Suc (length xs). fst (((Some (Seq Pa Q), t) # xs) ! i) \ Some Q)") + apply clarify + apply(rule_tac x="(Some Pa, sa) #(Some Pa, t) # zs" in exI,simp) + apply(rule conjI,erule CptnEnv) + apply(simp (no_asm_use) add:lift_def) + apply clarify + apply(erule_tac x="Suc i" in allE, simp) + apply(ind_cases "((Some (Seq Pa Q), sa), None, t) \ ctran" for Pa sa t) + apply(rule_tac x="(Some P, sa) # xs" in exI, simp add:cptn_iff_cptn_mod lift_def) +apply(erule_tac x="length xs" in allE, simp) +apply(simp only:Cons_lift_append) +apply(subgoal_tac "length xs < length ((Some P, sa) # xs)") + apply(simp only :nth_append length_map last_length nth_map) + apply(case_tac "last((Some P, sa) # xs)") + apply(simp add:lift_def) +apply simp +done +declare map_eq_Cons_conv [simp del] Cons_eq_map_conv [simp del] + +lemma Seq_sound2 [rule_format]: + "x \ cptn \ \s P i. x!0=(Some (Seq P Q), s) \ i fst(x!i)=Some Q \ + (\j(Some Q)) \ + (\xs ys. xs \ cp (Some P) s \ length xs=Suc i + \ ys \ cp (Some Q) (snd(xs !i)) \ x=(map (lift Q) xs)@tl ys)" +apply(erule cptn.induct) +apply(unfold cp_def) +apply safe +apply simp_all + apply(case_tac i,simp+) + apply(erule allE,erule impE,assumption,simp) + apply clarify + apply(subgoal_tac "(\j < nat. fst (((Some (Seq Pa Q), t) # xs) ! j) \ Some Q)",clarify) + prefer 2 + apply force + apply(case_tac xsa,simp,simp) + apply(rule_tac x="(Some Pa, sa) #(Some Pa, t) # list" in exI,simp) + apply(rule conjI,erule CptnEnv) + apply(simp (no_asm_use) add:lift_def) + apply(rule_tac x=ys in exI,simp) +apply(ind_cases "((Some (Seq Pa Q), sa), t) \ ctran" for Pa sa t) + apply simp + apply(rule_tac x="(Some Pa, sa)#[(None, ta)]" in exI,simp) + apply(rule conjI) + apply(drule_tac xs="[]" in CptnComp,force simp add:CptnOne,simp) + apply(case_tac i, simp+) + apply(case_tac nat,simp+) + apply(rule_tac x="(Some Q,ta)#xs" in exI,simp add:lift_def) + apply(case_tac nat,simp+) + apply(force) +apply(case_tac i, simp+) +apply(case_tac nat,simp+) +apply(erule_tac x="Suc nata" in allE,simp) +apply clarify +apply(subgoal_tac "(\j Some Q)",clarify) + prefer 2 + apply clarify + apply force +apply(rule_tac x="(Some Pa, sa)#(Some P2, ta)#(tl xsa)" in exI,simp) +apply(rule conjI,erule CptnComp) +apply(rule nth_tl_if,force,simp+) +apply(rule_tac x=ys in exI,simp) +apply(rule conjI) +apply(rule nth_tl_if,force,simp+) + apply(rule tl_zero,simp+) + apply force +apply(rule conjI,simp add:lift_def) +apply(subgoal_tac "lift Q (Some P2, ta) =(Some (Seq P2 Q), ta)") + apply(simp add:Cons_lift del:map.simps) + apply(rule nth_tl_if) + apply force + apply simp+ +apply(simp add:lift_def) +done +(* +lemma last_lift_not_None3: "fst (last (map (lift Q) (x#xs))) \ None" +apply(simp only:last_length [THEN sym]) +apply(subgoal_tac "length xs None" +apply(simp only:last_length [THEN sym]) +apply(subgoal_tac "length xs\ P sat [pre, rely, guar, mid]; \ Q sat [mid, rely, guar, post]\ + \ \ Seq P Q sat [pre, rely, guar, post]" +apply(unfold com_validity_def) +apply clarify +apply(case_tac "\i[]") + apply(drule last_conv_nth) + apply (simp del:map.simps) + apply(simp only:last_lift_not_None) + apply simp +--{* @{text "\i[]") + apply(drule last_conv_nth,simp) + apply(rule conjI) + apply(erule mp) + apply(case_tac "xs!m") + apply(case_tac "fst(xs!m)",simp) + apply(simp add:lift_def nth_append) + apply clarify + apply(erule_tac x="m+i" in allE) + back + back + apply(case_tac ys,(simp add:nth_append)+) + apply (case_tac i, (simp add:snd_lift)+) + apply(erule mp) + apply(case_tac "xs!m") + apply(force elim:etran.cases intro:Env simp add:lift_def) + apply simp +apply simp +apply clarify +apply(rule conjI,clarify) + apply(case_tac "i[]") + apply(drule last_conv_nth) + apply(simp add: snd_lift nth_append) + apply(rule conjI,clarify) + apply(case_tac ys,simp+) + apply clarify + apply(case_tac ys,simp+) +done + +subsubsection{* Soundness of the While rule *} + +lemma last_append[rule_format]: + "\xs. ys\[] \ ((xs@ys)!(length (xs@ys) - (Suc 0)))=(ys!(length ys - (Suc 0)))" +apply(induct ys) + apply simp +apply clarify +apply (simp add:nth_append length_append) +done + +lemma assum_after_body: + "\ \ P sat [pre \ b, rely, guar, pre]; + (Some P, s) # xs \ cptn_mod; fst (last ((Some P, s) # xs)) = None; s \ b; + (Some (While b P), s) # (Some (Seq P (While b P)), s) # + map (lift (While b P)) xs @ ys \ assum (pre, rely)\ + \ (Some (While b P), snd (last ((Some P, s) # xs))) # ys \ assum (pre, rely)" +apply(simp add:assum_def com_validity_def cp_def cptn_iff_cptn_mod) +apply clarify +apply(erule_tac x=s in allE) +apply(drule_tac c="(Some P, s) # xs" in subsetD,simp) + apply clarify + apply(erule_tac x="Suc i" in allE) + apply simp + apply(simp add:Cons_lift_append nth_append snd_lift del:map.simps) + apply(erule mp) + apply(erule etranE,simp) + apply(case_tac "fst(((Some P, s) # xs) ! i)") + apply(force intro:Env simp add:lift_def) + apply(force intro:Env simp add:lift_def) +apply(rule conjI) + apply clarify + apply(simp add:comm_def last_length) +apply clarify +apply(rule conjI) + apply(simp add:comm_def) +apply clarify +apply(erule_tac x="Suc(length xs + i)" in allE,simp) +apply(case_tac i, simp add:nth_append Cons_lift_append snd_lift del:map.simps) + apply(simp add:last_length) + apply(erule mp) + apply(case_tac "last xs") + apply(simp add:lift_def) +apply(simp add:Cons_lift_append nth_append snd_lift del:map.simps) +done + +lemma While_sound_aux [rule_format]: + "\ pre \ - b \ post; \ P sat [pre \ b, rely, guar, pre]; \s. (s, s) \ guar; + stable pre rely; stable post rely; x \ cptn_mod \ + \ \s xs. x=(Some(While b P),s)#xs \ x\assum(pre, rely) \ x \ comm (guar, post)" +apply(erule cptn_mod.induct) +apply safe +apply (simp_all del:last.simps) +--{* 5 subgoals left *} +apply(simp add:comm_def) +--{* 4 subgoals left *} +apply(rule etran_in_comm) +apply(erule mp) +apply(erule tl_of_assum_in_assum,simp) +--{* While-None *} +apply(ind_cases "((Some (While b P), s), None, t) \ ctran" for s t) +apply(simp add:comm_def) +apply(simp add:cptn_iff_cptn_mod [THEN sym]) +apply(rule conjI,clarify) + apply(force simp add:assum_def) +apply clarify +apply(rule conjI, clarify) + apply(case_tac i,simp,simp) + apply(force simp add:not_ctran_None2) +apply(subgoal_tac "\i. Suc i < length ((None, t) # xs) \ (((None, t) # xs) ! i, ((None, t) # xs) ! Suc i)\ etran") + prefer 2 + apply clarify + apply(rule_tac m="length ((None, s) # xs)" in etran_or_ctran,simp+) + apply(erule not_ctran_None2,simp) + apply simp+ +apply(frule_tac j="0" and k="length ((None, s) # xs) - 1" and p=post in stability,simp+) + apply(force simp add:assum_def subsetD) + apply(simp add:assum_def) + apply clarify + apply(erule_tac x="i" in allE,simp) + apply(erule_tac x="Suc i" in allE,simp) + apply simp +apply clarify +apply (simp add:last_length) +--{* WhileOne *} +apply(thin_tac "P = While b P \ ?Q") +apply(rule ctran_in_comm,simp) +apply(simp add:Cons_lift del:map.simps) +apply(simp add:comm_def del:map.simps) +apply(rule conjI) + apply clarify + apply(case_tac "fst(((Some P, sa) # xs) ! i)") + apply(case_tac "((Some P, sa) # xs) ! i") + apply (simp add:lift_def) + apply(ind_cases "(Some (While b P), ba) -c\ t" for ba t) + apply simp + apply simp + apply(simp add:snd_lift del:map.simps) + apply(simp only:com_validity_def cp_def cptn_iff_cptn_mod) + apply(erule_tac x=sa in allE) + apply(drule_tac c="(Some P, sa) # xs" in subsetD) + apply (simp add:assum_def del:map.simps) + apply clarify + apply(erule_tac x="Suc ia" in allE,simp add:snd_lift del:map.simps) + apply(erule mp) + apply(case_tac "fst(((Some P, sa) # xs) ! ia)") + apply(erule etranE,simp add:lift_def) + apply(rule Env) + apply(erule etranE,simp add:lift_def) + apply(rule Env) + apply (simp add:comm_def del:map.simps) + apply clarify + apply(erule allE,erule impE,assumption) + apply(erule mp) + apply(case_tac "((Some P, sa) # xs) ! i") + apply(case_tac "xs!i") + apply(simp add:lift_def) + apply(case_tac "fst(xs!i)") + apply force + apply force +--{* last=None *} +apply clarify +apply(subgoal_tac "(map (lift (While b P)) ((Some P, sa) # xs))\[]") + apply(drule last_conv_nth) + apply (simp del:map.simps) + apply(simp only:last_lift_not_None) +apply simp +--{* WhileMore *} +apply(thin_tac "P = While b P \ ?Q") +apply(rule ctran_in_comm,simp del:last.simps) +--{* metiendo la hipotesis antes de dividir la conclusion. *} +apply(subgoal_tac "(Some (While b P), snd (last ((Some P, sa) # xs))) # ys \ assum (pre, rely)") + apply (simp del:last.simps) + prefer 2 + apply(erule assum_after_body) + apply (simp del:last.simps)+ +--{* lo de antes. *} +apply(simp add:comm_def del:map.simps last.simps) +apply(rule conjI) + apply clarify + apply(simp only:Cons_lift_append) + apply(case_tac "i t" for ba t) + apply simp + apply simp + apply(simp add:snd_lift del:map.simps last.simps) + apply(thin_tac " \i. i < length ys \ ?P i") + apply(simp only:com_validity_def cp_def cptn_iff_cptn_mod) + apply(erule_tac x=sa in allE) + apply(drule_tac c="(Some P, sa) # xs" in subsetD) + apply (simp add:assum_def del:map.simps last.simps) + apply clarify + apply(erule_tac x="Suc ia" in allE,simp add:nth_append snd_lift del:map.simps last.simps, erule mp) + apply(case_tac "fst(((Some P, sa) # xs) ! ia)") + apply(erule etranE,simp add:lift_def) + apply(rule Env) + apply(erule etranE,simp add:lift_def) + apply(rule Env) + apply (simp add:comm_def del:map.simps) + apply clarify + apply(erule allE,erule impE,assumption) + apply(erule mp) + apply(case_tac "((Some P, sa) # xs) ! i") + apply(case_tac "xs!i") + apply(simp add:lift_def) + apply(case_tac "fst(xs!i)") + apply force + apply force +--{* @{text "i \ length xs"} *} +apply(subgoal_tac "i-length xs length xs"} *} +apply(case_tac "i-length xs") + apply arith +apply(simp add:nth_append del:map.simps last.simps) +apply(rotate_tac -3) +apply(subgoal_tac "i- Suc (length xs)=nat") + prefer 2 + apply arith +apply simp +--{* last=None *} +apply clarify +apply(case_tac ys) + apply(simp add:Cons_lift del:map.simps last.simps) + apply(subgoal_tac "(map (lift (While b P)) ((Some P, sa) # xs))\[]") + apply(drule last_conv_nth) + apply (simp del:map.simps) + apply(simp only:last_lift_not_None) + apply simp +apply(subgoal_tac "((Some (Seq P (While b P)), sa) # map (lift (While b P)) xs @ ys)\[]") + apply(drule last_conv_nth) + apply (simp del:map.simps last.simps) + apply(simp add:nth_append del:last.simps) + apply(subgoal_tac "((Some (While b P), snd (last ((Some P, sa) # xs))) # a # list)\[]") + apply(drule last_conv_nth) + apply (simp del:map.simps last.simps) + apply simp +apply simp +done + +lemma While_sound: + "\stable pre rely; pre \ - b \ post; stable post rely; + \ P sat [pre \ b, rely, guar, pre]; \s. (s,s)\guar\ + \ \ While b P sat [pre, rely, guar, post]" +apply(unfold com_validity_def) +apply clarify +apply(erule_tac xs="tl x" in While_sound_aux) + apply(simp add:com_validity_def) + apply force + apply simp_all +apply(simp add:cptn_iff_cptn_mod cp_def) +apply(simp add:cp_def) +apply clarify +apply(rule nth_equalityI) + apply simp_all + apply(case_tac x,simp+) +apply clarify +apply(case_tac i,simp+) +apply(case_tac x,simp+) +done + +subsubsection{* Soundness of the Rule of Consequence *} + +lemma Conseq_sound: + "\pre \ pre'; rely \ rely'; guar' \ guar; post' \ post; + \ P sat [pre', rely', guar', post']\ + \ \ P sat [pre, rely, guar, post]" +apply(simp add:com_validity_def assum_def comm_def) +apply clarify +apply(erule_tac x=s in allE) +apply(drule_tac c=x in subsetD) + apply force +apply force +done + +subsubsection {* Soundness of the system for sequential component programs *} + +theorem rgsound: + "\ P sat [pre, rely, guar, post] \ \ P sat [pre, rely, guar, post]" +apply(erule rghoare.induct) + apply(force elim:Basic_sound) + apply(force elim:Seq_sound) + apply(force elim:Cond_sound) + apply(force elim:While_sound) + apply(force elim:Await_sound) +apply(erule Conseq_sound,simp+) +done + +subsection {* Soundness of the System for Parallel Programs *} + +constdefs + ParallelCom :: "('a rgformula) list \ 'a par_com" + "ParallelCom Ps \ map (Some \ fst) Ps" + +lemma two: + "\ \i (\j\{j. j < length xs \ j \ i}. Guar (xs ! j)) + \ Rely (xs ! i); + pre \ (\i\{i. i < length xs}. Pre (xs ! i)); + \i Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)]; + length xs=length clist; x \ par_cp (ParallelCom xs) s; x\par_assum(pre, rely); + \icp (Some(Com(xs!i))) s; x \ clist \ + \ \j i. i Suc j (clist!i!j) -c\ (clist!i!Suc j) + \ (snd(clist!i!j), snd(clist!i!Suc j)) \ Guar(xs!i)" +apply(unfold par_cp_def) +apply (rule ccontr) +--{* By contradiction: *} +apply (simp del: Un_subset_iff) +apply(erule exE) +--{* the first c-tran that does not satisfy the guarantee-condition is from @{text "\_i"} at step @{text "m"}. *} +apply(drule_tac n=j and P="\j. \i. ?H i j" in Ex_first_occurrence) +apply(erule exE) +apply clarify +--{* @{text "\_i \ A(pre, rely_1)"} *} +apply(subgoal_tac "take (Suc (Suc m)) (clist!i) \ assum(Pre(xs!i), Rely(xs!i))") +--{* but this contradicts @{text "\ \_i sat [pre_i,rely_i,guar_i,post_i]"} *} + apply(erule_tac x=i and P="\i. ?H i \ \ (?J i) sat [?I i,?K i,?M i,?N i]" in allE,erule impE,assumption) + apply(simp add:com_validity_def) + apply(erule_tac x=s in allE) + apply(simp add:cp_def comm_def) + apply(drule_tac c="take (Suc (Suc m)) (clist ! i)" in subsetD) + apply simp + apply (blast intro: takecptn_is_cptn) + apply simp + apply clarify + apply(erule_tac x=m and P="\j. ?I j \ ?J j \ ?H j" in allE) + apply (simp add:conjoin_def same_length_def) +apply(simp add:assum_def del: Un_subset_iff) +apply(rule conjI) + apply(erule_tac x=i and P="\j. ?H j \ ?I j \cp (?K j) (?J j)" in allE) + apply(simp add:cp_def par_assum_def) + apply(drule_tac c="s" in subsetD,simp) + apply simp +apply clarify +apply(erule_tac x=i and P="\j. ?H j \ ?M \ UNION (?S j) (?T j) \ (?L j)" in allE) +apply(simp del: Un_subset_iff) +apply(erule subsetD) +apply simp +apply(simp add:conjoin_def compat_label_def) +apply clarify +apply(erule_tac x=ia and P="\j. ?H j \ (?P j) \ ?Q j" in allE,simp) +--{* each etran in @{text "\_1[0\m]"} corresponds to *} +apply(erule disjE) +--{* a c-tran in some @{text "\_{ib}"} *} + apply clarify + apply(case_tac "i=ib",simp) + apply(erule etranE,simp) + apply(erule_tac x="ib" and P="\i. ?H i \ (?I i) \ (?J i)" in allE) + apply (erule etranE) + apply(case_tac "ia=m",simp) + apply simp + apply(erule_tac x=ia and P="\j. ?H j \ (\ i. ?P i j)" in allE) + apply(subgoal_tac "ia"}, +therefore it satisfies @{text "rely \ guar_{ib}"} *} +apply (force simp add:par_assum_def same_state_def) +done + + +lemma three [rule_format]: + "\ xs\[]; \i (\j\{j. j < length xs \ j \ i}. Guar (xs ! j)) + \ Rely (xs ! i); + pre \ (\i\{i. i < length xs}. Pre (xs ! i)); + \i Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)]; + length xs=length clist; x \ par_cp (ParallelCom xs) s; x \ par_assum(pre, rely); + \icp (Some(Com(xs!i))) s; x \ clist \ + \ \j i. i Suc j (clist!i!j) -e\ (clist!i!Suc j) + \ (snd(clist!i!j), snd(clist!i!Suc j)) \ rely \ (\j\{j. j < length xs \ j \ i}. Guar (xs ! j))" +apply(drule two) + apply simp_all +apply clarify +apply(simp add:conjoin_def compat_label_def) +apply clarify +apply(erule_tac x=j and P="\j. ?H j \ (?J j \ (\i. ?P i j)) \ ?I j" in allE,simp) +apply(erule disjE) + prefer 2 + apply(force simp add:same_state_def par_assum_def) +apply clarify +apply(case_tac "i=ia",simp) + apply(erule etranE,simp) +apply(erule_tac x="ia" and P="\i. ?H i \ (?I i) \ (?J i)" in allE,simp) +apply(erule_tac x=j and P="\j. \i. ?S j i \ (?I j i, ?H j i)\ ctran \ (?P i j)" in allE) +apply(erule_tac x=ia and P="\j. ?S j \ (?I j, ?H j)\ ctran \ (?P j)" in allE) +apply(simp add:same_state_def) +apply(erule_tac x=i and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in all_dupE) +apply(erule_tac x=ia and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in allE,simp) +done + +lemma four: + "\xs\[]; \i < length xs. rely \ (\j\{j. j < length xs \ j \ i}. Guar (xs ! j)) + \ Rely (xs ! i); + (\j\{j. j < length xs}. Guar (xs ! j)) \ guar; + pre \ (\i\{i. i < length xs}. Pre (xs ! i)); + \i < length xs. + \ Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)]; + x \ par_cp (ParallelCom xs) s; x \ par_assum (pre, rely); Suc i < length x; + x ! i -pc\ x ! Suc i\ + \ (snd (x ! i), snd (x ! Suc i)) \ guar" +apply(simp add: ParallelCom_def del: Un_subset_iff) +apply(subgoal_tac "(map (Some \ fst) xs)\[]") + prefer 2 + apply simp +apply(frule rev_subsetD) + apply(erule one [THEN equalityD1]) +apply(erule subsetD) +apply (simp del: Un_subset_iff) +apply clarify +apply(drule_tac pre=pre and rely=rely and x=x and s=s and xs=xs and clist=clist in two) +apply(assumption+) + apply(erule sym) + apply(simp add:ParallelCom_def) + apply assumption + apply(simp add:Com_def) + apply assumption +apply(simp add:conjoin_def same_program_def) +apply clarify +apply(erule_tac x=i and P="\j. ?H j \ fst(?I j)=(?J j)" in all_dupE) +apply(erule_tac x="Suc i" and P="\j. ?H j \ fst(?I j)=(?J j)" in allE) +apply(erule par_ctranE,simp) +apply(erule_tac x=i and P="\j. \i. ?S j i \ (?I j i, ?H j i)\ ctran \ (?P i j)" in allE) +apply(erule_tac x=ia and P="\j. ?S j \ (?I j, ?H j)\ ctran \ (?P j)" in allE) +apply(rule_tac x=ia in exI) +apply(simp add:same_state_def) +apply(erule_tac x=ia and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in all_dupE,simp) +apply(erule_tac x=ia and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in allE,simp) +apply(erule_tac x=i and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in all_dupE) +apply(erule_tac x=i and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in all_dupE,simp) +apply(erule_tac x="Suc i" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) +apply(erule mp) +apply(subgoal_tac "r=fst(clist ! ia ! Suc i)",simp) +apply(drule_tac i=ia in list_eq_if) +back +apply simp_all +done + +lemma parcptn_not_empty [simp]:"[] \ par_cptn" +apply(force elim:par_cptn.cases) +done + +lemma five: + "\xs\[]; \i (\j\{j. j < length xs \ j \ i}. Guar (xs ! j)) + \ Rely (xs ! i); + pre \ (\i\{i. i < length xs}. Pre (xs ! i)); + (\i\{i. i < length xs}. Post (xs ! i)) \ post; + \i < length xs. + \ Com (xs ! i) sat [Pre (xs ! i), Rely (xs ! i), Guar (xs ! i), Post (xs ! i)]; + x \ par_cp (ParallelCom xs) s; x \ par_assum (pre, rely); + All_None (fst (last x)) \ \ snd (last x) \ post" +apply(simp add: ParallelCom_def del: Un_subset_iff) +apply(subgoal_tac "(map (Some \ fst) xs)\[]") + prefer 2 + apply simp +apply(frule rev_subsetD) + apply(erule one [THEN equalityD1]) +apply(erule subsetD) +apply(simp del: Un_subset_iff) +apply clarify +apply(subgoal_tac "\iassum(Pre(xs!i), Rely(xs!i))") + apply(erule_tac x=i and P="\i. ?H i \ \ (?J i) sat [?I i,?K i,?M i,?N i]" in allE,erule impE,assumption) + apply(simp add:com_validity_def) + apply(erule_tac x=s in allE) + apply(erule_tac x=i and P="\j. ?H j \ (?I j) \ cp (?J j) s" in allE,simp) + apply(drule_tac c="clist!i" in subsetD) + apply (force simp add:Com_def) + apply(simp add:comm_def conjoin_def same_program_def del:last.simps) + apply clarify + apply(erule_tac x="length x - 1" and P="\j. ?H j \ fst(?I j)=(?J j)" in allE) + apply (simp add:All_None_def same_length_def) + apply(erule_tac x=i and P="\j. ?H j \ length(?J j)=(?K j)" in allE) + apply(subgoal_tac "length x - 1 < length x",simp) + apply(case_tac "x\[]") + apply(simp add: last_conv_nth) + apply(erule_tac x="clist!i" in ballE) + apply(simp add:same_state_def) + apply(subgoal_tac "clist!i\[]") + apply(simp add: last_conv_nth) + apply(case_tac x) + apply (force simp add:par_cp_def) + apply (force simp add:par_cp_def) + apply force + apply (force simp add:par_cp_def) + apply(case_tac x) + apply (force simp add:par_cp_def) + apply (force simp add:par_cp_def) +apply clarify +apply(simp add:assum_def) +apply(rule conjI) + apply(simp add:conjoin_def same_state_def par_cp_def) + apply clarify + apply(erule_tac x=ia and P="\j. (?T j) \ (\i. (?H j i) \ (snd (?d j i))=(snd (?e j i)))" in allE,simp) + apply(erule_tac x=0 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(case_tac x,simp+) + apply (simp add:par_assum_def) + apply clarify + apply(drule_tac c="snd (clist ! ia ! 0)" in subsetD) + apply assumption + apply simp +apply clarify +apply(erule_tac x=ia in all_dupE) +apply(rule subsetD, erule mp, assumption) +apply(erule_tac pre=pre and rely=rely and x=x and s=s in three) + apply(erule_tac x=ic in allE,erule mp) + apply simp_all + apply(simp add:ParallelCom_def) + apply(force simp add:Com_def) +apply(simp add:conjoin_def same_length_def) +done + +lemma ParallelEmpty [rule_format]: + "\i s. x \ par_cp (ParallelCom []) s \ + Suc i < length x \ (x ! i, x ! Suc i) \ par_ctran" +apply(induct_tac x) + apply(simp add:par_cp_def ParallelCom_def) +apply clarify +apply(case_tac list,simp,simp) +apply(case_tac i) + apply(simp add:par_cp_def ParallelCom_def) + apply(erule par_ctranE,simp) +apply(simp add:par_cp_def ParallelCom_def) +apply clarify +apply(erule par_cptn.cases,simp) + apply simp +apply(erule par_ctranE) +back +apply simp +done + +theorem par_rgsound: + "\ c SAT [pre, rely, guar, post] \ + \ (ParallelCom c) SAT [pre, rely, guar, post]" +apply(erule par_rghoare.induct) +apply(case_tac xs,simp) + apply(simp add:par_com_validity_def par_comm_def) + apply clarify + apply(case_tac "post=UNIV",simp) + apply clarify + apply(drule ParallelEmpty) + apply assumption + apply simp + apply clarify + apply simp +apply(subgoal_tac "xs\[]") + prefer 2 + apply simp +apply(thin_tac "xs = a # list") +apply(simp add:par_com_validity_def par_comm_def) +apply clarify +apply(rule conjI) + apply clarify + apply(erule_tac pre=pre and rely=rely and guar=guar and x=x and s=s and xs=xs in four) + apply(assumption+) + apply clarify + apply (erule allE, erule impE, assumption,erule rgsound) + apply(assumption+) +apply clarify +apply(erule_tac pre=pre and rely=rely and post=post and x=x and s=s and xs=xs in five) + apply(assumption+) + apply clarify + apply (erule allE, erule impE, assumption,erule rgsound) + apply(assumption+) +done + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/RG_Syntax.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/RG_Syntax.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,95 @@ +header {* \section{Concrete Syntax} *} + +theory RG_Syntax +imports RG_Hoare Quote_Antiquote +begin + +syntax + "_Assign" :: "idt \ 'b \ 'a com" ("(\_ :=/ _)" [70, 65] 61) + "_skip" :: "'a com" ("SKIP") + "_Seq" :: "'a com \ 'a com \ 'a com" ("(_;;/ _)" [60,61] 60) + "_Cond" :: "'a bexp \ 'a com \ 'a com \ 'a com" ("(0IF _/ THEN _/ ELSE _/FI)" [0, 0, 0] 61) + "_Cond2" :: "'a bexp \ 'a com \ 'a com" ("(0IF _ THEN _ FI)" [0,0] 56) + "_While" :: "'a bexp \ 'a com \ 'a com" ("(0WHILE _ /DO _ /OD)" [0, 0] 61) + "_Await" :: "'a bexp \ 'a com \ 'a com" ("(0AWAIT _ /THEN /_ /END)" [0,0] 61) + "_Atom" :: "'a com \ 'a com" ("(\_\)" 61) + "_Wait" :: "'a bexp \ 'a com" ("(0WAIT _ END)" 61) + +translations + "\\x := a" \ "Basic \\\(_update_name x (\_. a))\" + "SKIP" \ "Basic id" + "c1;; c2" \ "Seq c1 c2" + "IF b THEN c1 ELSE c2 FI" \ "Cond .{b}. c1 c2" + "IF b THEN c FI" \ "IF b THEN c ELSE SKIP FI" + "WHILE b DO c OD" \ "While .{b}. c" + "AWAIT b THEN c END" \ "Await .{b}. c" + "\c\" \ "AWAIT True THEN c END" + "WAIT b END" \ "AWAIT b THEN SKIP END" + +nonterminals + prgs + +syntax + "_PAR" :: "prgs \ 'a" ("COBEGIN//_//COEND" 60) + "_prg" :: "'a \ prgs" ("_" 57) + "_prgs" :: "['a, prgs] \ prgs" ("_//\//_" [60,57] 57) + +translations + "_prg a" \ "[a]" + "_prgs a ps" \ "a # ps" + "_PAR ps" \ "ps" + +syntax + "_prg_scheme" :: "['a, 'a, 'a, 'a] \ prgs" ("SCHEME [_ \ _ < _] _" [0,0,0,60] 57) + +translations + "_prg_scheme j i k c" \ "(map (\i. c) [j.. 'a" ("\_") + "_after" :: "id \ 'a" ("\_") + +translations + "\x" \ "x \fst" + "\x" \ "x \snd" + +print_translation {* + let + fun quote_tr' f (t :: ts) = + Term.list_comb (f $ Syntax.quote_tr' "_antiquote" t, ts) + | quote_tr' _ _ = raise Match; + + val assert_tr' = quote_tr' (Syntax.const "_Assert"); + + fun bexp_tr' name ((Const ("Collect", _) $ t) :: ts) = + quote_tr' (Syntax.const name) (t :: ts) + | bexp_tr' _ _ = raise Match; + + fun upd_tr' (x_upd, T) = + (case try (unsuffix Record.updateN) x_upd of + SOME x => (x, if T = dummyT then T else Term.domain_type T) + | NONE => raise Match); + + fun update_name_tr' (Free x) = Free (upd_tr' x) + | update_name_tr' ((c as Const ("_free", _)) $ Free x) = + c $ Free (upd_tr' x) + | update_name_tr' (Const x) = Const (upd_tr' x) + | update_name_tr' _ = raise Match; + + fun K_tr' (Abs (_,_,t)) = if null (loose_bnos t) then t else raise Match + | K_tr' (Abs (_,_,Abs (_,_,t)$Bound 0)) = if null (loose_bnos t) then t else raise Match + | K_tr' _ = raise Match; + + fun assign_tr' (Abs (x, _, f $ k $ Bound 0) :: ts) = + quote_tr' (Syntax.const "_Assign" $ update_name_tr' f) + (Abs (x, dummyT, K_tr' k) :: ts) + | assign_tr' _ = raise Match; + in + [("Collect", assert_tr'), ("Basic", assign_tr'), + ("Cond", bexp_tr' "_Cond"), ("While", bexp_tr' "_While_inv")] + end +*} + +end \ No newline at end of file diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/RG_Tran.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/RG_Tran.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,1075 @@ +header {* \section{Operational Semantics} *} + +theory RG_Tran +imports RG_Com +begin + +subsection {* Semantics of Component Programs *} + +subsubsection {* Environment transitions *} + +types 'a conf = "(('a com) option) \ 'a" + +inductive_set + etran :: "('a conf \ 'a conf) set" + and etran' :: "'a conf \ 'a conf \ bool" ("_ -e\ _" [81,81] 80) +where + "P -e\ Q \ (P,Q) \ etran" +| Env: "(P, s) -e\ (P, t)" + +lemma etranE: "c -e\ c' \ (\P s t. c = (P, s) \ c' = (P, t) \ Q) \ Q" + by (induct c, induct c', erule etran.cases, blast) + +subsubsection {* Component transitions *} + +inductive_set + ctran :: "('a conf \ 'a conf) set" + and ctran' :: "'a conf \ 'a conf \ bool" ("_ -c\ _" [81,81] 80) + and ctrans :: "'a conf \ 'a conf \ bool" ("_ -c*\ _" [81,81] 80) +where + "P -c\ Q \ (P,Q) \ ctran" +| "P -c*\ Q \ (P,Q) \ ctran^*" + +| Basic: "(Some(Basic f), s) -c\ (None, f s)" + +| Seq1: "(Some P0, s) -c\ (None, t) \ (Some(Seq P0 P1), s) -c\ (Some P1, t)" + +| Seq2: "(Some P0, s) -c\ (Some P2, t) \ (Some(Seq P0 P1), s) -c\ (Some(Seq P2 P1), t)" + +| CondT: "s\b \ (Some(Cond b P1 P2), s) -c\ (Some P1, s)" +| CondF: "s\b \ (Some(Cond b P1 P2), s) -c\ (Some P2, s)" + +| WhileF: "s\b \ (Some(While b P), s) -c\ (None, s)" +| WhileT: "s\b \ (Some(While b P), s) -c\ (Some(Seq P (While b P)), s)" + +| Await: "\s\b; (Some P, s) -c*\ (None, t)\ \ (Some(Await b P), s) -c\ (None, t)" + +monos "rtrancl_mono" + +subsection {* Semantics of Parallel Programs *} + +types 'a par_conf = "('a par_com) \ 'a" + +inductive_set + par_etran :: "('a par_conf \ 'a par_conf) set" + and par_etran' :: "['a par_conf,'a par_conf] \ bool" ("_ -pe\ _" [81,81] 80) +where + "P -pe\ Q \ (P,Q) \ par_etran" +| ParEnv: "(Ps, s) -pe\ (Ps, t)" + +inductive_set + par_ctran :: "('a par_conf \ 'a par_conf) set" + and par_ctran' :: "['a par_conf,'a par_conf] \ bool" ("_ -pc\ _" [81,81] 80) +where + "P -pc\ Q \ (P,Q) \ par_ctran" +| ParComp: "\i (r, t)\ \ (Ps, s) -pc\ (Ps[i:=r], t)" + +lemma par_ctranE: "c -pc\ c' \ + (\i Ps s r t. c = (Ps, s) \ c' = (Ps[i := r], t) \ i < length Ps \ + (Ps ! i, s) -c\ (r, t) \ P) \ P" + by (induct c, induct c', erule par_ctran.cases, blast) + +subsection {* Computations *} + +subsubsection {* Sequential computations *} + +types 'a confs = "('a conf) list" + +inductive_set cptn :: "('a confs) set" +where + CptnOne: "[(P,s)] \ cptn" +| CptnEnv: "(P, t)#xs \ cptn \ (P,s)#(P,t)#xs \ cptn" +| CptnComp: "\(P,s) -c\ (Q,t); (Q, t)#xs \ cptn \ \ (P,s)#(Q,t)#xs \ cptn" + +constdefs + cp :: "('a com) option \ 'a \ ('a confs) set" + "cp P s \ {l. l!0=(P,s) \ l \ cptn}" + +subsubsection {* Parallel computations *} + +types 'a par_confs = "('a par_conf) list" + +inductive_set par_cptn :: "('a par_confs) set" +where + ParCptnOne: "[(P,s)] \ par_cptn" +| ParCptnEnv: "(P,t)#xs \ par_cptn \ (P,s)#(P,t)#xs \ par_cptn" +| ParCptnComp: "\ (P,s) -pc\ (Q,t); (Q,t)#xs \ par_cptn \ \ (P,s)#(Q,t)#xs \ par_cptn" + +constdefs + par_cp :: "'a par_com \ 'a \ ('a par_confs) set" + "par_cp P s \ {l. l!0=(P,s) \ l \ par_cptn}" + +subsection{* Modular Definition of Computation *} + +constdefs + lift :: "'a com \ 'a conf \ 'a conf" + "lift Q \ \(P, s). (if P=None then (Some Q,s) else (Some(Seq (the P) Q), s))" + +inductive_set cptn_mod :: "('a confs) set" +where + CptnModOne: "[(P, s)] \ cptn_mod" +| CptnModEnv: "(P, t)#xs \ cptn_mod \ (P, s)#(P, t)#xs \ cptn_mod" +| CptnModNone: "\(Some P, s) -c\ (None, t); (None, t)#xs \ cptn_mod \ \ (Some P,s)#(None, t)#xs \cptn_mod" +| CptnModCondT: "\(Some P0, s)#ys \ cptn_mod; s \ b \ \ (Some(Cond b P0 P1), s)#(Some P0, s)#ys \ cptn_mod" +| CptnModCondF: "\(Some P1, s)#ys \ cptn_mod; s \ b \ \ (Some(Cond b P0 P1), s)#(Some P1, s)#ys \ cptn_mod" +| CptnModSeq1: "\(Some P0, s)#xs \ cptn_mod; zs=map (lift P1) xs \ + \ (Some(Seq P0 P1), s)#zs \ cptn_mod" +| CptnModSeq2: + "\(Some P0, s)#xs \ cptn_mod; fst(last ((Some P0, s)#xs)) = None; + (Some P1, snd(last ((Some P0, s)#xs)))#ys \ cptn_mod; + zs=(map (lift P1) xs)@ys \ \ (Some(Seq P0 P1), s)#zs \ cptn_mod" + +| CptnModWhile1: + "\ (Some P, s)#xs \ cptn_mod; s \ b; zs=map (lift (While b P)) xs \ + \ (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \ cptn_mod" +| CptnModWhile2: + "\ (Some P, s)#xs \ cptn_mod; fst(last ((Some P, s)#xs))=None; s \ b; + zs=(map (lift (While b P)) xs)@ys; + (Some(While b P), snd(last ((Some P, s)#xs)))#ys \ cptn_mod\ + \ (Some(While b P), s)#(Some(Seq P (While b P)), s)#zs \ cptn_mod" + +subsection {* Equivalence of Both Definitions.*} + +lemma last_length: "((a#xs)!(length xs))=last (a#xs)" +apply simp +apply(induct xs,simp+) +apply(case_tac xs) +apply simp_all +done + +lemma div_seq [rule_format]: "list \ cptn_mod \ + (\s P Q zs. list=(Some (Seq P Q), s)#zs \ + (\xs. (Some P, s)#xs \ cptn_mod \ (zs=(map (lift Q) xs) \ + ( fst(((Some P, s)#xs)!length xs)=None \ + (\ys. (Some Q, snd(((Some P, s)#xs)!length xs))#ys \ cptn_mod + \ zs=(map (lift (Q)) xs)@ys)))))" +apply(erule cptn_mod.induct) +apply simp_all + apply clarify + apply(force intro:CptnModOne) + apply clarify + apply(erule_tac x=Pa in allE) + apply(erule_tac x=Q in allE) + apply simp + apply clarify + apply(erule disjE) + apply(rule_tac x="(Some Pa,t)#xsa" in exI) + apply(rule conjI) + apply clarify + apply(erule CptnModEnv) + apply(rule disjI1) + apply(simp add:lift_def) + apply clarify + apply(rule_tac x="(Some Pa,t)#xsa" in exI) + apply(rule conjI) + apply(erule CptnModEnv) + apply(rule disjI2) + apply(rule conjI) + apply(case_tac xsa,simp,simp) + apply(rule_tac x="ys" in exI) + apply(rule conjI) + apply simp + apply(simp add:lift_def) + apply clarify + apply(erule ctran.cases,simp_all) + apply clarify + apply(rule_tac x="xs" in exI) + apply simp + apply clarify +apply(rule_tac x="xs" in exI) +apply(simp add: last_length) +done + +lemma cptn_onlyif_cptn_mod_aux [rule_format]: + "\s Q t xs.((Some a, s), Q, t) \ ctran \ (Q, t) # xs \ cptn_mod + \ (Some a, s) # (Q, t) # xs \ cptn_mod" +apply(induct a) +apply simp_all +--{* basic *} +apply clarify +apply(erule ctran.cases,simp_all) +apply(rule CptnModNone,rule Basic,simp) +apply clarify +apply(erule ctran.cases,simp_all) +--{* Seq1 *} +apply(rule_tac xs="[(None,ta)]" in CptnModSeq2) + apply(erule CptnModNone) + apply(rule CptnModOne) + apply simp +apply simp +apply(simp add:lift_def) +--{* Seq2 *} +apply(erule_tac x=sa in allE) +apply(erule_tac x="Some P2" in allE) +apply(erule allE,erule impE, assumption) +apply(drule div_seq,simp) +apply force +apply clarify +apply(erule disjE) + apply clarify + apply(erule allE,erule impE, assumption) + apply(erule_tac CptnModSeq1) + apply(simp add:lift_def) +apply clarify +apply(erule allE,erule impE, assumption) +apply(erule_tac CptnModSeq2) + apply (simp add:last_length) + apply (simp add:last_length) +apply(simp add:lift_def) +--{* Cond *} +apply clarify +apply(erule ctran.cases,simp_all) +apply(force elim: CptnModCondT) +apply(force elim: CptnModCondF) +--{* While *} +apply clarify +apply(erule ctran.cases,simp_all) +apply(rule CptnModNone,erule WhileF,simp) +apply(drule div_seq,force) +apply clarify +apply (erule disjE) + apply(force elim:CptnModWhile1) +apply clarify +apply(force simp add:last_length elim:CptnModWhile2) +--{* await *} +apply clarify +apply(erule ctran.cases,simp_all) +apply(rule CptnModNone,erule Await,simp+) +done + +lemma cptn_onlyif_cptn_mod [rule_format]: "c \ cptn \ c \ cptn_mod" +apply(erule cptn.induct) + apply(rule CptnModOne) + apply(erule CptnModEnv) +apply(case_tac P) + apply simp + apply(erule ctran.cases,simp_all) +apply(force elim:cptn_onlyif_cptn_mod_aux) +done + +lemma lift_is_cptn: "c\cptn \ map (lift P) c \ cptn" +apply(erule cptn.induct) + apply(force simp add:lift_def CptnOne) + apply(force intro:CptnEnv simp add:lift_def) +apply(force simp add:lift_def intro:CptnComp Seq2 Seq1 elim:ctran.cases) +done + +lemma cptn_append_is_cptn [rule_format]: + "\b a. b#c1\cptn \ a#c2\cptn \ (b#c1)!length c1=a \ b#c1@c2\cptn" +apply(induct c1) + apply simp +apply clarify +apply(erule cptn.cases,simp_all) + apply(force intro:CptnEnv) +apply(force elim:CptnComp) +done + +lemma last_lift: "\xs\[]; fst(xs!(length xs - (Suc 0)))=None\ + \ fst((map (lift P) xs)!(length (map (lift P) xs)- (Suc 0)))=(Some P)" +apply(case_tac "(xs ! (length xs - (Suc 0)))") +apply (simp add:lift_def) +done + +lemma last_fst [rule_format]: "P((a#x)!length x) \ \P a \ P (x!(length x - (Suc 0)))" +apply(induct x,simp+) +done + +lemma last_fst_esp: + "fst(((Some a,s)#xs)!(length xs))=None \ fst(xs!(length xs - (Suc 0)))=None" +apply(erule last_fst) +apply simp +done + +lemma last_snd: "xs\[] \ + snd(((map (lift P) xs))!(length (map (lift P) xs) - (Suc 0)))=snd(xs!(length xs - (Suc 0)))" +apply(case_tac "(xs ! (length xs - (Suc 0)))",simp) +apply (simp add:lift_def) +done + +lemma Cons_lift: "(Some (Seq P Q), s) # (map (lift Q) xs) = map (lift Q) ((Some P, s) # xs)" +by(simp add:lift_def) + +lemma Cons_lift_append: + "(Some (Seq P Q), s) # (map (lift Q) xs) @ ys = map (lift Q) ((Some P, s) # xs)@ ys " +by(simp add:lift_def) + +lemma lift_nth: "i map (lift Q) xs ! i = lift Q (xs! i)" +by (simp add:lift_def) + +lemma snd_lift: "i< length xs \ snd(lift Q (xs ! i))= snd (xs ! i)" +apply(case_tac "xs!i") +apply(simp add:lift_def) +done + +lemma cptn_if_cptn_mod: "c \ cptn_mod \ c \ cptn" +apply(erule cptn_mod.induct) + apply(rule CptnOne) + apply(erule CptnEnv) + apply(erule CptnComp,simp) + apply(rule CptnComp) + apply(erule CondT,simp) + apply(rule CptnComp) + apply(erule CondF,simp) +--{* Seq1 *} +apply(erule cptn.cases,simp_all) + apply(rule CptnOne) + apply clarify + apply(drule_tac P=P1 in lift_is_cptn) + apply(simp add:lift_def) + apply(rule CptnEnv,simp) +apply clarify +apply(simp add:lift_def) +apply(rule conjI) + apply clarify + apply(rule CptnComp) + apply(rule Seq1,simp) + apply(drule_tac P=P1 in lift_is_cptn) + apply(simp add:lift_def) +apply clarify +apply(rule CptnComp) + apply(rule Seq2,simp) +apply(drule_tac P=P1 in lift_is_cptn) +apply(simp add:lift_def) +--{* Seq2 *} +apply(rule cptn_append_is_cptn) + apply(drule_tac P=P1 in lift_is_cptn) + apply(simp add:lift_def) + apply simp +apply(case_tac "xs\[]") + apply(drule_tac P=P1 in last_lift) + apply(rule last_fst_esp) + apply (simp add:last_length) + apply(simp add:Cons_lift del:map.simps) + apply(rule conjI, clarify, simp) + apply(case_tac "(((Some P0, s) # xs) ! length xs)") + apply clarify + apply (simp add:lift_def last_length) +apply (simp add:last_length) +--{* While1 *} +apply(rule CptnComp) +apply(rule WhileT,simp) +apply(drule_tac P="While b P" in lift_is_cptn) +apply(simp add:lift_def) +--{* While2 *} +apply(rule CptnComp) +apply(rule WhileT,simp) +apply(rule cptn_append_is_cptn) +apply(drule_tac P="While b P" in lift_is_cptn) + apply(simp add:lift_def) + apply simp +apply(case_tac "xs\[]") + apply(drule_tac P="While b P" in last_lift) + apply(rule last_fst_esp,simp add:last_length) + apply(simp add:Cons_lift del:map.simps) + apply(rule conjI, clarify, simp) + apply(case_tac "(((Some P, s) # xs) ! length xs)") + apply clarify + apply (simp add:last_length lift_def) +apply simp +done + +theorem cptn_iff_cptn_mod: "(c \ cptn) = (c \ cptn_mod)" +apply(rule iffI) + apply(erule cptn_onlyif_cptn_mod) +apply(erule cptn_if_cptn_mod) +done + +section {* Validity of Correctness Formulas*} + +subsection {* Validity for Component Programs. *} + +types 'a rgformula = "'a com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set" + +constdefs + assum :: "('a set \ ('a \ 'a) set) \ ('a confs) set" + "assum \ \(pre, rely). {c. snd(c!0) \ pre \ (\i. Suc i + c!i -e\ c!(Suc i) \ (snd(c!i), snd(c!Suc i)) \ rely)}" + + comm :: "(('a \ 'a) set \ 'a set) \ ('a confs) set" + "comm \ \(guar, post). {c. (\i. Suc i + c!i -c\ c!(Suc i) \ (snd(c!i), snd(c!Suc i)) \ guar) \ + (fst (last c) = None \ snd (last c) \ post)}" + + com_validity :: "'a com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set \ 'a set \ bool" + ("\ _ sat [_, _, _, _]" [60,0,0,0,0] 45) + "\ P sat [pre, rely, guar, post] \ + \s. cp (Some P) s \ assum(pre, rely) \ comm(guar, post)" + +subsection {* Validity for Parallel Programs. *} + +constdefs + All_None :: "(('a com) option) list \ bool" + "All_None xs \ \c\set xs. c=None" + + par_assum :: "('a set \ ('a \ 'a) set) \ ('a par_confs) set" + "par_assum \ \(pre, rely). {c. snd(c!0) \ pre \ (\i. Suc i + c!i -pe\ c!Suc i \ (snd(c!i), snd(c!Suc i)) \ rely)}" + + par_comm :: "(('a \ 'a) set \ 'a set) \ ('a par_confs) set" + "par_comm \ \(guar, post). {c. (\i. Suc i + c!i -pc\ c!Suc i \ (snd(c!i), snd(c!Suc i)) \ guar) \ + (All_None (fst (last c)) \ snd( last c) \ post)}" + + par_com_validity :: "'a par_com \ 'a set \ ('a \ 'a) set \ ('a \ 'a) set +\ 'a set \ bool" ("\ _ SAT [_, _, _, _]" [60,0,0,0,0] 45) + "\ Ps SAT [pre, rely, guar, post] \ + \s. par_cp Ps s \ par_assum(pre, rely) \ par_comm(guar, post)" + +subsection {* Compositionality of the Semantics *} + +subsubsection {* Definition of the conjoin operator *} + +constdefs + same_length :: "'a par_confs \ ('a confs) list \ bool" + "same_length c clist \ (\i ('a confs) list \ bool" + "same_state c clist \ (\i j ('a confs) list \ bool" + "same_program c clist \ (\jx. fst(nth x j)) clist)" + + compat_label :: "'a par_confs \ ('a confs) list \ bool" + "compat_label c clist \ (\j. Suc j + (c!j -pc\ c!Suc j \ (\i (clist!i)! Suc j \ + (\li \ (clist!l)!j -e\ (clist!l)! Suc j))) \ + (c!j -pe\ c!Suc j \ (\i (clist!i)! Suc j)))" + + conjoin :: "'a par_confs \ ('a confs) list \ bool" ("_ \ _" [65,65] 64) + "c \ clist \ (same_length c clist) \ (same_state c clist) \ (same_program c clist) \ (compat_label c clist)" + +subsubsection {* Some previous lemmas *} + +lemma list_eq_if [rule_format]: + "\ys. xs=ys \ (length xs = length ys) \ (\i (\i ys!0=a; ys\[] \ \ ys=(a#(tl ys))" +apply(case_tac ys) + apply simp+ +done + +lemma nth_tl_if [rule_format]: "ys\[] \ ys!0=a \ P ys \ P (a#(tl ys))" +apply(induct ys) + apply simp+ +done + +lemma nth_tl_onlyif [rule_format]: "ys\[] \ ys!0=a \ P (a#(tl ys)) \ P ys" +apply(induct ys) + apply simp+ +done + +lemma seq_not_eq1: "Seq c1 c2\c1" +apply(rule com.induct) +apply simp_all +apply clarify +done + +lemma seq_not_eq2: "Seq c1 c2\c2" +apply(rule com.induct) +apply simp_all +apply clarify +done + +lemma if_not_eq1: "Cond b c1 c2 \c1" +apply(rule com.induct) +apply simp_all +apply clarify +done + +lemma if_not_eq2: "Cond b c1 c2\c2" +apply(rule com.induct) +apply simp_all +apply clarify +done + +lemmas seq_and_if_not_eq [simp] = seq_not_eq1 seq_not_eq2 +seq_not_eq1 [THEN not_sym] seq_not_eq2 [THEN not_sym] +if_not_eq1 if_not_eq2 if_not_eq1 [THEN not_sym] if_not_eq2 [THEN not_sym] + +lemma prog_not_eq_in_ctran_aux: + assumes c: "(P,s) -c\ (Q,t)" + shows "P\Q" using c + by (induct x1 \ "(P,s)" x2 \ "(Q,t)" arbitrary: P s Q t) auto + +lemma prog_not_eq_in_ctran [simp]: "\ (P,s) -c\ (P,t)" +apply clarify +apply(drule prog_not_eq_in_ctran_aux) +apply simp +done + +lemma prog_not_eq_in_par_ctran_aux [rule_format]: "(P,s) -pc\ (Q,t) \ (P\Q)" +apply(erule par_ctran.induct) +apply(drule prog_not_eq_in_ctran_aux) +apply clarify +apply(drule list_eq_if) + apply simp_all +apply force +done + +lemma prog_not_eq_in_par_ctran [simp]: "\ (P,s) -pc\ (P,t)" +apply clarify +apply(drule prog_not_eq_in_par_ctran_aux) +apply simp +done + +lemma tl_in_cptn: "\ a#xs \cptn; xs\[] \ \ xs\cptn" +apply(force elim:cptn.cases) +done + +lemma tl_zero[rule_format]: + "P (ys!Suc j) \ Suc j ys\[] \ P (tl(ys)!j)" +apply(induct ys) + apply simp_all +done + +subsection {* The Semantics is Compositional *} + +lemma aux_if [rule_format]: + "\xs s clist. (length clist = length xs \ (\i cptn) + \ ((xs, s)#ys \ map (\i. (fst i,s)#snd i) (zip xs clist)) + \ (xs, s)#ys \ par_cptn)" +apply(induct ys) + apply(clarify) + apply(rule ParCptnOne) +apply(clarify) +apply(simp add:conjoin_def compat_label_def) +apply clarify +apply(erule_tac x="0" and P="\j. ?H j \ (?P j \ ?Q j)" in all_dupE,simp) +apply(erule disjE) +--{* first step is a Component step *} + apply clarify + apply simp + apply(subgoal_tac "a=(xs[i:=(fst(clist!i!0))])") + apply(subgoal_tac "b=snd(clist!i!0)",simp) + prefer 2 + apply(simp add: same_state_def) + apply(erule_tac x=i in allE,erule impE,assumption, + erule_tac x=1 and P="\j. (?H j) \ (snd (?d j))=(snd (?e j))" in allE,simp) + prefer 2 + apply(simp add:same_program_def) + apply(erule_tac x=1 and P="\j. ?H j \ (fst (?s j))=(?t j)" in allE,simp) + apply(rule nth_equalityI,simp) + apply clarify + apply(case_tac "i=ia",simp,simp) + apply(erule_tac x=ia and P="\j. ?H j \ ?I j \ ?J j" in allE) + apply(drule_tac t=i in not_sym,simp) + apply(erule etranE,simp) + apply(rule ParCptnComp) + apply(erule ParComp,simp) +--{* applying the induction hypothesis *} + apply(erule_tac x="xs[i := fst (clist ! i ! 0)]" in allE) + apply(erule_tac x="snd (clist ! i ! 0)" in allE) + apply(erule mp) + apply(rule_tac x="map tl clist" in exI,simp) + apply(rule conjI,clarify) + apply(case_tac "i=ia",simp) + apply(rule nth_tl_if) + apply(force simp add:same_length_def length_Suc_conv) + apply simp + apply(erule allE,erule impE,assumption,erule tl_in_cptn) + apply(force simp add:same_length_def length_Suc_conv) + apply(rule nth_tl_if) + apply(force simp add:same_length_def length_Suc_conv) + apply(simp add:same_state_def) + apply(erule_tac x=ia in allE, erule impE, assumption, + erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(erule_tac x=ia and P="\j. ?H j \ ?I j \ ?J j" in allE) + apply(drule_tac t=i in not_sym,simp) + apply(erule etranE,simp) + apply(erule allE,erule impE,assumption,erule tl_in_cptn) + apply(force simp add:same_length_def length_Suc_conv) + apply(simp add:same_length_def same_state_def) + apply(rule conjI) + apply clarify + apply(case_tac j,simp,simp) + apply(erule_tac x=ia in allE, erule impE, assumption, + erule_tac x="Suc(Suc nat)" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(force simp add:same_length_def length_Suc_conv) + apply(rule conjI) + apply(simp add:same_program_def) + apply clarify + apply(case_tac j,simp) + apply(rule nth_equalityI,simp) + apply clarify + apply(case_tac "i=ia",simp,simp) + apply(erule_tac x="Suc(Suc nat)" and P="\j. ?H j \ (fst (?s j))=(?t j)" in allE,simp) + apply(rule nth_equalityI,simp,simp) + apply(force simp add:length_Suc_conv) + apply(rule allI,rule impI) + apply(erule_tac x="Suc j" and P="\j. ?H j \ (?I j \ ?J j)" in allE,simp) + apply(erule disjE) + apply clarify + apply(rule_tac x=ia in exI,simp) + apply(case_tac "i=ia",simp) + apply(rule conjI) + apply(force simp add: length_Suc_conv) + apply clarify + apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE,assumption) + apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE,assumption) + apply simp + apply(case_tac j,simp) + apply(rule tl_zero) + apply(erule_tac x=l in allE, erule impE, assumption, + erule_tac x=1 and P="\j. (?H j) \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(force elim:etranE intro:Env) + apply force + apply force + apply simp + apply(rule tl_zero) + apply(erule tl_zero) + apply force + apply force + apply force + apply force + apply(rule conjI,simp) + apply(rule nth_tl_if) + apply force + apply(erule_tac x=ia in allE, erule impE, assumption, + erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(erule_tac x=ia and P="\j. ?H j \ ?I j \ ?J j" in allE) + apply(drule_tac t=i in not_sym,simp) + apply(erule etranE,simp) + apply(erule tl_zero) + apply force + apply force + apply clarify + apply(case_tac "i=l",simp) + apply(rule nth_tl_if) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply simp + apply(erule_tac P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE,assumption,erule impE,assumption) + apply(erule tl_zero,force) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply(rule nth_tl_if) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply(erule_tac x=l in allE, erule impE, assumption, + erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE, assumption,simp) + apply(erule etranE,simp) + apply(rule tl_zero) + apply force + apply force + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply(rule disjI2) + apply(case_tac j,simp) + apply clarify + apply(rule tl_zero) + apply(erule_tac x=ia and P="\j. ?H j \ ?I j\etran" in allE,erule impE, assumption) + apply(case_tac "i=ia",simp,simp) + apply(erule_tac x=ia in allE, erule impE, assumption, + erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(erule_tac x=ia and P="\j. ?H j \ ?I j \ ?J j" in allE,erule impE, assumption,simp) + apply(force elim:etranE intro:Env) + apply force + apply(erule_tac x=ia and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply simp + apply clarify + apply(rule tl_zero) + apply(rule tl_zero,force) + apply force + apply(erule_tac x=ia and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply force + apply(erule_tac x=ia and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) +--{* first step is an environmental step *} +apply clarify +apply(erule par_etran.cases) +apply simp +apply(rule ParCptnEnv) +apply(erule_tac x="Ps" in allE) +apply(erule_tac x="t" in allE) +apply(erule mp) +apply(rule_tac x="map tl clist" in exI,simp) +apply(rule conjI) + apply clarify + apply(erule_tac x=i and P="\j. ?H j \ (?I ?s j) \ cptn" in allE,simp) + apply(erule cptn.cases) + apply(simp add:same_length_def) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply(simp add:same_state_def) + apply(erule_tac x=i in allE, erule impE, assumption, + erule_tac x=1 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(erule_tac x=i and P="\j. ?H j \ ?J j \etran" in allE,simp) + apply(erule etranE,simp) +apply(simp add:same_state_def same_length_def) +apply(rule conjI,clarify) + apply(case_tac j,simp,simp) + apply(erule_tac x=i in allE, erule impE, assumption, + erule_tac x="Suc(Suc nat)" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(rule tl_zero) + apply(simp) + apply force + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) +apply(rule conjI) + apply(simp add:same_program_def) + apply clarify + apply(case_tac j,simp) + apply(rule nth_equalityI,simp) + apply clarify + apply simp + apply(erule_tac x="Suc(Suc nat)" and P="\j. ?H j \ (fst (?s j))=(?t j)" in allE,simp) + apply(rule nth_equalityI,simp,simp) + apply(force simp add:length_Suc_conv) +apply(rule allI,rule impI) +apply(erule_tac x="Suc j" and P="\j. ?H j \ (?I j \ ?J j)" in allE,simp) +apply(erule disjE) + apply clarify + apply(rule_tac x=i in exI,simp) + apply(rule conjI) + apply(erule_tac x=i and P="\i. ?H i \ ?J i \etran" in allE, erule impE, assumption) + apply(erule etranE,simp) + apply(erule_tac x=i in allE, erule impE, assumption, + erule_tac x=1 and P="\j. (?H j) \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(rule nth_tl_if) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply simp + apply(erule tl_zero,force) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply clarify + apply(erule_tac x=l and P="\i. ?H i \ ?J i \etran" in allE, erule impE, assumption) + apply(erule etranE,simp) + apply(erule_tac x=l in allE, erule impE, assumption, + erule_tac x=1 and P="\j. (?H j) \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(rule nth_tl_if) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply simp + apply(rule tl_zero,force) + apply force + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) +apply(rule disjI2) +apply simp +apply clarify +apply(case_tac j,simp) + apply(rule tl_zero) + apply(erule_tac x=i and P="\i. ?H i \ ?J i \etran" in allE, erule impE, assumption) + apply(erule_tac x=i and P="\i. ?H i \ ?J i \etran" in allE, erule impE, assumption) + apply(force elim:etranE intro:Env) + apply force + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) +apply simp +apply(rule tl_zero) + apply(rule tl_zero,force) + apply force + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply force +apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) +done + +lemma less_Suc_0 [iff]: "(n < Suc 0) = (n = 0)" +by auto + +lemma aux_onlyif [rule_format]: "\xs s. (xs, s)#ys \ par_cptn \ + (\clist. (length clist = length xs) \ + (xs, s)#ys \ map (\i. (fst i,s)#(snd i)) (zip xs clist) \ + (\i cptn))" +apply(induct ys) + apply(clarify) + apply(rule_tac x="map (\i. []) [0..[] \ (\ys. ((xs, s)#ys \ par_cptn) = + (\clist. length clist= length xs \ + ((xs, s)#ys \ map (\i. (fst i,s)#(snd i)) (zip xs clist)) \ + (\i cptn))) = + (par_cp (xs) s = {c. \clist. (length clist)=(length xs) \ + (\i cp(xs!i) s) \ c \ clist})" +apply (rule iffI) + apply(rule subset_antisym) + apply(rule subsetI) + apply(clarify) + apply(simp add:par_cp_def cp_def) + apply(case_tac x) + apply(force elim:par_cptn.cases) + apply simp + apply(erule_tac x="list" in allE) + apply clarify + apply simp + apply(rule_tac x="map (\i. (fst i, s) # snd i) (zip xs clist)" in exI,simp) + apply(rule subsetI) + apply(clarify) + apply(case_tac x) + apply(erule_tac x=0 in allE) + apply(simp add:cp_def conjoin_def same_length_def same_program_def same_state_def compat_label_def) + apply clarify + apply(erule cptn.cases,force,force,force) + apply(simp add:par_cp_def conjoin_def same_length_def same_program_def same_state_def compat_label_def) + apply clarify + apply(erule_tac x=0 and P="\j. ?H j \ (length (?s j) = ?t)" in all_dupE) + apply(subgoal_tac "a = xs") + apply(subgoal_tac "b = s",simp) + prefer 3 + apply(erule_tac x=0 and P="\j. ?H j \ (fst (?s j))=((?t j))" in allE) + apply (simp add:cp_def) + apply(rule nth_equalityI,simp,simp) + prefer 2 + apply(erule_tac x=0 in allE) + apply (simp add:cp_def) + apply(erule_tac x=0 and P="\j. ?H j \ (\i. ?T i \ (snd (?d j i))=(snd (?e j i)))" in allE,simp) + apply(erule_tac x=0 and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(erule_tac x=list in allE) + apply(rule_tac x="map tl clist" in exI,simp) + apply(rule conjI) + apply clarify + apply(case_tac j,simp) + apply(erule_tac x=i in allE, erule impE, assumption, + erule_tac x="0" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE,simp) + apply(erule_tac x=i in allE, erule impE, assumption, + erule_tac x="Suc nat" and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply(rule conjI) + apply clarify + apply(rule nth_equalityI,simp,simp) + apply(case_tac j) + apply clarify + apply(erule_tac x=i in allE) + apply(simp add:cp_def) + apply clarify + apply simp + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply(thin_tac "?H = (\i. ?J i)") + apply(rule conjI) + apply clarify + apply(erule_tac x=j in allE,erule impE, assumption,erule disjE) + apply clarify + apply(rule_tac x=i in exI,simp) + apply(case_tac j,simp) + apply(rule conjI) + apply(erule_tac x=i in allE) + apply(simp add:cp_def) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply clarify + apply(erule_tac x=l in allE) + apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE) + apply clarify + apply(simp add:cp_def) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!l",simp,simp) + apply simp + apply(rule conjI) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply clarify + apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!l",simp,simp) + apply clarify + apply(erule_tac x=i in allE) + apply(simp add:cp_def) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp) + apply(rule nth_tl_if,simp,simp) + apply(erule_tac x=i and P="\j. ?H j \ (?P j)\etran" in allE, erule impE, assumption,simp) + apply(simp add:cp_def) + apply clarify + apply(rule nth_tl_if) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply force + apply force +apply clarify +apply(rule iffI) + apply(simp add:par_cp_def) + apply(erule_tac c="(xs, s) # ys" in equalityCE) + apply simp + apply clarify + apply(rule_tac x="map tl clist" in exI) + apply simp + apply (rule conjI) + apply(simp add:conjoin_def cp_def) + apply(rule conjI) + apply clarify + apply(unfold same_length_def) + apply clarify + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,simp) + apply(rule conjI) + apply(simp add:same_state_def) + apply clarify + apply(erule_tac x=i in allE, erule impE, assumption, + erule_tac x=j and P="\j. ?H j \ (snd (?d j))=(snd (?e j))" in allE) + apply(case_tac j,simp) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply(rule conjI) + apply(simp add:same_program_def) + apply clarify + apply(rule nth_equalityI,simp,simp) + apply(case_tac j,simp) + apply clarify + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply clarify + apply(simp add:compat_label_def) + apply(rule allI,rule impI) + apply(erule_tac x=j in allE,erule impE, assumption) + apply(erule disjE) + apply clarify + apply(rule_tac x=i in exI,simp) + apply(rule conjI) + apply(erule_tac x=i in allE) + apply(case_tac j,simp) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!i",simp,simp) + apply clarify + apply(erule_tac x=l and P="\j. ?H j \ ?I j \ ?J j" in allE) + apply(erule_tac x=l and P="\j. ?H j \ (length (?s j) = ?t)" in allE) + apply(case_tac "clist!l",simp,simp) + apply(erule_tac x=l in allE,simp) + apply(rule disjI2) + apply clarify + apply(rule tl_zero) + apply(case_tac j,simp,simp) + apply(rule tl_zero,force) + apply force + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply force + apply(erule_tac x=i and P="\j. ?H j \ (length (?s j) = ?t)" in allE,force) + apply clarify + apply(erule_tac x=i in allE) + apply(simp add:cp_def) + apply(rule nth_tl_if) + apply(simp add:conjoin_def) + apply clarify + apply(simp add:same_length_def) + apply(erule_tac x=i in allE,simp) + apply simp + apply simp + apply simp +apply clarify +apply(erule_tac c="(xs, s) # ys" in equalityCE) + apply(simp add:par_cp_def) +apply simp +apply(erule_tac x="map (\i. (fst i, s) # snd i) (zip xs clist)" in allE) +apply simp +apply clarify +apply(simp add:cp_def) +done + +theorem one: "xs\[] \ + par_cp xs s = {c. \clist. (length clist)=(length xs) \ + (\i cp(xs!i) s) \ c \ clist}" +apply(frule one_iff_aux) +apply(drule sym) +apply(erule iffD2) +apply clarify +apply(rule iffI) + apply(erule aux_onlyif) +apply clarify +apply(force intro:aux_if) +done + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/ROOT.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,2 @@ + +use_thy "Hoare_Parallel"; diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/document/root.bib --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/document/root.bib Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,28 @@ +@inproceedings{NipkowP-FASE99,author={Tobias Nipkow and Prensa Nieto, Leonor}, +title={{Owicki/Gries} in {Isabelle/HOL}}, +booktitle={Fundamental Approaches to Software Engineering (FASE'99)}, +editor={J.-P. Finance},publisher="Springer",series="LNCS",volume=1577, +pages={188--203},year=1999} + +@InProceedings{PrenEsp00, + author = {Prensa Nieto, Leonor and Javier Esparza}, + title = {Verifying Single and Multi-mutator Garbage Collectors + with {Owicki/Gries} in {Isabelle/HOL}}, + booktitle = {Mathematical Foundations of Computer Science (MFCS 2000)}, + editor = {M. Nielsen and B. Rovan}, + publisher = {Springer-Verlag}, + series = {LNCS}, + volume = 1893, + pages = {619--628}, + year = 2000 +} + +@PhdThesis{Prensa-PhD,author={Leonor Prensa Nieto}, +title={Verification of Parallel Programs with the Owicki-Gries and +Rely-Guarantee Methods in Isabelle/HOL}, +school={Technische Universit{\"a}t M{\"u}nchen},year=2002} + +@inproceedings{Prensa-ESOP03,author={Prensa Nieto, Leonor}, +title={The {Rely-Guarantee} Method in {Isabelle/HOL}}, +booktitle={European Symposium on Programming (ESOP'03)},editor={P. Degano}, +publisher=Springer,series=LNCS,volume=2618,pages={348--362},year=2003} diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Hoare_Parallel/document/root.tex --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Hoare_Parallel/document/root.tex Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,62 @@ + +% $Id$ + +\documentclass[11pt,a4paper]{report} +\usepackage{graphicx} +\usepackage[english]{babel} +\usepackage{isabelle,isabellesym} +\usepackage{pdfsetup} + +\urlstyle{rm} +\isabellestyle{it} + +\renewcommand{\isamarkupheader}[1]{#1} + +\begin{document} + +\title{Hoare Logic for Parallel Programs} +\author{Leonor Prensa Nieto} +\maketitle + +\begin{abstract}\noindent + In the following theories a formalization of the Owicki-Gries and + the rely-guarantee methods is presented. These methods are widely + used for correctness proofs of parallel imperative programs with + shared variables. We define syntax, semantics and proof rules in + Isabelle/HOL. The proof rules also provide for programs + parameterized in the number of parallel components. Their + correctness w.r.t.\ the semantics is proven. Completeness proofs + for both methods are extended to the new case of parameterized + programs. (These proofs have not been formalized in Isabelle. They + can be found in~\cite{Prensa-PhD}.) Using this formalizations we + verify several non-trivial examples for parameterized and + non-parameterized programs. For the automatic generation of + verification conditions with the Owicki-Gries method we define a + tactic based on the proof rules. The most involved examples are the + verification of two garbage-collection algorithms, the second one + parameterized in the number of mutators. + +For excellent descriptions of this work see +\cite{NipkowP-FASE99,PrenEsp00,Prensa-PhD,Prensa-ESOP03}. + +\end{abstract} + +\pagestyle{plain} +\thispagestyle{empty} +\tableofcontents + +\clearpage + +\begin{center} + \includegraphics[scale=0.7]{session_graph} +\end{center} + +\newpage + +\parindent 0pt\parskip 0.5ex +\input{session} + +\bibliographystyle{plain} +\bibliography{root} + +\end{document} diff -r c876bcb601fc -r 5b65449d7669 src/HOL/IsaMakefile --- a/src/HOL/IsaMakefile Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/IsaMakefile Tue Sep 22 13:52:19 2009 +1000 @@ -18,7 +18,7 @@ HOL-Extraction \ HOL-Hahn_Banach \ HOL-Hoare \ - HOL-HoareParallel \ + HOL-Hoare_Parallel \ HOL-Import \ HOL-IMP \ HOL-IMPP \ @@ -40,6 +40,7 @@ HOL-Prolog \ HOL-SET-Protocol \ HOL-SizeChange \ + HOL-SMT \ HOL-Statespace \ HOL-Subst \ TLA-Buffer \ @@ -536,21 +537,22 @@ @$(ISABELLE_TOOL) usedir $(OUT)/HOL Hoare -## HOL-HoareParallel +## HOL-Hoare_Parallel -HOL-HoareParallel: HOL $(LOG)/HOL-HoareParallel.gz +HOL-Hoare_Parallel: HOL $(LOG)/HOL-Hoare_Parallel.gz -$(LOG)/HOL-HoareParallel.gz: $(OUT)/HOL HoareParallel/Gar_Coll.thy \ - HoareParallel/Graph.thy HoareParallel/Mul_Gar_Coll.thy \ - HoareParallel/OG_Com.thy HoareParallel/OG_Examples.thy \ - HoareParallel/OG_Hoare.thy HoareParallel/OG_Syntax.thy \ - HoareParallel/OG_Tactics.thy HoareParallel/OG_Tran.thy \ - HoareParallel/Quote_Antiquote.thy HoareParallel/RG_Com.thy \ - HoareParallel/RG_Examples.thy HoareParallel/RG_Hoare.thy \ - HoareParallel/RG_Syntax.thy HoareParallel/RG_Tran.thy \ - HoareParallel/ROOT.ML HoareParallel/document/root.tex \ - HoareParallel/document/root.bib - @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL HoareParallel +$(LOG)/HOL-Hoare_Parallel.gz: $(OUT)/HOL Hoare_Parallel/Gar_Coll.thy \ + Hoare_Parallel/Graph.thy Hoare_Parallel/Hoare_Parallel.thy \ + Hoare_Parallel/Mul_Gar_Coll.thy \ + Hoare_Parallel/OG_Com.thy Hoare_Parallel/OG_Examples.thy \ + Hoare_Parallel/OG_Hoare.thy Hoare_Parallel/OG_Syntax.thy \ + Hoare_Parallel/OG_Tactics.thy Hoare_Parallel/OG_Tran.thy \ + Hoare_Parallel/Quote_Antiquote.thy Hoare_Parallel/RG_Com.thy \ + Hoare_Parallel/RG_Examples.thy Hoare_Parallel/RG_Hoare.thy \ + Hoare_Parallel/RG_Syntax.thy Hoare_Parallel/RG_Tran.thy \ + Hoare_Parallel/ROOT.ML Hoare_Parallel/document/root.tex \ + Hoare_Parallel/document/root.bib + @$(ISABELLE_TOOL) usedir -g true $(OUT)/HOL Hoare_Parallel ## HOL-MetisExamples @@ -617,6 +619,10 @@ HOL-Auth: HOL $(LOG)/HOL-Auth.gz $(LOG)/HOL-Auth.gz: $(OUT)/HOL \ + Auth/Auth_Shared.thy Auth/Auth_Public.thy Auth/All_Symmetric.thy \ + Auth/Guard/Auth_Guard_Shared.thy \ + Auth/Guard/Auth_Guard_Public.thy \ + Auth/Smartcard/Auth_Smartcard.thy Auth/All_Symmetric.thy \ Auth/CertifiedEmail.thy Auth/Event.thy Auth/Message.thy \ Auth/NS_Public.thy Auth/NS_Public_Bad.thy Auth/NS_Shared.thy \ Auth/OtwayRees.thy Auth/OtwayReesBella.thy Auth/OtwayRees_AN.thy \ @@ -641,7 +647,7 @@ HOL-UNITY: HOL $(LOG)/HOL-UNITY.gz $(LOG)/HOL-UNITY.gz: $(OUT)/HOL Library/Multiset.thy UNITY/ROOT.ML \ - UNITY/UNITY_Main.thy UNITY/UNITY_tactics.ML UNITY/Comp.thy \ + UNITY/UNITY_Main.thy UNITY/UNITY_Examples.thy UNITY/UNITY_tactics.ML UNITY/Comp.thy \ UNITY/Constrains.thy UNITY/Detects.thy UNITY/ELT.thy \ UNITY/Extend.thy UNITY/FP.thy UNITY/Follows.thy UNITY/Guar.thy \ UNITY/Lift_prog.thy UNITY/ListOrder.thy UNITY/ProgressSets.thy \ @@ -830,7 +836,7 @@ HOL-Bali: HOL $(LOG)/HOL-Bali.gz -$(LOG)/HOL-Bali.gz: $(OUT)/HOL Bali/AxCompl.thy Bali/AxExample.thy \ +$(LOG)/HOL-Bali.gz: $(OUT)/HOL Bali/Bali.thy Bali/AxCompl.thy Bali/AxExample.thy \ Bali/AxSem.thy Bali/AxSound.thy Bali/Basis.thy Bali/Conform.thy \ Bali/Decl.thy Bali/DeclConcepts.thy Bali/Eval.thy Bali/Evaln.thy \ Bali/Example.thy Bali/Name.thy Bali/ROOT.ML Bali/State.thy \ @@ -1020,6 +1026,7 @@ HOL-Nominal-Examples: HOL-Nominal $(LOG)/HOL-Nominal-Examples.gz $(LOG)/HOL-Nominal-Examples.gz: $(OUT)/HOL-Nominal \ + Nominal/Examples/Nominal_Examples.thy \ Nominal/Examples/CK_Machine.thy \ Nominal/Examples/CR.thy \ Nominal/Examples/CR_Takahashi.thy \ @@ -1134,6 +1141,19 @@ @$(ISABELLE_TOOL) usedir $(OUT)/HOL Mirabelle +## HOL-SMT + +HOL-SMT: HOL-Word $(LOG)/HOL-SMT.gz + +$(LOG)/HOL-SMT.gz: $(OUT)/HOL-Word SMT/SMT_Definitions.thy SMT/SMT.thy \ + SMT/Tools/smt_normalize.ML SMT/Tools/smt_monomorph.ML \ + SMT/Tools/smt_translate.ML SMT/Tools/smt_builtin.ML \ + SMT/Tools/smtlib_interface.ML SMT/Tools/smt_solver.ML \ + SMT/Tools/cvc3_solver.ML SMT/Tools/yices_solver.ML \ + SMT/Tools/z3_interface.ML SMT/Tools/z3_solver.ML SMT/Tools/z3_model.ML + @cd SMT; $(ISABELLE_TOOL) usedir -b -g true $(OUT)/HOL-Word HOL-SMT + + ## clean clean: @@ -1143,7 +1163,7 @@ $(LOG)/HOL-Induct.gz $(LOG)/HOL-ex.gz \ $(LOG)/HOL-Subst.gz $(LOG)/HOL-IMP.gz \ $(LOG)/HOL-IMPP.gz $(LOG)/HOL-Hoare.gz \ - $(LOG)/HOL-HoareParallel.gz $(LOG)/HOL-Lex.gz \ + $(LOG)/HOL-Hoare_Parallel.gz $(LOG)/HOL-Lex.gz \ $(LOG)/HOL-Algebra.gz $(LOG)/HOL-Auth.gz \ $(LOG)/HOL-UNITY.gz $(LOG)/HOL-Modelcheck.gz \ $(LOG)/HOL-Lambda.gz $(LOG)/HOL-Bali.gz \ @@ -1156,4 +1176,4 @@ $(LOG)/HOL-Unix.gz $(OUT)/HOL-Word $(LOG)/HOL-Word.gz \ $(LOG)/HOL-Word-Examples.gz $(OUT)/HOL-NSA \ $(LOG)/HOL-NSA.gz $(LOG)/HOL-NSA-Examples.gz \ - $(LOG)/HOL-Mirabelle.gz + $(LOG)/HOL-Mirabelle.gz $(LOG)/HOL-SMT.gz diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Lambda/Eta.thy --- a/src/HOL/Lambda/Eta.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Lambda/Eta.thy Tue Sep 22 13:52:19 2009 +1000 @@ -1,5 +1,4 @@ (* Title: HOL/Lambda/Eta.thy - ID: $Id$ Author: Tobias Nipkow and Stefan Berghofer Copyright 1995, 2005 TU Muenchen *) @@ -87,7 +86,6 @@ lemma square_eta: "square eta eta (eta^==) (eta^==)" apply (unfold square_def id_def) apply (rule impI [THEN allI [THEN allI]]) - apply simp apply (erule eta.induct) apply (slowsimp intro: subst_not_free eta_subst free_eta [THEN iffD1]) apply safe diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Library/Executable_Set.thy --- a/src/HOL/Library/Executable_Set.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Library/Executable_Set.thy Tue Sep 22 13:52:19 2009 +1000 @@ -32,8 +32,8 @@ declare inter [code] -declare Inter_image_eq [symmetric, code] -declare Union_image_eq [symmetric, code] +declare Inter_image_eq [symmetric, code_unfold] +declare Union_image_eq [symmetric, code_unfold] declare List_Set.project_def [symmetric, code_unfold] diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML --- a/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Mirabelle/Tools/mirabelle_sledgehammer.ML Tue Sep 22 13:52:19 2009 +1000 @@ -38,9 +38,9 @@ } datatype min_data = MinData of { - calls: int, - ratios: int, - lemmas: int + succs: int, + ab_ratios: int, + it_ratios: int } (* The first me_data component is only used if "minimize" is on. @@ -52,8 +52,8 @@ ShData{calls=calls, success=success, lemmas=lemmas, time_isa=time_isa, time_atp=time_atp, time_atp_fail=time_atp_fail} -fun make_min_data (calls, ratios, lemmas) = - MinData{calls=calls, ratios=ratios, lemmas=lemmas} +fun make_min_data (succs, ab_ratios, it_ratios) = + MinData{succs=succs, ab_ratios=ab_ratios, it_ratios=it_ratios} fun make_me_data (calls, success, time, timeout, lemmas, posns) = MeData{calls=calls, success=success, time=time, timeout=timeout, lemmas=lemmas, posns=posns} @@ -61,7 +61,7 @@ val empty_data = Data(make_sh_data (0, 0, 0, 0, 0, 0), make_me_data(0, 0, 0, 0, 0, []), - MinData{calls=0, ratios=0, lemmas=0}, + MinData{succs=0, ab_ratios=0, it_ratios=0}, make_me_data(0, 0, 0, 0, 0, [])) fun map_sh_data f @@ -70,8 +70,8 @@ meda0, minda, meda) fun map_min_data f - (Data(shda, meda0, MinData{calls,ratios,lemmas}, meda)) = - Data(shda, meda0, make_min_data(f(calls,ratios,lemmas)), meda) + (Data(shda, meda0, MinData{succs,ab_ratios,it_ratios}, meda)) = + Data(shda, meda0, make_min_data(f(succs,ab_ratios,it_ratios)), meda) fun map_me_data0 f (Data (shda, MeData{calls,success,time,timeout,lemmas,posns}, minda, meda)) = Data(shda, make_me_data(f (calls,success,time,timeout,lemmas,posns)), minda, meda) @@ -103,11 +103,14 @@ map_sh_data (fn (calls, success, lemmas, time_isa, time_atp, time_atp_fail) => (calls, success, lemmas, time_isa, time_atp, time_atp_fail + t)) -val inc_min_calls = - map_min_data (fn (calls, ratios, lemmas) => (calls + 1, ratios, lemmas)) +val inc_min_succs = + map_min_data (fn (succs,ab_ratios,it_ratios) => (succs+1, ab_ratios, it_ratios)) -fun inc_min_ratios n = - map_min_data (fn (calls, ratios, lemmas) => (calls, ratios + n, lemmas)) +fun inc_min_ab_ratios r = + map_min_data (fn (succs, ab_ratios, it_ratios) => (succs, ab_ratios+r, it_ratios)) + +fun inc_min_it_ratios r = + map_min_data (fn (succs, ab_ratios, it_ratios) => (succs, ab_ratios, it_ratios+r)) val inc_metis_calls = map_me_data (fn (calls, success, time, timeout, lemmas,posns) @@ -192,8 +195,6 @@ (log ("Total number of " ^ tag ^ "metis calls: " ^ str metis_calls); log ("Number of successful " ^ tag ^ "metis calls: " ^ str metis_success); log ("Number of " ^ tag ^ "metis timeouts: " ^ str metis_timeout); - log ("Number of " ^ tag ^ "metis exceptions: " ^ - str (sh_success - metis_success - metis_timeout)); log ("Success rate: " ^ percentage metis_success sh_calls ^ "%"); log ("Number of successful " ^ tag ^ "metis lemmas: " ^ str metis_lemmas); log ("Total time for successful metis calls: " ^ str3 (time metis_time)); @@ -204,9 +205,10 @@ else () ) -fun log_min_data log calls ratios = - (log ("Number of minimizations: " ^ string_of_int calls); - log ("Minimization ratios: " ^ string_of_int ratios) +fun log_min_data log succs ab_ratios it_ratios = + (log ("Number of successful minimizations: " ^ string_of_int succs); + log ("After/before ratios: " ^ string_of_int ab_ratios); + log ("Iterations ratios: " ^ string_of_int it_ratios) ) in @@ -217,7 +219,7 @@ MeData{calls=metis_calls0, success=metis_success0, time=metis_time0, timeout=metis_timeout0, lemmas=metis_lemmas0,posns=metis_posns0}, - MinData{calls=min_calls, ratios=min_ratios, lemmas=min_lemmas}, + MinData{succs=min_succs, ab_ratios=ab_ratios, it_ratios=it_ratios}, MeData{calls=metis_calls, success=metis_success, time=metis_time, timeout=metis_timeout, lemmas=metis_lemmas,posns=metis_posns})) = @@ -230,7 +232,7 @@ metis_success metis_time metis_timeout metis_lemmas metis_posns else (); log ""; if metis_calls0 > 0 - then (log_min_data log min_calls min_ratios; log ""; + then (log_min_data log min_succs ab_ratios it_ratios; log ""; log_metis_data log "unminimized " sh_calls sh_success metis_calls0 metis_success0 metis_time0 metis_timeout0 metis_lemmas0 metis_posns0) else () @@ -362,8 +364,9 @@ in case minimize timeout st (these (!named_thms)) of (SOME (named_thms',its), msg) => - (change_data id inc_min_calls; - change_data id (inc_min_ratios ((100*its) div n0)); + (change_data id inc_min_succs; + change_data id (inc_min_ab_ratios ((100 * length named_thms') div n0)); + change_data id (inc_min_it_ratios ((100*its) div n0)); if length named_thms' = n0 then log (minimize_tag id ^ "already minimal") else (named_thms := SOME named_thms'; @@ -398,27 +401,29 @@ |> log o prefix (metis_tag id) end -fun sledgehammer_action args id (st as {log, ...}: Mirabelle.run_args) = +fun sledgehammer_action args id (st as {log, pre, ...}: Mirabelle.run_args) = + if can Logic.dest_conjunction (Thm.major_prem_of(snd(snd(Proof.get_goal pre)))) + then () else let val metis_fns = (inc_metis_calls, inc_metis_success, inc_metis_time, inc_metis_timeout, inc_metis_lemmas, inc_metis_posns) val metis0_fns = (inc_metis_calls0, inc_metis_success0, inc_metis_time0, inc_metis_timeout0, inc_metis_lemmas0, inc_metis_posns0) val named_thms = ref (NONE : (string * thm list) list option) - - fun if_enabled k f = - if AList.defined (op =) args k andalso is_some (!named_thms) - then f id st else () - - val _ = Mirabelle.catch sh_tag (run_sledgehammer args named_thms) id st - val _ = if_enabled minimizeK - (Mirabelle.catch metis_tag (run_metis metis0_fns args (these (!named_thms)))) - val _ = if_enabled minimizeK - (Mirabelle.catch minimize_tag (run_minimize args named_thms)) - val _ = if is_some (!named_thms) - then Mirabelle.catch metis_tag (run_metis metis_fns args (these (!named_thms))) id st - else () - in () end + val minimize = AList.defined (op =) args minimizeK + in + Mirabelle.catch sh_tag (run_sledgehammer args named_thms) id st; + if is_some (!named_thms) + then + (if minimize + then Mirabelle.catch metis_tag (run_metis metis0_fns args (these (!named_thms))) id st + else (); + if minimize andalso not(null(these(!named_thms))) + then Mirabelle.catch minimize_tag (run_minimize args named_thms) id st + else (); + Mirabelle.catch metis_tag (run_metis metis_fns args (these (!named_thms))) id st) + else () + end fun invoke args = let diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Nominal/Examples/Crary.thy --- a/src/HOL/Nominal/Examples/Crary.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Nominal/Examples/Crary.thy Tue Sep 22 13:52:19 2009 +1000 @@ -408,20 +408,7 @@ declare trm.inject[simp del] -lemma whn_eqvt[eqvt]: - fixes pi::"name prm" - assumes a: "t \ t'" - shows "(pi\t) \ (pi\t')" -using a -apply(induct) -apply(rule QAN_Reduce) -apply(rule whr_def.eqvt) -apply(assumption)+ -apply(rule QAN_Normal) -apply(auto) -apply(drule_tac pi="rev pi" in whr_def.eqvt) -apply(perm_simp) -done +equivariance whn_def lemma red_unicity : assumes a: "x \ a" @@ -631,6 +618,7 @@ apply (force) apply (rule ty_cases) done + termination by lexicographic_order lemma logical_monotonicity: @@ -968,7 +956,7 @@ then show "\ \ s \ t : T" using main_lemma(1) val by simp qed -text {* We leave soundness as an exercise - like in the book :-) \\ +text {* We leave soundness as an exercise - just like Crary in the ATS book :-) \\ @{prop[mode=IfThen] "\\ \ s \ t : T; \ \ t : T; \ \ s : T\ \ \ \ s \ t : T"} \\ @{prop "\\ \ s \ t : T; \ \ t : T; \ \ s : T\ \ \ \ s \ t : T"} *} diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Nominal/Examples/Nominal_Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/Nominal/Examples/Nominal_Examples.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,25 @@ +(* Author: Christian Urban TU Muenchen *) + +header {* Various examples involving nominal datatypes. *} + +theory Nominal_Examples +imports + CR + CR_Takahashi + Class + Compile + Fsub + Height + Lambda_mu + SN + Weakening + Crary + SOS + LocalWeakening + Support + Contexts + Standardization + W +begin + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Nominal/Examples/ROOT.ML --- a/src/HOL/Nominal/Examples/ROOT.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Nominal/Examples/ROOT.ML Tue Sep 22 13:52:19 2009 +1000 @@ -1,27 +1,4 @@ -(* Title: HOL/Nominal/Examples/ROOT.ML - ID: $Id$ - Author: Christian Urban, TU Muenchen - -Various examples involving nominal datatypes. -*) -use_thys [ - "CR", - "CR_Takahashi", - "Class", - "Compile", - "Fsub", - "Height", - "Lambda_mu", - "SN", - "Weakening", - "Crary", - "SOS", - "LocalWeakening", - "Support", - "Contexts", - "Standardization", - "W" -]; +use_thy "Nominal_Examples"; -setmp_noncritical quick_and_dirty true use_thy "VC_Condition"; +setmp_noncritical quick_and_dirty true use_thy "VC_Condition"; (*FIXME*) diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Nominal/Nominal.thy --- a/src/HOL/Nominal/Nominal.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Nominal/Nominal.thy Tue Sep 22 13:52:19 2009 +1000 @@ -2623,74 +2623,77 @@ avoiding a finitely supported c and there is a permutation which 'translates' between both sets. *} + lemma at_set_avoiding_aux: fixes Xs::"'a set" and As::"'a set" assumes at: "at TYPE('a)" - and a: "finite Xs" and b: "Xs \ As" and c: "finite As" and d: "finite ((supp c)::'a set)" - shows "\(Ys::'a set) (pi::'a prm). Ys\*c \ Ys \ As = {} \ (pi\Xs=Ys) \ - set pi \ Xs \ Ys \ finite Ys" -using a b c -proof (induct) - case empty - have "({}::'a set)\*c" by (simp add: fresh_star_def) - moreover - have "({}::'a set) \ As = {}" by simp - moreover - have "([]::'a prm)\{} = ({}::'a set)" - by (rule pt1[OF pt_fun_inst, OF at_pt_inst[OF at], OF pt_bool_inst, OF at]) - moreover - have "set ([]::'a prm) \ {} \ {}" by simp - moreover - have "finite ({}::'a set)" by simp - ultimately show ?case by blast -next - case (insert x Xs) - then have ih: "\Ys pi. Ys\*c \ Ys \ As = {} \ pi\Xs = Ys \ set pi \ Xs \ Ys \ finite Ys" by simp - then obtain Ys pi where a1: "Ys\*c" and a2: "Ys \ As = {}" and a3: "pi\Xs = Ys" and - a4: "set pi \ Xs \ Ys" and a5: "finite Ys" by blast - have b: "x\Xs" by fact - have d1: "finite As" by fact - have d2: "finite Xs" by fact - have d3: "insert x Xs \ As" by fact - have "\y::'a. y\(c,x,Ys,As)" using d d1 a5 - by (rule_tac at_exists_fresh'[OF at]) - (simp add: supp_prod at_supp[OF at] at_fin_set_supp[OF at]) - then obtain y::"'a" where e: "y\(c,x,Ys,As)" by blast - have "({y}\Ys)\*c" using a1 e by (simp add: fresh_star_def) - moreover - have "({y}\Ys)\As = {}" using a2 d1 e by (simp add: fresh_prod at_fin_set_fresh[OF at]) - moreover - have "(((pi\x,y)#pi)\(insert x Xs)) = {y}\Ys" - proof - - have eq: "[(pi\x,y)]\Ys = Ys" + shows "\(pi::'a prm). (pi\Xs)\*c \ (pi\Xs) \ As = {} \ set pi \ Xs \ (pi\Xs)" +proof - + from b c have "finite Xs" by (simp add: finite_subset) + then show ?thesis using b + proof (induct) + case empty + have "({}::'a set)\*c" by (simp add: fresh_star_def) + moreover + have "({}::'a set) \ As = {}" by simp + moreover + have "set ([]::'a prm) \ {} \ {}" by simp + moreover + have "([]::'a prm)\{} = ({}::'a set)" + by (rule pt1[OF pt_fun_inst, OF at_pt_inst[OF at] pt_bool_inst at]) + ultimately show ?case by simp + next + case (insert x Xs) + then have ih: "\pi. (pi\Xs)\*c \ (pi\Xs) \ As = {} \ set pi \ Xs \ (pi\Xs)" by simp + then obtain pi where a1: "(pi\Xs)\*c" and a2: "(pi\Xs) \ As = {}" and + a4: "set pi \ Xs \ (pi\Xs)" by blast + have b: "x\Xs" by fact + have d1: "finite As" by fact + have d2: "finite Xs" by fact + have d3: "({x} \ Xs) \ As" using insert(4) by simp + from d d1 d2 + obtain y::"'a" where fr: "y\(c,pi\Xs,As)" + apply(rule_tac at_exists_fresh[OF at, where x="(c,pi\Xs,As)"]) + apply(auto simp add: supp_prod at_supp[OF at] at_fin_set_supp[OF at] + pt_supp_finite_pi[OF pt_fun_inst[OF at_pt_inst[OF at] pt_bool_inst at] at]) + done + have "({y}\(pi\Xs))\*c" using a1 fr by (simp add: fresh_star_def) + moreover + have "({y}\(pi\Xs))\As = {}" using a2 d1 fr + by (simp add: fresh_prod at_fin_set_fresh[OF at]) + moreover + have "pi\x=x" using a4 b a2 d3 + by (rule_tac at_prm_fresh2[OF at]) (auto) + then have "set ((pi\x,y)#pi) \ ({x} \ Xs) \ ({y}\(pi\Xs))" using a4 by auto + moreover + have "(((pi\x,y)#pi)\({x} \ Xs)) = {y}\(pi\Xs)" proof - - have "(pi\x)\Ys" using a3[symmetric] b d2 - by(simp add: pt_fresh_bij[OF pt_fun_inst, OF at_pt_inst[OF at], OF pt_bool_inst, OF at, OF at] - at_fin_set_fresh[OF at]) - moreover - have "y\Ys" using e by simp - ultimately show "[(pi\x,y)]\Ys = Ys" - by (simp add: pt_fresh_fresh[OF pt_fun_inst, OF at_pt_inst[OF at], OF pt_bool_inst, OF at, OF at]) + have eq: "[(pi\x,y)]\(pi\Xs) = (pi\Xs)" + proof - + have "(pi\x)\(pi\Xs)" using b d2 + by(simp add: pt_fresh_bij[OF pt_fun_inst, OF at_pt_inst[OF at], + OF pt_bool_inst, OF at, OF at] + at_fin_set_fresh[OF at]) + moreover + have "y\(pi\Xs)" using fr by simp + ultimately show "[(pi\x,y)]\(pi\Xs) = (pi\Xs)" + by (simp add: pt_fresh_fresh[OF pt_fun_inst, + OF at_pt_inst[OF at], OF pt_bool_inst, OF at, OF at]) + qed + have "(((pi\x,y)#pi)\({x}\Xs)) = ([(pi\x,y)]\(pi\({x}\Xs)))" + by (simp add: pt2[symmetric, OF pt_fun_inst, OF at_pt_inst[OF at], + OF pt_bool_inst, OF at]) + also have "\ = {y}\([(pi\x,y)]\(pi\Xs))" + by (simp only: union_eqvt perm_set_eq[OF at_pt_inst[OF at], OF at] at_calc[OF at])(auto) + finally show "(((pi\x,y)#pi)\({x} \ Xs)) = {y}\(pi\Xs)" using eq by simp qed - have "(((pi\x,y)#pi)\({x}\Xs)) = ([(pi\x,y)]\(pi\({x}\Xs)))" - by (simp add: pt2[symmetric, OF pt_fun_inst, OF at_pt_inst[OF at], OF pt_bool_inst, OF at]) - also have "\ = {y}\([(pi\x,y)]\(pi\Xs))" - by (simp only: union_eqvt perm_set_eq[OF at_pt_inst[OF at], OF at] at_calc[OF at])(auto) - also have "\ = {y}\([(pi\x,y)]\Ys)" using a3 by simp - also have "\ = {y}\Ys" using eq by simp - finally show "(((pi\x,y)#pi)\(insert x Xs)) = {y}\Ys" by auto + ultimately + show ?case by (rule_tac x="(pi\x,y)#pi" in exI) (auto) qed - moreover - have "pi\x=x" using a4 b a2 a3 d3 by (rule_tac at_prm_fresh2[OF at]) (auto) - then have "set ((pi\x,y)#pi) \ (insert x Xs) \ ({y}\Ys)" using a4 by auto - moreover - have "finite ({y}\Ys)" using a5 by simp - ultimately - show ?case by blast qed lemma at_set_avoiding: @@ -2698,10 +2701,10 @@ assumes at: "at TYPE('a)" and a: "finite Xs" and b: "finite ((supp c)::'a set)" - obtains pi::"'a prm" where "(pi \ Xs) \* c" and "set pi \ Xs \ (pi \ Xs)" - using a b - by (frule_tac As="Xs" in at_set_avoiding_aux[OF at]) auto - + obtains pi::"'a prm" where "(pi\Xs)\*c" and "set pi \ Xs \ (pi\Xs)" +using a b at_set_avoiding_aux[OF at, where Xs="Xs" and As="Xs" and c="c"] +by (blast) + section {* composition instances *} (* ============================= *) diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Predicate.thy --- a/src/HOL/Predicate.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Predicate.thy Tue Sep 22 13:52:19 2009 +1000 @@ -75,29 +75,29 @@ subsubsection {* Binary union *} -lemma sup1_iff [simp]: "sup A B x \ A x | B x" +lemma sup1_iff: "sup A B x \ A x | B x" by (simp add: sup_fun_eq sup_bool_eq) -lemma sup2_iff [simp]: "sup A B x y \ A x y | B x y" +lemma sup2_iff: "sup A B x y \ A x y | B x y" by (simp add: sup_fun_eq sup_bool_eq) lemma sup_Un_eq [pred_set_conv]: "sup (\x. x \ R) (\x. x \ S) = (\x. x \ R \ S)" - by (simp add: expand_fun_eq) + by (simp add: sup1_iff expand_fun_eq) lemma sup_Un_eq2 [pred_set_conv]: "sup (\x y. (x, y) \ R) (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" - by (simp add: expand_fun_eq) + by (simp add: sup2_iff expand_fun_eq) lemma sup1I1 [elim?]: "A x \ sup A B x" - by simp + by (simp add: sup1_iff) lemma sup2I1 [elim?]: "A x y \ sup A B x y" - by simp + by (simp add: sup2_iff) lemma sup1I2 [elim?]: "B x \ sup A B x" - by simp + by (simp add: sup1_iff) lemma sup2I2 [elim?]: "B x y \ sup A B x y" - by simp + by (simp add: sup2_iff) text {* \medskip Classical introduction rule: no commitment to @{text A} vs @@ -105,115 +105,115 @@ *} lemma sup1CI [intro!]: "(~ B x ==> A x) ==> sup A B x" - by auto + by (auto simp add: sup1_iff) lemma sup2CI [intro!]: "(~ B x y ==> A x y) ==> sup A B x y" - by auto + by (auto simp add: sup2_iff) lemma sup1E [elim!]: "sup A B x ==> (A x ==> P) ==> (B x ==> P) ==> P" - by simp iprover + by (simp add: sup1_iff) iprover lemma sup2E [elim!]: "sup A B x y ==> (A x y ==> P) ==> (B x y ==> P) ==> P" - by simp iprover + by (simp add: sup2_iff) iprover subsubsection {* Binary intersection *} -lemma inf1_iff [simp]: "inf A B x \ A x \ B x" +lemma inf1_iff: "inf A B x \ A x \ B x" by (simp add: inf_fun_eq inf_bool_eq) -lemma inf2_iff [simp]: "inf A B x y \ A x y \ B x y" +lemma inf2_iff: "inf A B x y \ A x y \ B x y" by (simp add: inf_fun_eq inf_bool_eq) lemma inf_Int_eq [pred_set_conv]: "inf (\x. x \ R) (\x. x \ S) = (\x. x \ R \ S)" - by (simp add: expand_fun_eq) + by (simp add: inf1_iff expand_fun_eq) lemma inf_Int_eq2 [pred_set_conv]: "inf (\x y. (x, y) \ R) (\x y. (x, y) \ S) = (\x y. (x, y) \ R \ S)" - by (simp add: expand_fun_eq) + by (simp add: inf2_iff expand_fun_eq) lemma inf1I [intro!]: "A x ==> B x ==> inf A B x" - by simp + by (simp add: inf1_iff) lemma inf2I [intro!]: "A x y ==> B x y ==> inf A B x y" - by simp + by (simp add: inf2_iff) lemma inf1D1: "inf A B x ==> A x" - by simp + by (simp add: inf1_iff) lemma inf2D1: "inf A B x y ==> A x y" - by simp + by (simp add: inf2_iff) lemma inf1D2: "inf A B x ==> B x" - by simp + by (simp add: inf1_iff) lemma inf2D2: "inf A B x y ==> B x y" - by simp + by (simp add: inf2_iff) lemma inf1E [elim!]: "inf A B x ==> (A x ==> B x ==> P) ==> P" - by simp + by (simp add: inf1_iff) lemma inf2E [elim!]: "inf A B x y ==> (A x y ==> B x y ==> P) ==> P" - by simp + by (simp add: inf2_iff) subsubsection {* Unions of families *} -lemma SUP1_iff [simp]: "(SUP x:A. B x) b = (EX x:A. B x b)" +lemma SUP1_iff: "(SUP x:A. B x) b = (EX x:A. B x b)" by (simp add: SUPR_def Sup_fun_def Sup_bool_def) blast -lemma SUP2_iff [simp]: "(SUP x:A. B x) b c = (EX x:A. B x b c)" +lemma SUP2_iff: "(SUP x:A. B x) b c = (EX x:A. B x b c)" by (simp add: SUPR_def Sup_fun_def Sup_bool_def) blast lemma SUP1_I [intro]: "a : A ==> B a b ==> (SUP x:A. B x) b" - by auto + by (auto simp add: SUP1_iff) lemma SUP2_I [intro]: "a : A ==> B a b c ==> (SUP x:A. B x) b c" - by auto + by (auto simp add: SUP2_iff) lemma SUP1_E [elim!]: "(SUP x:A. B x) b ==> (!!x. x : A ==> B x b ==> R) ==> R" - by auto + by (auto simp add: SUP1_iff) lemma SUP2_E [elim!]: "(SUP x:A. B x) b c ==> (!!x. x : A ==> B x b c ==> R) ==> R" - by auto + by (auto simp add: SUP2_iff) lemma SUP_UN_eq: "(SUP i. (\x. x \ r i)) = (\x. x \ (UN i. r i))" - by (simp add: expand_fun_eq) + by (simp add: SUP1_iff expand_fun_eq) lemma SUP_UN_eq2: "(SUP i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (UN i. r i))" - by (simp add: expand_fun_eq) + by (simp add: SUP2_iff expand_fun_eq) subsubsection {* Intersections of families *} -lemma INF1_iff [simp]: "(INF x:A. B x) b = (ALL x:A. B x b)" +lemma INF1_iff: "(INF x:A. B x) b = (ALL x:A. B x b)" by (simp add: INFI_def Inf_fun_def Inf_bool_def) blast -lemma INF2_iff [simp]: "(INF x:A. B x) b c = (ALL x:A. B x b c)" +lemma INF2_iff: "(INF x:A. B x) b c = (ALL x:A. B x b c)" by (simp add: INFI_def Inf_fun_def Inf_bool_def) blast lemma INF1_I [intro!]: "(!!x. x : A ==> B x b) ==> (INF x:A. B x) b" - by auto + by (auto simp add: INF1_iff) lemma INF2_I [intro!]: "(!!x. x : A ==> B x b c) ==> (INF x:A. B x) b c" - by auto + by (auto simp add: INF2_iff) lemma INF1_D [elim]: "(INF x:A. B x) b ==> a : A ==> B a b" - by auto + by (auto simp add: INF1_iff) lemma INF2_D [elim]: "(INF x:A. B x) b c ==> a : A ==> B a b c" - by auto + by (auto simp add: INF2_iff) lemma INF1_E [elim]: "(INF x:A. B x) b ==> (B a b ==> R) ==> (a ~: A ==> R) ==> R" - by auto + by (auto simp add: INF1_iff) lemma INF2_E [elim]: "(INF x:A. B x) b c ==> (B a b c ==> R) ==> (a ~: A ==> R) ==> R" - by auto + by (auto simp add: INF2_iff) lemma INF_INT_eq: "(INF i. (\x. x \ r i)) = (\x. x \ (INT i. r i))" - by (simp add: expand_fun_eq) + by (simp add: INF1_iff expand_fun_eq) lemma INF_INT_eq2: "(INF i. (\x y. (x, y) \ r i)) = (\x y. (x, y) \ (INT i. r i))" - by (simp add: expand_fun_eq) + by (simp add: INF2_iff expand_fun_eq) subsection {* Predicates as relations *} @@ -429,7 +429,7 @@ by (auto simp add: bind_def sup_pred_def expand_fun_eq) lemma Sup_bind: "(\A \= f) = \((\x. x \= f) ` A)" - by (auto simp add: bind_def Sup_pred_def expand_fun_eq) + by (auto simp add: bind_def Sup_pred_def SUP1_iff expand_fun_eq) lemma pred_iffI: assumes "\x. eval A x \ eval B x" @@ -462,10 +462,10 @@ unfolding bot_pred_def by auto lemma supI1: "eval A x \ eval (A \ B) x" - unfolding sup_pred_def by simp + unfolding sup_pred_def by (simp add: sup1_iff) lemma supI2: "eval B x \ eval (A \ B) x" - unfolding sup_pred_def by simp + unfolding sup_pred_def by (simp add: sup1_iff) lemma supE: "eval (A \ B) x \ (eval A x \ P) \ (eval B x \ P) \ P" unfolding sup_pred_def by auto diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/Examples/SMT_Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Examples/SMT_Examples.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,183 @@ +(* Title: SMT_Examples.thy + Author: Sascha Boehme, TU Muenchen +*) + +header {* Examples for the 'smt' tactic. *} + +theory SMT_Examples +imports "../SMT" +begin + +declare [[smt_solver=z3, z3_proofs=false]] +declare [[smt_trace=false]] + + +section {* Propositional and first-order logic *} + +lemma "True" by smt +lemma "p \ \p" by smt +lemma "(p \ True) = p" by smt +lemma "(p \ q) \ \p \ q" by smt +lemma "(a \ b) \ (c \ d) \ (a \ b) \ (c \ d)" by smt +lemma "P=P=P=P=P=P=P=P=P=P" by smt + +axiomatization symm_f :: "'a \ 'a \ 'a" where + symm_f: "symm_f x y = symm_f y x" +lemma "a = a \ symm_f a b = symm_f b a" by (smt add: symm_f) + + +section {* Linear arithmetic *} + +lemma "(3::int) = 3" by smt +lemma "(3::real) = 3" by smt +lemma "(3 :: int) + 1 = 4" by smt +lemma "max (3::int) 8 > 5" by smt +lemma "abs (x :: real) + abs y \ abs (x + y)" by smt +lemma "let x = (2 :: int) in x + x \ 5" by smt + +text{* +The following example was taken from HOL/ex/PresburgerEx.thy, where it says: + + This following theorem proves that all solutions to the + recurrence relation $x_{i+2} = |x_{i+1}| - x_i$ are periodic with + period 9. The example was brought to our attention by John + Harrison. It does does not require Presburger arithmetic but merely + quantifier-free linear arithmetic and holds for the rationals as well. + + Warning: it takes (in 2006) over 4.2 minutes! + +There, it is proved by "arith". SMT is able to prove this within a fraction +of one second. +*} + +lemma "\ x3 = abs x2 - x1; x4 = abs x3 - x2; x5 = abs x4 - x3; + x6 = abs x5 - x4; x7 = abs x6 - x5; x8 = abs x7 - x6; + x9 = abs x8 - x7; x10 = abs x9 - x8; x11 = abs x10 - x9 \ + \ x1 = x10 & x2 = (x11::int)" + by smt + +lemma "\x::int. 0 < x" by smt +lemma "\x::real. 0 < x" by smt +lemma "\x y::int. x < y \ (2 * x + 1) < (2 * y)" by smt +lemma "\x y::int. (2 * x + 1) \ (2 * y)" by smt +lemma "~ (\x y z::int. 4 * x + -6 * y = (1::int))" by smt +lemma "~ (\x::int. False)" by smt + + +section {* Non-linear arithmetic *} + +lemma "((x::int) * (1 + y) - x * (1 - y)) = (2 * x * y)" by smt +lemma + "(U::int) + (1 + p) * (b + e) + p * d = + U + (2 * (1 + p) * (b + e) + (1 + p) * d + d * p) - (1 + p) * (b + d + e)" + by smt + + +section {* Linear arithmetic for natural numbers *} + +lemma "a < 3 \ (7::nat) > 2 * a" by smt +lemma "let x = (1::nat) + y in x - y > 0 * x" by smt +lemma + "let x = (1::nat) + y in + let P = (if x > 0 then True else False) in + False \ P = (x - 1 = y) \ (\P \ False)" + by smt + + +section {* Bitvectors *} + +locale bv +begin + +declare [[smt_solver=z3]] + +lemma "(27 :: 4 word) = -5" by smt +lemma "(27 :: 4 word) = 11" by smt +lemma "23 < (27::8 word)" by smt +lemma "27 + 11 = (6::5 word)" by smt +lemma "7 * 3 = (21::8 word)" by smt +lemma "11 - 27 = (-16::8 word)" by smt +lemma "- -11 = (11::5 word)" by smt +lemma "-40 + 1 = (-39::7 word)" by smt +lemma "a + 2 * b + c - b = (b + c) + (a :: 32 word)" by smt + +lemma "0b110 AND 0b101 = (0b100 :: 32 word)" by smt +lemma "0b110 OR 0b011 = (0b111 :: 8 word)" by smt +lemma "0xF0 XOR 0xFF = (0x0F :: 8 word)" by smt +lemma "NOT (0xF0 :: 16 word) = 0xFF0F" by smt + +lemma "word_cat (27::4 word) (27::8 word) = (2843::12 word)" by smt +lemma "word_cat (0b0011::4 word) (0b1111::6word) = (0b0011001111 :: 10 word)" + by smt + +lemma "slice 1 (0b10110 :: 4 word) = (0b11 :: 2 word)" by smt + +lemma "ucast (0b1010 :: 4 word) = (0b1010 :: 10 word)" by smt +lemma "scast (0b1010 :: 4 word) = (0b111010 :: 6 word)" by smt + +lemma "bv_lshr 0b10011 2 = (0b100::8 word)" by smt +lemma "bv_ashr 0b10011 2 = (0b100::8 word)" by smt + +lemma "word_rotr 2 0b0110 = (0b1001::4 word)" by smt +lemma "word_rotl 1 0b1110 = (0b1101::4 word)" by smt + +lemma "(x AND 0xff00) OR (x AND 0x00ff) = (x::16 word)" by smt + +lemma "w < 256 \ (w :: 16 word) AND 0x00FF = w" by smt + +end + + +section {* Pairs *} + +lemma "fst (x, y) = a \ x = a" by smt +lemma "p1 = (x, y) \ p2 = (y, x) \ fst p1 = snd p2" by smt + + +section {* Higher-order problems and recursion *} + +lemma "(f g x = (g x \ True)) \ (f g x = True) \ (g x = True)" by smt +lemma "P ((2::int) < 3) = P True" by smt +lemma "P ((2::int) < 3) = (P True :: bool)" by smt +lemma "P (0 \ (a :: 4 word)) = P True" using [[smt_solver=z3]] by smt +lemma "id 3 = 3 \ id True = True" by (smt add: id_def) +lemma "i \ i1 \ i \ i2 \ ((f (i1 := v1)) (i2 := v2)) i = f i" by smt +lemma "map (\i::nat. i + 1) [0, 1] = [1, 2]" by (smt add: map.simps) +lemma "(ALL x. P x) | ~ All P" by smt + +fun dec_10 :: "nat \ nat" where + "dec_10 n = (if n < 10 then n else dec_10 (n - 10))" +lemma "dec_10 (4 * dec_10 4) = 6" by (smt add: dec_10.simps) + +axiomatization + eval_dioph :: "int list \ nat list \ int" + where + eval_dioph_mod: + "eval_dioph ks xs mod int n = eval_dioph ks (map (\x. x mod n) xs) mod int n" + and + eval_dioph_div_mult: + "eval_dioph ks (map (\x. x div n) xs) * int n + + eval_dioph ks (map (\x. x mod n) xs) = eval_dioph ks xs" +lemma + "(eval_dioph ks xs = l) = + (eval_dioph ks (map (\x. x mod 2) xs) mod 2 = l mod 2 \ + eval_dioph ks (map (\x. x div 2) xs) = + (l - eval_dioph ks (map (\x. x mod 2) xs)) div 2)" + using [[smt_solver=z3]] + by (smt add: eval_dioph_mod[where n=2] eval_dioph_div_mult[where n=2]) + + +section {* Monomorphization examples *} + +definition P :: "'a \ bool" where "P x = True" +lemma poly_P: "P x \ (P [x] \ \P[x])" by (simp add: P_def) +lemma "P (1::int)" by (smt add: poly_P) + +consts g :: "'a \ nat" +axioms + g1: "g (Some x) = g [x]" + g2: "g None = g []" + g3: "g xs = length xs" +lemma "g (Some (3::int)) = g (Some True)" by (smt add: g1 g2 g3 list.size) + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/ROOT.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/ROOT.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,1 @@ +use_thy "SMT"; diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/SMT.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/SMT.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,51 @@ +(* Title: HOL/SMT/SMT.thy + Author: Sascha Boehme, TU Muenchen +*) + +header {* SMT method using external SMT solvers (CVC3, Yices, Z3) *} + +theory SMT +imports SMT_Definitions +uses + "Tools/smt_normalize.ML" + "Tools/smt_monomorph.ML" + "Tools/smt_translate.ML" + "Tools/smt_solver.ML" + "Tools/smtlib_interface.ML" + "Tools/cvc3_solver.ML" + "Tools/yices_solver.ML" + "Tools/z3_model.ML" + "Tools/z3_interface.ML" + "Tools/z3_solver.ML" +begin + +setup {* + SMT_Normalize.setup #> + SMT_Solver.setup #> + CVC3_Solver.setup #> + Yices_Solver.setup #> + Z3_Solver.setup +*} + +ML {* +OuterSyntax.improper_command "smt_status" + "Show the available SMT solvers and the currently selected solver." + OuterKeyword.diag + (Scan.succeed (Toplevel.no_timing o Toplevel.keep (fn state => + SMT_Solver.print_setup (Context.Proof (Toplevel.context_of state))))) +*} + +method_setup smt = {* + let fun solver thms ctxt = SMT_Solver.smt_tac ctxt thms + in + Scan.optional (Scan.lift (Args.add -- Args.colon) |-- Attrib.thms) [] >> + (Method.SIMPLE_METHOD' oo solver) + end +*} "Applies an SMT solver to the current goal." + +declare [[ smt_solver = z3, smt_timeout = 20, smt_trace = false ]] +declare [[ smt_unfold_defs = true ]] +declare [[ z3_proofs = false ]] + +end + diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/SMT_Definitions.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/SMT_Definitions.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,113 @@ +(* Title: HOL/SMT/SMT_Definitions.thy + Author: Sascha Boehme, TU Muenchen +*) + +header {* SMT-specific definitions *} + +theory SMT_Definitions +imports Real Word "~~/src/HOL/Decision_Procs/Dense_Linear_Order" +begin + +section {* Triggers for quantifier instantiation *} + +text {* +Some SMT solvers support triggers for quantifier instantiation. Each trigger +consists of one ore more patterns. A pattern may either be a list of positive +subterms (the first being tagged by "pat" and the consecutive subterms tagged +by "andpat"), or a list of negative subterms (the first being tagged by "nopat" +and the consecutive subterms tagged by "andpat"). +*} + +datatype pattern = Pattern + +definition pat :: "'a \ pattern" +where "pat _ = Pattern" + +definition nopat :: "bool \ pattern" +where "nopat _ = Pattern" + +definition andpat :: "pattern \ 'a \ pattern" (infixl "andpat" 60) +where "_ andpat _ = Pattern" + +definition trigger :: "pattern list \ bool \ bool" +where "trigger _ P = P" + + +section {* Arithmetic *} + +text {* +The sign of @{term "op mod :: int \ int \ int"} follows the sign of the +divisor. In contrast to that, the sign of the following operation is that of +the dividend. +*} + +definition rem :: "int \ int \ int" (infixl "rem" 70) +where "a rem b = + (if (a \ 0 \ b < 0) \ (a < 0 \ b \ 0) then - (a mod b) else a mod b)" + +text {* A decision procedure for linear real arithmetic: *} + +setup {* + Arith_Data.add_tactic "Ferrante-Rackoff" (K FerranteRackoff.dlo_tac) +*} + + +section {* Bitvectors *} + +text {* +The following definitions provide additional functions not found in HOL-Word. +*} + +definition sdiv :: "'a::len word \ 'a word \ 'a word" (infix "sdiv" 70) +where "w1 sdiv w2 = word_of_int (sint w1 div sint w2)" + +definition smod :: "'a::len word \ 'a word \ 'a word" (infix "smod" 70) + (* sign follows divisor *) +where "w1 smod w2 = word_of_int (sint w1 mod sint w2)" + +definition srem :: "'a::len word \ 'a word \ 'a word" (infix "srem" 70) + (* sign follows dividend *) +where "w1 srem w2 = word_of_int (sint w1 rem sint w2)" + +definition bv_shl :: "'a::len0 word \ 'a word \ 'a word" +where "bv_shl w1 w2 = (w1 << unat w2)" + +definition bv_lshr :: "'a::len0 word \ 'a word \ 'a word" +where "bv_lshr w1 w2 = (w1 >> unat w2)" + +definition bv_ashr :: "'a::len word \ 'a word \ 'a word" +where "bv_ashr w1 w2 = (w1 >>> unat w2)" + + +section {* Higher-order encoding *} + +definition "apply" where "apply f x = f x" + + +section {* First-order logic *} + +text {* +Some SMT solver formats require a strict separation between formulas and terms. +The following marker symbols are used internally to separate those categories: +*} + +definition formula :: "bool \ bool" where "formula x = x" +definition "term" where "term x = x" + +text {* +Predicate symbols also occurring as function symbols are turned into function +symbols by translating atomic formulas into terms: +*} + +abbreviation holds :: "bool \ bool" where "holds \ (\P. term P = term True)" + +text {* +The following constant represents equivalence, to be treated differently than +the (polymorphic) equality predicate: +*} + +definition iff :: "bool \ bool \ bool" (infix "iff" 50) where + "(x iff y) = (x = y)" + +end + diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/Tools/cvc3_solver.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/cvc3_solver.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,59 @@ +(* Title: HOL/SMT/Tools/cvc3_solver.ML + Author: Sascha Boehme, TU Muenchen + +Interface of the SMT solver CVC3. +*) + +signature CVC3_SOLVER = +sig + val setup: theory -> theory +end + +structure CVC3_Solver: CVC3_SOLVER = +struct + +val solver_name = "cvc3" +val env_var = "CVC3_SOLVER" + +val options = + ["+counterexample", "-lang", "smtlib", "-output-lang", "presentation"] + +val is_sat = String.isPrefix "Satisfiable." +val is_unsat = String.isPrefix "Unsatisfiable." +val is_unknown = String.isPrefix "Unknown." + +fun cex_kind true = "Counterexample" + | cex_kind false = "Possible counterexample" + +fun raise_cex real ctxt recon ls = + let + val start = String.isPrefix "%Satisfiable Variable Assignment: %" + val index = find_index start ls + val ls = if index > 0 then Library.drop (index + 1, ls) else [] + val p = Pretty.big_list (cex_kind real ^ " found:") (map Pretty.str ls) + in error (Pretty.string_of p) end + +fun core_oracle (SMT_Solver.ProofData {context, output, recon, ...}) = + let + val empty_line = (fn "" => true | _ => false) + val split_first = (fn [] => ("", []) | l :: ls => (l, ls)) + val (l, ls) = split_first (dropwhile empty_line output) + in + if is_unsat l then @{cprop False} + else if is_sat l then raise_cex true context recon ls + else if is_unknown l then raise_cex false context recon ls + else error (solver_name ^ " failed") + end + +fun smtlib_solver oracle _ = + SMT_Solver.SolverConfig { + name = {env_var=env_var, remote_name=solver_name}, + interface = SMTLIB_Interface.interface, + arguments = options, + reconstruct = oracle } + +val setup = + Thm.add_oracle (Binding.name solver_name, core_oracle) #-> (fn (_, oracle) => + SMT_Solver.add_solver (solver_name, smtlib_solver oracle)) + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/Tools/smt_builtin.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/smt_builtin.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,78 @@ +(* Title: HOL/SMT/Tools/smt_builtin.ML + Author: Sascha Boehme, TU Muenchen + +Tables for built-in symbols. +*) + +signature SMT_BUILTIN = +sig + type sterm = (SMT_Translate.sym, typ) sterm + type builtin_fun = typ -> sterm list -> (string * sterm list) option + type table = (typ * builtin_fun) list Symtab.table + + val make: (term * string) list -> table + val add: term * builtin_fun -> table -> table + val lookup: table -> theory -> string * typ -> sterm list -> + (string * sterm list) option + + val bv_rotate: (int -> string) -> builtin_fun + val bv_extend: (int -> string) -> builtin_fun + val bv_extract: (int -> int -> string) -> builtin_fun +end + +structure SMT_Builtin: SMT_BUILTIN = +struct + +structure T = SMT_Translate + +type sterm = (SMT_Translate.sym, typ) sterm +type builtin_fun = typ -> sterm list -> (string * sterm list) option +type table = (typ * builtin_fun) list Symtab.table + +fun make entries = + let + fun dest (t, bn) = + let val (n, T) = Term.dest_Const t + in (n, (Logic.varifyT T, K (pair bn))) end + in Symtab.make (AList.group (op =) (map dest entries)) end + +fun add (t, f) tab = + let val (n, T) = apsnd Logic.varifyT (Term.dest_Const t) + in Symtab.map_default (n, []) (AList.update (op =) (T, f)) tab end + +fun lookup tab thy (n, T) = + AList.lookup (Sign.typ_instance thy) (Symtab.lookup_list tab n) T T + + +fun dest_binT T = + (case T of + Type (@{type_name "Numeral_Type.num0"}, _) => 0 + | Type (@{type_name "Numeral_Type.num1"}, _) => 1 + | Type (@{type_name "Numeral_Type.bit0"}, [T]) => 2 * dest_binT T + | Type (@{type_name "Numeral_Type.bit1"}, [T]) => 1 + 2 * dest_binT T + | _ => raise TYPE ("dest_binT", [T], [])) + +fun dest_wordT T = + (case T of + Type (@{type_name "word"}, [T]) => dest_binT T + | _ => raise TYPE ("dest_wordT", [T], [])) + + +val dest_nat = (fn + SApp (SConst (@{const_name nat}, _), [SApp (SNum (i, _), _)]) => SOME i + | _ => NONE) + +fun bv_rotate mk_name T ts = + dest_nat (hd ts) |> Option.map (fn i => (mk_name i, tl ts)) + +fun bv_extend mk_name T ts = + (case (try dest_wordT (domain_type T), try dest_wordT (range_type T)) of + (SOME i, SOME j) => if j - i >= 0 then SOME (mk_name (j - i), ts) else NONE + | _ => NONE) + +fun bv_extract mk_name T ts = + (case (try dest_wordT (body_type T), dest_nat (hd ts)) of + (SOME i, SOME lb) => SOME (mk_name (i + lb - 1) lb, tl ts) + | _ => NONE) + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/Tools/smt_monomorph.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/smt_monomorph.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,120 @@ +(* Title: HOL/SMT/Tools/smt_monomorph.ML + Author: Sascha Boehme, TU Muenchen + +Monomorphization of terms, i.e., computation of all (necessary) instances. +*) + +signature SMT_MONOMORPH = +sig + val monomorph: theory -> term list -> term list +end + +structure SMT_Monomorph: SMT_MONOMORPH = +struct + +fun selection [] = [] + | selection (x :: xs) = (x, xs) :: map (apsnd (cons x)) (selection xs) + +fun permute [] = [] + | permute [x] = [[x]] + | permute xs = maps (fn (y, ys) => map (cons y) (permute ys)) (selection xs) + +fun fold_all f = fold (fn x => maps (f x)) + + +val typ_has_tvars = Term.exists_subtype (fn TVar _ => true | _ => false) +val term_has_tvars = Term.exists_type typ_has_tvars + +val ignored = member (op =) [ + @{const_name All}, @{const_name Ex}, @{const_name Let}, @{const_name If}, + @{const_name "op ="}, @{const_name zero_class.zero}, + @{const_name one_class.one}, @{const_name number_of}] +fun consts_of ts = AList.group (op =) (fold Term.add_consts ts []) + |> filter_out (ignored o fst) + +val join_consts = curry (AList.join (op =) (K (merge (op =)))) +fun diff_consts cs ds = + let fun diff (n, Ts) = + (case AList.lookup (op =) cs n of + NONE => SOME (n, Ts) + | SOME Us => + let val Ts' = fold (remove (op =)) Us Ts + in if null Ts' then NONE else SOME (n, Ts') end) + in map_filter diff ds end + +fun instances thy is (n, Ts) env = + let + val Us = these (AList.lookup (op =) is n) + val Ts' = filter typ_has_tvars (map (Envir.subst_type env) Ts) + in + (case map_product pair Ts' Us of + [] => [env] + | TUs => map_filter (try (fn TU => Sign.typ_match thy TU env)) TUs) + end + +fun proper_match ps env = + forall (forall (not o typ_has_tvars o Envir.subst_type env) o snd) ps + +val eq_tab = gen_eq_set (op =) o pairself Vartab.dest + +fun specialize thy cs is ((r, ps), ces) (ts, ns) = + let + val ps' = filter (AList.defined (op =) is o fst) ps + + val envs = permute ps' + |> maps (fn ps => fold_all (instances thy is) ps [Vartab.empty]) + |> filter (proper_match ps') + |> filter_out (member eq_tab ces) + |> distinct eq_tab + + val us = map (fn env => Envir.subst_term_types env r) envs + val ns' = join_consts (diff_consts is (diff_consts cs (consts_of us))) ns + in (envs @ ces, (fold (insert (op aconv)) us ts, ns')) end + + +fun incr_tvar_indices i t = + let + val incrT = Logic.incr_tvar i + + fun incr t = + (case t of + Const (n, T) => Const (n, incrT T) + | Free (n, T) => Free (n, incrT T) + | Abs (n, T, t1) => Abs (n, incrT T, incr t1) + | t1 $ t2 => incr t1 $ incr t2 + | _ => t) + in incr t end + + +val monomorph_limit = 10 + +(* Instantiate all polymorphic constants (i.e., constants occurring both with + ground types and type variables) with all (necessary) ground types; thereby + create copies of terms containing those constants. + To prevent non-termination, there is an upper limit for the number of + recursions involved in the fixpoint construction. *) +fun monomorph thy ts = + let + val (ps, ms) = List.partition term_has_tvars ts + + fun with_tvar (n, Ts) = + let val Ts' = filter typ_has_tvars Ts + in if null Ts' then NONE else SOME (n, Ts') end + fun incr t idx = (incr_tvar_indices idx t, idx + Term.maxidx_of_term t + 1) + val rps = fst (fold_map incr ps 0) + |> map (fn r => (r, map_filter with_tvar (consts_of [r]))) + + fun mono count is ces cs ts = + let + val spec = specialize thy cs is + val (ces', (ts', is')) = fold_map spec (rps ~~ ces) (ts, []) + val cs' = join_consts is cs + in + if null is' then ts' + else if count > monomorph_limit then + (Output.warning "monomorphization limit reached"; ts') + else mono (count + 1) is' ces' cs' ts' + end + in mono 0 (consts_of ms) (map (K []) rps) [] ms end + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/Tools/smt_normalize.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/smt_normalize.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,408 @@ +(* Title: HOL/SMT/Tools/smt_normalize.ML + Author: Sascha Boehme, TU Muenchen + +Normalization steps on theorems required by SMT solvers: + * unfold trivial let expressions, + * replace negative numerals by negated positive numerals, + * embed natural numbers into integers, + * add extra rules specifying types and constants which occur frequently, + * lift lambda terms, + * make applications explicit for functions with varying number of arguments, + * fully translate into object logic, add universal closure. +*) + +signature SMT_NORMALIZE = +sig + val normalize_rule: Proof.context -> thm -> thm + val instantiate_free: Thm.cterm * Thm.cterm -> thm -> thm + val discharge_definition: Thm.cterm -> thm -> thm + + val trivial_let: Proof.context -> thm list -> thm list + val positive_numerals: Proof.context -> thm list -> thm list + val nat_as_int: Proof.context -> thm list -> thm list + val unfold_defs: bool Config.T + val add_pair_rules: Proof.context -> thm list -> thm list + val add_fun_upd_rules: Proof.context -> thm list -> thm list + val add_abs_min_max_rules: Proof.context -> thm list -> thm list + + datatype config = + RewriteTrivialLets | + RewriteNegativeNumerals | + RewriteNaturalNumbers | + AddPairRules | + AddFunUpdRules | + AddAbsMinMaxRules + + val normalize: config list -> Proof.context -> thm list -> + Thm.cterm list * thm list + + val setup: theory -> theory +end + +structure SMT_Normalize: SMT_NORMALIZE = +struct + +val norm_binder_conv = Conv.try_conv (More_Conv.rewrs_conv [ + @{lemma "All P == ALL x. P x" by (rule reflexive)}, + @{lemma "Ex P == EX x. P x" by (rule reflexive)}, + @{lemma "Let c P == let x = c in P x" by (rule reflexive)}]) + +fun cert ctxt = Thm.cterm_of (ProofContext.theory_of ctxt) + +fun norm_meta_def cv thm = + let val thm' = Thm.combination thm (Thm.reflexive cv) + in Thm.transitive thm' (Thm.beta_conversion false (Thm.rhs_of thm')) end + +fun norm_def ctxt thm = + (case Thm.prop_of thm of + Const (@{const_name "=="}, _) $ _ $ Abs (_, T, _) => + let val v = Var ((Name.uu, #maxidx (Thm.rep_thm thm) + 1), T) + in norm_def ctxt (norm_meta_def (cert ctxt v) thm) end + | @{term Trueprop} $ (Const (@{const_name "op ="}, _) $ _ $ Abs _) => + norm_def ctxt (thm RS @{thm fun_cong}) + | _ => thm) + +fun normalize_rule ctxt = + Conv.fconv_rule ( + Thm.beta_conversion true then_conv + Thm.eta_conversion then_conv + More_Conv.bottom_conv (K norm_binder_conv) ctxt) #> + norm_def ctxt #> + Drule.forall_intr_vars #> + Conv.fconv_rule ObjectLogic.atomize + +fun instantiate_free (cv, ct) thm = + if Term.exists_subterm (equal (Thm.term_of cv)) (Thm.prop_of thm) + then Thm.forall_elim ct (Thm.forall_intr cv thm) + else thm + +fun discharge_definition ct thm = + let val (cv, cu) = Thm.dest_equals ct + in + Thm.implies_intr ct thm + |> instantiate_free (cv, cu) + |> (fn thm => Thm.implies_elim thm (Thm.reflexive cu)) + end + +fun if_conv c cv1 cv2 ct = (if c (Thm.term_of ct) then cv1 else cv2) ct +fun if_true_conv c cv = if_conv c cv Conv.all_conv + + +(* simplification of trivial let expressions (whose bound variables occur at + most once) *) + +local + fun count i (Bound j) = if j = i then 1 else 0 + | count i (t $ u) = count i t + count i u + | count i (Abs (_, _, t)) = count (i + 1) t + | count _ _ = 0 + + fun is_trivial_let (Const (@{const_name Let}, _) $ _ $ Abs (_, _, t)) = + (count 0 t <= 1) + | is_trivial_let _ = false + + fun let_conv _ = if_true_conv is_trivial_let (Conv.rewr_conv @{thm Let_def}) + + fun cond_let_conv ctxt = if_true_conv (Term.exists_subterm is_trivial_let) + (More_Conv.top_conv let_conv ctxt) +in +fun trivial_let ctxt = map (Conv.fconv_rule (cond_let_conv ctxt)) +end + + +(* rewriting of negative integer numerals into positive numerals *) + +local + fun neg_numeral @{term Int.Min} = true + | neg_numeral _ = false + fun is_number_sort thy T = Sign.of_sort thy (T, @{sort number_ring}) + fun is_neg_number ctxt (Const (@{const_name number_of}, T) $ t) = + Term.exists_subterm neg_numeral t andalso + is_number_sort (ProofContext.theory_of ctxt) (Term.body_type T) + | is_neg_number _ _ = false + fun has_neg_number ctxt = Term.exists_subterm (is_neg_number ctxt) + + val pos_numeral_ss = HOL_ss + addsimps [@{thm Int.number_of_minus}, @{thm Int.number_of_Min}] + addsimps [@{thm Int.numeral_1_eq_1}] + addsimps @{thms Int.pred_bin_simps} + addsimps @{thms Int.normalize_bin_simps} + addsimps @{lemma + "Int.Min = - Int.Bit1 Int.Pls" + "Int.Bit0 (- Int.Pls) = - Int.Pls" + "Int.Bit0 (- k) = - Int.Bit0 k" + "Int.Bit1 (- k) = - Int.Bit1 (Int.pred k)" + by simp_all (simp add: pred_def)} + + fun pos_conv ctxt = if_conv (is_neg_number ctxt) + (Simplifier.rewrite (Simplifier.context ctxt pos_numeral_ss)) + Conv.no_conv + + fun cond_pos_conv ctxt = if_true_conv (has_neg_number ctxt) + (More_Conv.top_sweep_conv pos_conv ctxt) +in +fun positive_numerals ctxt = map (Conv.fconv_rule (cond_pos_conv ctxt)) +end + + +(* embedding of standard natural number operations into integer operations *) + +local + val nat_embedding = @{lemma + "nat (int n) = n" + "i >= 0 --> int (nat i) = i" + "i < 0 --> int (nat i) = 0" + by simp_all} + + val nat_rewriting = @{lemma + "0 = nat 0" + "1 = nat 1" + "number_of i = nat (number_of i)" + "int (nat 0) = 0" + "int (nat 1) = 1" + "a < b = (int a < int b)" + "a <= b = (int a <= int b)" + "Suc a = nat (int a + 1)" + "a + b = nat (int a + int b)" + "a - b = nat (int a - int b)" + "a * b = nat (int a * int b)" + "a div b = nat (int a div int b)" + "a mod b = nat (int a mod int b)" + "int (nat (int a + int b)) = int a + int b" + "int (nat (int a * int b)) = int a * int b" + "int (nat (int a div int b)) = int a div int b" + "int (nat (int a mod int b)) = int a mod int b" + by (simp add: nat_mult_distrib nat_div_distrib nat_mod_distrib + int_mult[symmetric] zdiv_int[symmetric] zmod_int[symmetric])+} + + fun on_positive num f x = + (case try HOLogic.dest_number (Thm.term_of num) of + SOME (_, i) => if i >= 0 then SOME (f x) else NONE + | NONE => NONE) + + val cancel_int_nat_ss = HOL_ss + addsimps [@{thm Nat_Numeral.nat_number_of}] + addsimps [@{thm Nat_Numeral.int_nat_number_of}] + addsimps @{thms neg_simps} + + fun cancel_int_nat_simproc _ ss ct = + let + val num = Thm.dest_arg (Thm.dest_arg ct) + val goal = Thm.mk_binop @{cterm "op == :: int => _"} ct num + val simpset = Simplifier.inherit_context ss cancel_int_nat_ss + fun tac _ = Simplifier.simp_tac simpset 1 + in on_positive num (Goal.prove_internal [] goal) tac end + + val nat_ss = HOL_ss + addsimps nat_rewriting + addsimprocs [Simplifier.make_simproc { + name = "cancel_int_nat_num", lhss = [@{cpat "int (nat _)"}], + proc = cancel_int_nat_simproc, identifier = [] }] + + fun conv ctxt = Simplifier.rewrite (Simplifier.context ctxt nat_ss) + + val uses_nat_type = Term.exists_type (Term.exists_subtype (equal @{typ nat})) +in +fun nat_as_int ctxt thms = + let + fun norm thm uses_nat = + if not (uses_nat_type (Thm.prop_of thm)) then (thm, uses_nat) + else (Conv.fconv_rule (conv ctxt) thm, true) + val (thms', uses_nat) = fold_map norm thms false + in if uses_nat then nat_embedding @ thms' else thms' end +end + + +(* include additional rules *) + +val (unfold_defs, unfold_defs_setup) = + Attrib.config_bool "smt_unfold_defs" true + +local + val pair_rules = [@{thm fst_conv}, @{thm snd_conv}, @{thm pair_collapse}] + + val pair_type = (fn Type (@{type_name "*"}, _) => true | _ => false) + val exists_pair_type = Term.exists_type (Term.exists_subtype pair_type) + + val fun_upd_rules = [@{thm fun_upd_same}, @{thm fun_upd_apply}] + val is_fun_upd = (fn Const (@{const_name fun_upd}, _) => true | _ => false) + val exists_fun_upd = Term.exists_subterm is_fun_upd +in +fun add_pair_rules _ thms = + thms + |> exists (exists_pair_type o Thm.prop_of) thms ? append pair_rules + +fun add_fun_upd_rules _ thms = + thms + |> exists (exists_fun_upd o Thm.prop_of) thms ? append fun_upd_rules +end + + +local + fun mk_entry t thm = (Term.head_of t, (thm, thm RS @{thm eq_reflection})) + fun prepare_def thm = + (case HOLogic.dest_Trueprop (Thm.prop_of thm) of + Const (@{const_name "op ="}, _) $ t $ _ => mk_entry t thm + | t => raise TERM ("prepare_def", [t])) + + val defs = map prepare_def [ + @{thm abs_if[where 'a = int]}, @{thm abs_if[where 'a = real]}, + @{thm min_def[where 'a = int]}, @{thm min_def[where 'a = real]}, + @{thm max_def[where 'a = int]}, @{thm max_def[where 'a = real]}] + + fun add_sym t = if AList.defined (op =) defs t then insert (op =) t else I + fun add_syms thms = fold (Term.fold_aterms add_sym o Thm.prop_of) thms [] + + fun unfold_conv ctxt ct = + (case AList.lookup (op =) defs (Term.head_of (Thm.term_of ct)) of + SOME (_, eq) => Conv.rewr_conv eq + | NONE => Conv.all_conv) ct +in +fun add_abs_min_max_rules ctxt thms = + if Config.get ctxt unfold_defs + then map (Conv.fconv_rule (More_Conv.bottom_conv unfold_conv ctxt)) thms + else map fst (map_filter (AList.lookup (op =) defs) (add_syms thms)) @ thms +end + + +(* lift lambda terms into additional rules *) + +local + val meta_eq = @{cpat "op =="} + val meta_eqT = hd (Thm.dest_ctyp (Thm.ctyp_of_term meta_eq)) + fun inst_meta cT = Thm.instantiate_cterm ([(meta_eqT, cT)], []) meta_eq + fun mk_meta_eq ct cu = Thm.mk_binop (inst_meta (Thm.ctyp_of_term ct)) ct cu + + fun lambda_conv conv = + let + fun sub_conv cvs ctxt ct = + (case Thm.term_of ct of + Const (@{const_name All}, _) $ Abs _ => quant_conv cvs ctxt + | Const (@{const_name Ex}, _) $ Abs _ => quant_conv cvs ctxt + | Const _ $ Abs _ => Conv.arg_conv (at_lambda_conv cvs ctxt) + | Const (@{const_name Let}, _) $ _ $ Abs _ => Conv.combination_conv + (Conv.arg_conv (sub_conv cvs ctxt)) (abs_conv cvs ctxt) + | Abs _ => at_lambda_conv cvs ctxt + | _ $ _ => Conv.comb_conv (sub_conv cvs ctxt) + | _ => Conv.all_conv) ct + and abs_conv cvs = Conv.abs_conv (fn (cv, cx) => sub_conv (cv::cvs) cx) + and quant_conv cvs ctxt = Conv.arg_conv (abs_conv cvs ctxt) + and at_lambda_conv cvs ctxt = abs_conv cvs ctxt then_conv conv cvs ctxt + in sub_conv [] end + + fun used_vars cvs ct = + let + val lookup = AList.lookup (op aconv) (map (` Thm.term_of) cvs) + val add = (fn (SOME ct) => insert (op aconvc) ct | _ => I) + in Term.fold_aterms (add o lookup) (Thm.term_of ct) [] end + + val rev_int_fst_ord = rev_order o int_ord o pairself fst + fun ordered_values tab = + Termtab.fold (fn (_, x) => OrdList.insert rev_int_fst_ord x) tab [] + |> map snd +in +fun lift_lambdas ctxt thms = + let + val declare_frees = fold (Thm.fold_terms Term.declare_term_frees) + val names = ref (declare_frees thms (Name.make_context [])) + val fresh_name = change_result names o yield_singleton Name.variants + + val defs = ref (Termtab.empty : (int * thm) Termtab.table) + fun add_def t thm = change defs (Termtab.update (t, (serial (), thm))) + fun make_def cvs eq = Thm.symmetric (fold norm_meta_def cvs eq) + fun def_conv cvs ctxt ct = + let + val cvs' = used_vars cvs ct + val ct' = fold Thm.cabs cvs' ct + in + (case Termtab.lookup (!defs) (Thm.term_of ct') of + SOME (_, eq) => make_def cvs' eq + | NONE => + let + val {t, T, ...} = Thm.rep_cterm ct' + val eq = mk_meta_eq (cert ctxt (Free (fresh_name "", T))) ct' + val thm = Thm.assume eq + in (add_def t thm; make_def cvs' thm) end) + end + val thms' = map (Conv.fconv_rule (lambda_conv def_conv ctxt)) thms + val eqs = ordered_values (!defs) + in + (maps (#hyps o Thm.crep_thm) eqs, map (normalize_rule ctxt) eqs @ thms') + end +end + + +(* make application explicit for functions with varying number of arguments *) + +local + val const = prefix "c" and free = prefix "f" + fun min i (e as (_, j)) = if i <> j then (true, Int.min (i, j)) else e + fun add t i = Symtab.map_default (t, (false, i)) (min i) + fun traverse t = + (case Term.strip_comb t of + (Const (n, _), ts) => add (const n) (length ts) #> fold traverse ts + | (Free (n, _), ts) => add (free n) (length ts) #> fold traverse ts + | (Abs (_, _, u), ts) => fold traverse (u :: ts) + | (_, ts) => fold traverse ts) + val prune = (fn (n, (true, i)) => Symtab.update (n, i) | _ => I) + fun prune_tab tab = Symtab.fold prune tab Symtab.empty + + fun binop_conv cv1 cv2 = Conv.combination_conv (Conv.arg_conv cv1) cv2 + fun nary_conv conv1 conv2 ct = + (Conv.combination_conv (nary_conv conv1 conv2) conv2 else_conv conv1) ct + fun abs_conv conv tb = Conv.abs_conv (fn (cv, cx) => + let val n = fst (Term.dest_Free (Thm.term_of cv)) + in conv (Symtab.update (free n, 0) tb) cx end) + val apply_rule = @{lemma "f x == apply f x" by (simp add: apply_def)} +in +fun explicit_application ctxt thms = + let + fun sub_conv tb ctxt ct = + (case Term.strip_comb (Thm.term_of ct) of + (Const (n, _), ts) => app_conv tb (const n) (length ts) ctxt + | (Free (n, _), ts) => app_conv tb (free n) (length ts) ctxt + | (Abs _, ts) => nary_conv (abs_conv sub_conv tb ctxt) (sub_conv tb ctxt) + | (_, ts) => nary_conv Conv.all_conv (sub_conv tb ctxt)) ct + and app_conv tb n i ctxt = + (case Symtab.lookup tb n of + NONE => nary_conv Conv.all_conv (sub_conv tb ctxt) + | SOME j => apply_conv tb ctxt (i - j)) + and apply_conv tb ctxt i ct = ( + if i = 0 then nary_conv Conv.all_conv (sub_conv tb ctxt) + else + Conv.rewr_conv apply_rule then_conv + binop_conv (apply_conv tb ctxt (i-1)) (sub_conv tb ctxt)) ct + + val tab = prune_tab (fold (traverse o Thm.prop_of) thms Symtab.empty) + in map (Conv.fconv_rule (sub_conv tab ctxt)) thms end +end + + +(* combined normalization *) + +datatype config = + RewriteTrivialLets | + RewriteNegativeNumerals | + RewriteNaturalNumbers | + AddPairRules | + AddFunUpdRules | + AddAbsMinMaxRules + +fun normalize config ctxt thms = + let fun if_enabled c f = member (op =) config c ? f ctxt + in + thms + |> if_enabled RewriteTrivialLets trivial_let + |> if_enabled RewriteNegativeNumerals positive_numerals + |> if_enabled RewriteNaturalNumbers nat_as_int + |> if_enabled AddPairRules add_pair_rules + |> if_enabled AddFunUpdRules add_fun_upd_rules + |> if_enabled AddAbsMinMaxRules add_abs_min_max_rules + |> map (normalize_rule ctxt) + |> lift_lambdas ctxt + |> apsnd (explicit_application ctxt) + end + +val setup = unfold_defs_setup + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/Tools/smt_solver.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/smt_solver.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,246 @@ +(* Title: HOL/SMT/Tools/smt_solver.ML + Author: Sascha Boehme, TU Muenchen + +SMT solvers registry and SMT tactic. +*) + +signature SMT_SOLVER = +sig + exception SMT_COUNTEREXAMPLE of bool * term list + + datatype interface = Interface of { + normalize: SMT_Normalize.config list, + translate: SMT_Translate.config } + + datatype proof_data = ProofData of { + context: Proof.context, + output: string list, + recon: SMT_Translate.recon, + assms: thm list option } + + datatype solver_config = SolverConfig of { + name: {env_var: string, remote_name: string}, + interface: interface, + arguments: string list, + reconstruct: proof_data -> thm } + + (*options*) + val timeout: int Config.T + val with_timeout: Proof.context -> ('a -> 'b) -> 'a -> 'b + val trace: bool Config.T + val trace_msg: Proof.context -> ('a -> string) -> 'a -> unit + + (*solvers*) + type solver = Proof.context -> thm list -> thm + type solver_info = Context.generic -> Pretty.T list + val add_solver: string * (Proof.context -> solver_config) -> theory -> + theory + val all_solver_names_of: theory -> string list + val add_solver_info: string * solver_info -> theory -> theory + val solver_name_of: Context.generic -> string + val select_solver: string -> Context.generic -> Context.generic + val solver_of: Context.generic -> solver + + (*tactic*) + val smt_tac': bool -> Proof.context -> thm list -> int -> Tactical.tactic + val smt_tac: Proof.context -> thm list -> int -> Tactical.tactic + + (*setup*) + val setup: theory -> theory + val print_setup: Context.generic -> unit +end + +structure SMT_Solver: SMT_SOLVER = +struct + +exception SMT_COUNTEREXAMPLE of bool * term list + + +datatype interface = Interface of { + normalize: SMT_Normalize.config list, + translate: SMT_Translate.config } + +datatype proof_data = ProofData of { + context: Proof.context, + output: string list, + recon: SMT_Translate.recon, + assms: thm list option } + +datatype solver_config = SolverConfig of { + name: {env_var: string, remote_name: string}, + interface: interface, + arguments: string list, + reconstruct: proof_data -> thm } + + +(* SMT options *) + +val (timeout, setup_timeout) = Attrib.config_int "smt_timeout" 30 + +fun with_timeout ctxt f x = + TimeLimit.timeLimit (Time.fromSeconds (Config.get ctxt timeout)) f x + handle TimeLimit.TimeOut => error ("SMT: timeout") + +val (trace, setup_trace) = Attrib.config_bool "smt_trace" false + +fun trace_msg ctxt f x = + if Config.get ctxt trace then Output.tracing (f x) else () + + +(* interface to external solvers *) + +local + +fun with_tmp_files f x = + let + fun tmp_path () = File.tmp_path (Path.explode ("smt-" ^ serial_string ())) + val in_path = tmp_path () and out_path = tmp_path () + val y = Exn.capture (f in_path out_path) x + val _ = try File.rm in_path and _ = try File.rm out_path + in Exn.release y end + +fun run in_path out_path (ctxt, cmd, output) = + let + val x = File.open_output output in_path + val _ = trace_msg ctxt File.read in_path + + val _ = with_timeout ctxt system_out (cmd in_path out_path) + fun lines_of path = the_default [] (try (File.fold_lines cons out_path) []) + val ls = rev (dropwhile (equal "") (lines_of out_path)) + val _ = trace_msg ctxt cat_lines ls + in (x, ls) end + +in + +fun run_solver ctxt {env_var, remote_name} args output = + let + val qf = File.shell_path and qq = enclose "'" "'" + val path = getenv env_var and remote = getenv "REMOTE_SMT_SOLVER" + fun cmd f1 f2 = + if path <> "" + then map qq (path :: args) @ [qf f1, ">", qf f2] + else "perl -w" :: map qq (remote :: remote_name :: args) @ [qf f1, qf f2] + in with_tmp_files run (ctxt, space_implode " " oo cmd, output) end + +end + +fun make_proof_data ctxt ((recon, thms), ls) = + ProofData {context=ctxt, output=ls, recon=recon, assms=SOME thms} + +fun gen_solver solver ctxt prems = + let + val SolverConfig {name, interface, arguments, reconstruct} = solver ctxt + val Interface {normalize=nc, translate=tc} = interface + val thy = ProofContext.theory_of ctxt + in + SMT_Normalize.normalize nc ctxt prems + ||> run_solver ctxt name arguments o SMT_Translate.translate tc thy + ||> reconstruct o make_proof_data ctxt + |-> fold SMT_Normalize.discharge_definition + end + + +(* solver store *) + +type solver = Proof.context -> thm list -> thm +type solver_info = Context.generic -> Pretty.T list + +structure Solvers = TheoryDataFun +( + type T = ((Proof.context -> solver_config) * solver_info) Symtab.table + val empty = Symtab.empty + val copy = I + val extend = I + fun merge _ = Symtab.merge (K true) + handle Symtab.DUP name => error ("Duplicate SMT solver: " ^ quote name) +) + +val no_solver = "(none)" +val add_solver = Solvers.map o Symtab.update_new o apsnd (rpair (K [])) +val all_solver_names_of = Symtab.keys o Solvers.get +val lookup_solver = Symtab.lookup o Solvers.get +fun add_solver_info (n, i) = Solvers.map (Symtab.map_entry n (apsnd (K i))) + + +(* selected solver *) + +structure SelectedSolver = GenericDataFun +( + type T = serial * string + val empty = (serial (), no_solver) + val extend = I + fun merge _ (sl1 as (s1, _), sl2 as (s2, _)) = if s1 > s2 then sl1 else sl2 +) + +val solver_name_of = snd o SelectedSolver.get + +fun select_solver name gen = + if is_none (lookup_solver (Context.theory_of gen) name) + then error ("SMT solver not registered: " ^ quote name) + else SelectedSolver.map (K (serial (), name)) gen + +fun raw_solver_of gen = + (case lookup_solver (Context.theory_of gen) (solver_name_of gen) of + NONE => error "No SMT solver selected" + | SOME (s, _) => s) + +val solver_of = gen_solver o raw_solver_of + + +(* SMT tactic *) + +fun smt_unsat_tac solver ctxt rules = + Tactic.rtac @{thm ccontr} THEN' + SUBPROOF (fn {context, prems, ...} => + Tactic.rtac (solver context (rules @ prems)) 1) ctxt + +fun pretty_counterex ctxt (real, ex) = + let + val msg = if real then "Counterexample found:" + else "Potential counterexample found:" + val cex = if null ex then [Pretty.str "(no assignments)"] + else map (Syntax.pretty_term ctxt) ex + in Pretty.string_of (Pretty.big_list msg cex) end + +fun smt_tac' pass_smt_exns ctxt = + let + val solver = solver_of (Context.Proof ctxt) + fun safe_solver ctxt thms = solver ctxt thms + handle SMT_COUNTEREXAMPLE cex => error (pretty_counterex ctxt cex) + val solver' = if pass_smt_exns then solver else safe_solver + in smt_unsat_tac solver' ctxt end + +val smt_tac = smt_tac' false + + +(* setup *) + +val setup = + Attrib.setup (Binding.name "smt_solver") + (Scan.lift (OuterParse.$$$ "=" |-- Args.name) >> + (Thm.declaration_attribute o K o select_solver)) + "SMT solver configuration" #> + setup_timeout #> + setup_trace + +fun print_setup gen = + let + val t = string_of_int (Config.get_generic gen timeout) + val names = sort string_ord (all_solver_names_of (Context.theory_of gen)) + val ns = if null names then [no_solver] else names + val take_info = (fn (_, []) => NONE | info => SOME info) + val infos = + Context.theory_of gen + |> Symtab.dest o Solvers.get + |> map_filter (fn (n, (_, info)) => take_info (n, info gen)) + |> sort (prod_ord string_ord (K EQUAL)) + |> map (fn (n, ps) => Pretty.big_list (n ^ ":") ps) + in + Pretty.writeln (Pretty.big_list "SMT setup:" [ + Pretty.str ("Current SMT solver: " ^ solver_name_of gen), + Pretty.str_list "Available SMT solvers: " "" ns, + Pretty.str ("Current timeout: " ^ t ^ " seconds"), + Pretty.big_list "Solver-specific settings:" infos]) + end + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/Tools/smt_translate.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/smt_translate.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,507 @@ +(* Title: HOL/SMT/Tools/smt_translate.ML + Author: Sascha Boehme, TU Muenchen + +Translate theorems into an SMT intermediate format and serialize them, +depending on an SMT interface. +*) + +signature SMT_TRANSLATE = +sig + (* intermediate term structure *) + datatype sym = + SConst of string * typ | + SFree of string * typ | + SNum of int * typ + datatype squant = SForall | SExists + datatype 'a spattern = SPat of 'a list | SNoPat of 'a list + datatype ('a, 'b) sterm = + SVar of int | + SApp of 'a * ('a, 'b) sterm list | + SLet of (string * 'b) * ('a, 'b) sterm * ('a, 'b) sterm | + SQuant of squant * (string * 'b) list * ('a, 'b) sterm spattern list * + ('a, 'b) sterm + + (* table for built-in symbols *) + type builtin_fun = typ -> (sym, typ) sterm list -> + (string * (sym, typ) sterm list) option + type builtin_table = (typ * builtin_fun) list Symtab.table + val builtin_make: (term * string) list -> builtin_table + val builtin_add: term * builtin_fun -> builtin_table -> builtin_table + val builtin_lookup: builtin_table -> theory -> string * typ -> + (sym, typ) sterm list -> (string * (sym, typ) sterm list) option + val bv_rotate: (int -> string) -> builtin_fun + val bv_extend: (int -> string) -> builtin_fun + val bv_extract: (int -> int -> string) -> builtin_fun + + (* configuration options *) + datatype prefixes = Prefixes of { + var_prefix: string, + typ_prefix: string, + fun_prefix: string, + pred_prefix: string } + datatype markers = Markers of { + term_marker: string, + formula_marker: string } + datatype builtins = Builtins of { + builtin_typ: typ -> string option, + builtin_num: int * typ -> string option, + builtin_fun: bool -> builtin_table } + datatype sign = Sign of { + typs: string list, + funs: (string * (string list * string)) list, + preds: (string * string list) list } + datatype config = Config of { + strict: bool, + prefixes: prefixes, + markers: markers, + builtins: builtins, + serialize: sign -> (string, string) sterm list -> TextIO.outstream -> unit} + datatype recon = Recon of {typs: typ Symtab.table, terms: term Symtab.table} + + val translate: config -> theory -> thm list -> TextIO.outstream -> + recon * thm list + + val dest_binT: typ -> int + val dest_funT: int -> typ -> typ list * typ +end + +structure SMT_Translate: SMT_TRANSLATE = +struct + +(* Intermediate term structure *) + +datatype sym = + SConst of string * typ | + SFree of string * typ | + SNum of int * typ +datatype squant = SForall | SExists +datatype 'a spattern = SPat of 'a list | SNoPat of 'a list +datatype ('a, 'b) sterm = + SVar of int | + SApp of 'a * ('a, 'b) sterm list | + SLet of (string * 'b) * ('a, 'b) sterm * ('a, 'b) sterm | + SQuant of squant * (string * 'b) list * ('a, 'b) sterm spattern list * + ('a, 'b) sterm + +fun app c ts = SApp (c, ts) + +fun map_pat f (SPat ps) = SPat (map f ps) + | map_pat f (SNoPat ps) = SNoPat (map f ps) + +fun fold_map_pat f (SPat ps) = fold_map f ps #>> SPat + | fold_map_pat f (SNoPat ps) = fold_map f ps #>> SNoPat + +val make_sconst = SConst o Term.dest_Const + + +(* General type destructors. *) + +fun dest_binT T = + (case T of + Type (@{type_name "Numeral_Type.num0"}, _) => 0 + | Type (@{type_name "Numeral_Type.num1"}, _) => 1 + | Type (@{type_name "Numeral_Type.bit0"}, [T]) => 2 * dest_binT T + | Type (@{type_name "Numeral_Type.bit1"}, [T]) => 1 + 2 * dest_binT T + | _ => raise TYPE ("dest_binT", [T], [])) + +val dest_wordT = (fn + Type (@{type_name "word"}, [T]) => dest_binT T + | T => raise TYPE ("dest_wordT", [T], [])) + +val dest_funT = + let + fun dest Ts 0 T = (rev Ts, T) + | dest Ts i (Type ("fun", [T, U])) = dest (T::Ts) (i-1) U + | dest _ _ T = raise TYPE ("dest_funT", [T], []) + in dest [] end + + +(* Table for built-in symbols *) + +type builtin_fun = typ -> (sym, typ) sterm list -> + (string * (sym, typ) sterm list) option +type builtin_table = (typ * builtin_fun) list Symtab.table + +fun builtin_make entries = + let + fun dest (t, bn) = + let val (n, T) = Term.dest_Const t + in (n, (Logic.varifyT T, K (SOME o pair bn))) end + in Symtab.make (AList.group (op =) (map dest entries)) end + +fun builtin_add (t, f) tab = + let val (n, T) = apsnd Logic.varifyT (Term.dest_Const t) + in Symtab.map_default (n, []) (AList.update (op =) (T, f)) tab end + +fun builtin_lookup tab thy (n, T) ts = + AList.lookup (Sign.typ_instance thy) (Symtab.lookup_list tab n) T + |> (fn SOME f => f T ts | NONE => NONE) + +local + val dest_nat = (fn + SApp (SConst (@{const_name nat}, _), [SApp (SNum (i, _), _)]) => SOME i + | _ => NONE) +in +fun bv_rotate mk_name T ts = + dest_nat (hd ts) |> Option.map (fn i => (mk_name i, tl ts)) + +fun bv_extend mk_name T ts = + (case (try dest_wordT (domain_type T), try dest_wordT (range_type T)) of + (SOME i, SOME j) => if j - i >= 0 then SOME (mk_name (j - i), ts) else NONE + | _ => NONE) + +fun bv_extract mk_name T ts = + (case (try dest_wordT (body_type T), dest_nat (hd ts)) of + (SOME i, SOME lb) => SOME (mk_name (i + lb - 1) lb, tl ts) + | _ => NONE) +end + + +(* Configuration options *) + +datatype prefixes = Prefixes of { + var_prefix: string, + typ_prefix: string, + fun_prefix: string, + pred_prefix: string } +datatype markers = Markers of { + term_marker: string, + formula_marker: string } +datatype builtins = Builtins of { + builtin_typ: typ -> string option, + builtin_num: int * typ -> string option, + builtin_fun: bool -> builtin_table } +datatype sign = Sign of { + typs: string list, + funs: (string * (string list * string)) list, + preds: (string * string list) list } +datatype config = Config of { + strict: bool, + prefixes: prefixes, + markers: markers, + builtins: builtins, + serialize: sign -> (string, string) sterm list -> TextIO.outstream -> unit} +datatype recon = Recon of {typs: typ Symtab.table, terms: term Symtab.table} + + +(* Translate Isabelle/HOL terms into SMT intermediate terms. + We assume that lambda-lifting has been performed before, i.e., lambda + abstractions occur only at quantifiers and let expressions. +*) +local + val quantifier = (fn + @{const_name All} => SOME SForall + | @{const_name Ex} => SOME SExists + | _ => NONE) + + fun group_quant qname vs (t as Const (q, _) $ Abs (n, T, u)) = + if q = qname then group_quant qname ((n, T) :: vs) u else (vs, t) + | group_quant qname vs t = (vs, t) + + fun dest_trigger (@{term trigger} $ tl $ t) = (HOLogic.dest_list tl, t) + | dest_trigger t = ([], t) + + fun pat f ps (Const (@{const_name pat}, _) $ p) = SPat (rev (f p :: ps)) + | pat f ps (Const (@{const_name nopat}, _) $ p) = SNoPat (rev (f p :: ps)) + | pat f ps (Const (@{const_name andpat}, _) $ p $ t) = pat f (f p :: ps) t + | pat _ _ t = raise TERM ("pat", [t]) + + fun trans Ts t = + (case Term.strip_comb t of + (t1 as Const (qn, qT), [t2 as Abs (n, T, t3)]) => + (case quantifier qn of + SOME q => + let + val (vs, u) = group_quant qn [(n, T)] t3 + val Us = map snd vs @ Ts + val (ps, b) = dest_trigger u + in SQuant (q, rev vs, map (pat (trans Us) []) ps, trans Us b) end + | NONE => raise TERM ("intermediate", [t])) + | (Const (@{const_name Let}, _), [t1, Abs (n, T, t2)]) => + SLet ((n, T), trans Ts t1, trans (T :: Ts) t2) + | (Const (c as (@{const_name distinct}, _)), [t1]) => + (* this is not type-correct, but will be corrected at a later stage *) + SApp (SConst c, map (trans Ts) (HOLogic.dest_list t1)) + | (Const c, ts) => + (case try HOLogic.dest_number t of + SOME (T, i) => SApp (SNum (i, T), []) + | NONE => SApp (SConst c, map (trans Ts) ts)) + | (Free c, ts) => SApp (SFree c, map (trans Ts) ts) + | (Bound i, []) => SVar i + | _ => raise TERM ("intermediate", [t])) +in +fun intermediate ts = map (trans [] o HOLogic.dest_Trueprop) ts +end + + +(* Separate formulas from terms by adding special marker symbols ("term", + "formula"). Atoms "P" whose head symbol also occurs as function symbol are + rewritten to "term P = term True". Connectives and built-in predicates + occurring at term level are replaced by new constants, and theorems + specifying their meaning are added. +*) +local + (** Add the marker symbols "term" and "formulas" to separate formulas and + terms. **) + + val connectives = map make_sconst [@{term True}, @{term False}, + @{term Not}, @{term "op &"}, @{term "op |"}, @{term "op -->"}, + @{term "op = :: bool => _"}] + + fun note false c (ps, fs) = (insert (op =) c ps, fs) + | note true c (ps, fs) = (ps, insert (op =) c fs) + + val term_marker = SConst (@{const_name term}, Term.dummyT) + val formula_marker = SConst (@{const_name formula}, Term.dummyT) + fun mark f true t = f true t #>> app term_marker o single + | mark f false t = f false t #>> app formula_marker o single + fun mark' f false t = f true t #>> app term_marker o single + | mark' f true t = f true t + val mark_term = app term_marker o single + fun lift_term_marker c ts = + let val rem = (fn SApp (SConst (@{const_name term}, _), [t]) => t | t => t) + in mark_term (SApp (c, map rem ts)) end + fun is_term (SApp (SConst (@{const_name term}, _), _)) = true + | is_term _ = false + + fun either x = (fn y as SOME _ => y | _ => x) + fun get_loc loc i t = + (case t of + SVar j => if i = j then SOME loc else NONE + | SApp (SConst (@{const_name term}, _), us) => get_locs true i us + | SApp (SConst (@{const_name formula}, _), us) => get_locs false i us + | SApp (_, us) => get_locs loc i us + | SLet (_, u1, u2) => either (get_loc true i u1) (get_loc loc (i+1) u2) + | SQuant (_, vs, _, u) => get_loc loc (i + length vs) u) + and get_locs loc i ts = fold (either o get_loc loc i) ts NONE + + fun sep loc t = + (case t of + SVar i => pair t + | SApp (c as SConst (@{const_name If}, _), u :: us) => + mark sep false u ##>> fold_map (sep loc) us #>> app c o (op ::) + | SApp (c, us) => + if not loc andalso member (op =) connectives c + then fold_map (sep loc) us #>> app c + else note loc c #> fold_map (mark' sep loc) us #>> app c + | SLet (v, u1, u2) => + sep loc u2 #-> (fn u2' => + mark sep (the (get_loc loc 0 u2')) u1 #>> (fn u1' => + SLet (v, u1', u2'))) + | SQuant (q, vs, ps, u) => + fold_map (fold_map_pat (mark sep true)) ps ##>> + sep loc u #>> (fn (ps', u') => + SQuant (q, vs, ps', u'))) + + (** Rewrite atoms. **) + + val unterm_rule = @{lemma "term x == x" by (simp add: term_def)} + val unterm_conv = More_Conv.top_sweep_conv (K (Conv.rewr_conv unterm_rule)) + + val dest_word_type = (fn Type (@{type_name word}, [T]) => T | T => T) + fun instantiate [] _ = I + | instantiate (v :: _) T = + Term.subst_TVars [(v, dest_word_type (Term.domain_type T))] + + fun dest_alls (Const (@{const_name All}, _) $ Abs (_, _, t)) = dest_alls t + | dest_alls t = t + val dest_iff = (fn (Const (@{const_name iff}, _) $ t $ _ ) => t | t => t) + val dest_eq = (fn (Const (@{const_name "op ="}, _) $ t $ _ ) => t | t => t) + val dest_not = (fn (@{term Not} $ t) => t | t => t) + val head_of = HOLogic.dest_Trueprop #> dest_alls #> dest_iff #> dest_not #> + dest_eq #> Term.head_of + + fun prepare ctxt thm = + let + val rule = Conv.fconv_rule (unterm_conv ctxt) thm + val prop = Thm.prop_of thm + val inst = instantiate (Term.add_tvar_names prop []) + fun inst_for T = (singleton intermediate (inst T prop), rule) + in (make_sconst (head_of (Thm.prop_of rule)), inst_for) end + + val logicals = map (prepare @{context}) + @{lemma + "~ holds False" + "ALL p. holds (~ p) iff (~ holds p)" + "ALL p q. holds (p & q) iff (holds p & holds q)" + "ALL p q. holds (p | q) iff (holds p | holds q)" + "ALL p q. holds (p --> q) iff (holds p --> holds q)" + "ALL p q. holds (p iff q) iff (holds p iff holds q)" + "ALL p q. holds (p = q) iff (p = q)" + "ALL (a::int) b. holds (a < b) iff (a < b)" + "ALL (a::int) b. holds (a <= b) iff (a <= b)" + "ALL (a::real) b. holds (a < b) iff (a < b)" + "ALL (a::real) b. holds (a <= b) iff (a <= b)" + "ALL (a::'a::len0 word) b. holds (a < b) iff (a < b)" + "ALL (a::'a::len0 word) b. holds (a <= b) iff (a <= b)" + "ALL a b. holds (a Option.map (fn inst_for => inst_for T) + | lookup_logical _ _ = NONE + + val s_eq = make_sconst @{term "op = :: bool => _"} + val s_True = mark_term (SApp (make_sconst @{term True}, [])) + fun holds (SApp (c, ts)) = SApp (s_eq, [lift_term_marker c ts, s_True]) + | holds t = SApp (s_eq, [mark_term t, s_True]) + + val rewr_iff = (fn + SConst (@{const_name "op ="}, T as @{typ "bool => bool => bool"}) => + SConst (@{const_name iff}, T) + | c => c) + + fun rewrite ls = + let + fun rewr env loc t = + (case t of + SVar i => if not loc andalso nth env i then holds t else t + | SApp (c as SConst (@{const_name term}, _), [u]) => + SApp (c, [rewr env true u]) + | SApp (c as SConst (@{const_name formula}, _), [u]) => + SApp (c, [rewr env false u]) + | SApp (c, us) => + let val f = if not loc andalso member (op =) ls c then holds else I + in f (SApp (rewr_iff c, map (rewr env loc) us)) end + | SLet (v, u1, u2) => + SLet (v, rewr env loc u1, rewr (is_term u1 :: env) loc u2) + | SQuant (q, vs, ps, u) => + let val e = replicate (length vs) true @ env + in SQuant (q, vs, map (map_pat (rewr e loc)) ps, rewr e loc u) end) + in map (rewr [] false) end +in +fun separate thy ts = + let + val (ts', (ps, fs)) = fold_map (sep false) ts ([], []) + val eq_name = (fn + (SConst (n, _), SConst (m, _)) => n = m + | (SFree (n, _), SFree (m, _)) => n = m + | _ => false) + val ls = filter (member eq_name fs) ps + val (us, thms) = split_list (map_filter (lookup_logical thy) fs) + in (thms, us @ rewrite ls ts') end +end + + +(* Collect the signature of intermediate terms, identify built-in symbols, + rename uninterpreted symbols and types, make bound variables unique. + We require @{term distinct} to be a built-in constant of the SMT solver. +*) +local + fun empty_nctxt p = (p, 1) + fun make_nctxt (pT, pf, pp) = (empty_nctxt pT, empty_nctxt (pf, pp)) + fun fresh_name (p, i) = (p ^ string_of_int i, (p, i+1)) + fun fresh_typ (nT, nfp) = fresh_name nT ||> (fn nT' => (nT', nfp)) + fun fresh_fun loc (nT, ((pf, pp), i)) = + let val p = if loc then pf else pp + in fresh_name (p, i) ||> (fn (_, i') => (nT, ((pf, pp), i'))) end + + val empty_sign = (Typtab.empty, Termtab.empty, Termtab.empty) + fun lookup_typ (typs, _, _) = Typtab.lookup typs + fun lookup_fun true (_, funs, _) = Termtab.lookup funs + | lookup_fun false (_, _, preds) = Termtab.lookup preds + fun add_typ x (typs, funs, preds) = (Typtab.update x typs, funs, preds) + fun add_fun true x (typs, funs, preds) = (typs, Termtab.update x funs, preds) + | add_fun false x (typs, funs, preds) = (typs, funs, Termtab.update x preds) + fun make_sign (typs, funs, preds) = Sign { + typs = map snd (Typtab.dest typs), + funs = map snd (Termtab.dest funs), + preds = map (apsnd fst o snd) (Termtab.dest preds) } + fun make_rtab (typs, funs, preds) = + let + val rTs = Typtab.dest typs |> map swap |> Symtab.make + val rts = Termtab.dest funs @ Termtab.dest preds + |> map (apfst fst o swap) |> Symtab.make + in Recon {typs=rTs, terms=rts} end + + fun either f g x = (case f x of NONE => g x | y => y) + + fun rep_typ (Builtins {builtin_typ, ...}) T (st as (vars, ns, sgn)) = + (case either builtin_typ (lookup_typ sgn) T of + SOME n => (n, st) + | NONE => + let val (n, ns') = fresh_typ ns + in (n, (vars, ns', add_typ (T, n) sgn)) end) + + fun rep_var bs (n, T) (vars, ns, sgn) = + let val (n', vars') = fresh_name vars + in (vars', ns, sgn) |> rep_typ bs T |>> pair n' end + + fun rep_fun bs loc t T i (st as (_, _, sgn0)) = + (case lookup_fun loc sgn0 t of + SOME (n, _) => (n, st) + | NONE => + let + val (Us, U) = dest_funT i T + val (uns, (vars, ns, sgn)) = + st |> fold_map (rep_typ bs) Us ||>> rep_typ bs U + val (n, ns') = fresh_fun loc ns + in (n, (vars, ns', add_fun loc (t, (n, uns)) sgn)) end) + + fun rep_num (bs as Builtins {builtin_num, ...}) (i, T) st = + (case builtin_num (i, T) of + SOME n => (n, st) + | NONE => rep_fun bs true (HOLogic.mk_number T i) T 0 st) +in +fun signature_of prefixes markers builtins thy ts = + let + val Prefixes {var_prefix, typ_prefix, fun_prefix, pred_prefix} = prefixes + val Markers {formula_marker, term_marker} = markers + val Builtins {builtin_fun, ...} = builtins + + fun sign loc t = + (case t of + SVar i => pair (SVar i) + | SApp (c as SConst (@{const_name term}, _), [u]) => + sign true u #>> app term_marker o single + | SApp (c as SConst (@{const_name formula}, _), [u]) => + sign false u #>> app formula_marker o single + | SApp (SConst (c as (_, T)), ts) => + (case builtin_lookup (builtin_fun loc) thy c ts of + SOME (n, ts') => fold_map (sign loc) ts' #>> app n + | NONE => + rep_fun builtins loc (Const c) T (length ts) ##>> + fold_map (sign loc) ts #>> SApp) + | SApp (SFree (c as (_, T)), ts) => + rep_fun builtins loc (Free c) T (length ts) ##>> + fold_map (sign loc) ts #>> SApp + | SApp (SNum n, _) => rep_num builtins n #>> (fn n => SApp (n, [])) + | SLet (v, u1, u2) => + rep_var builtins v #-> (fn v' => + sign loc u1 ##>> sign loc u2 #>> (fn (u1', u2') => + SLet (v', u1', u2'))) + | SQuant (q, vs, ps, u) => + fold_map (rep_var builtins) vs ##>> + fold_map (fold_map_pat (sign loc)) ps ##>> + sign loc u #>> (fn ((vs', ps'), u') => + SQuant (q, vs', ps', u'))) + in + (empty_nctxt var_prefix, make_nctxt (typ_prefix, fun_prefix, pred_prefix), + empty_sign) + |> fold_map (sign false) ts + |> (fn (us, (_, _, sgn)) => (make_rtab sgn, (make_sign sgn, us))) + end +end + + +(* Combination of all translation functions and invocation of serialization. *) + +fun translate config thy thms stream = + let val Config {strict, prefixes, markers, builtins, serialize} = config + in + map Thm.prop_of thms + |> SMT_Monomorph.monomorph thy + |> intermediate + |> (if strict then separate thy else pair []) + ||>> signature_of prefixes markers builtins thy + ||> (fn (sgn, ts) => serialize sgn ts stream) + |> (fn ((thms', rtab), _) => (rtab, thms' @ thms)) + end + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/Tools/smtlib_interface.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/smtlib_interface.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,168 @@ +(* Title: HOL/SMT/Tools/smtlib_interface.ML + Author: Sascha Boehme, TU Muenchen + +Interface to SMT solvers based on the SMT-LIB format. +*) + +signature SMTLIB_INTERFACE = +sig + val basic_builtins: SMT_Translate.builtins + val default_attributes: string list + val gen_interface: SMT_Translate.builtins -> string list -> + SMT_Solver.interface + val interface: SMT_Solver.interface +end + +structure SMTLIB_Interface: SMTLIB_INTERFACE = +struct + +structure T = SMT_Translate + + +(* built-in types, functions and predicates *) + +val builtin_typ = (fn + @{typ int} => SOME "Int" + | @{typ real} => SOME "Real" + | _ => NONE) + +val builtin_num = (fn + (i, @{typ int}) => SOME (string_of_int i) + | (i, @{typ real}) => SOME (string_of_int i ^ ".0") + | _ => NONE) + +val builtin_funcs = T.builtin_make [ + (@{term If}, "ite"), + (@{term "uminus :: int => _"}, "~"), + (@{term "plus :: int => _"}, "+"), + (@{term "minus :: int => _"}, "-"), + (@{term "times :: int => _"}, "*"), + (@{term "uminus :: real => _"}, "~"), + (@{term "plus :: real => _"}, "+"), + (@{term "minus :: real => _"}, "-"), + (@{term "times :: real => _"}, "*") ] + +val builtin_preds = T.builtin_make [ + (@{term True}, "true"), + (@{term False}, "false"), + (@{term Not}, "not"), + (@{term "op &"}, "and"), + (@{term "op |"}, "or"), + (@{term "op -->"}, "implies"), + (@{term "op iff"}, "iff"), + (@{term If}, "if_then_else"), + (@{term distinct}, "distinct"), + (@{term "op ="}, "="), + (@{term "op < :: int => _"}, "<"), + (@{term "op <= :: int => _"}, "<="), + (@{term "op < :: real => _"}, "<"), + (@{term "op <= :: real => _"}, "<=") ] + + +(* serialization *) + +fun wr s stream = (TextIO.output (stream, s); stream) +fun wr_line f = f #> wr "\n" + +fun sep f = wr " " #> f +fun par f = sep (wr "(" #> f #> wr ")") + +fun wr1 s = sep (wr s) +fun wrn f n = (fn [] => wr1 n | ts => par (wr n #> fold f ts)) +fun ins s f = (fn [] => I | x::xs => f x #> fold (fn x => wr1 s #> f x) xs) + +val term_marker = "__term" +val formula_marker = "__form" +fun dest_marker (T.SApp ("__term", [t])) = SOME (true, t) + | dest_marker (T.SApp ("__form", [t])) = SOME (false, t) + | dest_marker _ = NONE + +val tvar = prefix "?" +val fvar = prefix "$" + +fun wr_expr loc env t = + (case t of + T.SVar i => wr1 (nth env i) + | T.SApp (n, ts) => + (case dest_marker t of + SOME (loc', t') => wr_expr loc' env t' + | NONE => wrn (wr_expr loc env) n ts) + | T.SLet ((v, _), t1, t2) => + if loc then raise TERM ("SMTLIB: let expression in term", []) + else + let + val (loc', t1') = the (dest_marker t1) + val (kind, v') = if loc' then ("let", tvar v) else ("flet", fvar v) + in + par (wr kind #> par (wr v' #> wr_expr loc' env t1') #> + wr_expr loc (v' :: env) t2) + end + | T.SQuant (q, vs, ps, b) => + let + val wr_quant = wr o (fn T.SForall => "forall" | T.SExists => "exists") + fun wr_var (n, s) = par (wr (tvar n) #> wr1 s) + + val wre = wr_expr loc (map (tvar o fst) (rev vs) @ env) + + fun wrp s ts = wr1 (":" ^ s ^ "{") #> ins "," wre ts #> wr1 "}" + fun wr_pat (T.SPat ts) = wrp "pat" ts + | wr_pat (T.SNoPat ts) = wrp "nopat" ts + in par (wr_quant q #> fold wr_var vs #> wre b #> fold wr_pat ps) end) + +fun serialize attributes (T.Sign {typs, funs, preds}) ts stream = + let + fun wr_decl (n, Ts) = wr_line (sep (par (wr n #> fold wr1 Ts))) + in + stream + |> wr_line (wr "(benchmark Isabelle") + |> fold (wr_line o wr) attributes + |> length typs > 0 ? + wr_line (wr ":extrasorts" #> par (fold wr1 typs)) + |> length funs > 0 ? ( + wr_line (wr ":extrafuns (") #> + fold (wr_decl o apsnd (fn (Ts, T) => Ts @ [T])) funs #> + wr_line (wr " )")) + |> length preds > 0 ? ( + wr_line (wr ":extrapreds (") #> + fold wr_decl preds #> + wr_line (wr " )")) + |> fold (fn t => wr ":assumption" #> wr_line (wr_expr false [] t)) ts + |> wr_line (wr ":formula true") + |> wr_line (wr ")") + |> K () + end + + +(* SMT solver interface using the SMT-LIB input format *) + +val basic_builtins = T.Builtins { + builtin_typ = builtin_typ, + builtin_num = builtin_num, + builtin_fun = (fn true => builtin_funcs | false => builtin_preds) } + +val default_attributes = [":logic AUFLIRA", ":status unknown"] + +fun gen_interface builtins attributes = SMT_Solver.Interface { + normalize = [ + SMT_Normalize.RewriteTrivialLets, + SMT_Normalize.RewriteNegativeNumerals, + SMT_Normalize.RewriteNaturalNumbers, + SMT_Normalize.AddAbsMinMaxRules, + SMT_Normalize.AddPairRules, + SMT_Normalize.AddFunUpdRules ], + translate = T.Config { + strict = true, + prefixes = T.Prefixes { + var_prefix = "x", + typ_prefix = "T", + fun_prefix = "uf_", + pred_prefix = "up_" }, + markers = T.Markers { + term_marker = term_marker, + formula_marker = formula_marker }, + builtins = builtins, + serialize = serialize attributes }} + +val interface = gen_interface basic_builtins default_attributes + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/Tools/yices_solver.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/yices_solver.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,52 @@ +(* Title: HOL/SMT/Tools/yices_solver.ML + Author: Sascha Boehme, TU Muenchen + +Interface of the SMT solver Yices. +*) + +signature YICES_SOLVER = +sig + val setup: theory -> theory +end + +structure Yices_Solver: YICES_SOLVER = +struct + +val solver_name = "yices" +val env_var = "YICES_SOLVER" + +val options = ["--evidence", "--smtlib"] + +fun cex_kind true = "Counterexample" + | cex_kind false = "Possible counterexample" + +fun raise_cex real ctxt rtab ls = + let val p = Pretty.big_list (cex_kind real ^ " found:") (map Pretty.str ls) + in error (Pretty.string_of p) end + +structure S = SMT_Solver + +fun core_oracle (SMT_Solver.ProofData {context, output, recon, ...}) = + let + val empty_line = (fn "" => true | _ => false) + val split_first = (fn [] => ("", []) | l :: ls => (l, ls)) + val (l, ls) = split_first (dropwhile empty_line output) + in + if String.isPrefix "unsat" l then @{cprop False} + else if String.isPrefix "sat" l then raise_cex true context recon ls + else if String.isPrefix "unknown" l then raise_cex false context recon ls + else error (solver_name ^ " failed") + end + +fun smtlib_solver oracle _ = + SMT_Solver.SolverConfig { + name = {env_var=env_var, remote_name=solver_name}, + interface = SMTLIB_Interface.interface, + arguments = options, + reconstruct = oracle } + +val setup = + Thm.add_oracle (Binding.name solver_name, core_oracle) #-> (fn (_, oracle) => + SMT_Solver.add_solver (solver_name, smtlib_solver oracle)) + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/Tools/z3_interface.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/z3_interface.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,98 @@ +(* Title: HOL/SMT/Tools/z3_interface.ML + Author: Sascha Boehme, TU Muenchen + +Interface to Z3 based on a relaxed version of SMT-LIB. +*) + +signature Z3_INTERFACE = +sig + val interface: SMT_Solver.interface +end + +structure Z3_Interface: Z3_INTERFACE = +struct + +structure T = SMT_Translate + +fun mk_name1 n i = n ^ "[" ^ string_of_int i ^ "]" +fun mk_name2 n i j = n ^ "[" ^ string_of_int i ^ ":" ^ string_of_int j ^ "]" + +val builtin_typ = (fn + @{typ int} => SOME "Int" + | @{typ real} => SOME "Real" + | Type (@{type_name word}, [T]) => + Option.map (mk_name1 "BitVec") (try T.dest_binT T) + | _ => NONE) + +val builtin_num = (fn + (i, @{typ int}) => SOME (string_of_int i) + | (i, @{typ real}) => SOME (string_of_int i ^ ".0") + | (i, Type (@{type_name word}, [T])) => + Option.map (mk_name1 ("bv" ^ string_of_int i)) (try T.dest_binT T) + | _ => NONE) + +val builtin_funcs = T.builtin_make [ + (@{term If}, "ite"), + (@{term "uminus :: int => _"}, "~"), + (@{term "plus :: int => _"}, "+"), + (@{term "minus :: int => _"}, "-"), + (@{term "times :: int => _"}, "*"), + (@{term "op div :: int => _"}, "div"), + (@{term "op mod :: int => _"}, "mod"), + (@{term "op rem"}, "rem"), + (@{term "uminus :: real => _"}, "~"), + (@{term "plus :: real => _"}, "+"), + (@{term "minus :: real => _"}, "-"), + (@{term "times :: real => _"}, "*"), + (@{term "op / :: real => _"}, "/"), + (@{term "bitNOT :: 'a::len0 word => _"}, "bvnot"), + (@{term "op AND :: 'a::len0 word => _"}, "bvand"), + (@{term "op OR :: 'a::len0 word => _"}, "bvor"), + (@{term "op XOR :: 'a::len0 word => _"}, "bvxor"), + (@{term "uminus :: 'a::len0 word => _"}, "bvneg"), + (@{term "op + :: 'a::len0 word => _"}, "bvadd"), + (@{term "op - :: 'a::len0 word => _"}, "bvsub"), + (@{term "op * :: 'a::len0 word => _"}, "bvmul"), + (@{term "op div :: 'a::len0 word => _"}, "bvudiv"), + (@{term "op mod :: 'a::len0 word => _"}, "bvurem"), + (@{term "op sdiv"}, "bvsdiv"), + (@{term "op smod"}, "bvsmod"), + (@{term "op srem"}, "bvsrem"), + (@{term word_cat}, "concat"), + (@{term bv_shl}, "bvshl"), + (@{term bv_lshr}, "bvlshr"), + (@{term bv_ashr}, "bvashr")] + |> T.builtin_add (@{term slice}, T.bv_extract (mk_name2 "extract")) + |> T.builtin_add (@{term ucast}, T.bv_extend (mk_name1 "zero_extend")) + |> T.builtin_add (@{term scast}, T.bv_extend (mk_name1 "sign_extend")) + |> T.builtin_add (@{term word_rotl}, T.bv_rotate (mk_name1 "rotate_left")) + |> T.builtin_add (@{term word_rotr}, T.bv_rotate (mk_name1 "rotate_right")) + +val builtin_preds = T.builtin_make [ + (@{term True}, "true"), + (@{term False}, "false"), + (@{term Not}, "not"), + (@{term "op &"}, "and"), + (@{term "op |"}, "or"), + (@{term "op -->"}, "implies"), + (@{term "op iff"}, "iff"), + (@{term If}, "if_then_else"), + (@{term distinct}, "distinct"), + (@{term "op ="}, "="), + (@{term "op < :: int => _"}, "<"), + (@{term "op <= :: int => _"}, "<="), + (@{term "op < :: real => _"}, "<"), + (@{term "op <= :: real => _"}, "<="), + (@{term "op < :: 'a::len0 word => _"}, "bvult"), + (@{term "op <= :: 'a::len0 word => _"}, "bvule"), + (@{term word_sless}, "bvslt"), + (@{term word_sle}, "bvsle")] + +val builtins = T.Builtins { + builtin_typ = builtin_typ, + builtin_num = builtin_num, + builtin_fun = (fn true => builtin_funcs | false => builtin_preds) } + +val interface = SMTLIB_Interface.gen_interface builtins [] + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/Tools/z3_model.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/z3_model.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,159 @@ +(* Title: HOL/SMT/Tools/z3_model.ML + Author: Sascha Boehme and Philipp Meyer, TU Muenchen + +Parser for counterexamples generated by Z3. +*) + +signature Z3_MODEL = +sig + val parse_counterex: SMT_Translate.recon -> string list -> term list +end + +structure Z3_Model: Z3_MODEL = +struct + +(* parsing primitives *) + +fun lift f (x, y) = apsnd (pair x) (f y) +fun lift' f v (x, y) = apsnd (rpair y) (f v x) + +fun $$ s = lift (Scan.$$ s) +fun this s = lift (Scan.this_string s) + +fun par scan = $$ "(" |-- scan --| $$ ")" + +val digit = (fn + "0" => SOME 0 | "1" => SOME 1 | "2" => SOME 2 | "3" => SOME 3 | + "4" => SOME 4 | "5" => SOME 5 | "6" => SOME 6 | "7" => SOME 7 | + "8" => SOME 8 | "9" => SOME 9 | _ => NONE) + +val nat_num = Scan.repeat1 (Scan.some digit) >> + (fn ds => fold (fn d => fn i => i * 10 + d) ds 0) +val int_num = Scan.optional (Scan.$$ "-" >> K (fn i => ~i)) I :|-- + (fn sign => nat_num >> sign) + +val is_char = Symbol.is_ascii_letter orf Symbol.is_ascii_digit orf + member (op =) (explode "_+*-/%~=<>$&|?!.@^#") +val name = Scan.many1 is_char >> implode + + +(* parsing counterexamples *) + +datatype context = Context of { + ttab: term Symtab.table, + nctxt: Name.context, + vtab: term Inttab.table } + +fun make_context (ttab, nctxt, vtab) = + Context {ttab=ttab, nctxt=nctxt, vtab=vtab} + +fun empty_context (SMT_Translate.Recon {terms, ...}) = + let + val ns = Symtab.fold (Term.add_free_names o snd) terms [] + val nctxt = Name.make_context ns + in make_context (terms, nctxt, Inttab.empty) end + +fun map_context f (Context {ttab, nctxt, vtab}) = + make_context (f (ttab, nctxt, vtab)) + +fun fresh_name (cx as Context {nctxt, ...}) = + let val (n, nctxt') = yield_singleton Name.variants "" nctxt + in (n, map_context (fn (ttab, _, vtab) => (ttab, nctxt', vtab)) cx) end + +fun ident name (cx as Context {ttab, ...}) = + (case Symtab.lookup ttab name of + SOME t => (t, cx) + | NONE => + let val (n, cx') = fresh_name cx + in (Free (n, Term.dummyT), cx) end) + +fun set_value t i = map_context (fn (ttab, nctxt, vtab) => + (ttab, nctxt, Inttab.update (i, t) vtab)) + +fun get_value T i (cx as Context {vtab, ...}) = + (case Inttab.lookup vtab i of + SOME t => (t, cx) + | _ => cx |> fresh_name |-> (fn n => + let val t = Free (n, T) + in set_value t i #> pair t end)) + + +fun space s = lift (Scan.many Symbol.is_ascii_blank) s +fun spaced p = p --| space + +val key = spaced (lift name) #-> lift' ident +val mapping = spaced (this "->") +fun in_braces p = spaced ($$ "{") |-- p --| spaced ($$ "}") + +val bool_expr = + this "true" >> K @{term True} || + this "false" >> K @{term False} + +fun number_expr T = + let + val num = lift int_num >> HOLogic.mk_number T + fun frac n d = Const (@{const_name divide}, T --> T --> T) $ n $ d + in num :|-- (fn n => Scan.optional ($$ "/" |-- num >> frac n) n) end + +val value = this "val!" |-- lift nat_num +fun value_expr T = value #-> lift' (get_value T) + +val domT = Term.domain_type +val ranT = Term.range_type +fun const_array T t = Abs ("x", T, t) +fun upd_array T ((a, t), u) = + Const (@{const_name fun_upd}, [T, domT T, ranT T] ---> T) $ a $ t $ u +fun array_expr T st = if not (can domT T) then Scan.fail st else st |> ( + par (spaced (this "const") |-- expr (ranT T)) >> const_array (domT T) || + par (spaced (this "store") |-- spaced (array_expr T) -- + expr (Term.domain_type T) -- expr (Term.range_type T)) >> upd_array T) + +and expr T st = + spaced (bool_expr || number_expr T || value_expr T || array_expr T) st + +fun const_val t = + let fun rep u = spaced value #-> apfst o set_value u #> pair [] + in + if can HOLogic.dest_number t then rep t + else if t = @{term TT} then rep @{term True} + else expr (Term.fastype_of t) >> (fn u => [HOLogic.mk_eq (t, u)]) + end + +fun func_value T = mapping |-- expr T + +fun first_pat T = + let + fun args T = if not (can domT T) then Scan.succeed [] else + expr (domT T) ::: args (ranT T) + fun value ts = func_value (snd (SMT_Translate.dest_funT (length ts) T)) + in args T :-- value end + +fun func_pat (Ts, T) = fold_map expr Ts -- func_value T +fun else_pat (Ts, T) = + let fun else_arg T cx = cx |> fresh_name |>> (fn n => Free (n, T)) + in + fold_map (lift' else_arg) Ts ##>> + spaced (this "else") |-- func_value T + end +fun next_pats T (fts as (ts, _)) = + let val Tps = SMT_Translate.dest_funT (length ts) T + in Scan.repeat (func_pat Tps) @@@ (else_pat Tps >> single) >> cons fts end + +fun mk_def' f (ts, t) = HOLogic.mk_eq (Term.list_comb (f, ts), t) +fun mk_def (Const (@{const_name apply}, _)) (u :: us, t) = mk_def' u (us, t) + | mk_def f (ts, t) = mk_def' f (ts, t) +fun func_pats t = + let val T = Term.fastype_of t + in first_pat T :|-- next_pats T >> map (mk_def t) end + +val assign = + key --| mapping :|-- (fn t => in_braces (func_pats t) || const_val t) + +val cex = space |-- Scan.repeat assign + +fun parse_counterex recon ls = + (empty_context recon, explode (cat_lines ls)) + |> Scan.catch (Scan.finite' Symbol.stopper (Scan.error cex)) + |> flat o fst + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/Tools/z3_solver.ML --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/Tools/z3_solver.ML Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,83 @@ +(* Title: HOL/SMT/Tools/z3_solver.ML + Author: Sascha Boehme, TU Muenchen + +Interface of the SMT solver Z3. +*) + +signature Z3_SOLVER = +sig + val proofs: bool Config.T + val options: string Config.T + + val setup: theory -> theory +end + +structure Z3_Solver: Z3_SOLVER = +struct + +val solver_name = "z3" +val env_var = "Z3_SOLVER" + +val (proofs, proofs_setup) = Attrib.config_bool "z3_proofs" false +val (options, options_setup) = Attrib.config_string "z3_options" "" + +fun add xs ys = ys @ xs + +fun get_options ctxt = + ["MODEL=true", "PRE_SIMPLIFY_EXPR=false", "CONTEXT_SIMPLIFIER=false"] + |> Config.get ctxt proofs ? add ["DISPLAY_PROOF=true", "PROOF_MODE=2"] + |> add (space_explode " " (Config.get ctxt options)) + +fun pretty_config context = [ + Pretty.str ("With proofs: " ^ + (if Config.get_generic context proofs then "true" else "false")), + Pretty.str ("Options: " ^ + space_implode " " (get_options (Context.proof_of context))) ] + +fun cmdline_options ctxt = + get_options ctxt + |> add ["-smt"] + +fun raise_cex real recon ls = + let val cex = Z3_Model.parse_counterex recon ls + in raise SMT_Solver.SMT_COUNTEREXAMPLE (real, cex) end + +fun check_unsat recon output = + let + val raw = not o String.isPrefix "WARNING" orf String.isPrefix "ERROR" + val (ls, l) = the_default ([], "") (try split_last (filter raw output)) + in + if String.isPrefix "unsat" l then ls + else if String.isPrefix "sat" l then raise_cex true recon ls + else if String.isPrefix "unknown" l then raise_cex false recon ls + else error (solver_name ^ " failed") + end + +fun core_oracle (SMT_Solver.ProofData {output, recon, ...}) = + check_unsat recon output + |> K @{cprop False} + +(* FIXME +fun prover (SMT_Solver.ProofData {context, output, recon, assms}) = + check_unsat recon output + |> Z3_Proof.reconstruct context assms recon +*) + +fun solver oracle ctxt = + let val with_proof = Config.get ctxt proofs + in + SMT_Solver.SolverConfig { + name = {env_var=env_var, remote_name=solver_name}, + interface = Z3_Interface.interface, + arguments = cmdline_options ctxt, + reconstruct = (*FIXME:if with_proof then prover else*) oracle } + end + +val setup = + proofs_setup #> + options_setup #> + Thm.add_oracle (Binding.name solver_name, core_oracle) #-> (fn (_, oracle) => + SMT_Solver.add_solver (solver_name, solver oracle)) #> + SMT_Solver.add_solver_info (solver_name, pretty_config) + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/etc/settings --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/etc/settings Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,13 @@ +ISABELLE_SMT="$COMPONENT" + +REMOTE_SMT_SOLVER="$ISABELLE_SMT/lib/scripts/remote_smt.pl" + +REMOTE_SMT_URL="http://www4.in.tum.de/smt/smt" + +# +# Paths to local SMT solvers: +# +# CVC_SOLVER=PATH +# YICES_SOLVER=PATH +# Z3_SOLVER=PATH + diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SMT/lib/scripts/remote_smt.pl --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/SMT/lib/scripts/remote_smt.pl Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,42 @@ +# +# Script to invoke remote SMT solvers. +# Author: Sascha Boehme, TU Muenchen +# + +use strict; +use LWP; + + +# environment + +my $remote_smt_url = $ENV{"REMOTE_SMT_URL"}; + + +# arguments + +my $solver = $ARGV[0]; +my @options = @ARGV[1 .. ($#ARGV - 2)]; +my $problem_file = $ARGV[-2]; +my $output_file = $ARGV[-1]; + + +# call solver + +my $agent = LWP::UserAgent->new; +$agent->agent("SMT-Request"); +$agent->timeout(180); +my $response = $agent->post($remote_smt_url, [ + "Solver" => $solver, + "Options" => join(" ", @options), + "Problem" => [$problem_file] ], + "Content_Type" => "form-data"); +if (not $response->is_success) { + print "HTTP-Error: " . $response->message . "\n"; + exit 1; +} +else { + open(FILE, ">$output_file"); + print FILE $response->content; + close(FILE); +} + diff -r c876bcb601fc -r 5b65449d7669 src/HOL/SetInterval.thy --- a/src/HOL/SetInterval.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/SetInterval.thy Tue Sep 22 13:52:19 2009 +1000 @@ -514,6 +514,30 @@ qed +subsubsection {* Proving Inclusions and Equalities between Unions *} + +lemma UN_UN_finite_eq: "(\n::nat. \i\{0..n. A n)" + by (auto simp add: atLeast0LessThan) + +lemma UN_finite_subset: "(!!n::nat. (\i\{0.. C) \ (\n. A n) \ C" + by (subst UN_UN_finite_eq [symmetric]) blast + +lemma UN_finite2_subset: + assumes sb: "!!n::nat. (\i\{0.. (\i\{0..n. A n) \ (\n. B n)" +proof (rule UN_finite_subset) + fix n + have "(\i\{0.. (\i\{0.. (\n::nat. \i\{0..n. B n)" by (simp add: UN_UN_finite_eq) + finally show "(\i\{0.. (\n. B n)" . +qed + +lemma UN_finite2_eq: + "(!!n::nat. (\i\{0..i\{0.. (\n. A n) = (\n. B n)" + by (iprover intro: subset_antisym UN_finite2_subset elim: equalityE) + + subsubsection {* Cardinality *} lemma card_lessThan [simp]: "card {..&1" fun split_time s = let val split = String.tokens (fn c => str c = "\n") val (proof, t) = s |> split |> split_last |> apfst cat_lines - val time = Scan.many1 Symbol.is_ascii_digit >> (fst o read_int) + fun as_num f = f >> (fst o read_int) + val num = as_num (Scan.many1 Symbol.is_ascii_digit) + val digit = Scan.one Symbol.is_ascii_digit + val num3 = as_num (digit ::: digit ::: (digit >> single)) + val time = num --| Scan.$$ "." -- num3 >> (fn (a, b) => a * 1000 + b) val as_time = the_default 0 o Scan.read Symbol.stopper time o explode in (proof, as_time t) end fun run_on probfile = @@ -97,7 +100,7 @@ else error ("Bad executable: " ^ Path.implode cmd) (* if problemfile has not been exported, delete problemfile; otherwise export proof, too *) - fun cleanup probfile = if destdir' = "" then File.rm probfile else () + fun cleanup probfile = if destdir' = "" then try File.rm probfile else NONE fun export probfile (((proof, _), _), _) = if destdir' = "" then () else File.write (Path.explode (Path.implode probfile ^ "_proof")) proof diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/ATP_Manager/lib/scripts/local_atp.pl --- a/src/HOL/Tools/ATP_Manager/lib/scripts/local_atp.pl Thu Sep 17 14:17:37 2009 +1000 +++ /dev/null Thu Jan 01 00:00:00 1970 +0000 @@ -1,15 +0,0 @@ -use POSIX qw(setsid); - -$SIG{'INT'} = "DEFAULT"; - -defined (my $pid = fork) or die "$!"; -if (not $pid) { - POSIX::setsid or die $!; - exec @ARGV; -} -else { - wait; - my ($user, $system, $cuser, $csystem) = times; - my $time = ($user + $cuser) * 1000; - print "$time\n"; -} diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/Function/decompose.ML --- a/src/HOL/Tools/Function/decompose.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/Function/decompose.ML Tue Sep 22 13:52:19 2009 +1000 @@ -29,7 +29,7 @@ fun prove_chain c1 c2 D = if is_some (Termination.get_chain D c1 c2) then D else let - val goal = HOLogic.mk_eq (HOLogic.mk_binop @{const_name "Relation.rel_comp"} (c1, c2), + val goal = HOLogic.mk_eq (HOLogic.mk_binop @{const_name Relation.rel_comp} (c1, c2), Const (@{const_name Set.empty}, fastype_of c1)) |> HOLogic.mk_Trueprop (* "C1 O C2 = {}" *) diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/Function/fundef_common.ML --- a/src/HOL/Tools/Function/fundef_common.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/Function/fundef_common.ML Tue Sep 22 13:52:19 2009 +1000 @@ -16,7 +16,7 @@ fun PROFILE msg = if !profile then timeap_msg msg else I -val acc_const_name = @{const_name "accp"} +val acc_const_name = @{const_name accp} fun mk_acc domT R = Const (acc_const_name, (domT --> domT --> HOLogic.boolT) --> domT --> HOLogic.boolT) $ R diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/Function/fundef_core.ML --- a/src/HOL/Tools/Function/fundef_core.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/Function/fundef_core.ML Tue Sep 22 13:52:19 2009 +1000 @@ -769,7 +769,7 @@ val Rrel = Free ("R", HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT))) val inrel_R = Const (@{const_name FunDef.in_rel}, HOLogic.mk_setT (HOLogic.mk_prodT (domT, domT)) --> fastype_of R) $ Rrel - val wfR' = cterm_of thy (HOLogic.mk_Trueprop (Const (@{const_name "Wellfounded.wfP"}, (domT --> domT --> boolT) --> boolT) $ R')) (* "wf R'" *) + val wfR' = cterm_of thy (HOLogic.mk_Trueprop (Const (@{const_name Wellfounded.wfP}, (domT --> domT --> boolT) --> boolT) $ R')) (* "wf R'" *) (* Inductive Hypothesis: !!z. (z,x):R' ==> z : acc R *) val ihyp = Term.all domT $ Abs ("z", domT, diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/Function/induction_scheme.ML --- a/src/HOL/Tools/Function/induction_scheme.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/Function/induction_scheme.ML Tue Sep 22 13:52:19 2009 +1000 @@ -152,7 +152,7 @@ end fun mk_wf ctxt R (IndScheme {T, ...}) = - HOLogic.Trueprop $ (Const (@{const_name "wf"}, mk_relT T --> HOLogic.boolT) $ R) + HOLogic.Trueprop $ (Const (@{const_name wf}, mk_relT T --> HOLogic.boolT) $ R) fun mk_ineqs R (IndScheme {T, cases, branches}) = let diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/Function/lexicographic_order.ML --- a/src/HOL/Tools/Function/lexicographic_order.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/Function/lexicographic_order.ML Tue Sep 22 13:52:19 2009 +1000 @@ -26,7 +26,7 @@ val mlexT = (domT --> HOLogic.natT) --> relT --> relT fun mk_ms [] = Const (@{const_name Set.empty}, relT) | mk_ms (f::fs) = - Const (@{const_name "mlex_prod"}, mlexT) $ f $ mk_ms fs + Const (@{const_name mlex_prod}, mlexT) $ f $ mk_ms fs in mk_ms mfuns end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/Function/measure_functions.ML --- a/src/HOL/Tools/Function/measure_functions.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/Function/measure_functions.ML Tue Sep 22 13:52:19 2009 +1000 @@ -22,7 +22,7 @@ val description = "Rules that guide the heuristic generation of measure functions" ); -fun mk_is_measures t = Const (@{const_name "is_measure"}, fastype_of t --> HOLogic.boolT) $ t +fun mk_is_measures t = Const (@{const_name is_measure}, fastype_of t --> HOLogic.boolT) $ t fun find_measures ctxt T = DEPTH_SOLVE (resolve_tac (Measure_Heuristic_Rules.get ctxt) 1) diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/Function/sum_tree.ML --- a/src/HOL/Tools/Function/sum_tree.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/Function/sum_tree.ML Tue Sep 22 13:52:19 2009 +1000 @@ -17,22 +17,22 @@ (* Sum types *) fun mk_sumT LT RT = Type ("+", [LT, RT]) -fun mk_sumcase TL TR T l r = Const (@{const_name "sum.sum_case"}, (TL --> T) --> (TR --> T) --> mk_sumT TL TR --> T) $ l $ r +fun mk_sumcase TL TR T l r = Const (@{const_name sum.sum_case}, (TL --> T) --> (TR --> T) --> mk_sumT TL TR --> T) $ l $ r val App = curry op $ fun mk_inj ST n i = access_top_down { init = (ST, I : term -> term), - left = (fn (T as Type ("+", [LT, RT]), inj) => (LT, inj o App (Const (@{const_name "Inl"}, LT --> T)))), - right =(fn (T as Type ("+", [LT, RT]), inj) => (RT, inj o App (Const (@{const_name "Inr"}, RT --> T))))} n i + left = (fn (T as Type ("+", [LT, RT]), inj) => (LT, inj o App (Const (@{const_name Inl}, LT --> T)))), + right =(fn (T as Type ("+", [LT, RT]), inj) => (RT, inj o App (Const (@{const_name Inr}, RT --> T))))} n i |> snd fun mk_proj ST n i = access_top_down { init = (ST, I : term -> term), - left = (fn (T as Type ("+", [LT, RT]), proj) => (LT, App (Const (@{const_name "Datatype.Projl"}, T --> LT)) o proj)), - right =(fn (T as Type ("+", [LT, RT]), proj) => (RT, App (Const (@{const_name "Datatype.Projr"}, T --> RT)) o proj))} n i + left = (fn (T as Type ("+", [LT, RT]), proj) => (LT, App (Const (@{const_name Datatype.Projl}, T --> LT)) o proj)), + right =(fn (T as Type ("+", [LT, RT]), proj) => (RT, App (Const (@{const_name Datatype.Projr}, T --> RT)) o proj))} n i |> snd fun mk_sumcases T fs = diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/Function/termination.ML --- a/src/HOL/Tools/Function/termination.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/Function/termination.ML Tue Sep 22 13:52:19 2009 +1000 @@ -79,14 +79,14 @@ (* concrete versions for sum types *) -fun is_inj (Const ("Sum_Type.Inl", _) $ _) = true - | is_inj (Const ("Sum_Type.Inr", _) $ _) = true +fun is_inj (Const (@{const_name Sum_Type.Inl}, _) $ _) = true + | is_inj (Const (@{const_name Sum_Type.Inr}, _) $ _) = true | is_inj _ = false -fun dest_inl (Const ("Sum_Type.Inl", _) $ t) = SOME t +fun dest_inl (Const (@{const_name Sum_Type.Inl}, _) $ t) = SOME t | dest_inl _ = NONE -fun dest_inr (Const ("Sum_Type.Inr", _) $ t) = SOME t +fun dest_inr (Const (@{const_name Sum_Type.Inr}, _) $ t) = SOME t | dest_inr _ = NONE @@ -145,8 +145,8 @@ fun mk_sum_skel rel = let - val cs = FundefLib.dest_binop_list @{const_name union} rel - fun collect_pats (Const ("Collect", _) $ Abs (_, _, c)) = + val cs = FundefLib.dest_binop_list @{const_name Set.union} rel + fun collect_pats (Const (@{const_name Collect}, _) $ Abs (_, _, c)) = let val (Const ("op &", _) $ (Const ("op =", _) $ _ $ (Const ("Pair", _) $ r $ l)) $ Gam) = Term.strip_qnt_body "Ex" c @@ -179,7 +179,7 @@ fun get_descent (_, _, _, _, D) c m1 m2 = Term3tab.lookup D (c, (m1, m2)) -fun dest_call D (Const ("Collect", _) $ Abs (_, _, c)) = +fun dest_call D (Const (@{const_name Collect}, _) $ Abs (_, _, c)) = let val n = get_num_points D val (sk, _, _, _, _) = D @@ -233,13 +233,13 @@ fun CALLS tac i st = if Thm.no_prems st then all_tac st else case Thm.term_of (Thm.cprem_of st i) of - (_ $ (_ $ rel)) => tac (FundefLib.dest_binop_list @{const_name union} rel, i) st + (_ $ (_ $ rel)) => tac (FundefLib.dest_binop_list @{const_name Set.union} rel, i) st |_ => no_tac st type ttac = (data -> int -> tactic) -> (data -> int -> tactic) -> data -> int -> tactic fun TERMINATION ctxt tac = - SUBGOAL (fn (_ $ (Const (@{const_name "wf"}, wfT) $ rel), i) => + SUBGOAL (fn (_ $ (Const (@{const_name wf}, wfT) $ rel), i) => let val (T, _) = HOLogic.dest_prodT (HOLogic.dest_setT (domain_type wfT)) in @@ -293,7 +293,7 @@ if null ineqs then Const (@{const_name Set.empty}, fastype_of rel) else - foldr1 (HOLogic.mk_binop @{const_name union}) (map mk_compr ineqs) + foldr1 (HOLogic.mk_binop @{const_name Set.union}) (map mk_compr ineqs) fun solve_membership_tac i = (EVERY' (replicate (i - 2) (rtac @{thm UnI2})) (* pick the right component of the union *) diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/Qelim/cooper.ML --- a/src/HOL/Tools/Qelim/cooper.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/Qelim/cooper.ML Tue Sep 22 13:52:19 2009 +1000 @@ -81,7 +81,7 @@ else if term_of x aconv z then Ge (Thm.dest_arg1 ct) else Nox | Const (@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name HOL.plus},_)$y$_) => if term_of x aconv y then Dvd (Thm.dest_binop ct ||> Thm.dest_arg) else Nox -| Const (@{const_name Not},_) $ (Const (@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name "HOL.plus"},_)$y$_)) => +| Const (@{const_name Not},_) $ (Const (@{const_name Ring_and_Field.dvd},_)$_$(Const(@{const_name HOL.plus},_)$y$_)) => if term_of x aconv y then NDvd (Thm.dest_binop (Thm.dest_arg ct) ||> Thm.dest_arg) else Nox | _ => Nox) diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/Qelim/presburger.ML --- a/src/HOL/Tools/Qelim/presburger.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/Qelim/presburger.ML Tue Sep 22 13:52:19 2009 +1000 @@ -52,18 +52,18 @@ local fun isnum t = case t of - Const(@{const_name "HOL.zero"},_) => true - | Const(@{const_name "HOL.one"},_) => true + Const(@{const_name HOL.zero},_) => true + | Const(@{const_name HOL.one},_) => true | @{term "Suc"}$s => isnum s | @{term "nat"}$s => isnum s | @{term "int"}$s => isnum s - | Const(@{const_name "HOL.uminus"},_)$s => isnum s - | Const(@{const_name "HOL.plus"},_)$l$r => isnum l andalso isnum r - | Const(@{const_name "HOL.times"},_)$l$r => isnum l andalso isnum r - | Const(@{const_name "HOL.minus"},_)$l$r => isnum l andalso isnum r - | Const(@{const_name "Power.power"},_)$l$r => isnum l andalso isnum r - | Const(@{const_name "Divides.mod"},_)$l$r => isnum l andalso isnum r - | Const(@{const_name "Divides.div"},_)$l$r => isnum l andalso isnum r + | Const(@{const_name HOL.uminus},_)$s => isnum s + | Const(@{const_name HOL.plus},_)$l$r => isnum l andalso isnum r + | Const(@{const_name HOL.times},_)$l$r => isnum l andalso isnum r + | Const(@{const_name HOL.minus},_)$l$r => isnum l andalso isnum r + | Const(@{const_name Power.power},_)$l$r => isnum l andalso isnum r + | Const(@{const_name Divides.mod},_)$l$r => isnum l andalso isnum r + | Const(@{const_name Divides.div},_)$l$r => isnum l andalso isnum r | _ => can HOLogic.dest_number t orelse can HOLogic.dest_nat t fun ty cts t = diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/Qelim/qelim.ML --- a/src/HOL/Tools/Qelim/qelim.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/Qelim/qelim.ML Tue Sep 22 13:52:19 2009 +1000 @@ -29,8 +29,8 @@ @{const_name "op -->"}, @{const_name "op ="}] s then binop_conv (conv env) p else atcv env p - | Const(@{const_name "Not"},_)$_ => arg_conv (conv env) p - | Const(@{const_name "Ex"},_)$Abs(s,_,_) => + | Const(@{const_name Not},_)$_ => arg_conv (conv env) p + | Const(@{const_name Ex},_)$Abs(s,_,_) => let val (e,p0) = Thm.dest_comb p val (x,p') = Thm.dest_abs (SOME s) p0 @@ -41,8 +41,8 @@ val (l,r) = Thm.dest_equals (cprop_of th') in if Thm.is_reflexive th' then Thm.transitive th (qcv env (Thm.rhs_of th)) else Thm.transitive (Thm.transitive th th') (conv env r) end - | Const(@{const_name "Ex"},_)$ _ => (Thm.eta_long_conversion then_conv conv env) p - | Const(@{const_name "All"},_)$_ => + | Const(@{const_name Ex},_)$ _ => (Thm.eta_long_conversion then_conv conv env) p + | Const(@{const_name All},_)$_ => let val p = Thm.dest_arg p val ([(_,T)],[]) = Thm.match (@{cpat "All"}, Thm.dest_fun p) diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/TFL/rules.ML --- a/src/HOL/Tools/TFL/rules.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/TFL/rules.ML Tue Sep 22 13:52:19 2009 +1000 @@ -456,7 +456,7 @@ fun is_cong thm = case (Thm.prop_of thm) of (Const("==>",_)$(Const("Trueprop",_)$ _) $ - (Const("==",_) $ (Const (@{const_name "Recdef.cut"},_) $ f $ R $ a $ x) $ _)) => false + (Const("==",_) $ (Const (@{const_name Recdef.cut},_) $ f $ R $ a $ x) $ _)) => false | _ => true; @@ -659,7 +659,7 @@ end; fun restricted t = isSome (S.find_term - (fn (Const(@{const_name "Recdef.cut"},_)) =>true | _ => false) + (fn (Const(@{const_name Recdef.cut},_)) =>true | _ => false) t) fun CONTEXT_REWRITE_RULE (func, G, cut_lemma, congs) th = diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/TFL/usyntax.ML --- a/src/HOL/Tools/TFL/usyntax.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/TFL/usyntax.ML Tue Sep 22 13:52:19 2009 +1000 @@ -398,7 +398,7 @@ end handle Bind => raise USYN_ERR "dest_relation" "unexpected term structure" else raise USYN_ERR "dest_relation" "not a boolean term"; -fun is_WFR (Const(@{const_name "Wellfounded.wf"},_)$_) = true +fun is_WFR (Const(@{const_name Wellfounded.wf},_)$_) = true | is_WFR _ = false; fun ARB ty = mk_select{Bvar=Free("v",ty), diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/float_syntax.ML --- a/src/HOL/Tools/float_syntax.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/float_syntax.ML Tue Sep 22 13:52:19 2009 +1000 @@ -27,10 +27,10 @@ fun mk_frac str = let val {mant=i, exp = n} = Syntax.read_float str; - val exp = Syntax.const @{const_name "Power.power"}; + val exp = Syntax.const @{const_name Power.power}; val ten = mk_number 10; val exp10 = if n=1 then ten else exp $ ten $ (mk_number n); - in (Syntax.const @{const_name "divide"}) $ (mk_number i) $ exp10 end + in (Syntax.const @{const_name divide}) $ (mk_number i) $ exp10 end in fun float_tr (*"_Float"*) [t as Const (str, _)] = mk_frac str diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/inductive.ML --- a/src/HOL/Tools/inductive.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/inductive.ML Tue Sep 22 13:52:19 2009 +1000 @@ -86,13 +86,13 @@ (** theory context references **) val inductive_forall_name = "HOL.induct_forall"; -val inductive_forall_def = thm "induct_forall_def"; +val inductive_forall_def = @{thm induct_forall_def}; val inductive_conj_name = "HOL.induct_conj"; -val inductive_conj_def = thm "induct_conj_def"; -val inductive_conj = thms "induct_conj"; -val inductive_atomize = thms "induct_atomize"; -val inductive_rulify = thms "induct_rulify"; -val inductive_rulify_fallback = thms "induct_rulify_fallback"; +val inductive_conj_def = @{thm induct_conj_def}; +val inductive_conj = @{thms induct_conj}; +val inductive_atomize = @{thms induct_atomize}; +val inductive_rulify = @{thms induct_rulify}; +val inductive_rulify_fallback = @{thms induct_rulify_fallback}; val notTrueE = TrueI RSN (2, notE); val notFalseI = Seq.hd (atac 1 notI); @@ -103,6 +103,7 @@ "(P & True) = P" "(True & P) = P" by (fact simp_thms)+}; +val simp_thms'' = inf_fun_eq :: inf_bool_eq :: simp_thms'; (** context data **) @@ -176,7 +177,7 @@ case concl of Const ("==", _) $ _ $ _ => eq2mono (thm RS meta_eq_to_obj_eq) | _ $ (Const ("op =", _) $ _ $ _) => eq2mono thm - | _ $ (Const ("HOL.ord_class.less_eq", _) $ _ $ _) => + | _ $ (Const (@{const_name HOL.less_eq}, _) $ _ $ _) => [dest_less_concl (Seq.hd (REPEAT (FIRSTGOAL (resolve_tac [le_funI, le_boolI'])) thm))] | _ => [thm] @@ -559,7 +560,7 @@ [rewrite_goals_tac [inductive_conj_def], DETERM (rtac raw_fp_induct 1), REPEAT (resolve_tac [le_funI, le_boolI] 1), - rewrite_goals_tac (inf_fun_eq :: inf_bool_eq :: simp_thms'), + rewrite_goals_tac simp_thms'', (*This disjE separates out the introduction rules*) REPEAT (FIRSTGOAL (eresolve_tac [disjE, exE, FalseE])), (*Now break down the individual cases. No disjE here in case @@ -568,7 +569,7 @@ REPEAT (FIRSTGOAL (resolve_tac [conjI, impI] ORELSE' (etac notE THEN' atac))), EVERY (map (fn prem => DEPTH_SOLVE_1 (ares_tac [rewrite_rule - (inductive_conj_def :: rec_preds_defs @ simp_thms') prem, + (inductive_conj_def :: rec_preds_defs @ simp_thms'') prem, conjI, refl] 1)) prems)]); val lemma = SkipProof.prove ctxt'' [] [] diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/inductive_set.ML --- a/src/HOL/Tools/inductive_set.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/inductive_set.ML Tue Sep 22 13:52:19 2009 +1000 @@ -74,8 +74,8 @@ in Drule.instantiate' [] (rev (map (SOME o cterm_of thy o Var) vs)) (p (fold (Logic.all o Var) vs t) f) end; - fun mkop "op &" T x = SOME (Const (@{const_name inter}, T --> T --> T), x) - | mkop "op |" T x = SOME (Const (@{const_name union}, T --> T --> T), x) + fun mkop "op &" T x = SOME (Const (@{const_name Set.inter}, T --> T --> T), x) + | mkop "op |" T x = SOME (Const (@{const_name Set.union}, T --> T --> T), x) | mkop _ _ _ = NONE; fun mk_collect p T t = let val U = HOLogic.dest_setT T diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/int_arith.ML --- a/src/HOL/Tools/int_arith.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/int_arith.ML Tue Sep 22 13:52:19 2009 +1000 @@ -49,13 +49,13 @@ make_simproc {lhss = lhss1, name = "one_to_of_int_one_simproc", proc = proc1, identifier = []}; -fun check (Const (@{const_name "HOL.one"}, @{typ int})) = false - | check (Const (@{const_name "HOL.one"}, _)) = true +fun check (Const (@{const_name HOL.one}, @{typ int})) = false + | check (Const (@{const_name HOL.one}, _)) = true | check (Const (s, _)) = member (op =) [@{const_name "op ="}, - @{const_name "HOL.times"}, @{const_name "HOL.uminus"}, - @{const_name "HOL.minus"}, @{const_name "HOL.plus"}, - @{const_name "HOL.zero"}, - @{const_name "HOL.less"}, @{const_name "HOL.less_eq"}] s + @{const_name HOL.times}, @{const_name HOL.uminus}, + @{const_name HOL.minus}, @{const_name HOL.plus}, + @{const_name HOL.zero}, + @{const_name HOL.less}, @{const_name HOL.less_eq}] s | check (a $ b) = check a andalso check b | check _ = false; diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Tools/lin_arith.ML --- a/src/HOL/Tools/lin_arith.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Tools/lin_arith.ML Tue Sep 22 13:52:19 2009 +1000 @@ -51,7 +51,7 @@ atomize (thm RS conjunct1) @ atomize (thm RS conjunct2) | _ => [thm]; -fun neg_prop ((TP as Const("Trueprop", _)) $ (Const (@{const_name "Not"}, _) $ t)) = TP $ t +fun neg_prop ((TP as Const("Trueprop", _)) $ (Const (@{const_name Not}, _) $ t)) = TP $ t | neg_prop ((TP as Const("Trueprop", _)) $ t) = TP $ (HOLogic.Not $t) | neg_prop t = raise TERM ("neg_prop", [t]); diff -r c876bcb601fc -r 5b65449d7669 src/HOL/Transitive_Closure.thy --- a/src/HOL/Transitive_Closure.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/Transitive_Closure.thy Tue Sep 22 13:52:19 2009 +1000 @@ -77,7 +77,7 @@ subsection {* Reflexive-transitive closure *} lemma reflcl_set_eq [pred_set_conv]: "(sup (\x y. (x, y) \ r) op =) = (\x y. (x, y) \ r Un Id)" - by (simp add: expand_fun_eq) + by (simp add: expand_fun_eq sup2_iff) lemma r_into_rtrancl [intro]: "!!p. p \ r ==> p \ r^*" -- {* @{text rtrancl} of @{text r} contains @{text r} *} diff -r c876bcb601fc -r 5b65449d7669 src/HOL/UNITY/ProgressSets.thy --- a/src/HOL/UNITY/ProgressSets.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/UNITY/ProgressSets.thy Tue Sep 22 13:52:19 2009 +1000 @@ -534,7 +534,7 @@ subsubsection{*Commutativity of Functions and Relation*} text{*Thesis, page 109*} -(*FIXME: this proof is an ungodly mess*) +(*FIXME: this proof is still an ungodly mess*) text{*From Meier's thesis, section 4.5.6*} lemma commutativity2_lemma: assumes dcommutes: @@ -548,36 +548,35 @@ and TL: "T \ L" and Fstable: "F \ stable T" shows "commutes F T B L" -apply (simp add: commutes_def del: Int_subset_iff, clarify) -apply (rename_tac t) -apply (subgoal_tac "\s. (s,t) \ relcl L & s \ T \ wp act M") - prefer 2 - apply (force simp add: cl_eq_Collect_relcl [OF lattice], simp, clarify) -apply (subgoal_tac "\u\L. s \ u --> t \ u") - prefer 2 - apply (intro ballI impI) - apply (subst cl_ident [symmetric], assumption) - apply (simp add: relcl_def) - apply (blast intro: cl_mono [THEN [2] rev_subsetD]) -apply (subgoal_tac "funof act s \ T\M") - prefer 2 - apply (cut_tac Fstable) - apply (force intro!: funof_in - simp add: wp_def stable_def constrains_def determ total) -apply (subgoal_tac "s \ B | t \ B | (funof act s, funof act t) \ relcl L") - prefer 2 - apply (rule dcommutes [rule_format], assumption+) -apply (subgoal_tac "t \ B | funof act t \ cl L (T\M)") - prefer 2 - apply (simp add: relcl_def) - apply (blast intro: BL cl_mono [THEN [2] rev_subsetD]) -apply (subgoal_tac "t \ B | t \ wp act (cl L (T\M))") - prefer 2 - apply (blast intro: funof_imp_wp determ) -apply (blast intro: TL cl_mono [THEN [2] rev_subsetD]) -done - - +apply (simp add: commutes_def del: Int_subset_iff le_inf_iff, clarify) +proof - + fix M and act and t + assume 1: "B \ M" "act \ Acts F" "t \ cl L (T \ wp act M)" + then have "\s. (s,t) \ relcl L \ s \ T \ wp act M" + by (force simp add: cl_eq_Collect_relcl [OF lattice]) + then obtain s where 2: "(s, t) \ relcl L" "s \ T" "s \ wp act M" + by blast + then have 3: "\u\L. s \ u --> t \ u" + apply (intro ballI impI) + apply (subst cl_ident [symmetric], assumption) + apply (simp add: relcl_def) + apply (blast intro: cl_mono [THEN [2] rev_subsetD]) + done + with 1 2 Fstable have 4: "funof act s \ T\M" + by (force intro!: funof_in + simp add: wp_def stable_def constrains_def determ total) + with 1 2 3 have 5: "s \ B | t \ B | (funof act s, funof act t) \ relcl L" + by (intro dcommutes [rule_format]) assumption+ + with 1 2 3 4 have "t \ B | funof act t \ cl L (T\M)" + by (simp add: relcl_def) (blast intro: BL cl_mono [THEN [2] rev_subsetD]) + with 1 2 3 4 5 have "t \ B | t \ wp act (cl L (T\M))" + by (blast intro: funof_imp_wp determ) + with 2 3 have "t \ T \ (t \ B \ t \ wp act (cl L (T \ M)))" + by (blast intro: TL cl_mono [THEN [2] rev_subsetD]) + then show "t \ T \ (B \ wp act (cl L (T \ M)))" + by simp +qed + text{*Version packaged with @{thm progress_set_Union}*} lemma commutativity2: assumes leadsTo: "F \ A leadsTo B" diff -r c876bcb601fc -r 5b65449d7669 src/HOL/UNITY/ROOT.ML --- a/src/HOL/UNITY/ROOT.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/UNITY/ROOT.ML Tue Sep 22 13:52:19 2009 +1000 @@ -1,50 +1,13 @@ -(* Title: HOL/UNITY/ROOT - ID: $Id$ - Author: Lawrence C Paulson, Cambridge University Computer Laboratory +(* Author: Lawrence C Paulson, Cambridge University Computer Laboratory Copyright 1998 University of Cambridge -Root file for UNITY proofs. *) (*Verifying security protocols using UNITY*) no_document use_thy "../Auth/Public"; -use_thys [ - (*Basic meta-theory*) - "UNITY_Main", - - (*Simple examples: no composition*) - "Simple/Deadlock", - "Simple/Common", - "Simple/Network", - "Simple/Token", - "Simple/Channel", - "Simple/Lift", - "Simple/Mutex", - "Simple/Reach", - "Simple/Reachability", - - (*Verifying security protocols using UNITY*) - "Simple/NSP_Bad", +(*Basic meta-theory*) +use_thy "UNITY_Main"; - (*Example of composition*) - "Comp/Handshake", - - (*Universal properties examples*) - "Comp/Counter", - "Comp/Counterc", - "Comp/Priority", - - "Comp/TimerArray", - "Comp/Progress", - - (*obsolete*) - "ELT" -]; - -(*Allocator example*) -(* FIXME some parts no longer work -- had been commented out for a long time *) -setmp_noncritical quick_and_dirty true - use_thy "Comp/Alloc"; - -use_thys ["Comp/AllocImpl", "Comp/Client"]; +(*Examples*) +use_thy "UNITY_Examples"; diff -r c876bcb601fc -r 5b65449d7669 src/HOL/UNITY/Simple/Common.thy --- a/src/HOL/UNITY/Simple/Common.thy Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/UNITY/Simple/Common.thy Tue Sep 22 13:52:19 2009 +1000 @@ -83,19 +83,24 @@ (*** Progress under weak fairness ***) -declare atMost_Int_atLeast [simp] - lemma leadsTo_common_lemma: - "[| \m. F \ {m} Co (maxfg m); - \m \ lessThan n. F \ {m} LeadsTo (greaterThan m); - n \ common |] - ==> F \ (atMost n) LeadsTo common" -apply (rule LeadsTo_weaken_R) -apply (rule_tac f = id and l = n in GreaterThan_bounded_induct) -apply (simp_all (no_asm_simp) del: Int_insert_right_if1) -apply (rule_tac [2] subset_refl) -apply (blast dest: PSP_Stable2 intro: common_stable LeadsTo_weaken_R) -done + assumes "\m. F \ {m} Co (maxfg m)" + and "\m \ lessThan n. F \ {m} LeadsTo (greaterThan m)" + and "n \ common" + shows "F \ (atMost n) LeadsTo common" +proof (rule LeadsTo_weaken_R) + show "F \ {..n} LeadsTo {..n} \ id -` {n..} \ common" + proof (induct rule: GreaterThan_bounded_induct [of n _ _ id]) + case 1 + from assms have "\m\{.. {..n} \ {m} LeadsTo {..n} \ {m<..} \ common" + by (blast dest: PSP_Stable2 intro: common_stable LeadsTo_weaken_R) + then show ?case by simp + qed +next + from `n \ common` + show "{..n} \ id -` {n..} \ common \ common" + by (simp add: atMost_Int_atLeast) +qed (*The "\m \ -common" form echoes CMT6.*) lemma leadsTo_common: diff -r c876bcb601fc -r 5b65449d7669 src/HOL/UNITY/UNITY_Examples.thy --- /dev/null Thu Jan 01 00:00:00 1970 +0000 +++ b/src/HOL/UNITY/UNITY_Examples.thy Tue Sep 22 13:52:19 2009 +1000 @@ -0,0 +1,45 @@ +(* Author: Lawrence C Paulson Cambridge University Computer Laboratory + Copyright 1998 University of Cambridge +*) + +header {* Various examples for UNITY *} + +theory UNITY_Examples +imports + UNITY_Main + + (*Simple examples: no composition*) + "Simple/Deadlock" + "Simple/Common" + "Simple/Network" + "Simple/Token" + "Simple/Channel" + "Simple/Lift" + "Simple/Mutex" + "Simple/Reach" + "Simple/Reachability" + + (*Verifying security protocols using UNITY*) + "Simple/NSP_Bad" + + (*Example of composition*) + "Comp/Handshake" + + (*Universal properties examples*) + "Comp/Counter" + "Comp/Counterc" + "Comp/Priority" + + "Comp/TimerArray" + "Comp/Progress" + + "Comp/Alloc" + "Comp/AllocImpl" + "Comp/Client" + + (*obsolete*) + "ELT" + +begin + +end diff -r c876bcb601fc -r 5b65449d7669 src/HOL/ex/ROOT.ML --- a/src/HOL/ex/ROOT.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/HOL/ex/ROOT.ML Tue Sep 22 13:52:19 2009 +1000 @@ -68,8 +68,7 @@ "Landau" ]; -Future.shutdown (); -(setmp_noncritical proofs 2 (setmp_noncritical Multithreading.max_threads 1 use_thy)) +(setmp_noncritical proofs 2 (setmp_noncritical Goal.parallel_proofs 0 use_thy)) "Hilbert_Classical"; use_thy "SVC_Oracle"; diff -r c876bcb601fc -r 5b65449d7669 src/Pure/Concurrent/future.ML --- a/src/Pure/Concurrent/future.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/Pure/Concurrent/future.ML Tue Sep 22 13:52:19 2009 +1000 @@ -237,7 +237,7 @@ val total = length (! workers); val active = count_active (); in - "SCHEDULE " ^ string_of_int (Time.toMilliseconds now) ^ ": " ^ + "SCHEDULE " ^ Time.toString now ^ ": " ^ string_of_int ready ^ " ready, " ^ string_of_int pending ^ " pending, " ^ string_of_int running ^ " running; " ^ @@ -257,7 +257,7 @@ "SCHEDULE: disposed " ^ string_of_int (length dead) ^ " dead worker threads"))); val m = if ! do_shutdown then 0 else Multithreading.max_threads_value (); - val mm = (m * 3) div 2; + val mm = if m = 9999 then 1 else (m * 3) div 2; val l = length (! workers); val _ = excessive := l - mm; val _ = diff -r c876bcb601fc -r 5b65449d7669 src/Pure/Concurrent/par_list.ML --- a/src/Pure/Concurrent/par_list.ML Thu Sep 17 14:17:37 2009 +1000 +++ b/src/Pure/Concurrent/par_list.ML Tue Sep 22 13:52:19 2009 +1000 @@ -27,8 +27,10 @@ struct fun raw_map f xs = - let val group = Task_Queue.new_group (Future.worker_group ()) - in Future.join_results (map (fn x => Future.fork_group group (fn () => f x)) xs) end; + if Multithreading.enabled () then + let val group = Task_Queue.new_group (Future.worker_group ()) + in Future.join_results (map (fn x => Future.fork_group group (fn () => f x)) xs) end + else map (Exn.capture f) xs; fun map f xs = Exn.release_first (raw_map f xs);