(* This file is generated by Why3's Coq driver *) (* Beware! Only edit allowed sections below *) Require Import BuiltIn. Require BuiltIn. Require HighOrd. Require int.Int. Require map.Map. Require bool.Bool. (* Why3 assumption *) Inductive datatype := | Tint : datatype | Tbool : datatype. Axiom datatype_WhyType : WhyType datatype. Existing Instance datatype_WhyType. (* Why3 assumption *) Inductive operator := | Oplus : operator | Ominus : operator | Omult : operator | Ole : operator. Axiom operator_WhyType : WhyType operator. Existing Instance operator_WhyType. (* Why3 assumption *) Definition ident := Numbers.BinNums.Z. (* Why3 assumption *) Inductive term := | Tconst : Numbers.BinNums.Z -> term | Tvar : Numbers.BinNums.Z -> term | Tderef : Numbers.BinNums.Z -> term | Tbin : term -> operator -> term -> term. Axiom term_WhyType : WhyType term. Existing Instance term_WhyType. (* Why3 assumption *) Inductive fmla := | Fterm : term -> fmla | Fand : fmla -> fmla -> fmla | Fnot : fmla -> fmla | Fimplies : fmla -> fmla -> fmla | Flet : Numbers.BinNums.Z -> term -> fmla -> fmla | Fforall : Numbers.BinNums.Z -> datatype -> fmla -> fmla. Axiom fmla_WhyType : WhyType fmla. Existing Instance fmla_WhyType. (* Why3 assumption *) Inductive value := | Vint : Numbers.BinNums.Z -> value | Vbool : Init.Datatypes.bool -> value. Axiom value_WhyType : WhyType value. Existing Instance value_WhyType. (* Why3 assumption *) Definition env := Numbers.BinNums.Z -> value. Parameter eval_bin: value -> operator -> value -> value. Axiom eval_bin_def : forall (x:value) (op:operator) (y:value), match (x, y) with | (Vint x1, Vint y1) => match op with | Oplus => ((eval_bin x op y) = (Vint (x1 + y1)%Z)) | Ominus => ((eval_bin x op y) = (Vint (x1 - y1)%Z)) | Omult => ((eval_bin x op y) = (Vint (x1 * y1)%Z)) | Ole => ((x1 <= y1)%Z -> ((eval_bin x op y) = (Vbool Init.Datatypes.true))) /\ (~ (x1 <= y1)%Z -> ((eval_bin x op y) = (Vbool Init.Datatypes.false))) end | (_, _) => ((eval_bin x op y) = (Vbool Init.Datatypes.false)) end. (* Why3 assumption *) Fixpoint eval_term (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (t:term) {struct t}: value := match t with | Tconst n => Vint n | Tvar id => pi id | Tderef id => sigma id | Tbin t1 op t2 => eval_bin (eval_term sigma pi t1) op (eval_term sigma pi t2) end. (* Why3 assumption *) Fixpoint eval_fmla (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (f:fmla) {struct f}: Prop := match f with | Fterm t => ((eval_term sigma pi t) = (Vbool Init.Datatypes.true)) | Fand f1 f2 => eval_fmla sigma pi f1 /\ eval_fmla sigma pi f2 | Fnot f1 => ~ eval_fmla sigma pi f1 | Fimplies f1 f2 => eval_fmla sigma pi f1 -> eval_fmla sigma pi f2 | Flet x t f1 => eval_fmla sigma (map.Map.set pi x (eval_term sigma pi t)) f1 | Fforall x Tint f1 => forall (n:Numbers.BinNums.Z), eval_fmla sigma (map.Map.set pi x (Vint n)) f1 | Fforall x Tbool f1 => forall (b:Init.Datatypes.bool), eval_fmla sigma (map.Map.set pi x (Vbool b)) f1 end. Parameter subst_term: term -> Numbers.BinNums.Z -> Numbers.BinNums.Z -> term. Axiom subst_term_def : forall (e:term) (r:Numbers.BinNums.Z) (v:Numbers.BinNums.Z), match e with | Tconst _ => ((subst_term e r v) = e) | Tvar _ => ((subst_term e r v) = e) | Tderef x => ((r = x) -> ((subst_term e r v) = (Tvar v))) /\ (~ (r = x) -> ((subst_term e r v) = e)) | Tbin e1 op e2 => ((subst_term e r v) = (Tbin (subst_term e1 r v) op (subst_term e2 r v))) end. (* Why3 assumption *) Fixpoint fresh_in_term (id:Numbers.BinNums.Z) (t:term) {struct t}: Prop := match t with | Tconst _ => True | Tvar v => ~ (id = v) | Tderef _ => True | Tbin t1 _ t2 => fresh_in_term id t1 /\ fresh_in_term id t2 end. Axiom eval_subst_term : forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (e:term) (x:Numbers.BinNums.Z) (v:Numbers.BinNums.Z), fresh_in_term v e -> ((eval_term sigma pi (subst_term e x v)) = (eval_term (map.Map.set sigma x (pi v)) pi e)). Axiom eval_term_change_free : forall (t:term) (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (id:Numbers.BinNums.Z) (v:value), fresh_in_term id t -> ((eval_term sigma (map.Map.set pi id v) t) = (eval_term sigma pi t)). (* Why3 assumption *) Fixpoint fresh_in_fmla (id:Numbers.BinNums.Z) (f:fmla) {struct f}: Prop := match f with | Fterm e => fresh_in_term id e | (Fand f1 f2)|(Fimplies f1 f2) => fresh_in_fmla id f1 /\ fresh_in_fmla id f2 | Fnot f1 => fresh_in_fmla id f1 | Flet y t f1 => ~ (id = y) /\ fresh_in_term id t /\ fresh_in_fmla id f1 | Fforall y _ f1 => ~ (id = y) /\ fresh_in_fmla id f1 end. (* Why3 assumption *) Fixpoint subst (f:fmla) (x:Numbers.BinNums.Z) (v:Numbers.BinNums.Z) {struct f}: fmla := match f with | Fterm e => Fterm (subst_term e x v) | Fand f1 f2 => Fand (subst f1 x v) (subst f2 x v) | Fnot f1 => Fnot (subst f1 x v) | Fimplies f1 f2 => Fimplies (subst f1 x v) (subst f2 x v) | Flet y t f1 => Flet y (subst_term t x v) (subst f1 x v) | Fforall y ty f1 => Fforall y ty (subst f1 x v) end. Axiom eval_subst : forall (f:fmla) (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (x:Numbers.BinNums.Z) (v:Numbers.BinNums.Z), fresh_in_fmla v f -> eval_fmla sigma pi (subst f x v) <-> eval_fmla (map.Map.set sigma x (pi v)) pi f. Axiom eval_swap : forall (f:fmla) (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (id1:Numbers.BinNums.Z) (id2:Numbers.BinNums.Z) (v1:value) (v2:value), ~ (id1 = id2) -> eval_fmla sigma (map.Map.set (map.Map.set pi id1 v1) id2 v2) f <-> eval_fmla sigma (map.Map.set (map.Map.set pi id2 v2) id1 v1) f. Axiom eval_change_free : forall (f:fmla) (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (id:Numbers.BinNums.Z) (v:value), fresh_in_fmla id f -> eval_fmla sigma (map.Map.set pi id v) f <-> eval_fmla sigma pi f. (* Why3 assumption *) Inductive stmt := | Sskip : stmt | Sassign : Numbers.BinNums.Z -> term -> stmt | Sseq : stmt -> stmt -> stmt | Sif : term -> stmt -> stmt -> stmt | Sassert : fmla -> stmt | Swhile : term -> fmla -> stmt -> stmt. Axiom stmt_WhyType : WhyType stmt. Existing Instance stmt_WhyType. Axiom check_skip : forall (s:stmt), (s = Sskip) \/ ~ (s = Sskip). (* Why3 assumption *) Inductive one_step: (Numbers.BinNums.Z -> value) -> (Numbers.BinNums.Z -> value) -> stmt -> (Numbers.BinNums.Z -> value) -> (Numbers.BinNums.Z -> value) -> stmt -> Prop := | one_step_assign : forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (x:Numbers.BinNums.Z) (e:term), one_step sigma pi (Sassign x e) (map.Map.set sigma x (eval_term sigma pi e)) pi Sskip | one_step_seq : forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (sigma':Numbers.BinNums.Z -> value) (pi':Numbers.BinNums.Z -> value) (i1:stmt) (i1':stmt) (i2:stmt), one_step sigma pi i1 sigma' pi' i1' -> one_step sigma pi (Sseq i1 i2) sigma' pi' (Sseq i1' i2) | one_step_seq_skip : forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (i:stmt), one_step sigma pi (Sseq Sskip i) sigma pi i | one_step_if_true : forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (e:term) (i1:stmt) (i2:stmt), ((eval_term sigma pi e) = (Vbool Init.Datatypes.true)) -> one_step sigma pi (Sif e i1 i2) sigma pi i1 | one_step_if_false : forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (e:term) (i1:stmt) (i2:stmt), ((eval_term sigma pi e) = (Vbool Init.Datatypes.false)) -> one_step sigma pi (Sif e i1 i2) sigma pi i2 | one_step_assert : forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (f:fmla), eval_fmla sigma pi f -> one_step sigma pi (Sassert f) sigma pi Sskip | one_step_while_true : forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (e:term) (inv:fmla) (i:stmt), eval_fmla sigma pi inv -> ((eval_term sigma pi e) = (Vbool Init.Datatypes.true)) -> one_step sigma pi (Swhile e inv i) sigma pi (Sseq i (Swhile e inv i)) | one_step_while_false : forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (e:term) (inv:fmla) (i:stmt), eval_fmla sigma pi inv -> ((eval_term sigma pi e) = (Vbool Init.Datatypes.false)) -> one_step sigma pi (Swhile e inv i) sigma pi Sskip. (* Why3 assumption *) Inductive many_steps: (Numbers.BinNums.Z -> value) -> (Numbers.BinNums.Z -> value) -> stmt -> (Numbers.BinNums.Z -> value) -> (Numbers.BinNums.Z -> value) -> stmt -> Numbers.BinNums.Z -> Prop := | many_steps_refl : forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (i:stmt), many_steps sigma pi i sigma pi i 0%Z | many_steps_trans : forall (sigma1:Numbers.BinNums.Z -> value) (pi1:Numbers.BinNums.Z -> value) (sigma2:Numbers.BinNums.Z -> value) (pi2:Numbers.BinNums.Z -> value) (sigma3:Numbers.BinNums.Z -> value) (pi3:Numbers.BinNums.Z -> value) (i1:stmt) (i2:stmt) (i3:stmt) (n:Numbers.BinNums.Z), one_step sigma1 pi1 i1 sigma2 pi2 i2 -> many_steps sigma2 pi2 i2 sigma3 pi3 i3 n -> many_steps sigma1 pi1 i1 sigma3 pi3 i3 (n + 1%Z)%Z. Axiom steps_non_neg : forall (sigma1:Numbers.BinNums.Z -> value) (pi1:Numbers.BinNums.Z -> value) (sigma2:Numbers.BinNums.Z -> value) (pi2:Numbers.BinNums.Z -> value) (i1:stmt) (i2:stmt) (n:Numbers.BinNums.Z), many_steps sigma1 pi1 i1 sigma2 pi2 i2 n -> (0%Z <= n)%Z. Axiom many_steps_seq : forall (sigma1:Numbers.BinNums.Z -> value) (pi1:Numbers.BinNums.Z -> value) (sigma3:Numbers.BinNums.Z -> value) (pi3:Numbers.BinNums.Z -> value) (i1:stmt) (i2:stmt) (n:Numbers.BinNums.Z), many_steps sigma1 pi1 (Sseq i1 i2) sigma3 pi3 Sskip n -> exists sigma2:Numbers.BinNums.Z -> value, exists pi2: Numbers.BinNums.Z -> value, exists n1:Numbers.BinNums.Z, exists n2: Numbers.BinNums.Z, many_steps sigma1 pi1 i1 sigma2 pi2 Sskip n1 /\ many_steps sigma2 pi2 i2 sigma3 pi3 Sskip n2 /\ (n = ((1%Z + n1)%Z + n2)%Z). (* Why3 assumption *) Definition valid_fmla (p:fmla) : Prop := forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value), eval_fmla sigma pi p. (* Why3 assumption *) Definition valid_triple (p:fmla) (i:stmt) (q:fmla) : Prop := forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value), eval_fmla sigma pi p -> forall (sigma':Numbers.BinNums.Z -> value) (pi':Numbers.BinNums.Z -> value) (n:Numbers.BinNums.Z), many_steps sigma pi i sigma' pi' Sskip n -> eval_fmla sigma' pi' q. Axiom fset : forall (a:Type), Type. Parameter fset_WhyType : forall (a:Type) {a_WT:WhyType a}, WhyType (fset a). Existing Instance fset_WhyType. Parameter mem: forall {a:Type} {a_WT:WhyType a}, a -> fset a -> Prop. (* Why3 assumption *) Definition infix_eqeq {a:Type} {a_WT:WhyType a} (s1:fset a) (s2:fset a) : Prop := forall (x:a), mem x s1 <-> mem x s2. Axiom extensionality : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), infix_eqeq s1 s2 -> (s1 = s2). (* Why3 assumption *) Definition subset {a:Type} {a_WT:WhyType a} (s1:fset a) (s2:fset a) : Prop := forall (x:a), mem x s1 -> mem x s2. Axiom subset_refl : forall {a:Type} {a_WT:WhyType a}, forall (s:fset a), subset s s. Axiom subset_trans : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a) (s3:fset a), subset s1 s2 -> subset s2 s3 -> subset s1 s3. (* Why3 assumption *) Definition is_empty {a:Type} {a_WT:WhyType a} (s:fset a) : Prop := forall (x:a), ~ mem x s. Parameter empty: forall {a:Type} {a_WT:WhyType a}, fset a. Axiom is_empty_empty : forall {a:Type} {a_WT:WhyType a}, is_empty (empty : fset a). Axiom empty_is_empty : forall {a:Type} {a_WT:WhyType a}, forall (s:fset a), is_empty s -> (s = (empty : fset a)). Parameter add: forall {a:Type} {a_WT:WhyType a}, a -> fset a -> fset a. Axiom add_def : forall {a:Type} {a_WT:WhyType a}, forall (x:a) (s:fset a) (y:a), mem y (add x s) <-> mem y s \/ (y = x). Axiom mem_singleton : forall {a:Type} {a_WT:WhyType a}, forall (x:a) (y:a), mem y (add x (empty : fset a)) -> (y = x). Parameter remove: forall {a:Type} {a_WT:WhyType a}, a -> fset a -> fset a. Axiom remove_def : forall {a:Type} {a_WT:WhyType a}, forall (x:a) (s:fset a) (y:a), mem y (remove x s) <-> mem y s /\ ~ (y = x). Axiom add_remove : forall {a:Type} {a_WT:WhyType a}, forall (x:a) (s:fset a), mem x s -> ((add x (remove x s)) = s). Axiom remove_add : forall {a:Type} {a_WT:WhyType a}, forall (x:a) (s:fset a), ((remove x (add x s)) = (remove x s)). Axiom subset_remove : forall {a:Type} {a_WT:WhyType a}, forall (x:a) (s:fset a), subset (remove x s) s. Parameter union: forall {a:Type} {a_WT:WhyType a}, fset a -> fset a -> fset a. Axiom union_def : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a) (x:a), mem x (union s1 s2) <-> mem x s1 \/ mem x s2. Axiom subset_union_1 : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), subset s1 (union s1 s2). Axiom subset_union_2 : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), subset s2 (union s1 s2). Parameter inter: forall {a:Type} {a_WT:WhyType a}, fset a -> fset a -> fset a. Axiom inter_def : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a) (x:a), mem x (inter s1 s2) <-> mem x s1 /\ mem x s2. Axiom subset_inter_1 : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), subset (inter s1 s2) s1. Axiom subset_inter_2 : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), subset (inter s1 s2) s2. Parameter diff: forall {a:Type} {a_WT:WhyType a}, fset a -> fset a -> fset a. Axiom diff_def : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a) (x:a), mem x (diff s1 s2) <-> mem x s1 /\ ~ mem x s2. Axiom subset_diff : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), subset (diff s1 s2) s1. Parameter pick: forall {a:Type} {a_WT:WhyType a}, fset a -> a. Axiom pick_def : forall {a:Type} {a_WT:WhyType a}, forall (s:fset a), ~ is_empty s -> mem (pick s) s. (* Why3 assumption *) Definition disjoint {a:Type} {a_WT:WhyType a} (s1:fset a) (s2:fset a) : Prop := forall (x:a), ~ mem x s1 \/ ~ mem x s2. Axiom disjoint_inter_empty : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), disjoint s1 s2 <-> is_empty (inter s1 s2). Axiom disjoint_diff_eq : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), disjoint s1 s2 <-> ((diff s1 s2) = s1). Axiom disjoint_diff_s2 : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), disjoint (diff s1 s2) s2. Parameter filter: forall {a:Type} {a_WT:WhyType a}, fset a -> (a -> Init.Datatypes.bool) -> fset a. Axiom filter_def : forall {a:Type} {a_WT:WhyType a}, forall (s:fset a) (p:a -> Init.Datatypes.bool) (x:a), mem x (filter s p) <-> mem x s /\ ((p x) = Init.Datatypes.true). Axiom subset_filter : forall {a:Type} {a_WT:WhyType a}, forall (s:fset a) (p:a -> Init.Datatypes.bool), subset (filter s p) s. Parameter map: forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b}, (a -> b) -> fset a -> fset b. Axiom map_def : forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b}, forall (f:a -> b) (u:fset a) (y:b), mem y (map f u) <-> (exists x:a, mem x u /\ (y = (f x))). Axiom mem_map : forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b}, forall (f:a -> b) (u:fset a), forall (x:a), mem x u -> mem (f x) (map f u). Parameter cardinal: forall {a:Type} {a_WT:WhyType a}, fset a -> Numbers.BinNums.Z. Axiom cardinal_nonneg : forall {a:Type} {a_WT:WhyType a}, forall (s:fset a), (0%Z <= (cardinal s))%Z. Axiom cardinal_empty : forall {a:Type} {a_WT:WhyType a}, forall (s:fset a), is_empty s <-> ((cardinal s) = 0%Z). Axiom cardinal_add : forall {a:Type} {a_WT:WhyType a}, forall (x:a), forall (s:fset a), (mem x s -> ((cardinal (add x s)) = (cardinal s))) /\ (~ mem x s -> ((cardinal (add x s)) = ((cardinal s) + 1%Z)%Z)). Axiom cardinal_remove : forall {a:Type} {a_WT:WhyType a}, forall (x:a), forall (s:fset a), (mem x s -> ((cardinal (remove x s)) = ((cardinal s) - 1%Z)%Z)) /\ (~ mem x s -> ((cardinal (remove x s)) = (cardinal s))). Axiom cardinal_subset : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), subset s1 s2 -> ((cardinal s1) <= (cardinal s2))%Z. Axiom subset_eq : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), subset s1 s2 -> ((cardinal s1) = (cardinal s2)) -> (s1 = s2). Axiom cardinal1 : forall {a:Type} {a_WT:WhyType a}, forall (s:fset a), ((cardinal s) = 1%Z) -> forall (x:a), mem x s -> (x = (pick s)). Axiom cardinal_union : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), ((cardinal (union s1 s2)) = (((cardinal s1) + (cardinal s2))%Z - (cardinal (inter s1 s2)))%Z). Axiom cardinal_inter_disjoint : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), disjoint s1 s2 -> ((cardinal (inter s1 s2)) = 0%Z). Axiom cardinal_diff : forall {a:Type} {a_WT:WhyType a}, forall (s1:fset a) (s2:fset a), ((cardinal (diff s1 s2)) = ((cardinal s1) - (cardinal (inter s1 s2)))%Z). Axiom cardinal_filter : forall {a:Type} {a_WT:WhyType a}, forall (s:fset a) (p:a -> Init.Datatypes.bool), ((cardinal (filter s p)) <= (cardinal s))%Z. Axiom cardinal_map : forall {a:Type} {a_WT:WhyType a} {b:Type} {b_WT:WhyType b}, forall (f:a -> b) (s:fset a), ((cardinal (map f s)) <= (cardinal s))%Z. Axiom set : Type. Parameter set_WhyType : WhyType set. Existing Instance set_WhyType. Parameter to_fset: set -> fset Numbers.BinNums.Z. Parameter choose: set -> Numbers.BinNums.Z. Axiom choose_spec : forall (s:set), ~ is_empty (to_fset s) -> mem (choose s) (to_fset s). (* Why3 assumption *) Definition assigns (sigma:Numbers.BinNums.Z -> value) (a:fset Numbers.BinNums.Z) (sigma':Numbers.BinNums.Z -> value) : Prop := forall (i:Numbers.BinNums.Z), ~ mem i a -> ((sigma i) = (sigma' i)). Axiom assigns_refl : forall (sigma:Numbers.BinNums.Z -> value) (a:fset Numbers.BinNums.Z), assigns sigma a sigma. Axiom assigns_trans : forall (sigma1:Numbers.BinNums.Z -> value) (sigma2:Numbers.BinNums.Z -> value) (sigma3:Numbers.BinNums.Z -> value) (a:fset Numbers.BinNums.Z), assigns sigma1 a sigma2 /\ assigns sigma2 a sigma3 -> assigns sigma1 a sigma3. Axiom assigns_union_left : forall (sigma:Numbers.BinNums.Z -> value) (sigma':Numbers.BinNums.Z -> value) (s1:fset Numbers.BinNums.Z) (s2:fset Numbers.BinNums.Z), assigns sigma s1 sigma' -> assigns sigma (union s1 s2) sigma'. Axiom assigns_union_right : forall (sigma:Numbers.BinNums.Z -> value) (sigma':Numbers.BinNums.Z -> value) (s1:fset Numbers.BinNums.Z) (s2:fset Numbers.BinNums.Z), assigns sigma s2 sigma' -> assigns sigma (union s1 s2) sigma'. (* Why3 assumption *) Fixpoint stmt_writes (i:stmt) (w:fset Numbers.BinNums.Z) {struct i}: Prop := match i with | Sskip|(Sassert _) => True | Sassign id _ => mem id w | (Sseq s1 s2)|(Sif _ s1 s2) => stmt_writes s1 w /\ stmt_writes s2 w | Swhile _ _ s => stmt_writes s w end. (* Why3 goal *) Theorem VC_compute_writes : forall (s:stmt), forall (result:set), (exists x:term, exists x1:fmla, exists x2:stmt, (s = (Swhile x x1 x2)) /\ (forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (sigma':Numbers.BinNums.Z -> value) (pi':Numbers.BinNums.Z -> value) (n:Numbers.BinNums.Z), many_steps sigma pi x2 sigma' pi' Sskip n -> assigns sigma (to_fset result) sigma')) -> forall (sigma:Numbers.BinNums.Z -> value) (pi:Numbers.BinNums.Z -> value) (sigma':Numbers.BinNums.Z -> value) (pi':Numbers.BinNums.Z -> value) (n:Numbers.BinNums.Z), many_steps sigma pi s sigma' pi' Sskip n -> assigns sigma (to_fset result) sigma'. Proof. intros s result (x,(x1,(x2,(h1,h2)))); rewrite h1 in *. intros. generalize sigma pi sigma' pi' H. generalize (steps_non_neg _ _ _ _ _ _ _ H). clear sigma pi sigma' pi' H. intro H_n_pos. pattern n. apply Z_lt_induction; auto. intros x3 Hind sigma pi sigma' qt' Hsteps. inversion Hsteps; subst; clear Hsteps. inversion H; subst; clear H. generalize (many_steps_seq _ _ _ _ _ _ _ H0). intros (sigma3&pi3&n1&n2&h3&h4&h5). apply assigns_trans with sigma3; split. eapply h2; eauto. eapply Hind; eauto. generalize (steps_non_neg _ _ _ _ _ _ _ h3). generalize (steps_non_neg _ _ _ _ _ _ _ h4). omega. inversion H0; subst; clear H0. apply assigns_refl. inversion H. Qed.