Mentions légales du service

Skip to content
Snippets Groups Projects
Commit 8a28f1b1 authored by Léon Gondelman's avatar Léon Gondelman
Browse files

mini-compiler: clean-up (removed all but one coq proofs)

parent 9ca92458
No related branches found
No related tags found
No related merge requests found
Showing
with 1636 additions and 3941 deletions
......@@ -3,6 +3,7 @@
(*Imp to ImpVm compiler *)
(**************************************************************************)
module Compile_aexpr
meta compute_max_steps 0x10000
......@@ -22,14 +23,16 @@ module Compile_aexpr
| VMS _ s m -> ms' = VMS (p + len) (push (aeval m a ) s) m
end
lemma aexpr_post_lemma:
meta rewrite_def function aexpr_post
(* lemma aexpr_post_lemma:
forall a len, x: 'a, p ms ms'.
aexpr_post a len x p ms ms' =
match ms with
| VMS _ s m -> ms' = VMS (p + len) (push (aeval m a ) s) m
end
meta rewrite prop aexpr_post_lemma
meta rewrite prop aexpr_post_lemma *)
let rec compile_aexpr (a : aexpr) : hl 'a
......@@ -61,7 +64,6 @@ module Compile_aexpr
end
module Compile_bexpr
use import int.Int
use import list.List
use import list.Length
......@@ -139,7 +141,7 @@ end
module Compile_com
meta compute_max_steps 65536
use import int.Int
use import list.List
use import list.Length
......@@ -157,11 +159,13 @@ module Compile_com
function com_pre (cmd: com) : 'a -> pos -> pred =
\x p ms. match ms with VMS p' _ m -> p = p' /\ exists m'. ceval m cmd m' end
lemma com_pre_lemma:
meta rewrite_def function com_pre
(*lemma com_pre_lemma:
forall cmd, x:'a, p ms. com_pre cmd x p ms =
match ms with VMS p' _ m -> p = p' /\ exists m'. ceval m cmd m' end
meta rewrite prop com_pre_lemma
meta rewrite prop com_pre_lemma*)
function com_post (cmd: com) (len:pos) : 'a -> pos -> post =
\x p ms ms'. match ms, ms' with
......@@ -169,14 +173,16 @@ module Compile_com
p' = p + len /\ s' = s /\ ceval m cmd m'
end
lemma com_post_lemma:
meta rewrite_def function com_post
(*lemma com_post_lemma:
forall cmd len, x:'a, p ms ms'. com_post cmd len x p ms ms' =
match ms, ms' with
| VMS p s m, VMS p' s' m' ->
p' = p + len /\ s' = s /\ ceval m cmd m'
end
meta rewrite prop com_post_lemma
meta rewrite prop com_post_lemma*)
function exn_bool_old (b1: bexpr) (cond: bool) : ('a,machine_state) -> pos -> pred =
\x p ms. match snd x with
......@@ -260,7 +266,10 @@ module Compile_com
(VMS p s m)
(VMS (p + length result) s m') }
= let res = compile_com com : hl unit in
assert { forall p s m m'. ceval m com m' -> res.pre () p (VMS p s m) };
assert {
forall c p s m m'. ceval m com m' -> codeseq_at c p res.code ->
res.pre () p (VMS p s m) && (forall ms'. res.post () p (VMS p s m) ms' ->
ms' = VMS (p + length res.code) s m')};
res.code
......
(* This file is generated by Why3's Coq driver *)
(* Beware! Only edit allowed sections below *)
Require Import BuiltIn.
Require BuiltIn.
Require int.Int.
Require int.Abs.
Require int.EuclideanDivision.
Require list.List.
Require list.Length.
Require list.Mem.
Require map.Map.
Require bool.Bool.
Require list.Append.
Axiom map : forall (a:Type) (b:Type), Type.
Parameter map_WhyType : forall (a:Type) {a_WT:WhyType a}
(b:Type) {b_WT:WhyType b}, WhyType (map a b).
Existing Instance map_WhyType.
Parameter get: forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
(map a b) -> a -> b.
Parameter set: forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
(map a b) -> a -> b -> (map a b).
Parameter const: forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
b -> (map a b).
(* Why3 assumption *)
Inductive id :=
| Id : Z -> id.
Axiom id_WhyType : WhyType id.
Existing Instance id_WhyType.
(* Why3 assumption *)
Inductive aexpr :=
| Anum : Z -> aexpr
| Avar : id -> aexpr
| Aadd : aexpr -> aexpr -> aexpr
| Asub : aexpr -> aexpr -> aexpr
| Amul : aexpr -> aexpr -> aexpr.
Axiom aexpr_WhyType : WhyType aexpr.
Existing Instance aexpr_WhyType.
(* Why3 assumption *)
Inductive bexpr :=
| Btrue : bexpr
| Bfalse : bexpr
| Band : bexpr -> bexpr -> bexpr
| Bnot : bexpr -> bexpr
| Beq : aexpr -> aexpr -> bexpr
| Ble : aexpr -> aexpr -> bexpr.
Axiom bexpr_WhyType : WhyType bexpr.
Existing Instance bexpr_WhyType.
(* Why3 assumption *)
Inductive com :=
| Cskip : com
| Cassign : id -> aexpr -> com
| Cseq : com -> com -> com
| Cif : bexpr -> com -> com -> com
| Cwhile : bexpr -> com -> com.
Axiom com_WhyType : WhyType com.
Existing Instance com_WhyType.
(* Why3 assumption *)
Fixpoint aeval (st:(map id Z)) (e:aexpr) {struct e}: Z :=
match e with
| (Anum n) => n
| (Avar x) => (get st x)
| (Aadd e1 e2) => ((aeval st e1) + (aeval st e2))%Z
| (Asub e1 e2) => ((aeval st e1) - (aeval st e2))%Z
| (Amul e1 e2) => ((aeval st e1) * (aeval st e2))%Z
end.
Parameter beval: (map id Z) -> bexpr -> bool.
Axiom beval_def : forall (st:(map id Z)) (b:bexpr),
match b with
| Btrue => ((beval st b) = true)
| Bfalse => ((beval st b) = false)
| (Bnot b') => ((beval st b) = (Init.Datatypes.negb (beval st b')))
| (Band b1 b2) => ((beval st b) = (Init.Datatypes.andb (beval st
b1) (beval st b2)))
| (Beq a1 a2) => (((aeval st a1) = (aeval st a2)) -> ((beval st
b) = true)) /\ ((~ ((aeval st a1) = (aeval st a2))) -> ((beval st
b) = false))
| (Ble a1 a2) => (((aeval st a1) <= (aeval st a2))%Z -> ((beval st
b) = true)) /\ ((~ ((aeval st a1) <= (aeval st a2))%Z) -> ((beval st
b) = false))
end.
(* Why3 assumption *)
Inductive ceval: (map id Z) -> com -> (map id Z) -> Prop :=
| E_Skip : forall (m:(map id Z)), (ceval m Cskip m)
| E_Ass : forall (m:(map id Z)) (a:aexpr) (n:Z) (x:id), ((aeval m
a) = n) -> (ceval m (Cassign x a) (set m x n))
| E_Seq : forall (cmd1:com) (cmd2:com) (m0:(map id Z)) (m1:(map id Z))
(m2:(map id Z)), (ceval m0 cmd1 m1) -> ((ceval m1 cmd2 m2) -> (ceval m0
(Cseq cmd1 cmd2) m2))
| E_IfTrue : forall (m0:(map id Z)) (m1:(map id Z)) (cond:bexpr) (cmd1:com)
(cmd2:com), ((beval m0 cond) = true) -> ((ceval m0 cmd1 m1) -> (ceval
m0 (Cif cond cmd1 cmd2) m1))
| E_IfFalse : forall (m0:(map id Z)) (m1:(map id Z)) (cond:bexpr)
(cmd1:com) (cmd2:com), ((beval m0 cond) = false) -> ((ceval m0 cmd2
m1) -> (ceval m0 (Cif cond cmd1 cmd2) m1))
| E_WhileEnd : forall (cond:bexpr) (m:(map id Z)) (body:com), ((beval m
cond) = false) -> (ceval m (Cwhile cond body) m)
| E_WhileLoop : forall (mi:(map id Z)) (mj:(map id Z)) (mf:(map id Z))
(cond:bexpr) (body:com), ((beval mi cond) = true) -> ((ceval mi body
mj) -> ((ceval mj (Cwhile cond body) mf) -> (ceval mi (Cwhile cond
body) mf))).
Axiom ceval_deterministic : forall (c:com) (mi:(map id Z)) (mf1:(map id Z))
(mf2:(map id Z)), (ceval mi c mf1) -> ((ceval mi c mf2) -> (mf1 = mf2)).
(* Why3 assumption *)
Inductive machine_state :=
| VMS : Z -> (list Z) -> (map id Z) -> machine_state.
Axiom machine_state_WhyType : WhyType machine_state.
Existing Instance machine_state_WhyType.
(* Why3 assumption *)
Inductive instr :=
| Iconst : Z -> instr
| Ivar : id -> instr
| Isetvar : id -> instr
| Ibranch : Z -> instr
| Iadd : instr
| Isub : instr
| Imul : instr
| Ibeq : Z -> instr
| Ibne : Z -> instr
| Ible : Z -> instr
| Ibgt : Z -> instr
| Ihalt : instr.
Axiom instr_WhyType : WhyType instr.
Existing Instance instr_WhyType.
(* Why3 assumption *)
Inductive codeseq_at: (list instr) -> Z -> (list instr) -> Prop :=
| codeseq_at_intro : forall (c1:(list instr)) (c2:(list instr))
(c3:(list instr)) (pc:Z), (pc = (list.Length.length c1)) -> (codeseq_at
(Init.Datatypes.app (Init.Datatypes.app c1 c2) c3) pc c2).
Parameter transition_star: (list instr) -> machine_state -> machine_state ->
Prop.
Axiom func : forall (a:Type) (b:Type), Type.
Parameter func_WhyType : forall (a:Type) {a_WT:WhyType a}
(b:Type) {b_WT:WhyType b}, WhyType (func a b).
Existing Instance func_WhyType.
Parameter infix_at: forall {a:Type} {a_WT:WhyType a}
{b:Type} {b_WT:WhyType b}, (func a b) -> a -> b.
(* Why3 assumption *)
Inductive hl
(a:Type) :=
| mk_hl : (list instr) -> (func a (func Z (func machine_state bool))) ->
(func a (func Z (func machine_state (func machine_state bool)))) -> hl
a.
Axiom hl_WhyType : forall (a:Type) {a_WT:WhyType a}, WhyType (hl a).
Existing Instance hl_WhyType.
Implicit Arguments mk_hl [[a]].
(* Why3 assumption *)
Definition post {a:Type} {a_WT:WhyType a} (v:(hl a)): (func a (func Z (func
machine_state (func machine_state bool)))) :=
match v with
| (mk_hl x x1 x2) => x2
end.
(* Why3 assumption *)
Definition pre {a:Type} {a_WT:WhyType a} (v:(hl a)): (func a (func Z (func
machine_state bool))) := match v with
| (mk_hl x x1 x2) => x1
end.
(* Why3 assumption *)
Definition code {a:Type} {a_WT:WhyType a} (v:(hl a)): (list instr) :=
match v with
| (mk_hl x x1 x2) => x
end.
(* Why3 assumption *)
Definition contextual_irrelevance (c:(list instr)) (p:Z) (ms1:machine_state)
(ms2:machine_state): Prop := forall (c_global:(list instr)), (codeseq_at
c_global p c) -> (transition_star c_global ms1 ms2).
(* Why3 assumption *)
Definition hl_correctness {a:Type} {a_WT:WhyType a} (cs:(hl a)): Prop :=
forall (x:a) (p:Z) (ms:machine_state),
((infix_at (infix_at (infix_at (pre cs) x) p) ms) = true) ->
exists ms':machine_state,
((infix_at (infix_at (infix_at (infix_at (post cs) x) p) ms)
ms') = true) /\ (contextual_irrelevance (code cs) p ms ms').
Parameter com_pre: forall {a:Type} {a_WT:WhyType a}, com -> (func a (func Z
(func machine_state bool))).
Axiom com_pre_def : forall {a:Type} {a_WT:WhyType a}, forall (cmd:com) (x:a)
(p:Z) (ms:machine_state),
((infix_at (infix_at (infix_at (com_pre cmd: (func a (func Z (func
machine_state bool)))) x) p) ms) = true) <->
match ms with
| (VMS p' _ m) => (p = p') /\ exists m':(map id Z), (ceval m cmd m')
end.
Parameter com_post: forall {a:Type} {a_WT:WhyType a}, com -> Z -> (func a
(func Z (func machine_state (func machine_state bool)))).
Axiom com_post_def : forall {a:Type} {a_WT:WhyType a}, forall (cmd:com)
(len:Z) (x:a) (p:Z) (ms:machine_state) (ms':machine_state),
((infix_at (infix_at (infix_at (infix_at (com_post cmd len: (func a (func Z
(func machine_state (func machine_state bool))))) x) p) ms)
ms') = true) <-> match (ms,
ms') with
| ((VMS p1 s m), (VMS p' s' m')) => (p' = (p1 + len)%Z) /\ ((s' = s) /\
(ceval m cmd m'))
end.
(* Why3 goal *)
Theorem WP_parameter_compile_com_natural : forall (com1:com),
forall (res:(list instr)) (res1:(func unit (func Z (func machine_state
bool)))) (res2:(func unit (func Z (func machine_state (func machine_state
bool))))), ((res1 = (com_pre com1: (func unit (func Z (func machine_state
bool))))) /\ ((res2 = (com_post com1 (list.Length.length res): (func unit
(func Z (func machine_state (func machine_state bool)))))) /\
(hl_correctness (mk_hl res res1 res2)))) -> ((forall (p:Z) (s:(list Z))
(m:(map id Z)) (m':(map id Z)), (ceval m com1 m') ->
((infix_at (infix_at (infix_at res1 tt) p) (VMS p s m)) = true)) ->
forall (c:(list instr)) (p:Z) (s:(list Z)) (m:(map id Z)) (m':(map id Z)),
(ceval m com1 m') -> ((codeseq_at c p res) -> (transition_star c (VMS p s
m) (VMS (p + (list.Length.length res))%Z s m')))).
(* Why3 intros com1 res res1 res2 (h1,(h2,h3)) h4 c p s m m' h5 h6. *)
intros com1 res res1 res2 (h1,(h2,h3)) h4 c p s m m' h5 h6.
unfold hl_correctness in *.
remember h5 as h7 eqn:eqn;clear eqn.
apply (h4 p s) in h5.
apply h3 in h5.
Require Import Why3.
simpl in *.
subst.
destruct h5.
rewrite com_post_def in H.
why3 "Alt-Ergo,0.95.2,".
Qed.
This diff is collapsed.
No preview for this file type
......@@ -51,13 +51,13 @@ theory Imp
if (aeval st a1) <= (aeval st a2) then True else False
end
lemma inversion_beval_t :
(* lemma inversion_beval_t :
forall a1 a2: aexpr, m: state.
beval m (Beq a1 a2) = True -> aeval m a1 = aeval m a2
lemma inversion_beval_f :
forall a1 a2: aexpr, m: state.
beval m (Beq a1 a2) = False -> aeval m a1 <> aeval m a2
beval m (Beq a1 a2) = False -> aeval m a1 <> aeval m a2 *)
inductive ceval state com state =
(* skip *)
......@@ -104,7 +104,9 @@ theory Imp
ceval mi (Cwhile cond body) mf
lemma ceval_deterministic:
forall c mi mf1 mf2. ceval mi c mf1 -> ceval mi c mf2 -> mf1 = mf2
lemma ceval_deterministic_aux :
forall c mi mf1. ceval mi c mf1 -> forall mf2. ("inversion" ceval mi c mf2) -> mf1 = mf2
lemma ceval_deterministic :
forall c mi mf1 mf2. ceval mi c mf1 -> ceval mi c mf2 -> mf1 = mf2
end
\ No newline at end of file
(* This file is generated by Why3's Coq driver *)
(* Beware! Only edit allowed sections below *)
Require Import BuiltIn.
Require BuiltIn.
Require map.Map.
Require int.Int.
Require bool.Bool.
Axiom map : forall (a:Type) (b:Type), Type.
Parameter map_WhyType : forall (a:Type) {a_WT:WhyType a}
(b:Type) {b_WT:WhyType b}, WhyType (map a b).
Existing Instance map_WhyType.
Parameter get: forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
(map a b) -> a -> b.
Parameter set: forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
(map a b) -> a -> b -> (map a b).
Axiom Select_eq : forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
forall (m:(map a b)), forall (a1:a) (a2:a), forall (b1:b), (a1 = a2) ->
((get (set m a1 b1) a2) = b1).
Axiom Select_neq : forall {a:Type} {a_WT:WhyType a}
{b:Type} {b_WT:WhyType b}, forall (m:(map a b)), forall (a1:a) (a2:a),
forall (b1:b), (~ (a1 = a2)) -> ((get (set m a1 b1) a2) = (get m a2)).
Parameter const: forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
b -> (map a b).
Axiom Const : forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
forall (b1:b) (a1:a), ((get (const b1: (map a b)) a1) = b1).
(* Why3 assumption *)
Inductive id :=
| Id : Z -> id.
Axiom id_WhyType : WhyType id.
Existing Instance id_WhyType.
(* Why3 assumption *)
Definition state := (map id Z).
(* Why3 assumption *)
Inductive aexpr :=
| Anum : Z -> aexpr
| Avar : id -> aexpr
| Aadd : aexpr -> aexpr -> aexpr
| Asub : aexpr -> aexpr -> aexpr
| Amul : aexpr -> aexpr -> aexpr.
Axiom aexpr_WhyType : WhyType aexpr.
Existing Instance aexpr_WhyType.
(* Why3 assumption *)
Inductive bexpr :=
| Btrue : bexpr
| Bfalse : bexpr
| Band : bexpr -> bexpr -> bexpr
| Bnot : bexpr -> bexpr
| Beq : aexpr -> aexpr -> bexpr
| Ble : aexpr -> aexpr -> bexpr.
Axiom bexpr_WhyType : WhyType bexpr.
Existing Instance bexpr_WhyType.
(* Why3 assumption *)
Inductive com :=
| Cskip : com
| Cassign : id -> aexpr -> com
| Cseq : com -> com -> com
| Cif : bexpr -> com -> com -> com
| Cwhile : bexpr -> com -> com.
Axiom com_WhyType : WhyType com.
Existing Instance com_WhyType.
(* Why3 assumption *)
Fixpoint aeval (st:(map id Z)) (e:aexpr) {struct e}: Z :=
match e with
| (Anum n) => n
| (Avar x) => (get st x)
| (Aadd e1 e2) => ((aeval st e1) + (aeval st e2))%Z
| (Asub e1 e2) => ((aeval st e1) - (aeval st e2))%Z
| (Amul e1 e2) => ((aeval st e1) * (aeval st e2))%Z
end.
Parameter beval: (map id Z) -> bexpr -> bool.
Axiom beval_def : forall (st:(map id Z)) (b:bexpr),
match b with
| Btrue => ((beval st b) = true)
| Bfalse => ((beval st b) = false)
| (Bnot b') => ((beval st b) = (Init.Datatypes.negb (beval st b')))
| (Band b1 b2) => ((beval st b) = (Init.Datatypes.andb (beval st
b1) (beval st b2)))
| (Beq a1 a2) => (((aeval st a1) = (aeval st a2)) -> ((beval st
b) = true)) /\ ((~ ((aeval st a1) = (aeval st a2))) -> ((beval st
b) = false))
| (Ble a1 a2) => (((aeval st a1) <= (aeval st a2))%Z -> ((beval st
b) = true)) /\ ((~ ((aeval st a1) <= (aeval st a2))%Z) -> ((beval st
b) = false))
end.
Axiom inversion_beval_t : forall (a1:aexpr) (a2:aexpr) (m:(map id Z)),
((beval m (Beq a1 a2)) = true) -> ((aeval m a1) = (aeval m a2)).
Axiom inversion_beval_f : forall (a1:aexpr) (a2:aexpr) (m:(map id Z)),
((beval m (Beq a1 a2)) = false) -> ~ ((aeval m a1) = (aeval m a2)).
(* Why3 assumption *)
Inductive ceval: (map id Z) -> com -> (map id Z) -> Prop :=
| E_Skip : forall (m:(map id Z)), (ceval m Cskip m)
| E_Ass : forall (m:(map id Z)) (a:aexpr) (n:Z) (x:id), ((aeval m
a) = n) -> (ceval m (Cassign x a) (set m x n))
| E_Seq : forall (cmd1:com) (cmd2:com) (m0:(map id Z)) (m1:(map id Z))
(m2:(map id Z)), (ceval m0 cmd1 m1) -> ((ceval m1 cmd2 m2) -> (ceval m0
(Cseq cmd1 cmd2) m2))
| E_IfTrue : forall (m0:(map id Z)) (m1:(map id Z)) (cond:bexpr) (cmd1:com)
(cmd2:com), ((beval m0 cond) = true) -> ((ceval m0 cmd1 m1) -> (ceval
m0 (Cif cond cmd1 cmd2) m1))
| E_IfFalse : forall (m0:(map id Z)) (m1:(map id Z)) (cond:bexpr)
(cmd1:com) (cmd2:com), ((beval m0 cond) = false) -> ((ceval m0 cmd2
m1) -> (ceval m0 (Cif cond cmd1 cmd2) m1))
| E_WhileEnd : forall (cond:bexpr) (m:(map id Z)) (body:com), ((beval m
cond) = false) -> (ceval m (Cwhile cond body) m)
| E_WhileLoop : forall (mi:(map id Z)) (mj:(map id Z)) (mf:(map id Z))
(cond:bexpr) (body:com), ((beval mi cond) = true) -> ((ceval mi body
mj) -> ((ceval mj (Cwhile cond body) mf) -> (ceval mi (Cwhile cond
body) mf))).
Require Import Why3.
Ltac cvc := why3 "CVC4,1.4,".
(* Why3 goal *)
Theorem ceval_deterministic : forall (c:com) (mi:(map id Z)) (mf1:(map id Z))
(mf2:(map id Z)), (ceval mi c mf1) -> ((ceval mi c mf2) -> (mf1 = mf2)).
intros c mi mf1 mf2 h1 h2.
generalize dependent mf2.
induction h1;intros mf2 h2; inversion h2; cvc.
Qed.
......@@ -2,18 +2,191 @@
<!DOCTYPE why3session PUBLIC "-//Why3//proof session v5//EN"
"http://why3.lri.fr/why3session.dtd">
<why3session shape_version="4">
<prover id="0" name="Coq" version="8.4pl2" timelimit="5" memlimit="1000"/>
<prover id="0" name="CVC4" version="1.4" timelimit="5" memlimit="1000"/>
<prover id="1" name="Alt-Ergo" version="0.95.2" timelimit="5" memlimit="1000"/>
<file name="../imp.why" expanded="true">
<theory name="Imp" sum="939397712f6415b10da2ec9b4ffeab30" expanded="true">
<goal name="inversion_beval_t" expanded="true">
<proof prover="1"><result status="valid" time="0.01"/></proof>
</goal>
<goal name="inversion_beval_f" expanded="true">
<proof prover="1"><result status="valid" time="0.01"/></proof>
<theory name="Imp" sum="8d9b512f0a2a1961484e5a0fa9ff9b68" expanded="true">
<goal name="ceval_deterministic_aux" expanded="true">
<transf name="induction_pr" expanded="true">
<goal name="ceval_deterministic_aux.1" expl="1.">
<transf name="inversion_pr">
<goal name="ceval_deterministic_aux.1.1" expl="1.">
<proof prover="1"><result status="valid" time="0.01"/></proof>
</goal>
<goal name="ceval_deterministic_aux.1.2" expl="2.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="ceval_deterministic_aux.1.3" expl="3.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.1.4" expl="4.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.1.5" expl="5.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.1.6" expl="6.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.1.7" expl="7.">
<proof prover="1"><result status="valid" time="0.01"/></proof>
</goal>
</transf>
</goal>
<goal name="ceval_deterministic_aux.2" expl="2.">
<transf name="inversion_pr">
<goal name="ceval_deterministic_aux.2.1" expl="1.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="ceval_deterministic_aux.2.2" expl="2.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.2.3" expl="3.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="ceval_deterministic_aux.2.4" expl="4.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="ceval_deterministic_aux.2.5" expl="5.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.2.6" expl="6.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.2.7" expl="7.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
</transf>
</goal>
<goal name="ceval_deterministic_aux.3" expl="3.">
<transf name="inversion_pr">
<goal name="ceval_deterministic_aux.3.1" expl="1.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.3.2" expl="2.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.3.3" expl="3.">
<proof prover="0"><result status="valid" time="0.11"/></proof>
</goal>
<goal name="ceval_deterministic_aux.3.4" expl="4.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.3.5" expl="5.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.3.6" expl="6.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="ceval_deterministic_aux.3.7" expl="7.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
</transf>
</goal>
<goal name="ceval_deterministic_aux.4" expl="4.">
<transf name="inversion_pr">
<goal name="ceval_deterministic_aux.4.1" expl="1.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.4.2" expl="2.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.4.3" expl="3.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="ceval_deterministic_aux.4.4" expl="4.">
<proof prover="0"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="ceval_deterministic_aux.4.5" expl="5.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="ceval_deterministic_aux.4.6" expl="6.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.4.7" expl="7.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
</transf>
</goal>
<goal name="ceval_deterministic_aux.5" expl="5.">
<transf name="inversion_pr">
<goal name="ceval_deterministic_aux.5.1" expl="1.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.5.2" expl="2.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.5.3" expl="3.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.5.4" expl="4.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.5.5" expl="5.">
<proof prover="0"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="ceval_deterministic_aux.5.6" expl="6.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.5.7" expl="7.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
</transf>
</goal>
<goal name="ceval_deterministic_aux.6" expl="6.">
<transf name="inversion_pr">
<goal name="ceval_deterministic_aux.6.1" expl="1.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.6.2" expl="2.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="ceval_deterministic_aux.6.3" expl="3.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="ceval_deterministic_aux.6.4" expl="4.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.6.5" expl="5.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.6.6" expl="6.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.6.7" expl="7.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
</transf>
</goal>
<goal name="ceval_deterministic_aux.7" expl="7.">
<transf name="inversion_pr">
<goal name="ceval_deterministic_aux.7.1" expl="1.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.7.2" expl="2.">
<proof prover="1"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="ceval_deterministic_aux.7.3" expl="3.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.7.4" expl="4.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.7.5" expl="5.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.7.6" expl="6.">
<proof prover="1"><result status="valid" time="0.02"/></proof>
</goal>
<goal name="ceval_deterministic_aux.7.7" expl="7.">
<proof prover="0"><result status="valid" time="0.08"/></proof>
</goal>
</transf>
</goal>
</transf>
</goal>
<goal name="ceval_deterministic" expanded="true">
<proof prover="0" edited="imp_Imp_ceval_deterministic_1.v"><result status="valid" time="2.36"/></proof>
<proof prover="1"><result status="valid" time="1.06"/></proof>
</goal>
</theory>
</file>
......
No preview for this file type
......@@ -179,7 +179,14 @@ module Compiler_logic
requires { c.post = loop_preservation inv var post }
ensures { result.pre = inv /\ result.post = forget_old post }
ensures { result.code.length = c.code.length /\ hl_correctness result }
= { code = c.code ; pre = inv ; post = forget_old post }
=
let res = { code = c.code ; pre = inv ; post = forget_old post } in
assert { forall x p ms. inv x p ms ->
acc (var x p) ms &&
exists ms'.
contextual_irrelevance res.code p ms ms' /\
forget_old post x p ms ms' };
res
......
(* This file is generated by Why3's Coq driver *)
(* Beware! Only edit allowed sections below *)
Require Import BuiltIn.
Require BuiltIn.
Require int.Int.
Require int.Abs.
Require int.EuclideanDivision.
Require list.List.
Require list.Length.
Require list.Mem.
Require map.Map.
Require list.Append.
(* Why3 assumption *)
Definition unit := unit.
Axiom qtmark : Type.
Parameter qtmark_WhyType : WhyType qtmark.
Existing Instance qtmark_WhyType.
Axiom map : forall (a:Type) (b:Type), Type.
Parameter map_WhyType : forall (a:Type) {a_WT:WhyType a}
(b:Type) {b_WT:WhyType b}, WhyType (map a b).
Existing Instance map_WhyType.
Parameter get: forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
(map a b) -> a -> b.
Parameter set: forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
(map a b) -> a -> b -> (map a b).
Axiom Select_eq : forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
forall (m:(map a b)), forall (a1:a) (a2:a), forall (b1:b), (a1 = a2) ->
((get (set m a1 b1) a2) = b1).
Axiom Select_neq : forall {a:Type} {a_WT:WhyType a}
{b:Type} {b_WT:WhyType b}, forall (m:(map a b)), forall (a1:a) (a2:a),
forall (b1:b), (~ (a1 = a2)) -> ((get (set m a1 b1) a2) = (get m a2)).
Parameter const: forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
b -> (map a b).
Axiom Const : forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b},
forall (b1:b) (a1:a), ((get (const b1: (map a b)) a1) = b1).
(* Why3 assumption *)
Inductive id :=
| Id : Z -> id.
Axiom id_WhyType : WhyType id.
Existing Instance id_WhyType.
(* Why3 assumption *)
Definition state := (map id Z).
(* Why3 assumption *)
Definition pos := Z.
(* Why3 assumption *)
Definition stack := (list Z).
(* Why3 assumption *)
Inductive machine_state :=
| VMS : Z -> (list Z) -> (map id Z) -> machine_state.
Axiom machine_state_WhyType : WhyType machine_state.
Existing Instance machine_state_WhyType.
(* Why3 assumption *)
Inductive instr :=
| Iconst : Z -> instr
| Ivar : id -> instr
| Isetvar : id -> instr
| Ibranch : Z -> instr
| Iadd : instr
| Isub : instr
| Imul : instr
| Ibeq : Z -> instr
| Ibne : Z -> instr
| Ible : Z -> instr
| Ibgt : Z -> instr
| Ihalt : instr.
Axiom instr_WhyType : WhyType instr.
Existing Instance instr_WhyType.
(* Why3 assumption *)
Definition code := (list instr).
(* Why3 assumption *)
Inductive codeseq_at: (list instr) -> Z -> (list instr) -> Prop :=
| codeseq_at_intro : forall (c1:(list instr)) (c2:(list instr))
(c3:(list instr)) (pc:Z), (pc = (list.Length.length c1)) -> (codeseq_at
(Init.Datatypes.app (Init.Datatypes.app c1 c2) c3) pc c2).
Axiom codeseq_at_app_right : forall (c:(list instr)) (c1:(list instr))
(c2:(list instr)) (p:Z), (codeseq_at c p (Init.Datatypes.app c1 c2)) ->
(codeseq_at c (p + (list.Length.length c1))%Z c2).
Axiom codeseq_at_app_left : forall (c:(list instr)) (c1:(list instr))
(c2:(list instr)) (p:Z), (codeseq_at c p (Init.Datatypes.app c1 c2)) ->
(codeseq_at c p c1).
(* Why3 assumption *)
Definition iconst (n:Z): (list instr) :=
(Init.Datatypes.cons (Iconst n) Init.Datatypes.nil).
(* Why3 assumption *)
Definition ivar (x:id): (list instr) :=
(Init.Datatypes.cons (Ivar x) Init.Datatypes.nil).
(* Why3 assumption *)
Definition isetvar (x:id): (list instr) :=
(Init.Datatypes.cons (Isetvar x) Init.Datatypes.nil).
(* Why3 assumption *)
Definition ibeq (ofs:Z): (list instr) :=
(Init.Datatypes.cons (Ibeq ofs) Init.Datatypes.nil).
(* Why3 assumption *)
Definition ible (ofs:Z): (list instr) :=
(Init.Datatypes.cons (Ible ofs) Init.Datatypes.nil).
(* Why3 assumption *)
Definition ibne (ofs:Z): (list instr) :=
(Init.Datatypes.cons (Ibne ofs) Init.Datatypes.nil).
(* Why3 assumption *)
Definition ibgt (ofs:Z): (list instr) :=
(Init.Datatypes.cons (Ibgt ofs) Init.Datatypes.nil).
(* Why3 assumption *)
Definition ibranch (ofs:Z): (list instr) :=
(Init.Datatypes.cons (Ibranch ofs) Init.Datatypes.nil).
(* Why3 assumption *)
Inductive transition: (list instr) -> machine_state -> machine_state ->
Prop :=
| trans_const : forall (c:(list instr)) (p:Z) (n:Z), (codeseq_at c p
(iconst n)) -> forall (m:(map id Z)) (s:(list Z)), (transition c (VMS p
s m) (VMS (p + 1%Z)%Z (Init.Datatypes.cons n s) m))
| trans_var : forall (c:(list instr)) (p:Z) (x:id), (codeseq_at c p
(ivar x)) -> forall (m:(map id Z)) (s:(list Z)), (transition c (VMS p s
m) (VMS (p + 1%Z)%Z (Init.Datatypes.cons (get m x) s) m))
| trans_set_var : forall (c:(list instr)) (p:Z) (x:id), (codeseq_at c p
(isetvar x)) -> forall (n:Z) (s:(list Z)) (m:(map id Z)), (transition c
(VMS p (Init.Datatypes.cons n s) m) (VMS (p + 1%Z)%Z s (set m x n)))
| trans_add : forall (c:(list instr)) (p:Z), (codeseq_at c p
(Init.Datatypes.cons Iadd Init.Datatypes.nil)) -> forall (n1:Z) (n2:Z)
(s:(list Z)) (m:(map id Z)), (transition c (VMS p
(Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m) (VMS (p + 1%Z)%Z
(Init.Datatypes.cons (n1 + n2)%Z s) m))
| trans_sub : forall (c:(list instr)) (p:Z), (codeseq_at c p
(Init.Datatypes.cons Isub Init.Datatypes.nil)) -> forall (n1:Z) (n2:Z)
(s:(list Z)) (m:(map id Z)), (transition c (VMS p
(Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m) (VMS (p + 1%Z)%Z
(Init.Datatypes.cons (n1 - n2)%Z s) m))
| trans_mul : forall (c:(list instr)) (p:Z), (codeseq_at c p
(Init.Datatypes.cons Imul Init.Datatypes.nil)) -> forall (n1:Z) (n2:Z)
(s:(list Z)) (m:(map id Z)), (transition c (VMS p
(Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m) (VMS (p + 1%Z)%Z
(Init.Datatypes.cons (n1 * n2)%Z s) m))
| trans_beq : forall (c:(list instr)) (p1:Z) (ofs:Z), (codeseq_at c p1
(ibeq ofs)) -> forall (s:(list Z)) (m:(map id Z)) (n1:Z) (n2:Z),
(n1 = n2) -> (transition c (VMS p1
(Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)
(VMS ((p1 + 1%Z)%Z + ofs)%Z s m))
| trans_beq1 : forall (c:(list instr)) (p1:Z) (ofs:Z), (codeseq_at c p1
(ibeq ofs)) -> forall (s:(list Z)) (m:(map id Z)) (n1:Z) (n2:Z),
(~ (n1 = n2)) -> (transition c (VMS p1
(Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)
(VMS (p1 + 1%Z)%Z s m))
| trans_bne : forall (c:(list instr)) (p1:Z) (ofs:Z), (codeseq_at c p1
(ibne ofs)) -> forall (s:(list Z)) (m:(map id Z)) (n1:Z) (n2:Z),
(n1 = n2) -> (transition c (VMS p1
(Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)
(VMS (p1 + 1%Z)%Z s m))
| trans_bne1 : forall (c:(list instr)) (p1:Z) (ofs:Z), (codeseq_at c p1
(ibne ofs)) -> forall (s:(list Z)) (m:(map id Z)) (n1:Z) (n2:Z),
(~ (n1 = n2)) -> (transition c (VMS p1
(Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)
(VMS ((p1 + 1%Z)%Z + ofs)%Z s m))
| trans_ble : forall (c:(list instr)) (p1:Z) (ofs:Z), (codeseq_at c p1
(ible ofs)) -> forall (s:(list Z)) (m:(map id Z)) (n1:Z) (n2:Z),
(n1 <= n2)%Z -> (transition c (VMS p1
(Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)
(VMS ((p1 + 1%Z)%Z + ofs)%Z s m))
| trans_ble1 : forall (c:(list instr)) (p1:Z) (ofs:Z), (codeseq_at c p1
(ible ofs)) -> forall (s:(list Z)) (m:(map id Z)) (n1:Z) (n2:Z),
(~ (n1 <= n2)%Z) -> (transition c (VMS p1
(Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)
(VMS (p1 + 1%Z)%Z s m))
| trans_bgt : forall (c:(list instr)) (p1:Z) (ofs:Z), (codeseq_at c p1
(ibgt ofs)) -> forall (s:(list Z)) (m:(map id Z)) (n1:Z) (n2:Z),
(n1 <= n2)%Z -> (transition c (VMS p1
(Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)
(VMS (p1 + 1%Z)%Z s m))
| trans_bgt1 : forall (c:(list instr)) (p1:Z) (ofs:Z), (codeseq_at c p1
(ibgt ofs)) -> forall (s:(list Z)) (m:(map id Z)) (n1:Z) (n2:Z),
(~ (n1 <= n2)%Z) -> (transition c (VMS p1
(Init.Datatypes.cons n2 (Init.Datatypes.cons n1 s)) m)
(VMS ((p1 + 1%Z)%Z + ofs)%Z s m))
| trans_branch : forall (c:(list instr)) (p:Z) (ofs:Z), (codeseq_at c p
(ibranch ofs)) -> forall (m:(map id Z)) (s:(list Z)), (transition c
(VMS p s m) (VMS ((p + 1%Z)%Z + ofs)%Z s m)).
(* Why3 assumption *)
Inductive trans :=
| TransZero : trans
| TransOne : machine_state -> trans -> trans.
Axiom trans_WhyType : WhyType trans.
Existing Instance trans_WhyType.
(* Why3 assumption *)
Fixpoint transition_star_proof (p:(list instr)) (s1:machine_state)
(s3:machine_state) (pi:trans) {struct pi}: Prop :=
match pi with
| TransZero => (s1 = s3)
| (TransOne s2 pi') => (transition p s1 s2) /\ (transition_star_proof p s2
s3 pi')
end.
Parameter transition_star: (list instr) -> machine_state -> machine_state ->
Prop.
Axiom transition_star_def : forall (p:(list instr)) (s1:machine_state)
(s3:machine_state), (transition_star p s1 s3) <-> exists pi:trans,
(transition_star_proof p s1 s3 pi).
Axiom transZero : forall (p:(list instr)) (s:machine_state), (transition_star
p s s).
Axiom transOne : forall (p:(list instr)) (s1:machine_state)
(s2:machine_state) (s3:machine_state), (transition p s1 s2) ->
((transition_star p s2 s3) -> (transition_star p s1 s3)).
Axiom transition_star_one : forall (p:(list instr)) (s1:machine_state)
(s2:machine_state), (transition p s1 s2) -> (transition_star p s1 s2).
Axiom trans_star : forall (p:(list instr)) (s1:machine_state)
(s2:machine_state) (s3:machine_state) (pi:trans), ((transition_star_proof p
s1 s2 pi) /\ (transition_star p s2 s3)) -> (transition_star p s1 s3).
Axiom transition_star_transitive : forall (p:(list instr)) (s1:machine_state)
(s2:machine_state) (s3:machine_state), (transition_star p s1 s2) ->
((transition_star p s2 s3) -> (transition_star p s1 s3)).
(* Why3 assumption *)
Definition vm_terminates (c:(list instr)) (mi:(map id Z)) (mf:(map id
Z)): Prop := exists p:Z, (codeseq_at c p
(Init.Datatypes.cons Ihalt Init.Datatypes.nil)) /\ (transition_star c
(VMS 0%Z Init.Datatypes.nil mi) (VMS p Init.Datatypes.nil mf)).
Axiom func : forall (a:Type) (b:Type), Type.
Parameter func_WhyType : forall (a:Type) {a_WT:WhyType a}
(b:Type) {b_WT:WhyType b}, WhyType (func a b).
Existing Instance func_WhyType.
(* Why3 assumption *)
Definition pred (a:Type) := (func a bool).
Parameter infix_at: forall {a:Type} {a_WT:WhyType a}
{b:Type} {b_WT:WhyType b}, (func a b) -> a -> b.
(* Why3 assumption *)
Definition fst {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b} (p:(a*
b)%type): a := match p with
| (x, _) => x
end.
(* Why3 assumption *)
Definition snd {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b} (p:(a*
b)%type): b := match p with
| (_, y) => y
end.
(* Why3 assumption *)
Definition pred1 := (func machine_state bool).
(* Why3 assumption *)
Definition post := (func machine_state (func machine_state bool)).
(* Why3 assumption *)
Inductive hl
(a:Type) :=
| mk_hl : (list instr) -> (func a (func Z (func machine_state bool))) ->
(func a (func Z (func machine_state (func machine_state bool)))) -> hl
a.
Axiom hl_WhyType : forall (a:Type) {a_WT:WhyType a}, WhyType (hl a).
Existing Instance hl_WhyType.
Implicit Arguments mk_hl [[a]].
(* Why3 assumption *)
Definition post1 {a:Type} {a_WT:WhyType a} (v:(hl a)): (func a (func Z (func
machine_state (func machine_state bool)))) :=
match v with
| (mk_hl x x1 x2) => x2
end.
(* Why3 assumption *)
Definition pre {a:Type} {a_WT:WhyType a} (v:(hl a)): (func a (func Z (func
machine_state bool))) := match v with
| (mk_hl x x1 x2) => x1
end.
(* Why3 assumption *)
Definition code1 {a:Type} {a_WT:WhyType a} (v:(hl a)): (list instr) :=
match v with
| (mk_hl x x1 x2) => x
end.
(* Why3 assumption *)
Inductive wp
(a:Type) :=
| mk_wp : (list instr) -> (func a (func Z (func (func machine_state bool)
(func machine_state bool)))) -> wp a.
Axiom wp_WhyType : forall (a:Type) {a_WT:WhyType a}, WhyType (wp a).
Existing Instance wp_WhyType.
Implicit Arguments mk_wp [[a]].
(* Why3 assumption *)
Definition wp1 {a:Type} {a_WT:WhyType a} (v:(wp a)): (func a (func Z (func
(func machine_state bool) (func machine_state bool)))) :=
match v with
| (mk_wp x x1) => x1
end.
(* Why3 assumption *)
Definition wcode {a:Type} {a_WT:WhyType a} (v:(wp a)): (list instr) :=
match v with
| (mk_wp x x1) => x
end.
(* Why3 assumption *)
Definition contextual_irrelevance (c:(list instr)) (p:Z) (ms1:machine_state)
(ms2:machine_state): Prop := forall (c_global:(list instr)), (codeseq_at
c_global p c) -> (transition_star c_global ms1 ms2).
(* Why3 assumption *)
Definition hl_correctness {a:Type} {a_WT:WhyType a} (cs:(hl a)): Prop :=
forall (x:a) (p:Z) (ms:machine_state),
((infix_at (infix_at (infix_at (pre cs) x) p) ms) = true) ->
exists ms':machine_state,
((infix_at (infix_at (infix_at (infix_at (post1 cs) x) p) ms)
ms') = true) /\ (contextual_irrelevance (code1 cs) p ms ms').
(* Why3 assumption *)
Definition wp_correctness {a:Type} {a_WT:WhyType a} (code2:(wp a)): Prop :=
forall (x:a) (p:Z) (post2:(func machine_state bool)) (ms:machine_state),
((infix_at (infix_at (infix_at (infix_at (wp1 code2) x) p) post2)
ms) = true) -> exists ms':machine_state, ((infix_at post2 ms') = true) /\
(contextual_irrelevance (wcode code2) p ms ms').
Parameter seq_wp: forall {a:Type} {a_WT:WhyType a}, (wp a) -> (wp (a*
machine_state)%type) -> (func a (func Z (func (func machine_state bool)
(func machine_state bool)))).
Axiom seq_wp_def : forall {a:Type} {a_WT:WhyType a}, forall (s1:(wp a))
(s2:(wp (a* machine_state)%type)) (x:a) (p:Z) (q:(func machine_state bool))
(ms:machine_state), ((infix_at (infix_at (infix_at (infix_at (seq_wp s1 s2)
x) p) q) ms) = (infix_at (infix_at (infix_at (infix_at (wp1 s1) x) p)
(infix_at (infix_at (infix_at (wp1 s2) (x, ms))
(p + (list.Length.length (wcode s1)))%Z) q)) ms)).
Axiom seq_wp_lemma : forall {a:Type} {a_WT:WhyType a}, forall (s1:(wp a))
(s2:(wp (a* machine_state)%type)) (x:a) (p:Z) (q:(func machine_state bool))
(ms:machine_state), ((infix_at (infix_at (infix_at (infix_at (seq_wp s1 s2)
x) p) q) ms) = (infix_at (infix_at (infix_at (infix_at (wp1 s1) x) p)
(infix_at (infix_at (infix_at (wp1 s2) (x, ms))
(p + (list.Length.length (wcode s1)))%Z) q)) ms)).
Parameter fork_wp: forall {a:Type} {a_WT:WhyType a}, (wp a) -> (func a (func
Z (func machine_state bool))) -> (func a (func Z (func (func machine_state
bool) (func machine_state bool)))).
Axiom fork_wp_def : forall {a:Type} {a_WT:WhyType a}, forall (s2:(wp a))
(exn:(func a (func Z (func machine_state bool)))) (x:a) (p:Z) (q:(func
machine_state bool)) (ms:machine_state),
((infix_at (infix_at (infix_at (infix_at (fork_wp s2 exn) x) p) q)
ms) = true) <-> ((((infix_at (infix_at (infix_at exn x) p) ms) = true) ->
((infix_at q ms) = true)) /\ ((~ ((infix_at (infix_at (infix_at exn x) p)
ms) = true)) -> ((infix_at (infix_at (infix_at (infix_at (wp1 s2) x) p) q)
ms) = true))).
Axiom fork_wp_lemma : forall {a:Type} {a_WT:WhyType a}, forall (s2:(wp a))
(exn:(func a (func Z (func machine_state bool)))) (x:a) (p:Z) (q:(func
machine_state bool)) (ms:machine_state),
((infix_at (infix_at (infix_at (infix_at (fork_wp s2 exn) x) p) q)
ms) = true) <-> ((((infix_at (infix_at (infix_at exn x) p) ms) = true) ->
((infix_at q ms) = true)) /\ ((~ ((infix_at (infix_at (infix_at exn x) p)
ms) = true)) -> ((infix_at (infix_at (infix_at (infix_at (wp1 s2) x) p) q)
ms) = true))).
Parameter towp_wp: forall {a:Type} {a_WT:WhyType a}, (hl a) -> (func a (func
Z (func (func machine_state bool) (func machine_state bool)))).
Axiom towp_wp_def : forall {a:Type} {a_WT:WhyType a}, forall (c:(hl a)) (x:a)
(p:Z) (q:(func machine_state bool)) (ms:machine_state),
((infix_at (infix_at (infix_at (infix_at (towp_wp c) x) p) q)
ms) = true) <-> (((infix_at (infix_at (infix_at (pre c) x) p)
ms) = true) /\ forall (ms':machine_state),
((infix_at (infix_at (infix_at (infix_at (post1 c) x) p) ms)
ms') = true) -> ((infix_at q ms') = true)).
Axiom towp_wp_lemma : forall {a:Type} {a_WT:WhyType a}, forall (c:(hl a))
(x:a) (p:Z) (q:(func machine_state bool)) (ms:machine_state),
((infix_at (infix_at (infix_at (infix_at (towp_wp c) x) p) q)
ms) = true) <-> (((infix_at (infix_at (infix_at (pre c) x) p)
ms) = true) /\ forall (ms':machine_state),
((infix_at (infix_at (infix_at (infix_at (post1 c) x) p) ms)
ms') = true) -> ((infix_at q ms') = true)).
Parameter trivial_pre: forall {a:Type} {a_WT:WhyType a}, (func a (func Z
(func machine_state bool))).
Axiom trivial_pre_def : forall {a:Type} {a_WT:WhyType a}, forall (x:a) (p:Z)
(ms:machine_state), ((infix_at (infix_at (infix_at (trivial_pre : (func a
(func Z (func machine_state bool)))) x) p) ms) = true) <->
match ms with
| (VMS p' _ _) => (p = p')
end.
Axiom trivial_pre_lemma : forall {a:Type} {a_WT:WhyType a}, forall (x:a)
(p:Z) (ms:machine_state),
((infix_at (infix_at (infix_at (trivial_pre : (func a (func Z (func
machine_state bool)))) x) p) ms) = true) <->
match ms with
| (VMS p' _ _) => (p = p')
end.
(* Why3 assumption *)
Inductive acc {a:Type} {a_WT:WhyType a}: (func a (func a bool)) -> a ->
Prop :=
| Acc : forall (r:(func a (func a bool))) (x:a), (forall (y:a),
((infix_at (infix_at r y) x) = true) -> (acc r y)) -> (acc r x).
Parameter loop_preservation: forall {a:Type} {a_WT:WhyType a}, (func a (func
Z (func machine_state bool))) -> (func a (func Z (func machine_state (func
machine_state bool)))) -> (func a (func Z (func machine_state bool))) ->
(func a (func Z (func machine_state (func machine_state bool)))).
Axiom loop_preservation_def : forall {a:Type} {a_WT:WhyType a},
forall (inv:(func a (func Z (func machine_state bool)))) (var:(func a (func
Z (func machine_state (func machine_state bool))))) (post2:(func a (func Z
(func machine_state bool)))) (x:a) (p:Z) (ms:machine_state)
(ms':machine_state),
((infix_at (infix_at (infix_at (infix_at (loop_preservation inv var post2)
x) p) ms) ms') = true) <-> ((((infix_at (infix_at (infix_at inv x) p)
ms') = true) /\ ((infix_at (infix_at (infix_at (infix_at var x) p) ms')
ms) = true)) \/ ((infix_at (infix_at (infix_at post2 x) p) ms') = true)).
Parameter forget_old: forall {a:Type} {a_WT:WhyType a}, (func a (func Z (func
machine_state bool))) -> (func a (func Z (func machine_state (func
machine_state bool)))).
Axiom forget_old_def : forall {a:Type} {a_WT:WhyType a}, forall (post2:(func
a (func Z (func machine_state bool)))) (x:a) (p:Z) (ms:machine_state),
((infix_at (infix_at (infix_at (forget_old post2) x) p)
ms) = (infix_at (infix_at post2 x) p)).
Require Import Why3.
Ltac ae := why3 "Alt-Ergo,0.95.2,".
Ltac cvc := why3 "CVC4,1.4,".
(* Why3 goal *)
Theorem WP_parameter_make_loop_hl : forall {a:Type} {a_WT:WhyType a},
forall (c:(list instr)) (c1:(func a (func Z (func machine_state bool))))
(c2:(func a (func Z (func machine_state (func machine_state bool)))))
(inv:(func a (func Z (func machine_state bool)))) (var:(func a (func Z
(func machine_state (func machine_state bool))))) (post2:(func a (func Z
(func machine_state bool)))), ((hl_correctness (mk_hl c c1 c2)) /\
((forall (x:a) (p:Z) (ms:machine_state), ((infix_at (infix_at (infix_at inv
x) p) ms) = true) -> (acc (infix_at (infix_at var x) p) ms)) /\
((c1 = inv) /\ (c2 = (loop_preservation inv var post2))))) ->
(hl_correctness (mk_hl c inv (forget_old post2))).
(* Why3 intros a a_WT c c1 c2 inv var post2 (h1,(h2,(h3,h4))). *)
intros a a_WT c c1 c2 inv var post2 (h1,(h2,(h3,h4))).
unfold hl_correctness in *.
intros.
simpl in *.
remember H as H' eqn : eqn; clear eqn.
apply h2 in H'.
remember (infix_at (infix_at var x) p) as R eqn : eqR.
induction H'.
ae.
Qed.
......@@ -2,26 +2,26 @@
<!DOCTYPE why3session PUBLIC "-//Why3//proof session v5//EN"
"http://why3.lri.fr/why3session.dtd">
<why3session shape_version="4">
<prover id="0" name="CVC4" version="1.4" timelimit="5" memlimit="1000"/>
<prover id="1" name="Coq" version="8.4pl2" timelimit="5" memlimit="1000"/>
<prover id="0" name="CVC3" version="2.4.1" timelimit="5" memlimit="1000"/>
<prover id="1" name="CVC4" version="1.4" timelimit="5" memlimit="1000"/>
<prover id="2" name="Alt-Ergo" version="0.95.2" timelimit="5" memlimit="1000"/>
<prover id="3" name="CVC4" version="1.3" timelimit="5" memlimit="1000"/>
<file name="../logic.mlw" expanded="true">
<theory name="Compiler_logic" sum="230cac94b76c20b9d89901f39b823027" expanded="true">
<theory name="Compiler_logic" sum="01f7b599ba74d4eafec3be887b1930e1" expanded="true">
<goal name="seq_wp_lemma">
<proof prover="2"><result status="valid" time="0.04"/></proof>
</goal>
<goal name="WP_parameter infix ~" expl="VC for infix ~" expanded="true">
<transf name="split_goal_wp" expanded="true">
<goal name="WP_parameter infix ~.1" expl="1. assertion" expanded="true">
<transf name="split_goal_wp" expanded="true">
<goal name="WP_parameter infix ~" expl="VC for infix ~">
<transf name="split_goal_wp">
<goal name="WP_parameter infix ~.1" expl="1. assertion">
<transf name="split_goal_wp">
<goal name="WP_parameter infix ~.1.1" expl="1. assertion">
<proof prover="2"><result status="valid" time="0.14"/></proof>
</goal>
<goal name="WP_parameter infix ~.1.2" expl="2. assertion" expanded="true">
<transf name="simplify_trivial_quantification_in_goal" expanded="true">
<goal name="WP_parameter infix ~.1.2.1" expl="1." expanded="true">
<transf name="compute_specified" expanded="true">
<goal name="WP_parameter infix ~.1.2" expl="2. assertion">
<transf name="simplify_trivial_quantification_in_goal">
<goal name="WP_parameter infix ~.1.2.1" expl="1.">
<transf name="compute_specified">
<goal name="WP_parameter infix ~.1.2.1.1" expl="1.">
<proof prover="2"><result status="valid" time="0.23"/></proof>
</goal>
......@@ -39,13 +39,13 @@
</goal>
</transf>
</goal>
<goal name="fork_wp_lemma" expanded="true">
<goal name="fork_wp_lemma">
<proof prover="2"><result status="valid" time="0.03"/></proof>
</goal>
<goal name="WP_parameter infix %" expl="VC for infix %">
<transf name="split_goal_wp">
<goal name="WP_parameter infix %.1" expl="1. postcondition">
<proof prover="0"><result status="valid" time="0.10"/></proof>
<proof prover="1"><result status="valid" time="0.10"/></proof>
</goal>
</transf>
</goal>
......@@ -58,7 +58,7 @@
<goal name="WP_parameter hoare" expl="VC for hoare">
<transf name="split_goal_wp">
<goal name="WP_parameter hoare.1" expl="1. postcondition">
<proof prover="0"><result status="valid" time="0.07"/></proof>
<proof prover="1"><result status="valid" time="0.07"/></proof>
<proof prover="3"><result status="valid" time="0.07"/></proof>
</goal>
</transf>
......@@ -66,10 +66,32 @@
<goal name="trivial_pre_lemma">
<proof prover="2"><result status="valid" time="0.04"/></proof>
</goal>
<goal name="WP_parameter make_loop_hl" expl="VC for make_loop_hl" expanded="true">
<transf name="split_goal_wp" expanded="true">
<goal name="WP_parameter make_loop_hl.1" expl="1. postcondition" expanded="true">
<proof prover="1" edited="logic_Compiler_logic_WP_parameter_make_loop_hl_1.v"><result status="valid" time="1.69"/></proof>
<goal name="WP_parameter make_loop_hl" expl="VC for make_loop_hl">
<transf name="split_goal_wp">
<goal name="WP_parameter make_loop_hl.1" expl="1. assertion">
<transf name="split_goal_wp">
<goal name="WP_parameter make_loop_hl.1.1" expl="1. assertion">
<proof prover="2"><result status="valid" time="0.04"/></proof>
</goal>
<goal name="WP_parameter make_loop_hl.1.2" expl="2. assertion">
<transf name="induction_pr">
<goal name="WP_parameter make_loop_hl.1.2.1" expl="1.">
<transf name="simplify_trivial_quantification_in_goal">
<goal name="WP_parameter make_loop_hl.1.2.1.1" expl="1.">
<transf name="compute_specified">
<goal name="WP_parameter make_loop_hl.1.2.1.1.1" expl="1.">
<proof prover="0"><result status="valid" time="0.09"/></proof>
</goal>
</transf>
</goal>
</transf>
</goal>
</transf>
</goal>
</transf>
</goal>
<goal name="WP_parameter make_loop_hl.2" expl="2. postcondition">
<proof prover="2"><result status="valid" time="0.06"/></proof>
</goal>
</transf>
</goal>
......
No preview for this file type
......@@ -61,25 +61,29 @@ module VM_arith_instr_spec
constant ibinop_pre : 'a -> pos -> pred =
\x p ms . exists n1 n2 s m. ms = VMS p (push n2 (push n1 s)) m
lemma ibinop_pre_lemma:
meta rewrite_def function ibinop_pre
(*lemma ibinop_pre_lemma:
forall x:'a, p ms. ibinop_pre x p ms =
exists n1 n2 s m. ms = VMS p (push n2 (push n1 s)) m
meta rewrite prop ibinop_pre_lemma
meta rewrite prop ibinop_pre_lemma*)
function ibinop_post (op : binop) : 'a -> pos -> machine_state -> pred =
\x p ms ms'. forall n1 n2 s m.
ms = VMS p (push n2 (push n1 s)) m ->
ms' = VMS (p + 1) (push (op n1 n2) s) m
lemma ibinop_post_lemma:
meta rewrite_def function ibinop_post
(*lemma ibinop_post_lemma:
forall op, x:'a, p ms ms'.
ibinop_post op x p ms ms' =
forall n1 n2 s m.
ms = VMS p (push n2 (push n1 s)) m ->
ms' = VMS (p + 1) (push (op n1 n2) s) m
meta rewrite prop ibinop_post_lemma
meta rewrite prop ibinop_post_lemma*)
let create_binop (code_binop: code) (ghost op : binop) : hl 'a
requires {
......@@ -181,6 +185,8 @@ module VM_bool_instr_spec
(cond n1 n2 -> ms' = VMS (p + ofs + 1) s m) /\
(not cond n1 n2 -> ms' = VMS (p+1) s m)
meta rewrite_def function icjump_post
(*
lemma icjump_post_lemma:
forall cond ofs, x:'a, p ms ms'.
icjump_post cond ofs x p ms ms' =
......@@ -189,25 +195,29 @@ module VM_bool_instr_spec
(cond n1 n2 -> ms' = VMS (p + ofs + 1) s m) /\
(not cond n1 n2 -> ms' = VMS (p+1) s m)
meta rewrite prop icjump_post_lemma
meta rewrite prop icjump_post_lemma *)
(* binary Boolean operators specification (Ibeq, Ibne, Ible, Ibgt) *)
constant beq : cond = \x y. x = y
lemma beq_lemma : forall x y. beq x y = (x = y)
meta rewrite prop beq_lemma
meta rewrite_def function beq
(*lemma beq_lemma : forall x y. beq x y = (x = y)
meta rewrite prop beq_lemma*)
constant bne : cond = \x y. x <> y
lemma bne_lemma : forall x y. bne x y = (x <> y)
meta rewrite prop bne_lemma
meta rewrite_def function bne
(*lemma bne_lemma : forall x y. bne x y = (x <> y)
meta rewrite prop bne_lemma*)
constant ble : cond = \x y. x <= y
lemma ble_lemma : forall x y. ble x y = (x <= y)
meta rewrite prop ble_lemma
meta rewrite_def function ble
(*lemma ble_lemma : forall x y. ble x y = (x <= y)
meta rewrite prop ble_lemma*)
constant bgt : cond = \x y. x > y
lemma bgt_lemma : forall x y. bgt x y = (x > y)
meta rewrite prop bgt_lemma
meta rewrite_def function bgt
(*lemma bgt_lemma : forall x y. bgt x y = (x > y)
meta rewrite prop bgt_lemma*)
let create_cjump (code_cond:code) (ghost cond:cond) (ghost ofs:pos) : hl 'a
......@@ -260,25 +270,29 @@ let ibgtf (ofs : pos) : hl 'a
constant isetvar_pre : 'a -> pos -> pred =
\x p ms . exists n s m. ms = VMS p (push n s) m
lemma isetvar_pre_lemma:
meta rewrite_def function isetvar_pre
(* lemma isetvar_pre_lemma:
forall x:'a, p ms. isetvar_pre x p ms =
exists n s m. ms = VMS p (push n s) m
meta rewrite prop isetvar_pre_lemma
meta rewrite prop isetvar_pre_lemma *)
function isetvar_post (x:id) : 'a -> pos -> post =
\a p ms ms'. forall s n m.
ms = VMS p (push n s) m ->
ms' = VMS (p+1) s m[ x <- n]
lemma isetvar_post_lemma:
meta rewrite_def function isetvar_post
(* lemma isetvar_post_lemma:
forall x, a: 'a, p ms ms'.
isetvar_post x a p ms ms' =
forall s n m.
ms = VMS p (push n s) m ->
ms' = VMS (p+1) s m[ x <- n]
meta rewrite prop isetvar_post_lemma
meta rewrite prop isetvar_post_lemma *)
let isetvarf (x: id) : hl 'a
ensures { result.pre = isetvar_pre /\ result.post = isetvar_post x }
......
......@@ -7,7 +7,7 @@
<prover id="2" name="Alt-Ergo" version="0.95.2" timelimit="5" memlimit="1000"/>
<prover id="3" name="CVC4" version="1.3" timelimit="5" memlimit="1000"/>
<file name="../specs.mlw" expanded="true">
<theory name="VM_arith_instr_spec" sum="89f02e7f809e0639fc6e34941844fd75">
<theory name="VM_arith_instr_spec" sum="12ca8ba8424f485f6e8e0827b562a356">
<goal name="WP_parameter iconstf" expl="VC for iconstf">
<transf name="split_goal_wp">
<goal name="WP_parameter iconstf.1" expl="1. assertion">
......@@ -72,12 +72,6 @@
</goal>
</transf>
</goal>
<goal name="ibinop_pre_lemma">
<proof prover="2"><result status="valid" time="0.04"/></proof>
</goal>
<goal name="ibinop_post_lemma">
<proof prover="2"><result status="valid" time="0.05"/></proof>
</goal>
<goal name="WP_parameter create_binop" expl="VC for create_binop">
<transf name="split_goal_wp">
<goal name="WP_parameter create_binop.1" expl="1. assertion">
......@@ -117,7 +111,7 @@
<proof prover="2"><result status="valid" time="0.05"/></proof>
</goal>
</theory>
<theory name="VM_bool_instr_spec" sum="32274015eab190519cb2670676a7edbf" expanded="true">
<theory name="VM_bool_instr_spec" sum="c8dd844b4da700cdce9948913b5d62fd" expanded="true">
<goal name="WP_parameter inil" expl="VC for inil">
<transf name="split_goal_wp">
<goal name="WP_parameter inil.1" expl="1. postcondition">
......@@ -164,21 +158,6 @@
</goal>
</transf>
</goal>
<goal name="icjump_post_lemma">
<proof prover="2"><result status="valid" time="0.05"/></proof>
</goal>
<goal name="beq_lemma">
<proof prover="2"><result status="valid" time="0.05"/></proof>
</goal>
<goal name="bne_lemma">
<proof prover="2"><result status="valid" time="0.05"/></proof>
</goal>
<goal name="ble_lemma">
<proof prover="2"><result status="valid" time="0.05"/></proof>
</goal>
<goal name="bgt_lemma">
<proof prover="2"><result status="valid" time="0.05"/></proof>
</goal>
<goal name="WP_parameter create_cjump" expl="VC for create_cjump">
<transf name="split_goal_wp">
<goal name="WP_parameter create_cjump.1" expl="1. assertion">
......@@ -230,12 +209,6 @@
<goal name="WP_parameter ibgtf" expl="VC for ibgtf">
<proof prover="2"><result status="valid" time="0.08"/></proof>
</goal>
<goal name="isetvar_pre_lemma">
<proof prover="2"><result status="valid" time="0.05"/></proof>
</goal>
<goal name="isetvar_post_lemma">
<proof prover="2"><result status="valid" time="0.07"/></proof>
</goal>
<goal name="WP_parameter isetvarf" expl="VC for isetvarf">
<transf name="split_goal_wp">
<goal name="WP_parameter isetvarf.1" expl="1. assertion">
......
No preview for this file type
......@@ -5,6 +5,20 @@ module ReflTransClosure
type state
predicate transition parameter state state
inductive transition_star parameter (x y: state) =
| Refl: forall p x. transition_star p x x
| Step: forall p x y z.
transition p x y -> transition_star p y z -> transition_star p x z
lemma transition_star_one:
forall p:parameter, s1 s2:state. transition p s1 s2 -> transition_star p s1 s2
lemma transition_star_transitive:
forall p: parameter, s1 s2 s3:state.
transition_star p s1 s2 -> transition_star p s2 s3 -> transition_star p s1 s3
(*
(* proof object *)
type trans =
| TransZero
......@@ -50,6 +64,8 @@ module ReflTransClosure
(*lemma transOneSnoc :
forall p: parameter, s1 s2 s3: state.
transition_star p s1 s2 -> transition p s2 s3 -> transition_star p s1 s3*)
*)
end
......
......@@ -2,33 +2,32 @@
<!DOCTYPE why3session PUBLIC "-//Why3//proof session v5//EN"
"http://why3.lri.fr/why3session.dtd">
<why3session shape_version="4">
<prover id="0" name="Z3" version="4.3.1" timelimit="5" memlimit="1000"/>
<prover id="1" name="Alt-Ergo" version="0.95.2" timelimit="5" memlimit="1000"/>
<prover id="2" name="Vampire" version="0.6" timelimit="5" memlimit="1000"/>
<prover id="0" name="CVC3" version="2.4.1" timelimit="5" memlimit="1000"/>
<prover id="1" name="Z3" version="4.3.1" timelimit="5" memlimit="1000"/>
<prover id="2" name="Alt-Ergo" version="0.95.2" timelimit="5" memlimit="1000"/>
<prover id="3" name="Vampire" version="0.6" timelimit="5" memlimit="1000"/>
<file name="../vm.mlw" expanded="true">
<theory name="ReflTransClosure" sum="858a40c1dde55e8befdaee8ae005e423">
<goal name="transZero">
<theory name="ReflTransClosure" sum="41f42e2f3926adcb698b8698151507f0" expanded="true">
<goal name="transition_star_one" expanded="true">
<proof prover="0"><result status="valid" time="0.00"/></proof>
</goal>
<goal name="transOne">
<proof prover="0"><result status="valid" time="0.04"/></proof>
</goal>
<goal name="transition_star_one">
<proof prover="0"><result status="valid" time="0.00"/></proof>
</goal>
<goal name="WP_parameter trans_star" expl="VC for trans_star">
<proof prover="1"><result status="valid" time="0.01"/></proof>
</goal>
<goal name="transition_star_transitive">
<proof prover="1"><result status="valid" time="0.02"/></proof>
<goal name="transition_star_transitive" expanded="true">
<transf name="induction_pr" expanded="true">
<goal name="transition_star_transitive.1" expl="1.">
<proof prover="2"><result status="valid" time="0.01"/></proof>
</goal>
<goal name="transition_star_transitive.2" expl="2." expanded="true">
<proof prover="0"><result status="valid" time="0.01"/></proof>
</goal>
</transf>
</goal>
</theory>
<theory name="Vm" sum="030cc7eb9172e9aca56370af11cf7236" expanded="true">
<theory name="Vm" sum="ff0e2f36dbba20610ead4bbdaf34a210" expanded="true">
<goal name="codeseq_at_app_right">
<proof prover="2"><result status="valid" time="0.28"/></proof>
<proof prover="3"><result status="valid" time="0.28"/></proof>
</goal>
<goal name="codeseq_at_app_left">
<proof prover="0"><result status="valid" time="0.01"/></proof>
<proof prover="1"><result status="valid" time="0.01"/></proof>
</goal>
</theory>
</file>
......
No preview for this file type
0% Loading or .
You are about to add 0 people to the discussion. Proceed with caution.
Please register or to comment